diff --git a/src/category_theory/generator.lean b/src/category_theory/generator.lean index 522630650fe36..23f20b7069223 100644 --- a/src/category_theory/generator.lean +++ b/src/category_theory/generator.lean @@ -6,6 +6,8 @@ Authors: Markus Himmel import category_theory.balanced import category_theory.limits.opposites import category_theory.limits.shapes.zero_morphisms +import category_theory.subobject.lattice +import category_theory.subobject.well_powered import data.set.opposite /-! @@ -237,6 +239,43 @@ lemma is_codetecting_empty_of_groupoid [βˆ€ {X Y : C} (f : X ⟢ Y), is_iso f] : end empty +section well_powered + +namespace subobject + +lemma eq_of_le_of_is_detecting {𝒒 : set C} (h𝒒 : is_detecting 𝒒) {X : C} (P Q : subobject X) + (h₁ : P ≀ Q) (hβ‚‚ : βˆ€ (G ∈ 𝒒) {f : G ⟢ X}, Q.factors f β†’ P.factors f) : P = Q := +begin + suffices : is_iso (of_le _ _ h₁), + { exactI le_antisymm h₁ (le_of_comm (inv (of_le _ _ h₁)) (by simp)) }, + refine h𝒒 _ (Ξ» G hG f, _), + have : P.factors (f ≫ Q.arrow) := hβ‚‚ _ hG ((factors_iff _ _).2 ⟨_, rfl⟩), + refine ⟨factor_thru _ _ this, _, Ξ» g (hg : g ≫ _ = f), _⟩, + { simp only [← cancel_mono Q.arrow, category.assoc, of_le_arrow, factor_thru_arrow] }, + { simp only [← cancel_mono (subobject.of_le _ _ h₁), ← cancel_mono Q.arrow, hg, + category.assoc, of_le_arrow, factor_thru_arrow] } +end + +lemma inf_eq_of_is_detecting [has_pullbacks C] {𝒒 : set C} (h𝒒 : is_detecting 𝒒) {X : C} + (P Q : subobject X) (h : βˆ€ (G ∈ 𝒒) {f : G ⟢ X}, P.factors f β†’ Q.factors f) : P βŠ“ Q = P := +eq_of_le_of_is_detecting h𝒒 _ _ _root_.inf_le_left (Ξ» G hG f hf, (inf_factors _).2 ⟨hf, h _ hG hf⟩) + +lemma eq_of_is_detecting [has_pullbacks C] {𝒒 : set C} (h𝒒 : is_detecting 𝒒) {X : C} + (P Q : subobject X) (h : βˆ€ (G ∈ 𝒒) {f : G ⟢ X}, P.factors f ↔ Q.factors f) : P = Q := +calc P = P βŠ“ Q : eq.symm $ inf_eq_of_is_detecting h𝒒 _ _ $ Ξ» G hG f hf, (h G hG).1 hf + ... = Q βŠ“ P : inf_comm + ... = Q : inf_eq_of_is_detecting h𝒒 _ _ $ Ξ» G hG f hf, (h G hG).2 hf + +end subobject + +/-- A category with pullbacks and a small detecting set is well-powered. -/ +lemma well_powered_of_is_detecting [has_pullbacks C] {𝒒 : set C} [small.{v} 𝒒] + (h𝒒 : is_detecting 𝒒) : well_powered C := +⟨λ X, @small_of_injective _ _ _ (Ξ» P : subobject X, { f : Ξ£ G : 𝒒, G.1 ⟢ X | P.factors f.2 }) $ + Ξ» P Q h, subobject.eq_of_is_detecting h𝒒 _ _ (by simpa [set.ext_iff] using h)⟩ + +end well_powered + /-- We say that `G` is a separator if the functor `C(G, -)` is faithful. -/ def is_separator (G : C) : Prop := is_separating ({G} : set C) @@ -442,4 +481,8 @@ begin rwa [is_iso_iff_bijective, function.bijective_iff_exists_unique] } end +lemma well_powered_of_is_detector [has_pullbacks C] (G : C) (hG : is_detector G) : + well_powered C := +well_powered_of_is_detecting hG + end category_theory