|
| 1 | +/- |
| 2 | +Copyright (c) 2020 Scott Morrison. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Scott Morrison, Bhavik Mehta |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module category_theory.limits.shapes.regular_mono |
| 7 | +! leanprover-community/mathlib commit 239d882c4fb58361ee8b3b39fb2091320edef10a |
| 8 | +! Please do not edit these lines, except to modify the commit id |
| 9 | +! if you have ported upstream changes. |
| 10 | +-/ |
| 11 | +import Mathlib.CategoryTheory.Limits.Shapes.Pullbacks |
| 12 | +import Mathlib.CategoryTheory.Limits.Shapes.StrongEpi |
| 13 | +import Mathlib.CategoryTheory.Limits.Shapes.Equalizers |
| 14 | +import Mathlib.Lean.Expr.Basic |
| 15 | + |
| 16 | +/-! |
| 17 | +# Definitions and basic properties of regular monomorphisms and epimorphisms. |
| 18 | +
|
| 19 | +A regular monomorphism is a morphism that is the equalizer of some parallel pair. |
| 20 | +
|
| 21 | +We give the constructions |
| 22 | +* `IsSplitMono → RegularMono` and |
| 23 | +* `RegularMono → Mono` |
| 24 | +as well as the dual constructions for regular epimorphisms. Additionally, we give the construction |
| 25 | +* `RegularEpi ⟶ StrongEpi`. |
| 26 | +
|
| 27 | +We also define classes `RegularMonoCategory` and `RegularEpiCategory` for categories in which |
| 28 | +every monomorphism or epimorphism is regular, and deduce that these categories are |
| 29 | +`StrongMonoCategory`s resp. `StrongEpiCategory`s. |
| 30 | +
|
| 31 | +-/ |
| 32 | + |
| 33 | + |
| 34 | +noncomputable section |
| 35 | + |
| 36 | +namespace CategoryTheory |
| 37 | + |
| 38 | +open CategoryTheory.Limits |
| 39 | + |
| 40 | +universe v₁ u₁ u₂ |
| 41 | + |
| 42 | +variable {C : Type u₁} [Category.{v₁} C] |
| 43 | + |
| 44 | +variable {X Y : C} |
| 45 | + |
| 46 | +/-- A regular monomorphism is a morphism which is the equalizer of some parallel pair. -/ |
| 47 | +class RegularMono (f : X ⟶ Y) where |
| 48 | + /-- An object in `C` -/ |
| 49 | + Z : C -- Porting note: violates naming but what is better? |
| 50 | + /-- A map from the codomain of `f` to `Z` -/ |
| 51 | + left : Y ⟶ Z |
| 52 | + /-- Another map from the codomain of `f` to `Z` -/ |
| 53 | + right : Y ⟶ Z |
| 54 | + /-- `f` equalizes the two maps -/ |
| 55 | + w : f ≫ left = f ≫ right |
| 56 | + /-- `f` is the equalizer of the two maps -/ |
| 57 | + isLimit : IsLimit (Fork.ofι f w) |
| 58 | +#align category_theory.regular_mono CategoryTheory.RegularMono |
| 59 | + |
| 60 | +attribute [reassoc] RegularMono.w |
| 61 | + |
| 62 | +/-- Every regular monomorphism is a monomorphism. -/ |
| 63 | +instance (priority := 100) RegularMono.mono (f : X ⟶ Y) [RegularMono f] : Mono f := |
| 64 | + mono_of_isLimit_fork RegularMono.isLimit |
| 65 | +#align category_theory.regular_mono.mono CategoryTheory.RegularMono.mono |
| 66 | + |
| 67 | +instance equalizerRegular (g h : X ⟶ Y) [HasLimit (parallelPair g h)] : |
| 68 | + RegularMono (equalizer.ι g h) where |
| 69 | + Z := Y |
| 70 | + left := g |
| 71 | + right := h |
| 72 | + w := equalizer.condition g h |
| 73 | + isLimit := |
| 74 | + Fork.IsLimit.mk _ (fun s => limit.lift _ s) (by simp) fun s m w => |
| 75 | + by |
| 76 | + apply equalizer.hom_ext |
| 77 | + simp [← w] |
| 78 | +#align category_theory.equalizer_regular CategoryTheory.equalizerRegular |
| 79 | + |
| 80 | +/-- Every split monomorphism is a regular monomorphism. -/ |
| 81 | +instance (priority := 100) RegularMono.ofIsSplitMono (f : X ⟶ Y) [IsSplitMono f] : RegularMono f |
| 82 | + where |
| 83 | + Z := Y |
| 84 | + left := 𝟙 Y |
| 85 | + right := retraction f ≫ f |
| 86 | + w := by aesop_cat |
| 87 | + isLimit := isSplitMonoEqualizes f |
| 88 | +#align category_theory.regular_mono.of_is_split_mono CategoryTheory.RegularMono.ofIsSplitMono |
| 89 | + |
| 90 | +/-- If `f` is a regular mono, then any map `k : W ⟶ Y` equalizing `RegularMono.left` and |
| 91 | + `RegularMono.right` induces a morphism `l : W ⟶ X` such that `l ≫ f = k`. -/ |
| 92 | +def RegularMono.lift' {W : C} (f : X ⟶ Y) [RegularMono f] (k : W ⟶ Y) |
| 93 | + (h : k ≫ (RegularMono.left : Y ⟶ @RegularMono.Z _ _ _ _ f _) = k ≫ RegularMono.right) : |
| 94 | + { l : W ⟶ X // l ≫ f = k } := |
| 95 | + Fork.IsLimit.lift' RegularMono.isLimit _ h |
| 96 | +#align category_theory.regular_mono.lift' CategoryTheory.RegularMono.lift' |
| 97 | + |
| 98 | +/-- The second leg of a pullback cone is a regular monomorphism if the right component is too. |
| 99 | +
|
| 100 | +See also `Pullback.sndOfMono` for the basic monomorphism version, and |
| 101 | +`regularOfIsPullbackFstOfRegular` for the flipped version. |
| 102 | +-/ |
| 103 | +def regularOfIsPullbackSndOfRegular {P Q R S : C} {f : P ⟶ Q} {g : P ⟶ R} {h : Q ⟶ S} {k : R ⟶ S} |
| 104 | + [hr : RegularMono h] (comm : f ≫ h = g ≫ k) (t : IsLimit (PullbackCone.mk _ _ comm)) : |
| 105 | + RegularMono g where |
| 106 | + Z := hr.Z |
| 107 | + left := k ≫ hr.left |
| 108 | + right := k ≫ hr.right |
| 109 | + w := by |
| 110 | + repeat (rw [← Category.assoc, ← eq_whisker comm]) |
| 111 | + simp only [Category.assoc, hr.w] |
| 112 | + isLimit := by |
| 113 | + apply Fork.IsLimit.mk' _ _ |
| 114 | + intro s |
| 115 | + have l₁ : (Fork.ι s ≫ k) ≫ RegularMono.left = (Fork.ι s ≫ k) ≫ hr.right |
| 116 | + rw [Category.assoc, s.condition, Category.assoc] |
| 117 | + obtain ⟨l, hl⟩ := Fork.IsLimit.lift' hr.isLimit _ l₁ |
| 118 | + obtain ⟨p, _, hp₂⟩ := PullbackCone.IsLimit.lift' t _ _ hl |
| 119 | + refine' ⟨p, hp₂, _⟩ |
| 120 | + intro m w |
| 121 | + have z : m ≫ g = p ≫ g := w.trans hp₂.symm |
| 122 | + apply t.hom_ext |
| 123 | + apply (PullbackCone.mk f g comm).equalizer_ext |
| 124 | + · erw [← cancel_mono h, Category.assoc, Category.assoc, comm] |
| 125 | + simp only [← Category.assoc, eq_whisker z] |
| 126 | + · exact z |
| 127 | +#align category_theory.regular_of_is_pullback_snd_of_regular CategoryTheory.regularOfIsPullbackSndOfRegular |
| 128 | + |
| 129 | +/-- The first leg of a pullback cone is a regular monomorphism if the left component is too. |
| 130 | +
|
| 131 | +See also `Pullback.fstOfMono` for the basic monomorphism version, and |
| 132 | +`regularOfIsPullbackSndOfRegular` for the flipped version. |
| 133 | +-/ |
| 134 | +def regularOfIsPullbackFstOfRegular {P Q R S : C} {f : P ⟶ Q} {g : P ⟶ R} {h : Q ⟶ S} {k : R ⟶ S} |
| 135 | + [RegularMono k] (comm : f ≫ h = g ≫ k) (t : IsLimit (PullbackCone.mk _ _ comm)) : |
| 136 | + RegularMono f := |
| 137 | + regularOfIsPullbackSndOfRegular comm.symm (PullbackCone.flipIsLimit t) |
| 138 | +#align category_theory.regular_of_is_pullback_fst_of_regular CategoryTheory.regularOfIsPullbackFstOfRegular |
| 139 | + |
| 140 | +instance (priority := 100) strongMono_of_regularMono (f : X ⟶ Y) [RegularMono f] : StrongMono f := |
| 141 | + StrongMono.mk' (by |
| 142 | + intro A B z hz u v sq |
| 143 | + have : v ≫ (RegularMono.left : Y ⟶ RegularMono.Z f) = v ≫ RegularMono.right := by |
| 144 | + apply (cancel_epi z).1 |
| 145 | + repeat (rw [← Category.assoc, ← eq_whisker sq.w]) |
| 146 | + simp only [Category.assoc, RegularMono.w] |
| 147 | + obtain ⟨t, ht⟩ := RegularMono.lift' _ _ this |
| 148 | + refine' CommSq.HasLift.mk' ⟨t, (cancel_mono f).1 _, ht⟩ |
| 149 | + simp only [Arrow.mk_hom, Arrow.homMk'_left, Category.assoc, ht, sq.w]) |
| 150 | +#align category_theory.strong_mono_of_regular_mono CategoryTheory.strongMono_of_regularMono |
| 151 | + |
| 152 | +/-- A regular monomorphism is an isomorphism if it is an epimorphism. -/ |
| 153 | +theorem isIso_of_regularMono_of_epi (f : X ⟶ Y) [RegularMono f] [Epi f] : IsIso f := |
| 154 | + isIso_of_epi_of_strongMono _ |
| 155 | +#align category_theory.is_iso_of_regular_mono_of_epi CategoryTheory.isIso_of_regularMono_of_epi |
| 156 | + |
| 157 | +section |
| 158 | + |
| 159 | +variable (C) |
| 160 | + |
| 161 | +/-- A regular mono category is a category in which every monomorphism is regular. -/ |
| 162 | +class RegularMonoCategory where |
| 163 | + /-- Every monomorphism is a regular monomorphism -/ |
| 164 | + regularMonoOfMono : ∀ {X Y : C} (f : X ⟶ Y) [Mono f], RegularMono f |
| 165 | +#align category_theory.regular_mono_category CategoryTheory.RegularMonoCategory |
| 166 | + |
| 167 | +end |
| 168 | + |
| 169 | +/-- In a category in which every monomorphism is regular, we can express every monomorphism as |
| 170 | + an equalizer. This is not an instance because it would create an instance loop. -/ |
| 171 | +def regularMonoOfMono [RegularMonoCategory C] (f : X ⟶ Y) [Mono f] : RegularMono f := |
| 172 | + RegularMonoCategory.regularMonoOfMono _ |
| 173 | +#align category_theory.regular_mono_of_mono CategoryTheory.regularMonoOfMono |
| 174 | + |
| 175 | +instance (priority := 100) regularMonoCategoryOfSplitMonoCategory [SplitMonoCategory C] : |
| 176 | + RegularMonoCategory C where |
| 177 | + regularMonoOfMono f _ := by |
| 178 | + haveI := isSplitMono_of_mono f |
| 179 | + infer_instance |
| 180 | +#align category_theory.regular_mono_category_of_split_mono_category CategoryTheory.regularMonoCategoryOfSplitMonoCategory |
| 181 | + |
| 182 | +instance (priority := 100) strongMonoCategory_of_regularMonoCategory [RegularMonoCategory C] : |
| 183 | + StrongMonoCategory C where |
| 184 | + strongMono_of_mono f _ := by |
| 185 | + haveI := regularMonoOfMono f |
| 186 | + infer_instance |
| 187 | +#align category_theory.strong_mono_category_of_regular_mono_category CategoryTheory.strongMonoCategory_of_regularMonoCategory |
| 188 | + |
| 189 | +/-- A regular epimorphism is a morphism which is the coequalizer of some parallel pair. -/ |
| 190 | +class RegularEpi (f : X ⟶ Y) where |
| 191 | + /-- An object from `C` -/ |
| 192 | + W : C -- Porting note: violates naming convention but what is better? |
| 193 | + /-- Two maps to the domain of `f` -/ |
| 194 | + (left right : W ⟶ X) |
| 195 | + /-- `f` coequalizes the two maps -/ |
| 196 | + w : left ≫ f = right ≫ f |
| 197 | + /-- `f` is the coequalizer -/ |
| 198 | + isColimit : IsColimit (Cofork.ofπ f w) |
| 199 | +#align category_theory.regular_epi CategoryTheory.RegularEpi |
| 200 | + |
| 201 | +attribute [reassoc] RegularEpi.w |
| 202 | + |
| 203 | +/-- Every regular epimorphism is an epimorphism. -/ |
| 204 | +instance (priority := 100) RegularEpi.epi (f : X ⟶ Y) [RegularEpi f] : Epi f := |
| 205 | + epi_of_isColimit_cofork RegularEpi.isColimit |
| 206 | +#align category_theory.regular_epi.epi CategoryTheory.RegularEpi.epi |
| 207 | + |
| 208 | +instance coequalizerRegular (g h : X ⟶ Y) [HasColimit (parallelPair g h)] : |
| 209 | + RegularEpi (coequalizer.π g h) where |
| 210 | + W := X |
| 211 | + left := g |
| 212 | + right := h |
| 213 | + w := coequalizer.condition g h |
| 214 | + isColimit := |
| 215 | + Cofork.IsColimit.mk _ (fun s => colimit.desc _ s) (by simp) fun s m w => by |
| 216 | + apply coequalizer.hom_ext; simp [← w] |
| 217 | +#align category_theory.coequalizer_regular CategoryTheory.coequalizerRegular |
| 218 | + |
| 219 | +/-- Every split epimorphism is a regular epimorphism. -/ |
| 220 | +instance (priority := 100) RegularEpi.ofSplitEpi (f : X ⟶ Y) [IsSplitEpi f] : RegularEpi f |
| 221 | + where |
| 222 | + W := X |
| 223 | + left := 𝟙 X |
| 224 | + right := f ≫ section_ f |
| 225 | + w := by aesop_cat |
| 226 | + isColimit := isSplitEpiCoequalizes f |
| 227 | +#align category_theory.regular_epi.of_split_epi CategoryTheory.RegularEpi.ofSplitEpi |
| 228 | + |
| 229 | +/-- If `f` is a regular epi, then every morphism `k : X ⟶ W` coequalizing `RegularEpi.left` and |
| 230 | + `RegularEpi.right` induces `l : Y ⟶ W` such that `f ≫ l = k`. -/ |
| 231 | +def RegularEpi.desc' {W : C} (f : X ⟶ Y) [RegularEpi f] (k : X ⟶ W) |
| 232 | + (h : (RegularEpi.left : RegularEpi.W f ⟶ X) ≫ k = RegularEpi.right ≫ k) : |
| 233 | + { l : Y ⟶ W // f ≫ l = k } := |
| 234 | + Cofork.IsColimit.desc' RegularEpi.isColimit _ h |
| 235 | +#align category_theory.regular_epi.desc' CategoryTheory.RegularEpi.desc' |
| 236 | + |
| 237 | +/-- The second leg of a pushout cocone is a regular epimorphism if the right component is too. |
| 238 | +
|
| 239 | +See also `Pushout.sndOfEpi` for the basic epimorphism version, and |
| 240 | +`regularOfIsPushoutFstOfRegular` for the flipped version. |
| 241 | +-/ |
| 242 | +def regularOfIsPushoutSndOfRegular {P Q R S : C} {f : P ⟶ Q} {g : P ⟶ R} {h : Q ⟶ S} {k : R ⟶ S} |
| 243 | + [gr : RegularEpi g] (comm : f ≫ h = g ≫ k) (t : IsColimit (PushoutCocone.mk _ _ comm)) : |
| 244 | + RegularEpi h where |
| 245 | + W := gr.W |
| 246 | + left := gr.left ≫ f |
| 247 | + right := gr.right ≫ f |
| 248 | + w := by rw [Category.assoc, Category.assoc, comm]; simp only [← Category.assoc, eq_whisker gr.w] |
| 249 | + isColimit := by |
| 250 | + apply Cofork.IsColimit.mk' _ _ |
| 251 | + intro s |
| 252 | + have l₁ : gr.left ≫ f ≫ s.π = gr.right ≫ f ≫ s.π |
| 253 | + rw [← Category.assoc, ← Category.assoc, s.condition] |
| 254 | + obtain ⟨l, hl⟩ := Cofork.IsColimit.desc' gr.isColimit (f ≫ Cofork.π s) l₁ |
| 255 | + obtain ⟨p, hp₁, _⟩ := PushoutCocone.IsColimit.desc' t _ _ hl.symm |
| 256 | + refine' ⟨p, hp₁, _⟩ |
| 257 | + intro m w |
| 258 | + have z := w.trans hp₁.symm |
| 259 | + apply t.hom_ext |
| 260 | + apply (PushoutCocone.mk _ _ comm).coequalizer_ext |
| 261 | + · exact z |
| 262 | + · erw [← cancel_epi g, ← Category.assoc, ← eq_whisker comm] |
| 263 | + erw [← Category.assoc, ← eq_whisker comm] |
| 264 | + dsimp at z; simp only [Category.assoc, z] |
| 265 | +#align category_theory.regular_of_is_pushout_snd_of_regular CategoryTheory.regularOfIsPushoutSndOfRegular |
| 266 | + |
| 267 | +/-- The first leg of a pushout cocone is a regular epimorphism if the left component is too. |
| 268 | +
|
| 269 | +See also `Pushout.fstOfEpi` for the basic epimorphism version, and |
| 270 | +`regularOfIsPushoutSndOfRegular` for the flipped version. |
| 271 | +-/ |
| 272 | +def regularOfIsPushoutFstOfRegular {P Q R S : C} {f : P ⟶ Q} {g : P ⟶ R} {h : Q ⟶ S} {k : R ⟶ S} |
| 273 | + [RegularEpi f] (comm : f ≫ h = g ≫ k) (t : IsColimit (PushoutCocone.mk _ _ comm)) : |
| 274 | + RegularEpi k := |
| 275 | + regularOfIsPushoutSndOfRegular comm.symm (PushoutCocone.flipIsColimit t) |
| 276 | +#align category_theory.regular_of_is_pushout_fst_of_regular CategoryTheory.regularOfIsPushoutFstOfRegular |
| 277 | + |
| 278 | +instance (priority := 100) strongEpi_of_regularEpi (f : X ⟶ Y) [RegularEpi f] : StrongEpi f := |
| 279 | + StrongEpi.mk' |
| 280 | + (by |
| 281 | + intro A B z hz u v sq |
| 282 | + have : (RegularEpi.left : RegularEpi.W f ⟶ X) ≫ u = RegularEpi.right ≫ u := |
| 283 | + by |
| 284 | + apply (cancel_mono z).1 |
| 285 | + simp only [Category.assoc, sq.w, RegularEpi.w_assoc] |
| 286 | + obtain ⟨t, ht⟩ := RegularEpi.desc' f u this |
| 287 | + exact |
| 288 | + CommSq.HasLift.mk' |
| 289 | + ⟨t, ht, |
| 290 | + (cancel_epi f).1 |
| 291 | + (by simp only [← Category.assoc, ht, ← sq.w, Arrow.mk_hom, Arrow.homMk'_right])⟩) |
| 292 | +#align category_theory.strong_epi_of_regular_epi CategoryTheory.strongEpi_of_regularEpi |
| 293 | + |
| 294 | +/-- A regular epimorphism is an isomorphism if it is a monomorphism. -/ |
| 295 | +theorem isIso_of_regularEpi_of_mono (f : X ⟶ Y) [RegularEpi f] [Mono f] : IsIso f := |
| 296 | + isIso_of_mono_of_strongEpi _ |
| 297 | +#align category_theory.is_iso_of_regular_epi_of_mono CategoryTheory.isIso_of_regularEpi_of_mono |
| 298 | + |
| 299 | +section |
| 300 | + |
| 301 | +variable (C) |
| 302 | + |
| 303 | +/-- A regular epi category is a category in which every epimorphism is regular. -/ |
| 304 | +class RegularEpiCategory where |
| 305 | + /-- Everyone epimorphism is a regular epimorphism -/ |
| 306 | + regularEpiOfEpi : ∀ {X Y : C} (f : X ⟶ Y) [Epi f], RegularEpi f |
| 307 | +#align category_theory.regular_epi_category CategoryTheory.RegularEpiCategory |
| 308 | + |
| 309 | +end |
| 310 | + |
| 311 | +/-- In a category in which every epimorphism is regular, we can express every epimorphism as |
| 312 | + a coequalizer. This is not an instance because it would create an instance loop. -/ |
| 313 | +def regularEpiOfEpi [RegularEpiCategory C] (f : X ⟶ Y) [Epi f] : RegularEpi f := |
| 314 | + RegularEpiCategory.regularEpiOfEpi _ |
| 315 | +#align category_theory.regular_epi_of_epi CategoryTheory.regularEpiOfEpi |
| 316 | + |
| 317 | +instance (priority := 100) regularEpiCategoryOfSplitEpiCategory [SplitEpiCategory C] : |
| 318 | + RegularEpiCategory C where |
| 319 | + regularEpiOfEpi f _ := |
| 320 | + by |
| 321 | + haveI := isSplitEpi_of_epi f |
| 322 | + infer_instance |
| 323 | +#align category_theory.regular_epi_category_of_split_epi_category CategoryTheory.regularEpiCategoryOfSplitEpiCategory |
| 324 | + |
| 325 | +instance (priority := 100) strongEpiCategory_of_regularEpiCategory [RegularEpiCategory C] : |
| 326 | + StrongEpiCategory C where |
| 327 | + strongEpi_of_epi f _ := by |
| 328 | + haveI := regularEpiOfEpi f |
| 329 | + infer_instance |
| 330 | +#align category_theory.strong_epi_category_of_regular_epi_category CategoryTheory.strongEpiCategory_of_regularEpiCategory |
| 331 | + |
| 332 | +end CategoryTheory |
| 333 | + |
0 commit comments