mirror of https://github.com/nmvdw/HITs-Examples
Some cleaning
This commit is contained in:
parent
a0844f6be4
commit
1bab2206a3
|
@ -9,8 +9,7 @@ Section operations_isIn.
|
||||||
|
|
||||||
Lemma union_idem : forall x: FSet A, U x x = x.
|
Lemma union_idem : forall x: FSet A, U x x = x.
|
||||||
Proof.
|
Proof.
|
||||||
hinduction;
|
hinduction ; try (intros ; apply set_path2).
|
||||||
try (intros ; apply set_path2).
|
|
||||||
- apply nl.
|
- apply nl.
|
||||||
- apply idem.
|
- apply idem.
|
||||||
- intros x y P Q.
|
- intros x y P Q.
|
||||||
|
@ -21,16 +20,16 @@ Section operations_isIn.
|
||||||
rewrite (comm y x).
|
rewrite (comm y x).
|
||||||
rewrite <- (assoc x y y).
|
rewrite <- (assoc x y y).
|
||||||
f_ap.
|
f_ap.
|
||||||
Qed.
|
Defined.
|
||||||
|
|
||||||
(** ** Properties about subset relation. *)
|
(** ** Properties about subset relation. *)
|
||||||
Lemma subset_union (X Y : FSet A) :
|
Lemma subset_union (X Y : FSet A) :
|
||||||
subset X Y -> U X Y = Y.
|
subset X Y -> U X Y = Y.
|
||||||
Proof.
|
Proof.
|
||||||
hinduction X; try (intros; apply path_forall; intro; apply set_path2).
|
hinduction X ; try (intros; apply path_forall; intro; apply set_path2).
|
||||||
- intros. apply nl.
|
- intros. apply nl.
|
||||||
- intros a. hinduction Y;
|
- intros a.
|
||||||
try (intros; apply path_forall; intro; apply set_path2).
|
hinduction Y ; try (intros; apply path_forall; intro; apply set_path2).
|
||||||
+ intro.
|
+ intro.
|
||||||
contradiction.
|
contradiction.
|
||||||
+ intro a0.
|
+ intro a0.
|
||||||
|
@ -57,30 +56,24 @@ Section operations_isIn.
|
||||||
Lemma subset_union_l (X : FSet A) :
|
Lemma subset_union_l (X : FSet A) :
|
||||||
forall Y, subset X (U X Y).
|
forall Y, subset X (U X Y).
|
||||||
Proof.
|
Proof.
|
||||||
hinduction X ;
|
hinduction X ; try (repeat (intro; intros; apply path_forall);
|
||||||
try (repeat (intro; intros; apply path_forall); intro; apply equiv_hprop_allpath ; apply _).
|
intro ; apply equiv_hprop_allpath ; apply _).
|
||||||
- apply (fun _ => tt).
|
- apply (fun _ => tt).
|
||||||
- intros a Y.
|
- intros a Y.
|
||||||
apply tr ; left ; apply tr ; reflexivity.
|
apply (tr(inl(tr idpath))).
|
||||||
- intros X1 X2 HX1 HX2 Y.
|
- intros X1 X2 HX1 HX2 Y.
|
||||||
split.
|
split.
|
||||||
* rewrite <- assoc. apply HX1.
|
* rewrite <- assoc. apply HX1.
|
||||||
* rewrite (comm X1 X2). rewrite <- assoc. apply HX2.
|
* rewrite (comm X1 X2). rewrite <- assoc. apply HX2.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
(* Union and membership *)
|
(* simplify it using extensionality *)
|
||||||
Lemma union_isIn (X Y : FSet A) (a : A) :
|
|
||||||
isIn a (U X Y) = isIn a X ∨ isIn a Y.
|
|
||||||
Proof.
|
|
||||||
reflexivity.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
|
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
|
||||||
comprehension (fun a => orb (ϕ a) (ψ a)) x = U (comprehension ϕ x)
|
comprehension (fun a => orb (ϕ a) (ψ a)) x = U (comprehension ϕ x)
|
||||||
(comprehension ψ x).
|
(comprehension ψ x).
|
||||||
Proof.
|
Proof.
|
||||||
intros ϕ ψ.
|
intros ϕ ψ.
|
||||||
hinduction; try (intros; apply set_path2).
|
hinduction ; try (intros; apply set_path2).
|
||||||
- apply (union_idem _)^.
|
- apply (union_idem _)^.
|
||||||
- intros.
|
- intros.
|
||||||
destruct (ϕ a) ; destruct (ψ a) ; symmetry.
|
destruct (ϕ a) ; destruct (ψ a) ; symmetry.
|
||||||
|
@ -108,78 +101,12 @@ Section properties.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
(** isIn properties *)
|
(** isIn properties *)
|
||||||
Lemma singleton_isIn (a b: A) : isIn a (L b) -> Trunc (-1) (a = b).
|
Definition empty_isIn (a: A) : isIn a E -> Empty := idmap.
|
||||||
Proof.
|
|
||||||
apply idmap.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma empty_isIn (a: A) : isIn a E -> Empty.
|
Definition singleton_isIn (a b: A) : isIn a (L b) -> Trunc (-1) (a = b) := idmap.
|
||||||
Proof.
|
|
||||||
apply idmap.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
(** comprehension properties *)
|
Definition union_isIn (X Y : FSet A) (a : A)
|
||||||
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
|
: isIn a (U X Y) = isIn a X ∨ isIn a Y := idpath.
|
||||||
Proof.
|
|
||||||
hrecursion Y; try (intros; apply set_path2).
|
|
||||||
- reflexivity.
|
|
||||||
- reflexivity.
|
|
||||||
- intros x y IHa IHb.
|
|
||||||
rewrite IHa.
|
|
||||||
rewrite IHb.
|
|
||||||
apply union_idem.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
|
||||||
U (comprehension ϕ X) X = X.
|
|
||||||
Proof.
|
|
||||||
intros ϕ.
|
|
||||||
hrecursion; try (intros ; apply set_path2) ; cbn.
|
|
||||||
- apply union_idem.
|
|
||||||
- intro a.
|
|
||||||
destruct (ϕ a).
|
|
||||||
* apply union_idem.
|
|
||||||
* apply nl.
|
|
||||||
- intros X Y P Q.
|
|
||||||
rewrite assoc.
|
|
||||||
rewrite <- (assoc (comprehension ϕ X) (comprehension ϕ Y) X).
|
|
||||||
rewrite (comm (comprehension ϕ Y) X).
|
|
||||||
rewrite assoc.
|
|
||||||
rewrite P.
|
|
||||||
rewrite <- assoc.
|
|
||||||
rewrite Q.
|
|
||||||
reflexivity.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma merely_choice : forall X : FSet A, hor (X = E) (hexists (fun a => isIn a X)).
|
|
||||||
Proof.
|
|
||||||
hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
|
|
||||||
- apply (tr (inl idpath)).
|
|
||||||
- intro a.
|
|
||||||
refine (tr (inr (tr (a ; tr idpath)))).
|
|
||||||
- intros X Y TX TY.
|
|
||||||
strip_truncations.
|
|
||||||
destruct TX as [XE | HX] ; destruct TY as [YE | HY] ; try(strip_truncations ; apply tr).
|
|
||||||
* apply tr ; left.
|
|
||||||
rewrite XE, YE.
|
|
||||||
apply (union_idem E).
|
|
||||||
* right.
|
|
||||||
destruct HY as [a Ya].
|
|
||||||
apply tr.
|
|
||||||
exists a.
|
|
||||||
apply (tr (inr Ya)).
|
|
||||||
* right.
|
|
||||||
destruct HX as [a Xa].
|
|
||||||
apply tr.
|
|
||||||
exists a.
|
|
||||||
apply (tr (inl Xa)).
|
|
||||||
* right.
|
|
||||||
destruct HX as [a Xa].
|
|
||||||
destruct HY as [b Yb].
|
|
||||||
apply tr.
|
|
||||||
exists a.
|
|
||||||
apply (tr (inl Xa)).
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
|
Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
|
||||||
isIn a (comprehension ϕ X) = if ϕ a then isIn a X else False_hp.
|
isIn a (comprehension ϕ X) = if ϕ a then isIn a X else False_hp.
|
||||||
|
@ -210,4 +137,64 @@ Section properties.
|
||||||
** intros ; apply tr ; right ; assumption.
|
** intros ; apply tr ; right ; assumption.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
(* The proof can be simplified using extensionality *)
|
||||||
|
(** comprehension properties *)
|
||||||
|
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
|
||||||
|
Proof.
|
||||||
|
hrecursion Y; try (intros; apply set_path2).
|
||||||
|
- reflexivity.
|
||||||
|
- reflexivity.
|
||||||
|
- intros x y IHa IHb.
|
||||||
|
rewrite IHa, IHb.
|
||||||
|
apply union_idem.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
(* Can be simplified using extensionality *)
|
||||||
|
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
||||||
|
U (comprehension ϕ X) X = X.
|
||||||
|
Proof.
|
||||||
|
intros ϕ.
|
||||||
|
hrecursion; try (intros ; apply set_path2) ; cbn.
|
||||||
|
- apply union_idem.
|
||||||
|
- intro a.
|
||||||
|
destruct (ϕ a).
|
||||||
|
* apply union_idem.
|
||||||
|
* apply nl.
|
||||||
|
- intros X Y P Q.
|
||||||
|
rewrite assoc.
|
||||||
|
rewrite <- (assoc (comprehension ϕ X) (comprehension ϕ Y) X).
|
||||||
|
rewrite (comm (comprehension ϕ Y) X).
|
||||||
|
rewrite assoc.
|
||||||
|
rewrite P.
|
||||||
|
rewrite <- assoc.
|
||||||
|
rewrite Q.
|
||||||
|
reflexivity.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Lemma merely_choice : forall X : FSet A, hor (X = E) (hexists (fun a => isIn a X)).
|
||||||
|
Proof.
|
||||||
|
hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
|
||||||
|
- apply (tr (inl idpath)).
|
||||||
|
- intro a.
|
||||||
|
refine (tr (inr (tr (a ; tr idpath)))).
|
||||||
|
- intros X Y TX TY.
|
||||||
|
strip_truncations.
|
||||||
|
destruct TX as [XE | HX] ; destruct TY as [YE | HY] ; try(strip_truncations ; apply tr).
|
||||||
|
* refine (tr (inl _)).
|
||||||
|
rewrite XE, YE.
|
||||||
|
apply (union_idem E).
|
||||||
|
* destruct HY as [a Ya].
|
||||||
|
refine (inr (tr _)).
|
||||||
|
exists a.
|
||||||
|
apply (tr (inr Ya)).
|
||||||
|
* destruct HX as [a Xa].
|
||||||
|
refine (inr (tr _)).
|
||||||
|
exists a.
|
||||||
|
apply (tr (inl Xa)).
|
||||||
|
* destruct (HX, HY) as [[a Xa] [b Yb]].
|
||||||
|
refine (inr (tr _)).
|
||||||
|
exists a.
|
||||||
|
apply (tr (inl Xa)).
|
||||||
|
Defined.
|
||||||
|
|
||||||
End properties.
|
End properties.
|
||||||
|
|
|
@ -42,19 +42,22 @@ Module Export FSet.
|
||||||
Arguments nl {_} _.
|
Arguments nl {_} _.
|
||||||
Arguments nr {_} _.
|
Arguments nr {_} _.
|
||||||
Arguments idem {_} _.
|
Arguments idem {_} _.
|
||||||
|
Notation "{| x |}" := (L x).
|
||||||
|
Infix "∪" := U (at level 8, right associativity).
|
||||||
|
Notation "∅" := E.
|
||||||
|
|
||||||
Section FSet_induction.
|
Section FSet_induction.
|
||||||
Variable A: Type.
|
Variable A: Type.
|
||||||
Variable (P : FSet A -> Type).
|
Variable (P : FSet A -> Type).
|
||||||
Variable (eP : P E).
|
Variable (eP : P ∅).
|
||||||
Variable (lP : forall a: A, P (L a)).
|
Variable (lP : forall a: A, P {|a |}).
|
||||||
Variable (uP : forall (x y: FSet A), P x -> P y -> P (U x y)).
|
Variable (uP : forall (x y: FSet A), P x -> P y -> P (x ∪ y)).
|
||||||
Variable (assocP : forall (x y z : FSet A)
|
Variable (assocP : forall (x y z : FSet A)
|
||||||
(px: P x) (py: P y) (pz: P z),
|
(px: P x) (py: P y) (pz: P z),
|
||||||
assoc x y z #
|
assoc x y z #
|
||||||
(uP x (U y z) px (uP y z py pz))
|
(uP x (y ∪ z) px (uP y z py pz))
|
||||||
=
|
=
|
||||||
(uP (U x y) z (uP x y px py) pz)).
|
(uP (x ∪ y) z (uP x y px py) pz)).
|
||||||
Variable (commP : forall (x y: FSet A) (px: P x) (py: P y),
|
Variable (commP : forall (x y: FSet A) (px: P x) (py: P y),
|
||||||
comm x y #
|
comm x y #
|
||||||
uP x y px py = uP y x py px).
|
uP x y px py = uP y x py px).
|
||||||
|
@ -71,26 +74,24 @@ Module Export FSet.
|
||||||
{struct x}
|
{struct x}
|
||||||
: P x
|
: P x
|
||||||
:= (match x return _ -> _ -> _ -> _ -> _ -> P x with
|
:= (match x return _ -> _ -> _ -> _ -> _ -> P x with
|
||||||
| E => fun _ => fun _ => fun _ => fun _ => fun _ => eP
|
| ∅ => fun _ _ _ _ _ => eP
|
||||||
| L a => fun _ => fun _ => fun _ => fun _ => fun _ => lP a
|
| {|a|} => fun _ _ _ _ _ => lP a
|
||||||
| U y z => fun _ => fun _ => fun _ => fun _ => fun _ => uP y z
|
| y ∪ z => fun _ _ _ _ _ => uP y z (FSet_ind y) (FSet_ind z)
|
||||||
(FSet_ind y)
|
|
||||||
(FSet_ind z)
|
|
||||||
end) assocP commP nlP nrP idemP.
|
end) 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)).
|
||||||
|
|
||||||
Axiom FSet_ind_beta_comm : forall (x y : FSet A),
|
Axiom FSet_ind_beta_comm : forall (x y : FSet A),
|
||||||
apD FSet_ind (comm x y) = (commP x y (FSet_ind x) (FSet_ind y)).
|
apD FSet_ind (comm x y) = commP x y (FSet_ind x) (FSet_ind y).
|
||||||
|
|
||||||
Axiom FSet_ind_beta_nl : forall (x : FSet A),
|
Axiom FSet_ind_beta_nl : forall (x : FSet A),
|
||||||
apD FSet_ind (nl x) = (nlP x (FSet_ind x)).
|
apD FSet_ind (nl x) = nlP x (FSet_ind x).
|
||||||
|
|
||||||
Axiom FSet_ind_beta_nr : forall (x : FSet A),
|
Axiom FSet_ind_beta_nr : forall (x : FSet A),
|
||||||
apD FSet_ind (nr x) = (nrP x (FSet_ind x)).
|
apD FSet_ind (nr x) = nrP x (FSet_ind x).
|
||||||
|
|
||||||
Axiom FSet_ind_beta_idem : forall (x : A), apD FSet_ind (idem x) = idemP x.
|
Axiom FSet_ind_beta_idem : forall (x : A), apD FSet_ind (idem x) = idemP x.
|
||||||
End FSet_induction.
|
End FSet_induction.
|
||||||
|
@ -183,9 +184,11 @@ Module Export FSet.
|
||||||
|
|
||||||
End FSet_recursion.
|
End FSet_recursion.
|
||||||
|
|
||||||
Instance FSet_recursion A : HitRecursion (FSet A) := {
|
Instance FSet_recursion A : HitRecursion (FSet A) :=
|
||||||
indTy := _; recTy := _;
|
{
|
||||||
H_inductor := FSet_ind A; H_recursor := FSet_rec A }.
|
indTy := _; recTy := _;
|
||||||
|
H_inductor := FSet_ind A; H_recursor := FSet_rec A
|
||||||
|
}.
|
||||||
|
|
||||||
End FSet.
|
End FSet.
|
||||||
|
|
||||||
|
@ -195,10 +198,12 @@ Notation "∅" := E.
|
||||||
|
|
||||||
Section set_sphere.
|
Section set_sphere.
|
||||||
From HoTT.HIT Require Import Circle.
|
From HoTT.HIT Require Import Circle.
|
||||||
From HoTT Require Import UnivalenceAxiom.
|
Context `{Univalence}.
|
||||||
Instance S1_recursion : HitRecursion S1 := {
|
Instance S1_recursion : HitRecursion S1 :=
|
||||||
indTy := _; recTy := _;
|
{
|
||||||
H_inductor := S1_ind; H_recursor := S1_rec }.
|
indTy := _; recTy := _;
|
||||||
|
H_inductor := S1_ind; H_recursor := S1_rec
|
||||||
|
}.
|
||||||
|
|
||||||
Variable A : Type.
|
Variable A : Type.
|
||||||
|
|
||||||
|
@ -206,8 +211,7 @@ Section set_sphere.
|
||||||
Proof.
|
Proof.
|
||||||
hrecursion x.
|
hrecursion x.
|
||||||
- exact loop.
|
- exact loop.
|
||||||
- etransitivity.
|
- refine (transport_paths_FlFr _ _ @ _).
|
||||||
eapply (@transport_paths_FlFr S1 S1 idmap idmap).
|
|
||||||
hott_simpl.
|
hott_simpl.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
@ -225,11 +229,10 @@ Section set_sphere.
|
||||||
Proof.
|
Proof.
|
||||||
hrecursion x.
|
hrecursion x.
|
||||||
- exact loop.
|
- exact loop.
|
||||||
- etransitivity.
|
- refine (transport_paths_FlFr loop _ @ _).
|
||||||
apply (@transport_paths_FlFr _ _ (fun x => S1op base x) idmap _ _ loop loop).
|
|
||||||
hott_simpl.
|
hott_simpl.
|
||||||
apply moveR_pM. apply moveR_pM. hott_simpl.
|
apply moveR_pM. apply moveR_pM. hott_simpl.
|
||||||
etransitivity. apply (ap_V (S1op base) loop).
|
refine (ap_V _ _ @ _).
|
||||||
f_ap. apply S1_rec_beta_loop.
|
f_ap. apply S1_rec_beta_loop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
@ -237,8 +240,7 @@ Section set_sphere.
|
||||||
Proof.
|
Proof.
|
||||||
hrecursion z.
|
hrecursion z.
|
||||||
- reflexivity.
|
- reflexivity.
|
||||||
- etransitivity.
|
- refine (transport_paths_FlFr loop _ @ _).
|
||||||
apply (@transport_paths_FlFr _ _ (fun z => S1op x (S1op y z)) (S1op (S1op x y)) _ _ loop idpath).
|
|
||||||
hott_simpl.
|
hott_simpl.
|
||||||
apply moveR_Mp. hott_simpl.
|
apply moveR_Mp. hott_simpl.
|
||||||
rewrite S1_rec_beta_loop.
|
rewrite S1_rec_beta_loop.
|
||||||
|
@ -274,7 +276,7 @@ Section set_sphere.
|
||||||
|
|
||||||
Lemma FSet_S_ap : (nl (@E A)) = (nr E) -> idpath = loop.
|
Lemma FSet_S_ap : (nl (@E A)) = (nr E) -> idpath = loop.
|
||||||
Proof.
|
Proof.
|
||||||
intros H.
|
intros H1.
|
||||||
enough (ap FSet_to_S (nl E) = ap FSet_to_S (nr E)) as H'.
|
enough (ap FSet_to_S (nl E) = ap FSet_to_S (nr E)) as H'.
|
||||||
- rewrite FSet_rec_beta_nl in H'.
|
- rewrite FSet_rec_beta_nl in H'.
|
||||||
rewrite FSet_rec_beta_nr in H'.
|
rewrite FSet_rec_beta_nr in H'.
|
||||||
|
@ -285,7 +287,7 @@ Section set_sphere.
|
||||||
|
|
||||||
Lemma FSet_not_hset : IsHSet (FSet A) -> False.
|
Lemma FSet_not_hset : IsHSet (FSet A) -> False.
|
||||||
Proof.
|
Proof.
|
||||||
intros H.
|
intros H1.
|
||||||
enough (idpath = loop).
|
enough (idpath = loop).
|
||||||
- assert (S1_encode _ idpath = S1_encode _ (loopexp loop (pos Int.one))) as H' by f_ap.
|
- assert (S1_encode _ idpath = S1_encode _ (loopexp loop (pos Int.one))) as H' by f_ap.
|
||||||
rewrite S1_encode_loopexp in H'. simpl in H'. symmetry in H'.
|
rewrite S1_encode_loopexp in H'. simpl in H'. symmetry in H'.
|
||||||
|
|
Loading…
Reference in New Issue