mirror of https://github.com/nmvdw/HITs-Examples
Shortened T.v
This commit is contained in:
parent
e8560a0c07
commit
40e1f45cfa
|
@ -1,330 +1,68 @@
|
||||||
(* Type which proves that all types have merely decidable equality implies LEM *)
|
(* Type which proves that all types have merely decidable equality implies LEM *)
|
||||||
Require Import HoTT HitTactics Sub.
|
Require Import HoTT HitTactics Sub.
|
||||||
|
|
||||||
Module Export T.
|
Section TR.
|
||||||
Section HIT.
|
|
||||||
Variable A : Type.
|
|
||||||
|
|
||||||
Private Inductive T (B : Type) : Type :=
|
|
||||||
| b : T B
|
|
||||||
| c : T B.
|
|
||||||
|
|
||||||
Axiom p : A -> b A = c A.
|
|
||||||
Axiom setT : IsHSet (T A).
|
|
||||||
|
|
||||||
End HIT.
|
|
||||||
|
|
||||||
Arguments p {_} _.
|
|
||||||
|
|
||||||
Section T_induction.
|
|
||||||
Variable A : Type.
|
|
||||||
Variable (P : (T A) -> Type).
|
|
||||||
Variable (H : forall x, IsHSet (P x)).
|
|
||||||
Variable (bP : P (b A)).
|
|
||||||
Variable (cP : P (c A)).
|
|
||||||
Variable (pP : forall a : A, (p a) # bP = cP).
|
|
||||||
|
|
||||||
(* Induction principle *)
|
|
||||||
Fixpoint T_ind
|
|
||||||
(x : T A)
|
|
||||||
{struct x}
|
|
||||||
: P x
|
|
||||||
:= (match x return _ -> _ -> P x with
|
|
||||||
| b => fun _ _ => bP
|
|
||||||
| c => fun _ _ => cP
|
|
||||||
end) pP H.
|
|
||||||
|
|
||||||
Axiom T_ind_beta_p : forall (a : A),
|
|
||||||
apD T_ind (p a) = pP a.
|
|
||||||
End T_induction.
|
|
||||||
|
|
||||||
Section T_recursion.
|
|
||||||
|
|
||||||
Variable A : Type.
|
|
||||||
Variable P : Type.
|
|
||||||
Variable H : IsHSet P.
|
|
||||||
Variable bP : P.
|
|
||||||
Variable cP : P.
|
|
||||||
Variable pP : A -> bP = cP.
|
|
||||||
|
|
||||||
Definition T_rec : T A -> P.
|
|
||||||
Proof.
|
|
||||||
simple refine (T_ind A _ _ _ _ _) ; cbn.
|
|
||||||
- apply bP.
|
|
||||||
- apply cP.
|
|
||||||
- intro a.
|
|
||||||
simple refine ((transport_const _ _) @ (pP a)).
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition T_rec_beta_p : forall (a : A),
|
|
||||||
ap T_rec (p a) = pP a.
|
|
||||||
Proof.
|
|
||||||
intros.
|
|
||||||
unfold T_rec.
|
|
||||||
eapply (cancelL (transport_const (p a) _)).
|
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
|
||||||
apply T_ind_beta_p.
|
|
||||||
Defined.
|
|
||||||
End T_recursion.
|
|
||||||
|
|
||||||
Instance T_recursion A : HitRecursion (T A)
|
|
||||||
:= {indTy := _; recTy := _;
|
|
||||||
H_inductor := T_ind A; H_recursor := T_rec A }.
|
|
||||||
|
|
||||||
End T.
|
|
||||||
|
|
||||||
Section merely_dec_lem.
|
|
||||||
Variable A : hProp.
|
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
Variable A : hProp.
|
||||||
|
|
||||||
Definition code_b : T A -> hProp.
|
Definition T := Unit + Unit.
|
||||||
|
|
||||||
|
Definition R (x y : T) : hProp :=
|
||||||
|
match x, y with
|
||||||
|
| inl tt, inl tt => Unit_hp
|
||||||
|
| inl tt, inr tt => A
|
||||||
|
| inr tt, inl tt => A
|
||||||
|
| inr tt, inr tt => Unit_hp
|
||||||
|
end.
|
||||||
|
|
||||||
|
Global Instance R_mere : is_mere_relation _ R.
|
||||||
Proof.
|
Proof.
|
||||||
hrecursion.
|
intros x y ; destruct x ; destruct y ; apply _.
|
||||||
- apply Unit_hp.
|
|
||||||
- apply A.
|
|
||||||
- intro a.
|
|
||||||
apply path_iff_hprop.
|
|
||||||
* apply (fun _ => a).
|
|
||||||
* apply (fun _ => tt).
|
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition code_c : T A -> hProp.
|
Global Instance R_refl : Reflexive R.
|
||||||
Proof.
|
Proof.
|
||||||
hrecursion.
|
intro x ; destruct x as [[ ] | [ ]] ; apply tt.
|
||||||
- apply A.
|
|
||||||
- apply Unit_hp.
|
|
||||||
- intro a.
|
|
||||||
apply path_iff_hprop.
|
|
||||||
* apply (fun _ => tt).
|
|
||||||
* apply (fun _ => a).
|
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition code : T A -> T A -> hProp.
|
Global Instance R_sym : Symmetric R.
|
||||||
Proof.
|
Proof.
|
||||||
simple refine (T_rec _ _ _ _ _ _).
|
repeat (let x := fresh in intro x ; destruct x as [[ ] | [ ]])
|
||||||
- exact code_b.
|
; auto ; apply tt.
|
||||||
- exact code_c.
|
|
||||||
- intro a.
|
|
||||||
apply path_forall.
|
|
||||||
intro z.
|
|
||||||
hinduction z.
|
|
||||||
* apply path_iff_hprop.
|
|
||||||
** apply (fun _ => a).
|
|
||||||
** apply (fun _ => tt).
|
|
||||||
* apply path_iff_hprop.
|
|
||||||
** apply (fun _ => tt).
|
|
||||||
** apply (fun _ => a).
|
|
||||||
* intros. apply set_path2.
|
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Local Ltac f_prop := apply path_forall ; intro ; apply path_ishprop.
|
Global Instance R_trans : Transitive R.
|
||||||
|
|
||||||
Lemma transport_code_b_x (a : A) :
|
|
||||||
transport code_b (p a) = fun _ => a.
|
|
||||||
Proof.
|
Proof.
|
||||||
f_prop.
|
repeat (let x := fresh in intro x ; destruct x as [[ ] | [ ]]) ; intros
|
||||||
|
; auto ; apply tt.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma transport_code_c_x (a : A) :
|
Definition TR : Type := quotient R.
|
||||||
transport code_c (p a) = fun _ => tt.
|
Definition TR_zero : TR := class_of R (inl tt).
|
||||||
Proof.
|
Definition TR_one : TR := class_of R (inr tt).
|
||||||
f_prop.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma transport_code_c_x_V (a : A) :
|
Definition equiv_pathspace_T : (TR_zero = TR_one) = (R (inl tt) (inr tt))
|
||||||
transport code_c (p a)^ = fun _ => a.
|
:= path_universe (sets_exact R (inl tt) (inr tt)).
|
||||||
Proof.
|
|
||||||
f_prop.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma transport_code_x_b (a : A) :
|
Global Instance quotientB_recursion : HitRecursion TR :=
|
||||||
transport (fun x => code x (b A)) (p a) = fun _ => a.
|
|
||||||
Proof.
|
|
||||||
f_prop.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma transport_code_x_c (a : A) :
|
|
||||||
transport (fun x => code x (c A)) (p a) = fun _ => tt.
|
|
||||||
Proof.
|
|
||||||
f_prop.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma transport_code_x_c_V (a : A) :
|
|
||||||
transport (fun x => code x (c A)) (p a)^ = fun _ => a.
|
|
||||||
Proof.
|
|
||||||
f_prop.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma ap_diag {B : Type} {x y : B} (p : x = y) :
|
|
||||||
ap (fun x : B => (x, x)) p = path_prod' p p.
|
|
||||||
Proof.
|
|
||||||
by path_induction.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma transport_code_diag (a : A) z :
|
|
||||||
(transport (fun i : (T A) => code i i) (p a)) z = tt.
|
|
||||||
Proof.
|
|
||||||
apply path_ishprop.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition r : forall (x : T A), code x x.
|
|
||||||
Proof.
|
|
||||||
simple refine (T_ind _ _ _ _ _ _); simpl.
|
|
||||||
- exact tt.
|
|
||||||
- exact tt.
|
|
||||||
- intro a.
|
|
||||||
apply transport_code_diag.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition encode_pre : forall (x y : T A), x = y -> code x y
|
|
||||||
:= fun x y p => transport (fun z => code x z) p (r x).
|
|
||||||
|
|
||||||
Definition encode : forall x y, x = y -> code x y.
|
|
||||||
Proof.
|
|
||||||
intros x y.
|
|
||||||
intro p.
|
|
||||||
refine (transport (fun z => code x z) p _). clear p.
|
|
||||||
revert x. simple refine (T_ind _ _ _ _ _ _); simpl.
|
|
||||||
- exact tt.
|
|
||||||
- exact tt.
|
|
||||||
- intro a.
|
|
||||||
apply path_ishprop.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition decode_b : forall (y : T A), code_b y -> (b A) = y.
|
|
||||||
Proof.
|
|
||||||
simple refine (T_ind _ _ _ _ _ _) ; simpl.
|
|
||||||
- exact (fun _ => idpath).
|
|
||||||
- exact (fun a => p a).
|
|
||||||
- intro a.
|
|
||||||
apply path_forall.
|
|
||||||
intro t.
|
|
||||||
refine (transport_arrow _ _ _ @ _).
|
|
||||||
refine (transport_paths_FlFr _ _ @ _).
|
|
||||||
hott_simpl.
|
|
||||||
f_ap.
|
|
||||||
apply path_ishprop.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition decode_c : forall (y : T A), code_c y -> (c A) = y.
|
|
||||||
Proof.
|
|
||||||
simple refine (T_ind _ _ _ _ _ _); simpl.
|
|
||||||
- exact (fun a => (p a)^).
|
|
||||||
- exact (fun _ => idpath).
|
|
||||||
- intro a.
|
|
||||||
apply path_forall.
|
|
||||||
intro t.
|
|
||||||
refine (transport_arrow _ _ _ @ _).
|
|
||||||
refine (transport_paths_FlFr _ _ @ _).
|
|
||||||
rewrite transport_code_c_x_V.
|
|
||||||
hott_simpl.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma transport_paths_FlFr_trunc :
|
|
||||||
forall {X Y : Type} (f g : X -> Y) {x1 x2 : X} (q : x1 = x2)
|
|
||||||
(r : f x1 = g x1),
|
|
||||||
transport (fun x : X => Trunc 0 (f x = g x)) q (tr r) = tr (((ap f q)^ @ r) @ ap g q).
|
|
||||||
Proof.
|
|
||||||
destruct q; simpl. intro r.
|
|
||||||
refine (ap tr _).
|
|
||||||
exact ((concat_1p r)^ @ (concat_p1 (1 @ r))^).
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition decode : forall (x y : T A), code x y -> x = y.
|
|
||||||
Proof.
|
|
||||||
simple refine (T_ind _ _ _ _ _ _); simpl.
|
|
||||||
- intro y. exact (decode_b y).
|
|
||||||
- intro y. exact (decode_c y).
|
|
||||||
- intro a.
|
|
||||||
apply path_forall. intro z.
|
|
||||||
rewrite transport_forall_constant.
|
|
||||||
apply path_forall. intros c.
|
|
||||||
rewrite transport_arrow.
|
|
||||||
hott_simpl.
|
|
||||||
rewrite (transport_paths_FlFr _ _).
|
|
||||||
revert z c. simple refine (T_ind _ _ _ _ _ _) ; simpl.
|
|
||||||
+ intro.
|
|
||||||
hott_simpl.
|
|
||||||
f_ap.
|
|
||||||
refine (ap (fun x => (p x)) _).
|
|
||||||
apply path_ishprop.
|
|
||||||
+ intro.
|
|
||||||
rewrite transport_code_x_c_V.
|
|
||||||
hott_simpl.
|
|
||||||
+ intro b.
|
|
||||||
apply path_forall.
|
|
||||||
intro z.
|
|
||||||
rewrite transport_forall.
|
|
||||||
apply set_path2.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma decode_encode (u v : T A) : forall (p : u = v),
|
|
||||||
decode u v (encode u v p) = p.
|
|
||||||
Proof.
|
|
||||||
intros p. induction p.
|
|
||||||
simpl. revert u. simple refine (T_ind _ _ _ _ _ _).
|
|
||||||
+ simpl. reflexivity.
|
|
||||||
+ simpl. reflexivity.
|
|
||||||
+ intro a.
|
|
||||||
apply set_path2.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma encode_decode : forall (u v : T A) (c : code u v),
|
|
||||||
encode u v (decode u v c) = c.
|
|
||||||
Proof.
|
|
||||||
simple refine (T_ind _ _ _ _ _ _).
|
|
||||||
- simple refine (T_ind _ _ _ _ _ _).
|
|
||||||
+ simpl. apply path_ishprop.
|
|
||||||
+ simpl. intro a. apply path_ishprop.
|
|
||||||
+ intro a. apply path_forall; intros ?. apply set_path2.
|
|
||||||
- simple refine (T_ind _ _ _ _ _ _).
|
|
||||||
+ simpl. intro a. apply path_ishprop.
|
|
||||||
+ simpl. apply path_ishprop.
|
|
||||||
+ intro a. apply path_forall; intros ?. apply set_path2.
|
|
||||||
- intro a. repeat (apply path_forall; intros ?). apply set_path2.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
|
|
||||||
Instance T_hprop (a : A) : IsHProp (b A = c A).
|
|
||||||
Proof.
|
|
||||||
apply hprop_allpath.
|
|
||||||
intros x y.
|
|
||||||
pose (encode (b A) _ x) as t1.
|
|
||||||
pose (encode (b A) _ y) as t2.
|
|
||||||
assert (t1 = t2).
|
|
||||||
{
|
{
|
||||||
unfold t1, t2.
|
indTy := _;
|
||||||
apply path_ishprop.
|
recTy :=
|
||||||
}
|
forall (P : Type) (HP: IsHSet P) (u : T -> P),
|
||||||
pose (decode _ _ t1) as t3.
|
(forall x y : T, R x y -> u x = u y) -> TR -> P;
|
||||||
pose (decode _ _ t2) as t4.
|
H_inductor := quotient_ind R ;
|
||||||
assert (t3 = t4) as H1.
|
H_recursor := @quotient_rec _ R _
|
||||||
{
|
}.
|
||||||
unfold t3, t4.
|
|
||||||
f_ap.
|
|
||||||
}
|
|
||||||
unfold t3, t4, t1, t2 in H1.
|
|
||||||
rewrite ?decode_encode in H1.
|
|
||||||
apply H1.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Lemma equiv_pathspace_T : (b A = c A) = A.
|
End TR.
|
||||||
Proof.
|
|
||||||
apply path_iff_ishprop.
|
|
||||||
- intro x.
|
|
||||||
apply (encode (b A) (c A) x).
|
|
||||||
- apply p.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
End merely_dec_lem.
|
|
||||||
|
|
||||||
Theorem merely_dec `{Univalence} : (forall (A : Type) (a b : A), hor (a = b) (~a = b))
|
Theorem merely_dec `{Univalence} : (forall (A : Type) (a b : A), hor (a = b) (~a = b))
|
||||||
->
|
->
|
||||||
forall (A : hProp), A + (~A).
|
forall (A : hProp), A + (~A).
|
||||||
Proof.
|
Proof.
|
||||||
intros.
|
intros X A.
|
||||||
specialize (X (T A) (b A) (c A)).
|
specialize (X (TR A) (TR_zero A) (TR_one A)).
|
||||||
rewrite (equiv_pathspace_T A) in X.
|
rewrite equiv_pathspace_T in X.
|
||||||
strip_truncations.
|
strip_truncations.
|
||||||
apply X.
|
apply X.
|
||||||
Defined.
|
Defined.
|
|
@ -74,20 +74,6 @@ Module Export FSet.
|
||||||
| 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),
|
|
||||||
apD FSet_ind (assoc x y z) =
|
|
||||||
(assocP x y z (FSet_ind x) (FSet_ind y) (FSet_ind z)).
|
|
||||||
|
|
||||||
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)).
|
|
||||||
|
|
||||||
Axiom FSet_ind_beta_nl : forall (x : FSet A),
|
|
||||||
apD FSet_ind (nl x) = (nlP x (FSet_ind x)).
|
|
||||||
|
|
||||||
Axiom FSet_ind_beta_nr : forall (x : FSet A),
|
|
||||||
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.
|
|
||||||
End FSet_induction.
|
End FSet_induction.
|
||||||
|
|
||||||
Section FSet_recursion.
|
Section FSet_recursion.
|
||||||
|
@ -118,66 +104,6 @@ Module Export FSet.
|
||||||
- apply idemP.
|
- apply idemP.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition FSet_rec_beta_assoc : forall (x y z : FSet A),
|
|
||||||
ap FSet_rec (assoc x y z)
|
|
||||||
=
|
|
||||||
assocP (FSet_rec x) (FSet_rec y) (FSet_rec z).
|
|
||||||
Proof.
|
|
||||||
intros.
|
|
||||||
unfold FSet_rec.
|
|
||||||
eapply (cancelL (transport_const (assoc x y z) _)).
|
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
|
||||||
apply FSet_ind_beta_assoc.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition FSet_rec_beta_comm : forall (x y : FSet A),
|
|
||||||
ap FSet_rec (comm x y)
|
|
||||||
=
|
|
||||||
commP (FSet_rec x) (FSet_rec y).
|
|
||||||
Proof.
|
|
||||||
intros.
|
|
||||||
unfold FSet_rec.
|
|
||||||
eapply (cancelL (transport_const (comm x y) _)).
|
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
|
||||||
apply FSet_ind_beta_comm.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition FSet_rec_beta_nl : forall (x : FSet A),
|
|
||||||
ap FSet_rec (nl x)
|
|
||||||
=
|
|
||||||
nlP (FSet_rec x).
|
|
||||||
Proof.
|
|
||||||
intros.
|
|
||||||
unfold FSet_rec.
|
|
||||||
eapply (cancelL (transport_const (nl x) _)).
|
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
|
||||||
apply FSet_ind_beta_nl.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition FSet_rec_beta_nr : forall (x : FSet A),
|
|
||||||
ap FSet_rec (nr x)
|
|
||||||
=
|
|
||||||
nrP (FSet_rec x).
|
|
||||||
Proof.
|
|
||||||
intros.
|
|
||||||
unfold FSet_rec.
|
|
||||||
eapply (cancelL (transport_const (nr x) _)).
|
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
|
||||||
apply FSet_ind_beta_nr.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Definition FSet_rec_beta_idem : forall (a : A),
|
|
||||||
ap FSet_rec (idem a)
|
|
||||||
=
|
|
||||||
idemP a.
|
|
||||||
Proof.
|
|
||||||
intros.
|
|
||||||
unfold FSet_rec.
|
|
||||||
eapply (cancelL (transport_const (idem a) _)).
|
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
|
||||||
apply FSet_ind_beta_idem.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
End FSet_recursion.
|
End FSet_recursion.
|
||||||
|
|
||||||
Instance FSet_recursion A : HitRecursion (FSet A) :=
|
Instance FSet_recursion A : HitRecursion (FSet A) :=
|
||||||
|
|
|
@ -87,35 +87,41 @@ End k_fin_lem_projective.
|
||||||
Section k_fin_projective_lem.
|
Section k_fin_projective_lem.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
Variable (P : Type).
|
Variable (P : Type).
|
||||||
Context `{Hprop : IsHProp P}.
|
Context `{IsHProp P}.
|
||||||
|
|
||||||
Definition X : Type := T P.
|
Definition X : Type := TR (BuildhProp P).
|
||||||
Instance X_set : IsHSet X.
|
Instance X_set : IsHSet X.
|
||||||
Proof. apply _. Defined.
|
Proof.
|
||||||
|
apply _.
|
||||||
|
Defined.
|
||||||
|
|
||||||
Definition X_fin : Kf X.
|
Definition X_fin : Kf X.
|
||||||
Proof.
|
Proof.
|
||||||
apply Kf_unfold.
|
apply Kf_unfold.
|
||||||
exists ({|b P|} ∪ {|c P|}).
|
exists ({|TR_zero _|} ∪ {|TR_one _|}).
|
||||||
hinduction.
|
hinduction.
|
||||||
- apply (tr (inl (tr idpath))).
|
- destruct x as [ [ ] | [ ] ].
|
||||||
- apply (tr (inr (tr idpath))).
|
* apply (tr (inl (tr idpath))).
|
||||||
- intros. apply path_ishprop.
|
* apply (tr (inr (tr idpath))).
|
||||||
|
- intros.
|
||||||
|
apply path_ishprop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition p (a : Unit + Unit) : X :=
|
Definition p (a : Unit + Unit) : X :=
|
||||||
match a with
|
match a with
|
||||||
| inl _ => b P
|
| inl _ => TR_zero _
|
||||||
| inr _ => c P
|
| inr _ => TR_one _
|
||||||
end.
|
end.
|
||||||
|
|
||||||
Instance p_surj : IsSurjection p.
|
Instance p_surj : IsSurjection p.
|
||||||
Proof.
|
Proof.
|
||||||
apply BuildIsSurjection.
|
apply BuildIsSurjection.
|
||||||
hinduction.
|
hinduction.
|
||||||
- apply tr. exists (inl tt). reflexivity.
|
- destruct x as [[ ] | [ ]].
|
||||||
- apply tr. exists (inr tt). reflexivity.
|
* apply tr. exists (inl tt). reflexivity.
|
||||||
- intros. apply path_ishprop.
|
* apply tr. exists (inr tt). reflexivity.
|
||||||
|
- intros.
|
||||||
|
apply path_ishprop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma LEM `{IsProjective X} : P + ~P.
|
Lemma LEM `{IsProjective X} : P + ~P.
|
||||||
|
@ -123,16 +129,17 @@ Section k_fin_projective_lem.
|
||||||
pose (k := projective p idmap _).
|
pose (k := projective p idmap _).
|
||||||
unfold hexists in k.
|
unfold hexists in k.
|
||||||
simple refine (Trunc_rec _ k); clear k; intros [g Hg].
|
simple refine (Trunc_rec _ k); clear k; intros [g Hg].
|
||||||
destruct (dec (g (b P) = g (c P))) as [Hℵ | Hℵ].
|
destruct (dec (g (TR_zero _) = g (TR_one _))) as [Hℵ | Hℵ].
|
||||||
- left.
|
- left.
|
||||||
assert (b P = c P) as Hbc.
|
assert (TR_zero (BuildhProp P) = TR_one _) as Hbc.
|
||||||
{ pose (ap p Hℵ) as Hα.
|
{ pose (ap p Hℵ) as Hα.
|
||||||
rewrite (ap (fun h => h (b P)) Hg) in Hα.
|
rewrite (ap (fun h => h (TR_zero _)) Hg) in Hα.
|
||||||
rewrite (ap (fun h => h (c P)) Hg) in Hα.
|
rewrite (ap (fun h => h (TR_one _)) Hg) in Hα.
|
||||||
assumption. }
|
assumption. }
|
||||||
apply (encode _ (b P) (c P) Hbc).
|
refine (classes_eq_related _ _ _ Hbc).
|
||||||
- right. intros HP.
|
- right. intros HP.
|
||||||
apply Hℵ.
|
apply Hℵ.
|
||||||
apply (ap g (T.p HP)).
|
refine (ap g (related_classes_eq _ _)).
|
||||||
|
apply HP.
|
||||||
Defined.
|
Defined.
|
||||||
End k_fin_projective_lem.
|
End k_fin_projective_lem.
|
||||||
|
|
Loading…
Reference in New Issue