diff --git a/FiniteSets/Sub.v b/FiniteSets/Sub.v index 29bffa8..ee299a3 100644 --- a/FiniteSets/Sub.v +++ b/FiniteSets/Sub.v @@ -100,144 +100,3 @@ Section intersect. apply (inl (tr (t2^ @ t1))). Defined. End intersect. - -Section finite_hott. - Variable A : Type. - Context `{Univalence} `{IsHSet A}. - - Definition finite (X : Sub A) : hProp := BuildhProp (Finite {a : A & a ∈ X}). - - Definition singleton : closedSingleton finite. - Proof. - intros a. - simple refine (Build_Finite _ _ _). - - apply 1. - - apply tr. - simple refine (BuildEquiv _ _ _ _). - * apply (fun _ => inr tt). - * simple refine (BuildIsEquiv _ _ _ _ _ _ _) ; unfold Sect in *. - ** apply (fun _ => (a;tr idpath)). - ** intros x ; destruct x as [ | x] ; try contradiction. - destruct x ; reflexivity. - ** intros [b bp] ; simpl. - strip_truncations. - simple refine (path_sigma _ _ _ _ _). - *** apply bp^. - *** apply path_ishprop. - ** intros. - apply path_ishprop. - Defined. - - Definition empty_finite : closedEmpty finite. - Proof. - simple refine (Build_Finite _ _ _). - - apply 0. - - apply tr. - simple refine (BuildEquiv _ _ _ _). - intros [a p] ; apply p. - Defined. - - Definition decidable_empty_finite : hasDecidableEmpty finite. - Proof. - intros X Y. - destruct Y as [n Xn]. - strip_truncations. - simpl in Xn. - destruct Xn as [f [g fg gf adj]]. - destruct n. - - refine (tr(inl _)). - unfold empty. - apply path_forall. - intro z. - apply path_iff_hprop. - * intros p. - contradiction (f(z;p)). - * contradiction. - - refine (tr(inr _)). - apply (tr(g(inr tt))). - Defined. - - Lemma no_union - (f : forall (X Y : Sub A), - Finite {a : A & X a} -> Finite {a : A & Y a} - -> Finite ({a : A & (X ∪ Y) a})) - (a b : A) - : - hor (a = b) (a = b -> Empty). - Proof. - specialize (f {|a|} {|b|} (singleton a) (singleton b)). - destruct f as [n pn]. - strip_truncations. - destruct pn as [f [g fg gf adj]]. - unfold Sect in *. - destruct n. - - cbn in *. contradiction f. - exists a. - apply (tr(inl(tr idpath))). - - destruct n ; cbn in *. - -- pose ((a;tr(inl(tr idpath))) - : {a0 : A & Trunc (-1) (Trunc (-1) (a0 = a) + Trunc (-1) (a0 = b))}) - as s1. - pose ((b;tr(inr(tr idpath))) - : {a0 : A & Trunc (-1) (Trunc (-1) (a0 = a) + Trunc (-1) (a0 = b))}) - as s2. - pose (f s1) as fs1. - pose (f s2) as fs2. - assert (fs1 = fs2) as fs_eq. - { apply path_ishprop. } - pose (g fs1) as gfs1. - pose (g fs2) as gfs2. - refine (tr(inl _)). - refine (ap (fun x => x.1) (gf s1)^ @ _ @ (ap (fun x => x.1) (gf s2))). - unfold fs1, fs2 in fs_eq. rewrite fs_eq. - reflexivity. - -- refine (tr(inr _)). - intros p. - pose (inl(inr tt) : Fin n + Unit + Unit) as s1. - pose (inr tt : Fin n + Unit + Unit) as s2. - pose (g s1) as gs1. - pose (c := g s1). - assert (c = gs1) as ps1. reflexivity. - pose (g s2) as gs2. - pose (d := g s2). - assert (d = gs2) as ps2. reflexivity. - pose (f gs1) as gfs1. - pose (f gs2) as gfs2. - destruct c as [x px] ; destruct d as [y py]. - simple refine (Trunc_ind _ _ px) ; intros p1. - simple refine (Trunc_ind _ _ py) ; intros p2. - simpl. - assert (x = y -> Empty) as X1. - { - enough (s1 = s2) as X. - { - intros. - unfold s1, s2 in X. - refine (not_is_inl_and_inr' (inl(inr tt)) _ _). - + apply tt. - + rewrite X ; apply tt. - } - transitivity gfs1. - { unfold gfs1, s1. apply (fg s1)^. } - symmetry ; transitivity gfs2. - { unfold gfs2, s2. apply (fg s2)^. } - unfold gfs2, gfs1. - rewrite <- ps1, <- ps2. - f_ap. - simple refine (path_sigma _ _ _ _ _). - * cbn. - destruct p1 as [p1 | p1] ; destruct p2 as [p2 | p2] ; strip_truncations. - ** apply (p2 @ p1^). - ** refine (p2 @ _^ @ p1^). auto. - ** refine (p2 @ _ @ p1^). auto. - ** apply (p2 @ p1^). - * apply path_ishprop. - } - apply X1. - destruct p1 as [p1 | p1] ; destruct p2 as [p2 | p2] ; strip_truncations. - ** apply (p1 @ p2^). - ** refine (p1 @ _ @ p2^). auto. - ** refine (p1 @ _ @ p2^). symmetry. auto. - ** apply (p1 @ p2^). - Defined. -End finite_hott. \ No newline at end of file diff --git a/FiniteSets/_CoqProject b/FiniteSets/_CoqProject index d789fa2..5749d82 100644 --- a/FiniteSets/_CoqProject +++ b/FiniteSets/_CoqProject @@ -23,5 +23,6 @@ implementations/interface.v implementations/lists.v variations/enumerated.v variations/k_finite.v +variations/b_finite.v #empty_set.v ordered.v diff --git a/FiniteSets/variations/b_finite.v b/FiniteSets/variations/b_finite.v new file mode 100644 index 0000000..b8be2c3 --- /dev/null +++ b/FiniteSets/variations/b_finite.v @@ -0,0 +1,135 @@ +(* Bishop-finiteness is that "default" notion of finiteness in the HoTT library *) +Require Import HoTT. +Require Import Sub notation. + +Section finite_hott. + Variable A : Type. + Context `{Univalence} `{IsHSet A}. + + (* A subobject is B-finite if its extension is B-finite as a type *) + Definition Bfin (X : Sub A) : hProp := BuildhProp (Finite {a : A & a ∈ X}). + + Global Instance singleton_contr a : Contr {b : A & b ∈ {|a|}}. + Proof. + exists (a; tr idpath). + intros [b p]. + simple refine (Trunc_ind (fun p => (a; tr 1%path) = (b; p)) _ p). + clear p; intro p. simpl. + apply path_sigma' with (p^). + apply path_ishprop. + Defined. + + Definition singleton_fin_equiv' a : Fin 1 -> {b : A & b ∈ {|a|}}. + Proof. + intros _. apply (center {b : A & b ∈ {|a|}}). + Defined. + + Global Instance singleton_fin_equiv a : IsEquiv (singleton_fin_equiv' a). + Proof. apply _. Defined. + + Definition singleton : closedSingleton Bfin. + Proof. + intros a. + simple refine (Build_Finite _ 1 _). + apply tr. + symmetry. + refine (BuildEquiv _ _ (singleton_fin_equiv' a) _). + Defined. + + Definition empty_finite : closedEmpty Bfin. + Proof. + simple refine (Build_Finite _ 0 _). + apply tr. + simple refine (BuildEquiv _ _ _ _). + intros [a p]; apply p. + Defined. + + Definition decidable_empty_finite : hasDecidableEmpty Bfin. + Proof. + intros X Y. + destruct Y as [n Xn]. + strip_truncations. + destruct Xn as [f [g fg gf adj]]. + destruct n. + - refine (tr(inl _)). + apply path_forall. intro z. + apply path_iff_hprop. + * intros p. + contradiction (f (z;p)). + * contradiction. + - refine (tr(inr _)). + apply (tr(g(inr tt))). + Defined. + + Lemma no_union + (f : forall (X Y : Sub A), + Bfin X -> Bfin Y -> Bfin (X ∪ Y)) + (a b : A) : + hor (a = b) (a = b -> Empty). + Proof. + specialize (f {|a|} {|b|} (singleton a) (singleton b)). + unfold Bfin in f. + destruct f as [n pn]. + strip_truncations. + destruct pn as [f [g fg gf _]]. + destruct n as [|n]. + unfold Sect in *. + - contradiction f. + exists a. apply (tr(inl(tr idpath))). + - destruct n as [|n]. + + (* If the size of the union is 1, then (a = b) *) + refine (tr (inl _)). + pose (s1 := (a;tr(inl(tr idpath))) + : {c : A & Trunc (-1) (Trunc (-1) (c = a) + Trunc (-1) (c = b))}). + pose (s2 := (b;tr(inr(tr idpath))) + : {c : A & Trunc (-1) (Trunc (-1) (c = a) + Trunc (-1) (c = b))}). + refine (ap (fun x => x.1) (gf s1)^ @ _ @ (ap (fun x => x.1) (gf s2))). + assert (fs_eq : f s1 = f s2). + { by apply path_ishprop. } + refine (ap (fun x => (g x).1) fs_eq). + + (* Otherwise, ¬(a = b) *) + refine (tr (inr _)). + intros p. + pose (s1 := inl (inr tt) : Fin n + Unit + Unit). + pose (s2 := inr tt : Fin n + Unit + Unit). + pose (gs1 := g s1). + pose (c := g s1). + pose (gs2 := g s2). + pose (d := g s2). + assert (Hgs1 : gs1 = c) by reflexivity. + assert (Hgs2 : gs2 = d) by reflexivity. + destruct c as [x px']. + destruct d as [y py']. + simple refine (Trunc_ind _ _ px') ; intros px. + simple refine (Trunc_ind _ _ py') ; intros py. + simpl. + cut (x = y). + { + enough (s1 = s2) as X. + { + intros. + unfold s1, s2 in X. + refine (not_is_inl_and_inr' (inl(inr tt)) _ _). + + apply tt. + + rewrite X ; apply tt. + } + transitivity (f gs1). + { apply (fg s1)^. } + symmetry ; transitivity (f gs2). + { apply (fg s2)^. } + rewrite Hgs1, Hgs2. + f_ap. + simple refine (path_sigma _ _ _ _ _); [ | apply path_ishprop ]; simpl. + destruct px as [p1 | p1] ; destruct py as [p2 | p2] ; strip_truncations. + * apply (p2 @ p1^). + * refine (p2 @ _^ @ p1^). auto. + * refine (p2 @ _ @ p1^). auto. + * apply (p2 @ p1^). + } + destruct px as [px | px] ; destruct py as [py | py]; strip_truncations. + ** apply (px @ py^). + ** refine (px @ _ @ py^). auto. + ** refine (px @ _ @ py^). symmetry. auto. + ** apply (px @ py^). + Defined. +End finite_hott.