B-fin => K-fin if the underlying type has decidable paths

This commit is contained in:
Dan Frumin 2017-08-16 15:59:36 +02:00
parent 920fdd91ab
commit 57a4535f87
1 changed files with 243 additions and 110 deletions

View File

@ -4,7 +4,7 @@ Require Import Sub notation variations.k_finite.
Require Import fsets.properties. Require Import fsets.properties.
Section finite_hott. Section finite_hott.
Variable A : Type. Variable (A : Type).
Context `{Univalence} `{IsHSet A}. Context `{Univalence} `{IsHSet A}.
(* A subobject is B-finite if its extension is B-finite as a type *) (* A subobject is B-finite if its extension is B-finite as a type *)
@ -243,7 +243,140 @@ Section finite_hott.
Defined. Defined.
End split. End split.
End finite_hott.
Arguments Bfin {_} _.
Section dec_membership.
Variable (A : Type).
Context `{DecidablePaths A} `{Univalence}.
Global Instance DecidableMembership (P : Sub A) (Hfin : Bfin P) (a : A) :
Decidable (a P).
Proof.
destruct Hfin as [n Hequiv].
strip_truncations.
revert Hequiv.
revert P.
induction n.
- intros.
pose (X_empty _ P Hequiv) as p.
rewrite p.
apply _.
- intros.
pose (new_el _ P n Hequiv) as b.
destruct b as [b HX'].
destruct (split _ P n Hequiv) as [X' X'equiv]. simpl in HX'.
unfold member, sub_membership.
rewrite (HX' a).
pose (IHn X' X'equiv) as IH.
destruct IH as [IH | IH].
+ left. apply (tr (inl IH)).
+ destruct (dec (a = b)) as [Hab | Hab].
left. apply (tr (inr (tr Hab))).
right. intros α. strip_truncations.
destruct α as [β | γ]; [ | strip_truncations]; contradiction.
Defined.
End dec_membership.
Section cowd.
Variable (A : Type).
Context `{DecidablePaths A} `{Univalence}.
Definition cow := { X : Sub A | Bfin X}.
Definition empty_cow : cow.
Proof.
exists empty. apply empty_finite.
Defined.
Definition add_cow : forall a : A, cow -> cow.
Proof.
intros a [X PX].
exists (fun z => lor (X z) (merely (z = a))).
destruct (dec (a X)) as [Ha | Ha];
destruct PX as [n PX];
strip_truncations.
- (* a ∈ X *)
exists n. apply tr.
transitivity ({a : A & a X}); [ | apply PX ].
apply equiv_functor_sigma_id.
intro a'. eapply equiv_iff_hprop_uncurried ; split; cbn.
+ intros Ha'. strip_truncations.
destruct Ha' as [HXa' | Haa'].
* assumption.
* strip_truncations. rewrite Haa'. assumption.
+ intros HXa'. apply tr.
left. assumption.
- (* a ∉ X *)
exists (S n). apply tr.
destruct PX as [f [g Hfg Hgf adj]].
unshelve esplit.
+ intros [a' Ha']. cbn in Ha'.
destruct (dec (a' = a)) as [Haa' | Haa'].
* right. apply tt.
* assert (X a') as HXa'.
{ strip_truncations.
destruct Ha' as [Ha' | Ha']; [ assumption | ].
strip_truncations. by (contradiction (Haa' Ha')). }
apply (inl (f (a';HXa'))).
+ apply isequiv_biinv; simpl.
unshelve esplit; simpl.
* unfold Sect; simpl.
simple refine (_;_).
{ destruct 1 as [M | ?].
- destruct (g M) as [a' Ha'].
exists a'. apply tr.
by left.
- exists a. apply (tr (inr (tr idpath))). }
simpl. intros [a' Ha'].
strip_truncations.
destruct Ha' as [HXa' | Haa']; simpl;
destruct (dec (a' = a)); simpl.
** apply path_sigma' with p^. apply path_ishprop.
** rewrite Hgf; cbn. done.
** apply path_sigma' with p^. apply path_ishprop.
** rewrite Hgf; cbn. apply path_sigma' with idpath. apply path_ishprop.
* unfold Sect; simpl.
simple refine (_;_).
{ destruct 1 as [M | ?].
- destruct (g M) as [a' Ha'].
exists a'. apply tr.
by left.
- exists a. apply (tr (inr (tr idpath))). }
simpl. intros [M | [] ].
** destruct (dec (_ = a)) as [Haa' | Haa']; simpl.
{ destruct (g M) as [a' Ha']. rewrite Haa' in Ha'. by contradiction. }
{ f_ap. }
** destruct (dec (a = a)); try by contradiction.
reflexivity.
Defined.
Theorem cowy
(P : cow -> hProp)
(doge : P empty_cow)
(koeientaart : forall a c, P c -> P (add_cow a c))
:
forall X : cow, P X.
Proof.
intros [X [n FX]]. strip_truncations.
revert X FX.
induction n; intros X FX.
- pose (HX_emp:= X_empty _ X FX).
assert ((X; Build_Finite _ 0 (tr FX)) = empty_cow) as HX.
{ apply path_sigma' with HX_emp. apply path_ishprop. }
rewrite HX. assumption.
- pose (a' := new_el _ X n FX).
destruct a' as [a' Ha'].
destruct (split _ X n FX) as [X' FX'].
pose (X'cow := (X'; Build_Finite _ n (tr FX')) : cow).
assert ((X; Build_Finite _ (n.+1) (tr FX)) = add_cow a' X'cow) as .
{ simple refine (path_sigma' _ _ _); [ | apply path_ishprop ].
apply path_forall. intros a.
unfold X'cow.
specialize (Ha' a). rewrite Ha'. simpl. reflexivity. }
rewrite .
apply koeientaart.
apply IHn.
Defined.
Definition bfin_to_kfin : forall (X : Sub A), Bfin X -> Kf_sub _ X. Definition bfin_to_kfin : forall (X : Sub A), Bfin X -> Kf_sub _ X.
Proof. Proof.
@ -263,8 +396,8 @@ Section finite_hott.
apply (fun Xz => f(z;Xz)). apply (fun Xz => f(z;Xz)).
- intros. - intros.
simpl in *. simpl in *.
destruct (new_el X n iso) as [a HXX']. destruct (new_el _ X n iso) as [a HXX'].
destruct (split X n iso) as [X' X'equiv]. destruct (split _ X n iso) as [X' X'equiv].
destruct (IHn X' X'equiv) as [Y HY]. destruct (IHn X' X'equiv) as [Y HY].
exists (Y {|a|}). exists (Y {|a|}).
unfold map in *. unfold map in *.
@ -277,111 +410,111 @@ Section finite_hott.
reflexivity. reflexivity.
Defined. Defined.
Context `{A_deceq : DecidablePaths A}. Lemma kfin_is_bfin : @closedUnion A Bfin.
(*
Lemma kfin_is_bfin : closedUnion Bfin.
Proof. Proof.
intros X Y HX HY. intros X Y HX HY.
unfold Bfin in *. pose (Xcow := (X; HX) : cow).
destruct HX as [n Xequiv]. pose (Ycow := (Y; HY) : cow).
revert X Xequiv. simple refine (cowy (fun C => Bfin (C.1 Y)) _ _ Xcow); simpl.
induction n. - assert ((fun a => Trunc (-1) (Empty + Y a)) = (fun a => Y a)) as Help.
- intros. { apply path_forall. intros z; simpl.
strip_truncations. apply path_iff_ishprop.
rewrite (X_empty X Xequiv). + intros; strip_truncations; auto.
assert( Y = Y). destruct X0; auto. destruct e.
{ apply path_forall ; intro z. + intros ?. apply tr. right; assumption.
compute-[lor]. (* TODO FIX THIS with sum_empty_l *)
eauto with lattice_hints typeclass_instances.
} }
rewrite X0. rewrite Help. apply HY.
apply HY. - intros a [X' HX'] [n FX'Y]. strip_truncations.
- simpl in *. destruct (dec(a X')) as [HaX' | HaX'].
intros. * exists n. apply tr.
destruct HY as [m Yequiv]. transitivity ({a : A & Trunc (-1) (X' a + Y a)}); [| assumption ].
strip_truncations. apply equiv_functor_sigma_id. intro a'.
destruct (new_el X n Xequiv) as [a HXX']. apply equiv_iff_hprop.
destruct (split X n Xequiv) as [X' X'equiv]. { intros Q. strip_truncations.
destruct (IHn X' (tr X'equiv)) as [k Hk]. destruct Q as [Q | Q].
strip_truncations. - strip_truncations.
cbn in *. apply tr. left.
rewrite (path_forall _ _ HXX'). destruct Q ; auto.
assert strip_truncations. rewrite t; assumption.
(forall a0, - apply (tr (inr Q)). }
BuildhProp (Trunc (-1) (X' a0 merely (a0 = a) + Y a0)) { intros Q. strip_truncations.
= destruct Q as [Q | Q]; apply tr.
BuildhProp (hor (Trunc (-1) (X' a0 + Y a0)) (merely (a0 = a))) - left. apply tr. left. done.
). - right. done. }
{ * destruct (dec (a Y)) as [HaY | HaY ].
intros. ** exists n. apply tr.
apply path_iff_hprop. transitivity ({a : A & Trunc (-1) (X' a + Y a)}); [| assumption ].
* intros X0. apply equiv_functor_sigma_id. intro a'.
strip_truncations. apply equiv_iff_hprop.
destruct X0 as [X0 | X0]. { intros Q. strip_truncations.
** strip_truncations. destruct Q as [Q | Q].
destruct X0 as [X0 | X0]. - strip_truncations.
*** refine (tr(inl(tr _))). apply tr.
apply (inl X0). destruct Q.
*** refine (tr(inr X0)). left. auto.
** refine (tr(inl(tr _))). right. strip_truncations. rewrite t; assumption.
apply (inr X0). - apply (tr (inr Q)). }
* intros X0. { intros Q. strip_truncations.
strip_truncations. destruct Q as [Q | Q]; apply tr.
destruct X0 as [X0 | X0]. - left. apply tr. left. done.
** strip_truncations. - right. done. }
destruct X0 as [X0 | X0]. ** exists (n.+1). apply tr.
*** refine (tr(inl(tr(inl X0)))). destruct FX'Y as [f [g Hfg Hgf adj]].
*** refine (tr(inr X0)). unshelve esplit.
** refine (tr(inl(tr(inr X0)))). { intros [a' Ha']. cbn in Ha'.
} destruct (dec (BuildhProp (a' = a))) as [Ha'a | Ha'a].
(* rewrite (path_forall _ _ X0). *) - right. apply tt.
assert - left. refine (f (a';_)).
( strip_truncations.
{a0 : A & hor (Trunc (-1) (X' a0 + Y a0)) (merely (a0 = a))} destruct Ha' as [Ha' | Ha'].
= + strip_truncations.
{a0 : A & Trunc (-1) (X' a0 + Y a0)} destruct Ha' as [Ha' | Ha'].
+ * apply (tr (inl Ha')).
{a0 : A & (merely (a0 = a))} * strip_truncations. contradiction.
). + apply (tr (inr Ha')). }
{ { apply isequiv_biinv; simpl.
assert ({a0 : A & Trunc (-1) (X' a0 + Y a0)} + {a0 : A & merely (a0 = a)} -> unshelve esplit; simpl.
{a0 : A & hor (Trunc (-1) (X' a0 + Y a0)) (merely (a0 = a))}). - unfold Sect; simpl.
{ simple refine (_;_).
intros. { destruct 1 as [M | ?].
destruct X1. - destruct (g M) as [a' Ha'].
* destruct s as [c p]. exists a'.
exists c. strip_truncations; apply tr.
apply tr. destruct Ha' as [Ha' | Ha'].
left. + left. apply (tr (inl Ha')).
apply p. + right. done.
* destruct s as [c p]. - exists a. apply (tr (inl (tr (inr (tr idpath))))). }
exists c. { intros [a' Ha']; simpl.
apply tr. strip_truncations.
right. apply p. destruct Ha' as [HXa' | Haa']; simpl;
destruct (dec (a' = a)); simpl.
simple refine (path_universe _). ** apply path_sigma' with p^. apply path_ishprop.
* intros [a0 p]. ** rewrite Hgf; cbn. apply path_sigma' with idpath. apply path_ishprop.
destruct (dec (a0 = a)). ** apply path_sigma' with p^. apply path_ishprop.
** right. exists a0. apply (tr p0). ** rewrite Hgf; cbn. done. }
** left. - unfold Sect; simpl.
exists a0. simple refine (_;_).
strip_truncations. { destruct 1 as [M | ?].
destruct p ; strip_truncations. - (* destruct (g M) as [a' Ha']. *)
*** apply (tr t). exists (g M).1.
*** contradiction (n0 t). simple refine (Trunc_rec _ (g M).2).
* apply isequiv_biinv. intros Ha'.
unfold BiInv. apply tr.
split. (* strip_truncations; apply tr. *)
** destruct Ha' as [Ha' | Ha'].
+ left. apply (tr (inl Ha')).
exists a0 + right. done.
} - exists a. apply (tr (inl (tr (inr (tr idpath))))). }
rewrite X1. simpl. intros [M | [] ].
apply finite_sum. ** destruct (dec (_ = a)) as [Haa' | Haa']; simpl.
* simple refine (Build_Finite _ k (tr Hk)). { destruct (g M) as [a' Ha']. simpl in Haa'.
* apply singleton. strip_truncations.
Admitted. rewrite Haa' in Ha'. destruct Ha'; by contradiction. }
*) { f_ap. transitivity (f (g M)); [ | apply Hfg].
f_ap. apply path_sigma' with idpath.
End finite_hott. apply path_ishprop. }
** destruct (dec (a = a)); try by contradiction.
reflexivity. }
Defined.
End cowd.