Further cleaning

This commit is contained in:
Niels 2017-09-01 17:08:00 +02:00
parent dfd590724b
commit 4ab70ae1fe
3 changed files with 10 additions and 18 deletions

View File

@ -1,5 +1,5 @@
Require Import HoTT. Require Import HoTT.
Require Import disjunction lattice notation. Require Import disjunction lattice notation plumbing.
Section subobjects. Section subobjects.
Variable A : Type. Variable A : Type.
@ -57,17 +57,7 @@ End isIn.
Section intersect. Section intersect.
Variable A : Type. Variable A : Type.
Variable C : (Sub A) -> hProp. Variable C : (Sub A) -> hProp.
Context `{Univalence}. Context `{Univalence}
Global Instance hprop_lem : forall (T : Type) (Ttrunc : IsHProp T), IsHProp (T + ~T).
Proof.
intros.
apply (equiv_hprop_allpath _)^-1.
intros [x | nx] [y | ny] ; try f_ap ; try (apply Ttrunc) ; try contradiction.
- apply equiv_hprop_allpath. apply _.
Defined.
Context
{HI : closedIntersection C} {HE : closedEmpty C} {HI : closedIntersection C} {HE : closedEmpty C}
{HS : closedSingleton C} {HDE : hasDecidableEmpty C}. {HS : closedSingleton C} {HDE : hasDecidableEmpty C}.
@ -85,4 +75,4 @@ Section intersect.
strip_truncations. strip_truncations.
apply (inl (tr (t2^ @ t1))). apply (inl (tr (t2^ @ t1))).
Defined. Defined.
End intersect. End intersect.

View File

@ -14,3 +14,10 @@ Defined.
Lemma ap_equiv {A B} (f : A <~> B) {x y : A} (p : x = y) : Lemma ap_equiv {A B} (f : A <~> B) {x y : A} (p : x = y) :
ap (f^-1 o f) p = eissect f x @ p @ (eissect f y)^. ap (f^-1 o f) p = eissect f x @ p @ (eissect f y)^.
Proof. destruct p. hott_simpl. Defined. Proof. destruct p. hott_simpl. Defined.
Global Instance hprop_lem `{Univalence} (T : Type) (Ttrunc : IsHProp T) : IsHProp (T + ~T).
Proof.
apply (equiv_hprop_allpath _)^-1.
intros [x | nx] [y | ny] ; try f_ap ; try (apply Ttrunc) ; try contradiction.
- apply equiv_hprop_allpath. apply _.
Defined.

View File

@ -15,11 +15,6 @@ Section TR.
| inr tt, inr tt => Unit_hp | inr tt, inr tt => Unit_hp
end. end.
Global Instance R_mere : is_mere_relation _ R.
Proof.
intros x y ; destruct x ; destruct y ; apply _.
Defined.
Global Instance R_refl : Reflexive R. Global Instance R_refl : Reflexive R.
Proof. Proof.
intro x ; destruct x as [[ ] | [ ]] ; apply tt. intro x ; destruct x as [[ ] | [ ]] ; apply tt.