HITs-Examples/FiniteSets/list_representation/length.v

40 lines
1.2 KiB
Coq
Raw Permalink Normal View History

2017-08-02 11:40:03 +02:00
(* The length function for finite sets *)
Require Import HoTT HitTactics.
From representations Require Import cons_repr definition.
From fsets Require Import operations_decidable isomorphism properties_decidable.
Section Length.
Context {A : Type}.
Context {A_deceq : DecidablePaths A}.
Context `{Univalence}.
2017-08-07 16:49:46 +02:00
Definition length (x : FSetC A) : nat.
2017-08-02 11:40:03 +02:00
Proof.
simple refine (FSetC_ind A _ _ _ _ _ _ x ); simpl.
- exact 0.
- intros a y n.
pose (y' := FSetC_to_FSet y).
2017-08-08 17:00:30 +02:00
exact (if a _d y' then n else (S n)).
- intros. rewrite transport_const. simpl.
simplify_isIn_b. reflexivity.
- intros. rewrite transport_const. simpl.
simplify_isIn_b.
2017-08-02 11:40:03 +02:00
destruct (dec (a = b)) as [Hab | Hab].
2017-08-08 17:00:30 +02:00
+ rewrite Hab. simplify_isIn_b. reflexivity.
+ rewrite ?L_isIn_b_false; auto.
++ simpl.
2017-08-08 17:00:30 +02:00
destruct (a _d (FSetC_to_FSet x0)), (b _d (FSetC_to_FSet x0))
; reflexivity.
++ intro p. contradiction (Hab p^).
2017-08-02 11:40:03 +02:00
Defined.
Definition length_FSet (x: FSet A) := length (FSet_to_FSetC x).
2017-08-07 16:49:46 +02:00
Lemma length_singleton: forall (a: A), length_FSet ({|a|}) = 1.
2017-08-02 11:40:03 +02:00
Proof.
intro a.
cbn. reflexivity.
Defined.
End Length.