mirror of
https://github.com/nmvdw/HITs-Examples
synced 2025-11-03 15:13:51 +01:00
Splitted cons_repr
This commit is contained in:
@@ -4,108 +4,108 @@ From representations Require Import definition.
|
||||
From fsets Require Import operations properties.
|
||||
|
||||
Section ext.
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Lemma subset_union_equiv
|
||||
: forall X Y : FSet A, subset X Y <~> U X Y = Y.
|
||||
Proof.
|
||||
intros X Y.
|
||||
eapply equiv_iff_hprop_uncurried.
|
||||
split.
|
||||
- apply subset_union.
|
||||
- intro HXY.
|
||||
rewrite <- HXY.
|
||||
apply subset_union_l.
|
||||
Defined.
|
||||
|
||||
Lemma subset_isIn (X Y : FSet A) :
|
||||
(forall (a : A), isIn a X -> isIn a Y)
|
||||
<~> (subset X Y).
|
||||
Proof.
|
||||
eapply equiv_iff_hprop_uncurried.
|
||||
split.
|
||||
- hinduction X ;
|
||||
try (intros; repeat (apply path_forall; intro); apply equiv_hprop_allpath ; apply _).
|
||||
+ intros ; reflexivity.
|
||||
+ intros a f.
|
||||
apply f.
|
||||
apply tr ; reflexivity.
|
||||
+ intros X1 X2 H1 H2 f.
|
||||
enough (subset X1 Y).
|
||||
enough (subset X2 Y).
|
||||
{ split. apply X. apply X0. }
|
||||
* apply H2.
|
||||
intros a Ha.
|
||||
apply f.
|
||||
apply tr.
|
||||
right.
|
||||
apply Ha.
|
||||
* apply H1.
|
||||
intros a Ha.
|
||||
apply f.
|
||||
apply tr.
|
||||
left.
|
||||
apply Ha.
|
||||
- hinduction X ;
|
||||
try (intros; repeat (apply path_forall; intro); apply equiv_hprop_allpath ; apply _).
|
||||
+ intros. contradiction.
|
||||
+ intros b f a.
|
||||
simple refine (Trunc_ind _ _) ; cbn.
|
||||
intro p.
|
||||
rewrite p^ in f.
|
||||
apply f.
|
||||
+ intros X1 X2 IH1 IH2 [H1 H2] a.
|
||||
simple refine (Trunc_ind _ _) ; cbn.
|
||||
intros [C1 | C2].
|
||||
++ apply (IH1 H1 a C1).
|
||||
++ apply (IH2 H2 a C2).
|
||||
Defined.
|
||||
|
||||
(** ** Extensionality proof *)
|
||||
|
||||
Lemma eq_subset' (X Y : FSet A) : X = Y <~> (U Y X = X) * (U X Y = Y).
|
||||
Proof.
|
||||
unshelve eapply BuildEquiv.
|
||||
{ intro H'. rewrite H'. split; apply union_idem. }
|
||||
unshelve esplit.
|
||||
{ intros [H1 H2]. etransitivity. apply H1^.
|
||||
rewrite comm. apply H2. }
|
||||
intro; apply path_prod; apply set_path2.
|
||||
all: intro; apply set_path2.
|
||||
Defined.
|
||||
|
||||
Lemma eq_subset (X Y : FSet A) :
|
||||
X = Y <~> (subset Y X * subset X Y).
|
||||
Proof.
|
||||
transitivity ((U Y X = X) * (U X Y = Y)).
|
||||
apply eq_subset'.
|
||||
symmetry.
|
||||
eapply equiv_functor_prod'; apply subset_union_equiv.
|
||||
Defined.
|
||||
|
||||
Theorem fset_ext (X Y : FSet A) :
|
||||
X = Y <~> (forall (a : A), isIn a X = isIn a Y).
|
||||
Proof.
|
||||
refine (@equiv_compose' _ _ _ _ _) ; [ | apply eq_subset ].
|
||||
refine (@equiv_compose' _ ((forall a, isIn a Y -> isIn a X)
|
||||
*(forall a, isIn a X -> isIn a Y)) _ _ _).
|
||||
- apply equiv_iff_hprop_uncurried.
|
||||
Lemma subset_union_equiv
|
||||
: forall X Y : FSet A, subset X Y <~> U X Y = Y.
|
||||
Proof.
|
||||
intros X Y.
|
||||
eapply equiv_iff_hprop_uncurried.
|
||||
split.
|
||||
* intros [H1 H2 a].
|
||||
specialize (H1 a) ; specialize (H2 a).
|
||||
apply path_iff_hprop.
|
||||
apply H2.
|
||||
apply H1.
|
||||
* intros H1.
|
||||
split ; intro a ; intro H2.
|
||||
- apply subset_union.
|
||||
- intro HXY.
|
||||
rewrite <- HXY.
|
||||
apply subset_union_l.
|
||||
Defined.
|
||||
|
||||
Lemma subset_isIn (X Y : FSet A) :
|
||||
(forall (a : A), isIn a X -> isIn a Y)
|
||||
<~> (subset X Y).
|
||||
Proof.
|
||||
eapply equiv_iff_hprop_uncurried.
|
||||
split.
|
||||
- hinduction X ;
|
||||
try (intros; repeat (apply path_forall; intro); apply equiv_hprop_allpath ; apply _).
|
||||
+ intros ; reflexivity.
|
||||
+ intros a f.
|
||||
apply f.
|
||||
apply tr ; reflexivity.
|
||||
+ intros X1 X2 H1 H2 f.
|
||||
enough (subset X1 Y).
|
||||
enough (subset X2 Y).
|
||||
{ split. apply X. apply X0. }
|
||||
* apply H2.
|
||||
intros a Ha.
|
||||
apply f.
|
||||
apply tr.
|
||||
right.
|
||||
apply Ha.
|
||||
* apply H1.
|
||||
intros a Ha.
|
||||
apply f.
|
||||
apply tr.
|
||||
left.
|
||||
apply Ha.
|
||||
- hinduction X ;
|
||||
try (intros; repeat (apply path_forall; intro); apply equiv_hprop_allpath ; apply _).
|
||||
+ intros. contradiction.
|
||||
+ intros b f a.
|
||||
simple refine (Trunc_ind _ _) ; cbn.
|
||||
intro p.
|
||||
rewrite p^ in f.
|
||||
apply f.
|
||||
+ intros X1 X2 IH1 IH2 [H1 H2] a.
|
||||
simple refine (Trunc_ind _ _) ; cbn.
|
||||
intros [C1 | C2].
|
||||
++ apply (IH1 H1 a C1).
|
||||
++ apply (IH2 H2 a C2).
|
||||
Defined.
|
||||
|
||||
(** ** Extensionality proof *)
|
||||
|
||||
Lemma eq_subset' (X Y : FSet A) : X = Y <~> (U Y X = X) * (U X Y = Y).
|
||||
Proof.
|
||||
unshelve eapply BuildEquiv.
|
||||
{ intro H'. rewrite H'. split; apply union_idem. }
|
||||
unshelve esplit.
|
||||
{ intros [H1 H2]. etransitivity. apply H1^.
|
||||
rewrite comm. apply H2. }
|
||||
intro; apply path_prod; apply set_path2.
|
||||
all: intro; apply set_path2.
|
||||
Defined.
|
||||
|
||||
Lemma eq_subset (X Y : FSet A) :
|
||||
X = Y <~> (subset Y X * subset X Y).
|
||||
Proof.
|
||||
transitivity ((U Y X = X) * (U X Y = Y)).
|
||||
apply eq_subset'.
|
||||
symmetry.
|
||||
eapply equiv_functor_prod'; apply subset_union_equiv.
|
||||
Defined.
|
||||
|
||||
Theorem fset_ext (X Y : FSet A) :
|
||||
X = Y <~> (forall (a : A), isIn a X = isIn a Y).
|
||||
Proof.
|
||||
refine (@equiv_compose' _ _ _ _ _) ; [ | apply eq_subset ].
|
||||
refine (@equiv_compose' _ ((forall a, isIn a Y -> isIn a X)
|
||||
*(forall a, isIn a X -> isIn a Y)) _ _ _).
|
||||
- apply equiv_iff_hprop_uncurried.
|
||||
split.
|
||||
* intros [H1 H2 a].
|
||||
specialize (H1 a) ; specialize (H2 a).
|
||||
apply path_iff_hprop.
|
||||
apply H2.
|
||||
apply H1.
|
||||
* intros H1.
|
||||
split ; intro a ; intro H2.
|
||||
+ rewrite (H1 a).
|
||||
apply H2.
|
||||
+ rewrite <- (H1 a).
|
||||
apply H2.
|
||||
- eapply equiv_functor_prod' ;
|
||||
apply equiv_iff_hprop_uncurried ;
|
||||
split ; apply subset_isIn.
|
||||
Defined.
|
||||
- eapply equiv_functor_prod' ;
|
||||
apply equiv_iff_hprop_uncurried ;
|
||||
split ; apply subset_isIn.
|
||||
Defined.
|
||||
|
||||
End ext.
|
||||
|
||||
69
FiniteSets/fsets/isomorphism.v
Normal file
69
FiniteSets/fsets/isomorphism.v
Normal file
@@ -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.
|
||||
40
FiniteSets/fsets/length.v
Normal file
40
FiniteSets/fsets/length.v
Normal file
@@ -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.
|
||||
19
FiniteSets/fsets/operations_cons_repr.v
Normal file
19
FiniteSets/fsets/operations_cons_repr.v
Normal file
@@ -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.
|
||||
75
FiniteSets/fsets/properties_cons_repr.v
Normal file
75
FiniteSets/fsets/properties_cons_repr.v
Normal file
@@ -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.
|
||||
Reference in New Issue
Block a user