mirror of https://github.com/nmvdw/HITs-Examples
Merge branch 'master' of https://github.com/nmvdw/HITs-Examples
This commit is contained in:
commit
35f0452a6a
|
@ -144,13 +144,27 @@ Section k_properties.
|
||||||
apply (tr (idpath)).
|
apply (tr (idpath)).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
Lemma Kf_surjection {X Y : Type} (f : X -> Y) `{IsSurjection f} :
|
||||||
|
Kf X -> Kf Y.
|
||||||
|
Proof.
|
||||||
|
intros HX. apply Kf_unfold. apply Kf_unfold in HX.
|
||||||
|
destruct HX as [Xf HXf].
|
||||||
|
exists (fmap FSet f Xf).
|
||||||
|
intro y.
|
||||||
|
pose (x' := center (merely (hfiber f y))).
|
||||||
|
simple refine (@Trunc_rec (-1) (hfiber f y) _ _ _ x'). clear x'; intro x.
|
||||||
|
destruct x as [x Hfx]. rewrite <- Hfx.
|
||||||
|
apply fmap_isIn.
|
||||||
|
apply (HXf x).
|
||||||
|
Defined.
|
||||||
|
|
||||||
Lemma Kf_sum {A B : Type} : Kf A -> Kf B -> Kf (A + B).
|
Lemma Kf_sum {A B : Type} : Kf A -> Kf B -> Kf (A + B).
|
||||||
Proof.
|
Proof.
|
||||||
intros HA HB.
|
intros HA HB.
|
||||||
kf_unfold.
|
kf_unfold.
|
||||||
destruct HA as [X HX].
|
destruct HA as [X HX].
|
||||||
destruct HB as [Y HY].
|
destruct HB as [Y HY].
|
||||||
exists ((fset_fmap inl X) ∪ (fset_fmap inr Y)).
|
exists (disjoint_sum X Y).
|
||||||
intros [a | b]; simpl; apply tr; [ left | right ];
|
intros [a | b]; simpl; apply tr; [ left | right ];
|
||||||
apply fmap_isIn.
|
apply fmap_isIn.
|
||||||
+ apply (HX a).
|
+ apply (HX a).
|
||||||
|
@ -194,20 +208,6 @@ Section k_properties.
|
||||||
- apply (HB b).
|
- apply (HB b).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma Kf_surjection {X Y : Type} (f : X -> Y) `{IsSurjection f} :
|
|
||||||
Kf X -> Kf Y.
|
|
||||||
Proof.
|
|
||||||
intros HX. apply Kf_unfold. apply Kf_unfold in HX.
|
|
||||||
destruct HX as [Xf HXf].
|
|
||||||
exists (fmap FSet f Xf).
|
|
||||||
intro y.
|
|
||||||
pose (x' := center (merely (hfiber f y))).
|
|
||||||
simple refine (@Trunc_rec (-1) (hfiber f y) _ _ _ x'). clear x'; intro x.
|
|
||||||
destruct x as [x Hfx]. rewrite <- Hfx.
|
|
||||||
apply fmap_isIn.
|
|
||||||
apply (HXf x).
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma S1_Kfinite : Kf S1.
|
Lemma S1_Kfinite : Kf S1.
|
||||||
Proof.
|
Proof.
|
||||||
apply Kf_unfold.
|
apply Kf_unfold.
|
||||||
|
|
Loading…
Reference in New Issue