mirror of https://github.com/nmvdw/HITs-Examples
Small improvements
This commit is contained in:
parent
30004e1c8b
commit
76fe6faff2
|
@ -139,35 +139,29 @@ Section properties.
|
||||||
|
|
||||||
Context {B : Type}.
|
Context {B : Type}.
|
||||||
|
|
||||||
Lemma isIn_singleproduct : forall (a : A) (b : B) (c : A) (Y : FSet B),
|
Lemma isIn_singleproduct (a : A) (b : B) (c : A) : forall (Y : FSet B),
|
||||||
isIn (a, b) (single_product c Y) = land (BuildhProp (Trunc (-1) (a = c))) (isIn b Y).
|
isIn (a, b) (single_product c Y) = land (BuildhProp (Trunc (-1) (a = c))) (isIn b Y).
|
||||||
Proof.
|
Proof.
|
||||||
intros a b c.
|
|
||||||
hinduction ; try (intros ; apply path_ishprop).
|
hinduction ; try (intros ; apply path_ishprop).
|
||||||
- apply path_hprop. symmetry. apply prod_empty_r.
|
- apply path_hprop ; symmetry ; apply prod_empty_r.
|
||||||
- intros d.
|
- intros d.
|
||||||
apply path_iff_hprop.
|
apply path_iff_hprop.
|
||||||
* intros.
|
* intros.
|
||||||
strip_truncations.
|
strip_truncations.
|
||||||
split ; apply tr ; try (apply (ap fst X)) ; try (apply (ap snd X)).
|
split ; apply tr ; try (apply (ap fst X)) ; try (apply (ap snd X)).
|
||||||
* intros [Z1 Z2].
|
* intros [Z1 Z2].
|
||||||
strip_truncations.
|
strip_truncations.
|
||||||
rewrite Z1, Z2.
|
rewrite Z1, Z2.
|
||||||
apply (tr idpath).
|
apply (tr idpath).
|
||||||
- intros X1 X2 HX1 HX2.
|
- intros X1 X2 HX1 HX2.
|
||||||
unfold lor.
|
|
||||||
apply path_iff_hprop.
|
apply path_iff_hprop.
|
||||||
* intros X.
|
* intros X.
|
||||||
strip_truncations.
|
strip_truncations.
|
||||||
destruct X as [H1 | H1].
|
destruct X as [H1 | H1] ; rewrite ?HX1, ?HX2 in H1 ; destruct H1 as [H1 H2].
|
||||||
** rewrite HX1 in H1.
|
** split.
|
||||||
destruct H1 as [H1 H2].
|
|
||||||
split.
|
|
||||||
*** apply H1.
|
*** apply H1.
|
||||||
*** apply (tr(inl H2)).
|
*** apply (tr(inl H2)).
|
||||||
** rewrite HX2 in H1.
|
** split.
|
||||||
destruct H1 as [H1 H2].
|
|
||||||
split.
|
|
||||||
*** apply H1.
|
*** apply H1.
|
||||||
*** apply (tr(inr H2)).
|
*** apply (tr(inr H2)).
|
||||||
* intros [H1 H2].
|
* intros [H1 H2].
|
||||||
|
@ -176,15 +170,14 @@ Section properties.
|
||||||
rewrite HX1, HX2.
|
rewrite HX1, HX2.
|
||||||
destruct H2 as [Hb1 | Hb2].
|
destruct H2 as [Hb1 | Hb2].
|
||||||
** left.
|
** left.
|
||||||
split ; try (apply (tr H1)) ; try (apply Hb1).
|
split ; try (apply (tr H1)) ; try (apply Hb1).
|
||||||
** right.
|
** right.
|
||||||
split ; try (apply (tr H1)) ; try (apply Hb2).
|
split ; try (apply (tr H1)) ; try (apply Hb2).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition isIn_product : forall (a : A) (b : B) (X : FSet A) (Y : FSet B),
|
Definition isIn_product (a : A) (b : B) (X : FSet A) (Y : FSet B) :
|
||||||
isIn (a,b) (product X Y) = land (isIn a X) (isIn b Y).
|
isIn (a,b) (product X Y) = land (isIn a X) (isIn b Y).
|
||||||
Proof.
|
Proof.
|
||||||
intros a b X Y.
|
|
||||||
hinduction X ; try (intros ; apply path_ishprop).
|
hinduction X ; try (intros ; apply path_ishprop).
|
||||||
- apply path_hprop ; symmetry ; apply prod_empty_l.
|
- apply path_hprop ; symmetry ; apply prod_empty_l.
|
||||||
- intros.
|
- intros.
|
||||||
|
@ -194,18 +187,14 @@ Section properties.
|
||||||
apply path_iff_hprop.
|
apply path_iff_hprop.
|
||||||
* intros X.
|
* intros X.
|
||||||
strip_truncations.
|
strip_truncations.
|
||||||
destruct X as [[H3 H4] | [H3 H4]].
|
destruct X as [[H3 H4] | [H3 H4]] ; split ; try (apply H4).
|
||||||
** split.
|
** apply (tr(inl H3)).
|
||||||
*** apply (tr(inl H3)).
|
** apply (tr(inr H3)).
|
||||||
*** apply H4.
|
|
||||||
** split.
|
|
||||||
*** apply (tr(inr H3)).
|
|
||||||
*** apply H4.
|
|
||||||
* intros [H1 H2].
|
* intros [H1 H2].
|
||||||
strip_truncations.
|
strip_truncations.
|
||||||
destruct H1 as [H1 | H1].
|
destruct H1 as [H1 | H1] ; apply tr.
|
||||||
** apply tr ; left ; split ; assumption.
|
** left ; split ; assumption.
|
||||||
** apply tr ; right ; split ; assumption.
|
** right ; split ; assumption.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
(* The proof can be simplified using extensionality *)
|
(* The proof can be simplified using extensionality *)
|
||||||
|
|
Loading…
Reference in New Issue