Further simplifications in interface

This commit is contained in:
Niels 2017-08-15 22:05:31 +02:00
parent e29e978218
commit e1bc900abc
1 changed files with 15 additions and 23 deletions

View File

@ -33,14 +33,10 @@ Section interface.
* refine (tr({|a|};_)).
apply f_singleton.
* intros ; apply path_ishprop.
- intros Y1 Y2 HY1 HY2.
destruct HY1 as [X1' HX1].
destruct HY2 as [X2' HX2].
- intros Y1 Y2 [X1' ?] [X2' ?].
simple refine (BuildContr _ _ _).
* simple refine (Trunc_rec _ X1') ; intro X1.
simple refine (Trunc_rec _ X2') ; intro X2.
destruct X1 as [X1 fX1].
destruct X2 as [X2 fX2].
* simple refine (Trunc_rec _ X1') ; intros [X1 fX1].
simple refine (Trunc_rec _ X2') ; intros [X2 fX2].
refine (tr(X1 X2;_)).
rewrite f_union, fX1, fX2.
reflexivity.
@ -97,7 +93,6 @@ Section quotient_surjection.
Proof.
apply isembedding_isinj_hset.
unfold isinj.
simpl.
hrecursion ; [ | intros; apply path_ishprop ].
intro.
hrecursion ; [ | intros; apply path_ishprop ].
@ -110,10 +105,8 @@ Section quotient_surjection.
apply BuildIsSurjection.
intros Y.
destruct (H Y).
simple refine (Trunc_rec _ center) ; intro X.
apply tr.
destruct X as [a fa].
apply (class_of _ a;fa).
simple refine (Trunc_rec _ center) ; intros [a fa].
apply (tr(class_of _ a;fa)).
Defined.
Definition quotient_iso : quotientB <~> B.
@ -210,8 +203,7 @@ Section quotient_properties.
- apply (member a).
- intros X Y HXY.
reduce T.
rewrite HXY.
reflexivity.
apply (ap _ HXY).
Defined.
Global Instance View_empty A: hasEmpty (View A).