A simpler `split` fn for B-finite subobjects.

Allows for shortening of some proofs
This commit is contained in:
Dan Frumin 2017-08-23 22:23:28 +02:00
parent 8a1405a1d8
commit e1a8220ba0
1 changed files with 75 additions and 119 deletions

View File

@ -151,95 +151,65 @@ Section finite_hott.
End empty. End empty.
Section split. Section split.
Variable (X : A -> hProp) Variable (P : A -> hProp)
(n : nat) (n : nat)
(Xequiv : {a : A & a X} <~> Fin n + Unit). (Xequiv : {a : A & P a } <~> Fin n + Unit).
Definition split : {X' : A -> hProp & {a : A & a X'} <~> Fin n}. Definition split : exists P' : Sub A, exists b : A,
({a : A & P' a} <~> Fin n) * (forall x, P x = P' x merely (x = b)).
Proof. Proof.
destruct Xequiv as [f [g fg gf adj]]. destruct Xequiv as [f [g fg gf adj]].
unfold Sect in *. unfold Sect in *.
pose (fun x : A => sig (fun y : Fin n => x = (g(inl y)).1 )) as X'. pose (fun x : A => sig (fun y : Fin n => x = (g (inl y)).1)) as P'.
assert (forall a : A, IsHProp (X' a)). assert (forall x, IsHProp (P' x)).
{ {
intros. intros a. unfold P'.
unfold X'.
apply hprop_allpath. apply hprop_allpath.
intros [x px] [y py]. intros [x px] [y py].
simple refine (path_sigma _ _ _ _ _). simple refine (path_sigma _ _ _ _ _); [ simpl | apply path_ishprop ].
* cbn. apply path_sum_inl with Unit.
pose (f(g(inl x))) as fgx. cut (g (inl x) = g (inl y)).
cut (g(inl x) = g(inl y)). { intros p.
{ pose (ap f p) as Hp.
intros q. by rewrite !fg in Hp. }
pose (ap f q). apply path_sigma with (px^ @ py).
rewrite !fg in p. apply path_ishprop.
refine (path_sum_inl _ p).
}
apply path_sigma with (px^ @ py).
apply path_ishprop.
* apply path_ishprop.
} }
pose (fun a => BuildhProp(X' a)) as Y. exists (fun a => BuildhProp (P' a)).
exists Y. exists (g (inr tt)).1.
unfold Y, X'. split.
cbn. { unshelve eapply BuildEquiv.
unshelve esplit. { refine (fun x => x.2.1). }
- intros [a [y p]]. apply y. apply isequiv_biinv.
- apply isequiv_biinv. unshelve esplit;
unshelve esplit. exists (fun x => (((g (inl x)).1; (x; idpath)))).
* exists (fun x => (( (g(inl x)).1 ;(x;idpath)))). - intros [a [y p]]; cbn.
unfold Sect. eapply path_sigma with p^.
intros [a [y p]]. apply path_ishprop.
apply path_sigma with p^. - intros x; cbn.
simpl. reflexivity. }
rewrite transport_sigma. { intros a.
simple refine (path_sigma _ _ _ _ _) ; simpl. unfold P'.
** rewrite transport_const ; reflexivity. apply path_iff_hprop.
** apply path_ishprop. - intros Ha.
* exists (fun x => (( (g(inl x)).1 ;(x;idpath)))). pose (y := f (a;Ha)).
unfold Sect. assert (Hy : y = f (a; Ha)) by reflexivity.
intros x. destruct y as [y | []].
reflexivity. + refine (tr (inl _)).
Defined. exists y.
rewrite Hy.
Definition new_el : {a' : A & forall z, X z = by rewrite gf.
lor (split.1 z) (merely (z = a'))}. + refine (tr (inr (tr _))).
Proof. rewrite Hy.
exists ((Xequiv^-1 (inr tt)).1). by rewrite gf.
intros. - intros Hstuff.
apply path_iff_hprop. strip_truncations.
- intros Xz. destruct Hstuff as [[y Hy] | Ha].
pose (Xequiv (z;Xz)) as fz. + rewrite Hy.
pose (c := fz). apply (g (inl y)).2.
assert (c = fz). reflexivity. + strip_truncations.
destruct c as [fz1 | fz2]. rewrite Ha.
* refine (tr(inl _)). apply (g (inr tt)).2. }
unfold split ; simpl.
exists fz1.
rewrite X0.
unfold fz.
destruct Xequiv as [? [? ? sect ?]].
compute.
rewrite sect.
reflexivity.
* refine (tr(inr(tr _))).
destruct fz2.
rewrite X0.
unfold fz.
rewrite eissect.
reflexivity.
- intros X0.
strip_truncations.
destruct X0 as [Xl | Xr].
* unfold split in Xl ; simpl in Xl.
destruct Xequiv as [f [g fg gf adj]].
destruct Xl as [m p].
rewrite p.
apply (g (inl m)).2.
* strip_truncations.
rewrite Xr.
apply ((Xequiv^-1(inr tt)).2).
Defined. Defined.
End split. End split.
@ -263,13 +233,11 @@ Section dec_membership.
rewrite p. rewrite p.
apply _. apply _.
- intros. - intros.
pose (new_el _ P n Hequiv) as b. destruct (split _ P n Hequiv) as
destruct b as [b HX']. (P' & b & HP' & HP).
destruct (split _ P n Hequiv) as [X' X'equiv]. simpl in HX'.
unfold member, sub_membership. unfold member, sub_membership.
rewrite (HX' a). rewrite (HP a).
pose (IHn X' X'equiv) as IH. destruct (IHn P' HP') as [IH | IH].
destruct IH as [IH | IH].
+ left. apply (tr (inl IH)). + left. apply (tr (inl IH)).
+ destruct (dec (a = b)) as [Hab | Hab]. + destruct (dec (a = b)) as [Hab | Hab].
left. apply (tr (inr (tr Hab))). left. apply (tr (inr (tr Hab))).
@ -364,50 +332,38 @@ Section cowd.
assert ((X; Build_Finite _ 0 (tr FX)) = empty_cow) as HX. assert ((X; Build_Finite _ 0 (tr FX)) = empty_cow) as HX.
{ apply path_sigma' with HX_emp. apply path_ishprop. } { apply path_sigma' with HX_emp. apply path_ishprop. }
rewrite HX. assumption. rewrite HX. assumption.
- pose (a' := new_el _ X n FX). - destruct (split _ X n FX) as
destruct a' as [a' Ha']. (X' & b & FX' & HX).
destruct (split _ X n FX) as [X' FX'].
pose (X'cow := (X'; Build_Finite _ n (tr FX')) : cow). pose (X'cow := (X'; Build_Finite _ n (tr FX')) : cow).
assert ((X; Build_Finite _ (n.+1) (tr FX)) = add_cow a' X'cow) as . assert ((X; Build_Finite _ (n.+1) (tr FX)) = add_cow b X'cow) as .
{ simple refine (path_sigma' _ _ _); [ | apply path_ishprop ]. { simple refine (path_sigma' _ _ _); [ | apply path_ishprop ].
apply path_forall. intros a. apply path_forall. intros a.
unfold X'cow. unfold X'cow.
specialize (Ha' a). rewrite Ha'. simpl. reflexivity. } rewrite (HX a). simpl. reflexivity. }
rewrite . rewrite .
apply koeientaart. apply koeientaart.
apply IHn. apply IHn.
Defined. Defined.
Definition bfin_to_kfin : forall (X : Sub A), Bfin X -> Kf_sub _ X. Definition bfin_to_kfin : forall (P : Sub A), Bfin P -> Kf_sub _ P.
Proof. Proof.
intros X BFinX. intros P [n f].
unfold Bfin in BFinX.
destruct BFinX as [n iso].
strip_truncations. strip_truncations.
revert iso ; revert X. unfold Kf_sub, Kf_sub_intern.
induction n ; unfold Kf_sub, Kf_sub_intern. revert f. revert P.
- intros. induction n; intros P f.
exists . - exists .
apply path_forall. apply path_forall; intro a; simpl.
intro z. apply path_iff_hprop; [ | contradiction ].
simpl in *. intros p.
apply path_iff_hprop ; try contradiction. apply (f (a;p)).
destruct iso as [f f_equiv]. - destruct (split _ P n f) as
apply (fun Xz => f(z;Xz)). (P' & b & HP' & HP).
- intros. destruct (IHn P' HP') as [Y HY].
simpl in *. exists (Y {|b|}).
destruct (new_el _ X n iso) as [a HXX']. apply path_forall; intro a. simpl.
destruct (split _ X n iso) as [X' X'equiv]. rewrite <- HY.
destruct (IHn X' X'equiv) as [Y HY]. apply HP.
exists (Y {|a|}).
unfold map in *.
apply path_forall.
intro z.
rewrite union_isIn.
rewrite <- (ap (fun h => h z) HY).
rewrite HXX'.
cbn.
reflexivity.
Defined. Defined.
Lemma kfin_is_bfin : @closedUnion A Bfin. Lemma kfin_is_bfin : @closedUnion A Bfin.