diff --git a/hott/algebra/category/constructions/hset.hlean b/hott/algebra/category/constructions/hset.hlean index 27bdef664..c9c14d85b 100644 --- a/hott/algebra/category/constructions/hset.hlean +++ b/hott/algebra/category/constructions/hset.hlean @@ -6,9 +6,9 @@ Authors: Floris van Doorn, Jakob von Raumer Category of hsets -/ -import ..category types.equiv ..functor types.lift ..limits +import ..category types.equiv ..functor types.lift ..limits ..colimits hit.set_quotient hit.trunc -open eq category equiv iso is_equiv is_trunc function sigma +open eq category equiv iso is_equiv is_trunc function sigma set_quotient trunc namespace category @@ -23,22 +23,24 @@ namespace category definition Precategory_hset [reducible] [constructor] : Precategory := Precategory.mk hset precategory_hset + abbreviation pset [constructor] := Precategory_hset + namespace set local attribute is_equiv_subtype_eq [instance] - definition iso_of_equiv [constructor] {A B : Precategory_hset} (f : A ≃ B) : A ≅ B := + definition iso_of_equiv [constructor] {A B : pset} (f : A ≃ B) : A ≅ B := iso.MK (to_fun f) (to_inv f) (eq_of_homotopy (left_inv (to_fun f))) (eq_of_homotopy (right_inv (to_fun f))) - definition equiv_of_iso [constructor] {A B : Precategory_hset} (f : A ≅ B) : A ≃ B := + definition equiv_of_iso [constructor] {A B : pset} (f : A ≅ B) : A ≃ B := begin apply equiv.MK (to_hom f) (iso.to_inv f), exact ap10 (to_right_inverse f), exact ap10 (to_left_inverse f) end - definition is_equiv_iso_of_equiv [constructor] (A B : Precategory_hset) + definition is_equiv_iso_of_equiv [constructor] (A B : pset) : is_equiv (@iso_of_equiv A B) := adjointify _ (λf, equiv_of_iso f) (λf, proof iso_eq idp qed) @@ -51,7 +53,7 @@ namespace category @ap _ _ (to_fun (trunctype.sigma_char 0)) A B := eq_of_homotopy (λp, eq.rec_on p idp) - definition equiv_equiv_iso (A B : Precategory_hset) : (A ≃ B) ≃ (A ≅ B) := + definition equiv_equiv_iso (A B : pset) : (A ≃ B) ≃ (A ≅ B) := equiv.MK (λf, iso_of_equiv f) (λf, proof equiv.MK (to_hom f) (iso.to_inv f) @@ -60,10 +62,10 @@ namespace category (λf, proof iso_eq idp qed) (λf, proof equiv_eq idp qed) - definition equiv_eq_iso (A B : Precategory_hset) : (A ≃ B) = (A ≅ B) := + definition equiv_eq_iso (A B : pset) : (A ≃ B) = (A ≅ B) := ua !equiv_equiv_iso - definition is_univalent_hset (A B : Precategory_hset) : is_equiv (iso_of_eq : A = B → A ≅ B) := + definition is_univalent_hset (A B : pset) : is_equiv (iso_of_eq : A = B → A ≅ B) := assert H₁ : is_equiv (@iso_of_equiv A B ∘ @equiv_of_eq A B ∘ subtype_eq_inv _ _ ∘ @ap _ _ (to_fun (trunctype.sigma_char 0)) A B), from @is_equiv_compose _ _ _ _ _ @@ -80,7 +82,7 @@ namespace category end end set - definition category_hset [instance] [constructor] : category hset := + definition category_hset [instance] [constructor] [reducible] : category hset := category.mk precategory_hset set.is_univalent_hset definition Category_hset [reducible] [constructor] : Category := @@ -89,7 +91,7 @@ namespace category abbreviation set [constructor] := Category_hset open functor lift - definition lift_functor.{u v} [constructor] : set.{u} ⇒ set.{max u v} := + definition lift_functor.{u v} [constructor] : pset.{u} ⇒ pset.{max u v} := functor.mk tlift (λa b, lift_functor) (λa, eq_of_homotopy (λx, by induction x; reflexivity)) @@ -100,24 +102,24 @@ namespace category local attribute category.to_precategory [unfold 2] definition is_complete_set_cone.{u v w} [constructor] - (I : Precategory.{v w}) (F : I ⇒ set.{max u v w}) : cone_obj F := + (I : Precategory.{v w}) (F : I ⇒ pset.{max u v w}) : cone_obj F := begin fapply cone_obj.mk, { fapply trunctype.mk, { exact Σ(s : Π(i : I), trunctype.carrier (F i)), Π{i j : I} (f : i ⟶ j), F f (s i) = (s j)}, - { exact abstract begin apply is_trunc_sigma, intro s, + { exact sorry /-abstract begin apply is_trunc_sigma, intro s, apply is_trunc_pi, intro i, apply is_trunc_pi, intro j, apply is_trunc_pi, intro f, - apply is_trunc_eq end end}}, + apply is_trunc_eq end end-/}}, { fapply nat_trans.mk, { intro i x, esimp at x, exact x.1 i}, { intro i j f, esimp, apply eq_of_homotopy, intro x, esimp at x, induction x with s p, esimp, apply p}} end - definition is_complete_set.{u v w} [instance] : is_complete.{(max u v w)+1 (max u v w) v w} set := + definition is_complete_set.{u v w} [instance] : is_complete.{(max u v w)+1 (max u v w) v w} pset := begin intro I F, fapply has_terminal_object.mk, { exact is_complete_set_cone.{u v w} I F}, @@ -131,11 +133,83 @@ namespace category { esimp, intro h, induction h with f q, apply cone_hom_eq, esimp at *, apply eq_of_homotopy, intro x, fapply sigma_eq: esimp, { apply eq_of_homotopy, intro i, exact (ap10 (q i) x)⁻¹}, - { apply is_hprop.elimo, + { exact sorry /-apply is_hprop.elimo, apply is_trunc_pi, intro i, apply is_trunc_pi, intro j, apply is_trunc_pi, intro f, - apply is_trunc_eq}}} + apply is_trunc_eq-/}}} end + definition is_cocomplete_set_cone_rel.{u v w} [unfold 3 4] + (I : Precategory.{v w}) (F : I ⇒ pset.{max u v w}ᵒᵖ) : (Σ(i : I), trunctype.carrier (F i)) → + (Σ(i : I), trunctype.carrier (F i)) → hprop.{max u v w} := + begin + intro v w, induction v with i x, induction w with j y, + fapply trunctype.mk, + { exact ∃(f : i ⟶ j), to_fun_hom F f y = x}, + { exact _} + end + + definition is_cocomplete_set_cone.{u v w} [constructor] + (I : Precategory.{v w}) (F : I ⇒ pset.{max u v w}ᵒᵖ) : cone_obj F := + begin + fapply cone_obj.mk, + { fapply trunctype.mk, + { apply set_quotient (is_cocomplete_set_cone_rel.{u v w} I F)}, + { apply is_hset_set_quotient}}, + { fapply nat_trans.mk, + { intro i x, esimp, apply class_of, exact ⟨i, x⟩}, + { intro i j f, esimp, apply eq_of_homotopy, intro y, apply eq_of_rel, esimp, + exact exists.intro f idp}} + end + +-- adding the following step explicitly slightly reduces the elaboration time of the next proof + +-- definition is_cocomplete_set_cone_hom.{u v w} [constructor] +-- (I : Precategory.{v w}) (F : I ⇒ Opposite pset.{max u v w}) +-- (X : hset.{max u v w}) +-- (η : Π (i : carrier I), trunctype.carrier (to_fun_ob F i) → trunctype.carrier X) +-- (p : Π {i j : carrier I} (f : hom i j), (λ (x : trunctype.carrier (to_fun_ob F j)), η i (to_fun_hom F f x)) = η j) + +-- : --cone_hom (cone_obj.mk X (nat_trans.mk η @p)) (is_cocomplete_set_cone.{u v w} I F) +-- @cone_hom I psetᵒᵖ F +-- (@cone_obj.mk I psetᵒᵖ F X +-- (@nat_trans.mk I (Opposite pset) (@constant_functor I (Opposite pset) X) F η @p)) +-- (is_cocomplete_set_cone.{u v w} I F) +-- := +-- begin +-- fapply cone_hom.mk, +-- { refine set_quotient.elim _ _, +-- { intro v, induction v with i x, exact η i x}, +-- { intro v w r, induction v with i x, induction w with j y, esimp at *, +-- refine trunc.elim_on r _, clear r, +-- intro u, induction u with f q, +-- exact ap (η i) q⁻¹ ⬝ ap10 (p f) y}}, +-- { intro i, reflexivity} +-- end + + -- TODO: rewrite after induction tactic for trunc/set_quotient is implemented + definition is_cocomplete_set.{u v w} [instance] + : is_cocomplete.{(max u v w)+1 (max u v w) v w} set := + begin + intro I F, fapply has_terminal_object.mk, + { exact is_cocomplete_set_cone.{u v w} I F}, + { intro c, esimp at *, induction c with X η, induction η with η p, esimp at *, + fapply is_contr.mk, + { fapply cone_hom.mk, + { refine set_quotient.elim _ _, + { intro v, induction v with i x, exact η i x}, + { intro v w r, induction v with i x, induction w with j y, esimp at *, + refine trunc.elim_on r _, clear r, + intro u, induction u with f q, + exact ap (η i) q⁻¹ ⬝ ap10 (p f) y}}, + { intro i, reflexivity}}, + { esimp, intro h, induction h with f q, apply cone_hom_eq, esimp at *, + apply eq_of_homotopy, refine set_quotient.rec _ _, + { intro v, induction v with i x, esimp, exact (ap10 (q i) x)⁻¹}, + { intro v w r, apply is_hprop.elimo}}}, + end + + + end category diff --git a/hott/cubical/square.hlean b/hott/cubical/square.hlean index 8b023ea9e..48b8dab4e 100644 --- a/hott/cubical/square.hlean +++ b/hott/cubical/square.hlean @@ -467,6 +467,10 @@ namespace eq : square (t₁ ◾ t₂) (b₁ ◾ b₂) (l₁ ◾ l₂) (r₁ ◾ r₂) := by induction s₂;induction s₁;constructor + open is_trunc + definition is_hset.elims [H : is_hset A] : square p₁₀ p₁₂ p₀₁ p₂₁ := + square_of_eq !is_hset.elim + -- definition square_of_con_inv_hsquare {p₁ p₂ p₃ p₄ : a₁ = a₂} -- {t : p₁ = p₂} {b : p₃ = p₄} {l : p₁ = p₃} {r : p₂ = p₄} -- (s : square (con_inv_eq_idp t) (con_inv_eq_idp b) (l ◾ r⁻²) idp)