HITs-Examples/FiniteSets/representations/T.v

63 lines
1.7 KiB
Coq
Raw Normal View History

2017-08-03 23:01:57 +02:00
(* Type which proves that all types have merely decidable equality implies LEM *)
Require Import HoTT HitTactics Sub.
2017-09-01 16:29:48 +02:00
Section TR.
2017-08-03 23:01:57 +02:00
Context `{Univalence}.
2017-09-01 16:29:48 +02:00
Variable A : hProp.
2017-08-03 23:01:57 +02:00
2017-09-01 16:29:48 +02:00
Definition T := Unit + Unit.
2017-08-03 23:01:57 +02:00
2017-09-01 16:29:48 +02:00
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.
2017-08-03 23:01:57 +02:00
2017-09-01 16:29:48 +02:00
Global Instance R_refl : Reflexive R.
2017-08-03 23:01:57 +02:00
Proof.
2017-09-01 16:29:48 +02:00
intro x ; destruct x as [[ ] | [ ]] ; apply tt.
2017-08-03 23:01:57 +02:00
Defined.
2017-09-01 16:29:48 +02:00
Global Instance R_sym : Symmetric R.
2017-08-03 23:01:57 +02:00
Proof.
2017-09-01 16:29:48 +02:00
repeat (let x := fresh in intro x ; destruct x as [[ ] | [ ]])
; auto ; apply tt.
2017-08-03 23:01:57 +02:00
Defined.
2017-09-01 16:29:48 +02:00
Global Instance R_trans : Transitive R.
2017-08-03 23:01:57 +02:00
Proof.
2017-09-01 16:29:48 +02:00
repeat (let x := fresh in intro x ; destruct x as [[ ] | [ ]]) ; intros
; auto ; apply tt.
2017-08-03 23:01:57 +02:00
Defined.
2017-09-01 16:29:48 +02:00
Definition TR : Type := quotient R.
Definition TR_zero : TR := class_of R (inl tt).
Definition TR_one : TR := class_of R (inr tt).
2017-09-01 16:29:48 +02:00
Definition equiv_pathspace_T : (TR_zero = TR_one) = (R (inl tt) (inr tt))
:= path_universe (sets_exact R (inl tt) (inr tt)).
2017-08-03 23:01:57 +02:00
2017-09-01 16:29:48 +02:00
Global Instance quotientB_recursion : HitRecursion TR :=
2017-08-03 23:01:57 +02:00
{
2017-09-01 16:29:48 +02:00
indTy := _;
recTy :=
forall (P : Type) (HP: IsHSet P) (u : T -> P),
(forall x y : T, R x y -> u x = u y) -> TR -> P;
H_inductor := quotient_ind R ;
H_recursor := @quotient_rec _ R _
}.
2017-08-03 23:01:57 +02:00
2017-09-01 16:29:48 +02:00
End TR.
2017-08-03 23:01:57 +02:00
Theorem merely_dec `{Univalence} : (forall (A : Type) (a b : A), hor (a = b) (~a = b))
->
forall (A : hProp), A + (~A).
Proof.
2017-09-01 16:29:48 +02:00
intros X A.
specialize (X (TR A) (TR_zero A) (TR_one A)).
rewrite equiv_pathspace_T in X.
2017-08-03 23:01:57 +02:00
strip_truncations.
apply X.
2017-09-01 16:29:48 +02:00
Defined.