mirror of https://github.com/nmvdw/HITs-Examples
Splitted cons_repr
This commit is contained in:
parent
5ee7053631
commit
2ccece3225
|
@ -4,14 +4,18 @@ lattice.v
|
||||||
disjunction.v
|
disjunction.v
|
||||||
representations/bad.v
|
representations/bad.v
|
||||||
representations/definition.v
|
representations/definition.v
|
||||||
|
representations/cons_repr.v
|
||||||
|
fsets/operations_cons_repr.v
|
||||||
|
fsets/properties_cons_repr.v
|
||||||
|
fsets/isomorphism.v
|
||||||
fsets/operations.v
|
fsets/operations.v
|
||||||
fsets/properties.v
|
fsets/properties.v
|
||||||
fsets/operations_decidable.v
|
fsets/operations_decidable.v
|
||||||
fsets/extensionality.v
|
fsets/extensionality.v
|
||||||
fsets/properties_decidable.v
|
fsets/properties_decidable.v
|
||||||
|
fsets/length.v
|
||||||
fsets/monad.v
|
fsets/monad.v
|
||||||
FSets.v
|
FSets.v
|
||||||
representations/cons_repr.v
|
|
||||||
implementations/lists.v
|
implementations/lists.v
|
||||||
variations/enumerated.v
|
variations/enumerated.v
|
||||||
#empty_set.v
|
#empty_set.v
|
||||||
|
|
|
@ -0,0 +1,69 @@
|
||||||
|
(* The representations [FSet A] and [FSetC A] are isomorphic for every A *)
|
||||||
|
Require Import HoTT HitTactics.
|
||||||
|
From representations Require Import cons_repr definition.
|
||||||
|
From fsets Require Import operations_cons_repr properties_cons_repr.
|
||||||
|
|
||||||
|
Section Iso.
|
||||||
|
|
||||||
|
Context {A : Type}.
|
||||||
|
Context `{Univalence}.
|
||||||
|
|
||||||
|
Definition FSetC_to_FSet: FSetC A -> FSet A.
|
||||||
|
Proof.
|
||||||
|
hrecursion.
|
||||||
|
- apply E.
|
||||||
|
- intros a x. apply (U (L a) x).
|
||||||
|
- intros. cbn.
|
||||||
|
etransitivity. apply assoc.
|
||||||
|
apply (ap (fun y => U y x)).
|
||||||
|
apply idem.
|
||||||
|
- intros. cbn.
|
||||||
|
etransitivity. apply assoc.
|
||||||
|
etransitivity. refine (ap (fun y => U y x) _ ).
|
||||||
|
apply FSet.comm.
|
||||||
|
symmetry.
|
||||||
|
apply assoc.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Definition FSet_to_FSetC: FSet A -> FSetC A :=
|
||||||
|
FSet_rec A (FSetC A) (FSetC.trunc A) Nil singleton append append_assoc
|
||||||
|
append_comm append_nl append_nr singleton_idem.
|
||||||
|
|
||||||
|
Lemma append_union: forall (x y: FSetC A),
|
||||||
|
FSetC_to_FSet (append x y) = U (FSetC_to_FSet x) (FSetC_to_FSet y).
|
||||||
|
Proof.
|
||||||
|
intros x.
|
||||||
|
hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
|
||||||
|
- intros. symmetry. apply nl.
|
||||||
|
- intros a x HR y. rewrite HR. apply assoc.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Lemma repr_iso_id_l: forall (x: FSet A), FSetC_to_FSet (FSet_to_FSetC x) = x.
|
||||||
|
Proof.
|
||||||
|
hinduction; try (intros; apply set_path2).
|
||||||
|
- reflexivity.
|
||||||
|
- intro. apply nr.
|
||||||
|
- intros x y p q. rewrite append_union, p, q. reflexivity.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
|
||||||
|
Lemma repr_iso_id_r: forall (x: FSetC A), FSet_to_FSetC (FSetC_to_FSet x) = x.
|
||||||
|
Proof.
|
||||||
|
hinduction; try (intros; apply set_path2).
|
||||||
|
- reflexivity.
|
||||||
|
- intros a x HR. rewrite HR. reflexivity.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
|
||||||
|
Theorem repr_iso: FSet A <~> FSetC A.
|
||||||
|
Proof.
|
||||||
|
simple refine (@BuildEquiv (FSet A) (FSetC A) FSet_to_FSetC _ ).
|
||||||
|
apply isequiv_biinv.
|
||||||
|
unfold BiInv. split.
|
||||||
|
exists FSetC_to_FSet.
|
||||||
|
unfold Sect. apply repr_iso_id_l.
|
||||||
|
exists FSetC_to_FSet.
|
||||||
|
unfold Sect. apply repr_iso_id_r.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
End Iso.
|
|
@ -0,0 +1,40 @@
|
||||||
|
(* 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}.
|
||||||
|
|
||||||
|
Opaque isIn_b.
|
||||||
|
|
||||||
|
Definition length (x: FSetC A) : nat.
|
||||||
|
Proof.
|
||||||
|
simple refine (FSetC_ind A _ _ _ _ _ _ x ); simpl.
|
||||||
|
- exact 0.
|
||||||
|
- intros a y n.
|
||||||
|
pose (y' := FSetC_to_FSet y).
|
||||||
|
exact (if isIn_b a y' then n else (S n)).
|
||||||
|
- intros. rewrite transport_const. cbn.
|
||||||
|
simplify_isIn. simpl. reflexivity.
|
||||||
|
- intros. rewrite transport_const. cbn.
|
||||||
|
simplify_isIn.
|
||||||
|
destruct (dec (a = b)) as [Hab | Hab].
|
||||||
|
+ rewrite Hab. simplify_isIn. simpl. reflexivity.
|
||||||
|
+ rewrite ?L_isIn_b_false; auto. simpl.
|
||||||
|
destruct (isIn_b a (FSetC_to_FSet x0)), (isIn_b b (FSetC_to_FSet x0)) ; reflexivity.
|
||||||
|
intro p. contradiction (Hab p^).
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Definition length_FSet (x: FSet A) := length (FSet_to_FSetC x).
|
||||||
|
|
||||||
|
Lemma length_singleton: forall (a: A), length_FSet (L a) = 1.
|
||||||
|
Proof.
|
||||||
|
intro a.
|
||||||
|
cbn. reflexivity.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
End Length.
|
|
@ -0,0 +1,19 @@
|
||||||
|
(* Operations on [FSetC A] *)
|
||||||
|
Require Import HoTT HitTactics.
|
||||||
|
Require Import representations.cons_repr.
|
||||||
|
|
||||||
|
Section operations.
|
||||||
|
|
||||||
|
Context {A : Type}.
|
||||||
|
|
||||||
|
Definition append (x y: FSetC A) : FSetC A.
|
||||||
|
hinduction x.
|
||||||
|
- apply y.
|
||||||
|
- apply Cns.
|
||||||
|
- apply dupl.
|
||||||
|
- apply comm.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Definition singleton (a:A) : FSetC A := a;;∅.
|
||||||
|
|
||||||
|
End operations.
|
|
@ -0,0 +1,75 @@
|
||||||
|
(* Properties of the operations on [FSetC A] *)
|
||||||
|
Require Import HoTT HitTactics.
|
||||||
|
Require Import representations.cons_repr.
|
||||||
|
From fsets Require Import operations_cons_repr.
|
||||||
|
|
||||||
|
Section properties.
|
||||||
|
Context {A : Type}.
|
||||||
|
|
||||||
|
Lemma append_nl:
|
||||||
|
forall (x: FSetC A), append ∅ x = x.
|
||||||
|
Proof.
|
||||||
|
intro. reflexivity.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Lemma append_nr:
|
||||||
|
forall (x: FSetC A), append x ∅ = x.
|
||||||
|
Proof.
|
||||||
|
hinduction; try (intros; apply set_path2).
|
||||||
|
- reflexivity.
|
||||||
|
- intros. apply (ap (fun y => a ;; y) X).
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Lemma append_assoc {H: Funext}:
|
||||||
|
forall (x y z: FSetC A), append x (append y z) = append (append x y) z.
|
||||||
|
Proof.
|
||||||
|
intro x; hinduction x; try (intros; apply set_path2).
|
||||||
|
- reflexivity.
|
||||||
|
- intros a x HR y z.
|
||||||
|
specialize (HR y z).
|
||||||
|
apply (ap (fun y => a ;; y) HR).
|
||||||
|
- intros. apply path_forall.
|
||||||
|
intro. apply path_forall.
|
||||||
|
intro. apply set_path2.
|
||||||
|
- intros. apply path_forall.
|
||||||
|
intro. apply path_forall.
|
||||||
|
intro. apply set_path2.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Lemma append_singleton: forall (a: A) (x: FSetC A),
|
||||||
|
a ;; x = append x (a ;; ∅).
|
||||||
|
Proof.
|
||||||
|
intro a. hinduction; try (intros; apply set_path2).
|
||||||
|
- reflexivity.
|
||||||
|
- intros b x HR. refine (_ @ _).
|
||||||
|
+ apply comm.
|
||||||
|
+ apply (ap (fun y => b ;; y) HR).
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Lemma append_comm {H: Funext}:
|
||||||
|
forall (x1 x2: FSetC A), append x1 x2 = append x2 x1.
|
||||||
|
Proof.
|
||||||
|
intro x1.
|
||||||
|
hinduction x1; try (intros; apply set_path2).
|
||||||
|
- intros. symmetry. apply append_nr.
|
||||||
|
- intros a x1 HR x2.
|
||||||
|
etransitivity.
|
||||||
|
apply (ap (fun y => a;;y) (HR x2)).
|
||||||
|
transitivity (append (append x2 x1) (a;;∅)).
|
||||||
|
+ apply append_singleton.
|
||||||
|
+ etransitivity.
|
||||||
|
* symmetry. apply append_assoc.
|
||||||
|
* simple refine (ap (fun y => append x2 y) (append_singleton _ _)^).
|
||||||
|
- intros. apply path_forall.
|
||||||
|
intro. apply set_path2.
|
||||||
|
- intros. apply path_forall.
|
||||||
|
intro. apply set_path2.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Lemma singleton_idem: forall (a: A),
|
||||||
|
append (singleton a) (singleton a) = singleton a.
|
||||||
|
Proof.
|
||||||
|
unfold singleton. intro. cbn. apply dupl.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
End properties.
|
|
@ -1,5 +1,7 @@
|
||||||
|
(* Implementation of [FSet A] using lists *)
|
||||||
Require Import HoTT HitTactics.
|
Require Import HoTT HitTactics.
|
||||||
Require Import representations.cons_repr FSets.
|
Require Import representations.cons_repr FSets.
|
||||||
|
From fsets Require Import operations_cons_repr isomorphism length.
|
||||||
|
|
||||||
Section Operations.
|
Section Operations.
|
||||||
Variable A : Type.
|
Variable A : Type.
|
||||||
|
@ -80,7 +82,7 @@ Section ListToSet.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma append_FSetCappend (l1 l2 : list A) :
|
Lemma append_FSetCappend (l1 l2 : list A) :
|
||||||
list_to_setC (append l1 l2) = FSetC.append (list_to_setC l1) (list_to_setC l2).
|
list_to_setC (append l1 l2) = operations_cons_repr.append (list_to_setC l1) (list_to_setC l2).
|
||||||
Proof.
|
Proof.
|
||||||
induction l1 ; simpl in *.
|
induction l1 ; simpl in *.
|
||||||
- reflexivity.
|
- reflexivity.
|
||||||
|
@ -117,7 +119,7 @@ Section ListToSet.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma length_sizeC (l : list A) :
|
Lemma length_sizeC (l : list A) :
|
||||||
cardinality l = cons_repr.length (list_to_setC l).
|
cardinality l = length (list_to_setC l).
|
||||||
Proof.
|
Proof.
|
||||||
induction l.
|
induction l.
|
||||||
- cbn.
|
- cbn.
|
||||||
|
|
|
@ -1,6 +1,5 @@
|
||||||
|
(* Definition of Finite Sets as via cons lists *)
|
||||||
Require Import HoTT HitTactics.
|
Require Import HoTT HitTactics.
|
||||||
Require Import representations.definition.
|
|
||||||
From fsets Require Import operations_decidable properties_decidable.
|
|
||||||
|
|
||||||
Module Export FSetC.
|
Module Export FSetC.
|
||||||
|
|
||||||
|
@ -90,7 +89,6 @@ Definition FSetC_rec_beta_dupl : forall (a: A) (x : FSetC A),
|
||||||
ap FSetC_rec (dupl a x) = duplP a (FSetC_rec x).
|
ap FSetC_rec (dupl a x) = duplP a (FSetC_rec x).
|
||||||
Proof.
|
Proof.
|
||||||
intros.
|
intros.
|
||||||
unfold FSet_rec.
|
|
||||||
eapply (cancelL (transport_const (dupl a x) _)).
|
eapply (cancelL (transport_const (dupl a x) _)).
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
simple refine ((apD_const _ _)^ @ _).
|
||||||
apply FSetC_ind_beta_dupl.
|
apply FSetC_ind_beta_dupl.
|
||||||
|
@ -100,7 +98,6 @@ Definition FSetC_rec_beta_comm : forall (a b: A) (x : FSetC A),
|
||||||
ap FSetC_rec (comm a b x) = commP a b (FSetC_rec x).
|
ap FSetC_rec (comm a b x) = commP a b (FSetC_rec x).
|
||||||
Proof.
|
Proof.
|
||||||
intros.
|
intros.
|
||||||
unfold FSet_rec.
|
|
||||||
eapply (cancelL (transport_const (comm a b x) _)).
|
eapply (cancelL (transport_const (comm a b x) _)).
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
simple refine ((apD_const _ _)^ @ _).
|
||||||
apply FSetC_ind_beta_comm.
|
apply FSetC_ind_beta_comm.
|
||||||
|
@ -113,211 +110,7 @@ Instance FSetC_recursion A : HitRecursion (FSetC A) := {
|
||||||
indTy := _; recTy := _;
|
indTy := _; recTy := _;
|
||||||
H_inductor := FSetC_ind A; H_recursor := FSetC_rec A }.
|
H_inductor := FSetC_ind A; H_recursor := FSetC_rec A }.
|
||||||
|
|
||||||
|
End FSetC.
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
Section Append.
|
|
||||||
|
|
||||||
Context {A : Type}.
|
|
||||||
Context {A_deceq : DecidablePaths A}.
|
|
||||||
|
|
||||||
Definition append (x y: FSetC A) : FSetC A.
|
|
||||||
hinduction x.
|
|
||||||
- apply y.
|
|
||||||
- apply Cns.
|
|
||||||
- apply dupl.
|
|
||||||
- apply comm.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
|
|
||||||
Lemma append_nl:
|
|
||||||
forall (x: FSetC A), append ∅ x = x.
|
|
||||||
Proof.
|
|
||||||
intro. reflexivity.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma append_nr:
|
|
||||||
forall (x: FSetC A), append x ∅ = x.
|
|
||||||
Proof.
|
|
||||||
hinduction; try (intros; apply set_path2).
|
|
||||||
- reflexivity.
|
|
||||||
- intros. apply (ap (fun y => a ;; y) X).
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma append_assoc {H: Funext}:
|
|
||||||
forall (x y z: FSetC A), append x (append y z) = append (append x y) z.
|
|
||||||
Proof.
|
|
||||||
intro x; hinduction x; try (intros; apply set_path2).
|
|
||||||
- reflexivity.
|
|
||||||
- intros a x HR y z.
|
|
||||||
specialize (HR y z).
|
|
||||||
apply (ap (fun y => a ;; y) HR).
|
|
||||||
- intros. apply path_forall.
|
|
||||||
intro. apply path_forall.
|
|
||||||
intro. apply set_path2.
|
|
||||||
- intros. apply path_forall.
|
|
||||||
intro. apply path_forall.
|
|
||||||
intro. apply set_path2.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma aux: forall (a: A) (x: FSetC A),
|
|
||||||
a ;; x = append x (a ;; ∅).
|
|
||||||
Proof.
|
|
||||||
intro a. hinduction; try (intros; apply set_path2).
|
|
||||||
- reflexivity.
|
|
||||||
- intros b x HR. refine (_ @ _).
|
|
||||||
+ apply comm.
|
|
||||||
+ apply (ap (fun y => b ;; y) HR).
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
|
|
||||||
Lemma append_comm {H: Funext}:
|
|
||||||
forall (x1 x2: FSetC A), append x1 x2 = append x2 x1.
|
|
||||||
Proof.
|
|
||||||
intro x1.
|
|
||||||
hinduction x1; try (intros; apply set_path2).
|
|
||||||
- intros. symmetry. apply append_nr.
|
|
||||||
- intros a x1 HR x2.
|
|
||||||
etransitivity.
|
|
||||||
apply (ap (fun y => a;;y) (HR x2)).
|
|
||||||
transitivity (append (append x2 x1) (a;;∅)).
|
|
||||||
+ apply aux.
|
|
||||||
+ etransitivity.
|
|
||||||
* symmetry. apply append_assoc.
|
|
||||||
* simple refine (ap (fun y => append x2 y) (aux _ _)^).
|
|
||||||
- intros. apply path_forall.
|
|
||||||
intro. apply set_path2.
|
|
||||||
- intros. apply path_forall.
|
|
||||||
intro. apply set_path2.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
|
|
||||||
End Append.
|
|
||||||
|
|
||||||
|
|
||||||
Section Singleton.
|
|
||||||
|
|
||||||
Context {A : Type}.
|
|
||||||
Context {A_deceq : DecidablePaths A}.
|
|
||||||
|
|
||||||
Definition singleton (a:A) : FSetC A := a;;∅.
|
|
||||||
|
|
||||||
Lemma singleton_idem: forall (a: A),
|
|
||||||
append (singleton a) (singleton a) = singleton a.
|
|
||||||
Proof.
|
|
||||||
unfold singleton. intro. cbn. apply dupl.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
End Singleton.
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
Infix ";;" := Cns (at level 8, right associativity).
|
Infix ";;" := Cns (at level 8, right associativity).
|
||||||
Notation "∅" := Nil.
|
Notation "∅" := Nil.
|
||||||
|
|
||||||
End FSetC.
|
|
||||||
|
|
||||||
|
|
||||||
Section Iso.
|
|
||||||
|
|
||||||
Context {A : Type}.
|
|
||||||
Context {A_deceq : DecidablePaths A}.
|
|
||||||
Context {H : Funext}.
|
|
||||||
|
|
||||||
|
|
||||||
Definition FSetC_to_FSet: FSetC A -> FSet A.
|
|
||||||
Proof.
|
|
||||||
hrecursion.
|
|
||||||
- apply E.
|
|
||||||
- intros a x. apply (U (L a) x).
|
|
||||||
- intros. cbn.
|
|
||||||
etransitivity. apply assoc.
|
|
||||||
apply (ap (fun y => U y x)).
|
|
||||||
apply idem.
|
|
||||||
- intros. cbn.
|
|
||||||
etransitivity. apply assoc.
|
|
||||||
etransitivity. refine (ap (fun y => U y x) _ ).
|
|
||||||
apply FSet.comm.
|
|
||||||
symmetry.
|
|
||||||
apply assoc.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition FSet_to_FSetC: FSet A -> FSetC A :=
|
|
||||||
FSet_rec A (FSetC A) (trunc A) ∅ singleton append append_assoc
|
|
||||||
append_comm append_nl append_nr singleton_idem.
|
|
||||||
|
|
||||||
|
|
||||||
Lemma append_union: forall (x y: FSetC A),
|
|
||||||
FSetC_to_FSet (append x y) = U (FSetC_to_FSet x) (FSetC_to_FSet y).
|
|
||||||
Proof.
|
|
||||||
intros x.
|
|
||||||
hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
|
|
||||||
- intros. symmetry. apply nl.
|
|
||||||
- intros a x HR y. rewrite HR. apply assoc.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma repr_iso_id_l: forall (x: FSet A), FSetC_to_FSet (FSet_to_FSetC x) = x.
|
|
||||||
Proof.
|
|
||||||
hinduction; try (intros; apply set_path2).
|
|
||||||
- reflexivity.
|
|
||||||
- intro. apply nr.
|
|
||||||
- intros x y p q. rewrite append_union, p, q. reflexivity.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
|
|
||||||
Lemma repr_iso_id_r: forall (x: FSetC A), FSet_to_FSetC (FSetC_to_FSet x) = x.
|
|
||||||
Proof.
|
|
||||||
hinduction; try (intros; apply set_path2).
|
|
||||||
- reflexivity.
|
|
||||||
- intros a x HR. rewrite HR. reflexivity.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
|
|
||||||
Theorem repr_iso: FSet A <~> FSetC A.
|
|
||||||
Proof.
|
|
||||||
simple refine (@BuildEquiv (FSet A) (FSetC A) FSet_to_FSetC _ ).
|
|
||||||
apply isequiv_biinv.
|
|
||||||
unfold BiInv. split.
|
|
||||||
exists FSetC_to_FSet.
|
|
||||||
unfold Sect. apply repr_iso_id_l.
|
|
||||||
exists FSetC_to_FSet.
|
|
||||||
unfold Sect. apply repr_iso_id_r.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
End Iso.
|
|
||||||
|
|
||||||
Section Length.
|
|
||||||
|
|
||||||
Context {A : Type}.
|
|
||||||
Context {A_deceq : DecidablePaths A}.
|
|
||||||
Context `{Univalence}.
|
|
||||||
|
|
||||||
Opaque isIn_b.
|
|
||||||
Definition length (x: FSetC A) : nat.
|
|
||||||
Proof.
|
|
||||||
simple refine (FSetC_ind A _ _ _ _ _ _ x ); simpl.
|
|
||||||
- exact 0.
|
|
||||||
- intros a y n.
|
|
||||||
pose (y' := FSetC_to_FSet y).
|
|
||||||
exact (if isIn_b a y' then n else (S n)).
|
|
||||||
- intros. rewrite transport_const. cbn.
|
|
||||||
simplify_isIn. simpl. reflexivity.
|
|
||||||
- intros. rewrite transport_const. cbn.
|
|
||||||
simplify_isIn.
|
|
||||||
destruct (dec (a = b)) as [Hab | Hab].
|
|
||||||
+ rewrite Hab. simplify_isIn. simpl. reflexivity.
|
|
||||||
+ rewrite ?L_isIn_b_false; auto. simpl.
|
|
||||||
destruct (isIn_b a (FSetC_to_FSet x0)), (isIn_b b (FSetC_to_FSet x0)) ; reflexivity.
|
|
||||||
intro p. contradiction (Hab p^).
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition length_FSet (x: FSet A) := length (FSet_to_FSetC x).
|
|
||||||
|
|
||||||
Lemma length_singleton: forall (a: A), length_FSet (L a) = 1.
|
|
||||||
Proof.
|
|
||||||
intro a.
|
|
||||||
cbn. reflexivity.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
End Length.
|
|
||||||
|
|
|
@ -3,7 +3,6 @@ Require Import HoTT.
|
||||||
Require Import HitTactics.
|
Require Import HitTactics.
|
||||||
|
|
||||||
Module Export FSet.
|
Module Export FSet.
|
||||||
|
|
||||||
Section FSet.
|
Section FSet.
|
||||||
Variable A : Type.
|
Variable A : Type.
|
||||||
|
|
||||||
|
@ -78,7 +77,6 @@ Fixpoint FSet_ind
|
||||||
| U y z => fun _ _ _ _ _ _ => uP y z (FSet_ind y) (FSet_ind z)
|
| U y z => fun _ _ _ _ _ _ => uP y z (FSet_ind y) (FSet_ind z)
|
||||||
end) H assocP commP nlP nrP idemP.
|
end) H assocP commP nlP nrP idemP.
|
||||||
|
|
||||||
|
|
||||||
Axiom FSet_ind_beta_assoc : forall (x y z : FSet A),
|
Axiom FSet_ind_beta_assoc : forall (x y z : FSet A),
|
||||||
apD FSet_ind (assoc x y z) =
|
apD FSet_ind (assoc x y z) =
|
||||||
(assocP x y z (FSet_ind x) (FSet_ind y) (FSet_ind z)).
|
(assocP x y z (FSet_ind x) (FSet_ind y) (FSet_ind z)).
|
||||||
|
@ -111,7 +109,8 @@ Variable idemP : forall (x : A), u (l x) (l x) = l x.
|
||||||
|
|
||||||
Definition FSet_rec : FSet A -> P.
|
Definition FSet_rec : FSet A -> P.
|
||||||
Proof.
|
Proof.
|
||||||
simple refine (FSet_ind A _ _ _ _ _ _ _ _ _ _) ; try (intros ; simple refine ((transport_const _ _) @ _)) ; cbn.
|
simple refine (FSet_ind A _ _ _ _ _ _ _ _ _ _)
|
||||||
|
; try (intros ; simple refine ((transport_const _ _) @ _)) ; cbn.
|
||||||
- apply e.
|
- apply e.
|
||||||
- apply l.
|
- apply l.
|
||||||
- intros x y ; apply u.
|
- intros x y ; apply u.
|
||||||
|
|
Loading…
Reference in New Issue