|
| 1 | +/- |
| 2 | +Copyright (c) 2020 Markus Himmel. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Markus Himmel |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module category_theory.limits.shapes.strong_epi |
| 7 | +! leanprover-community/mathlib commit 32253a1a1071173b33dc7d6a218cf722c6feb514 |
| 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.Balanced |
| 12 | +import Mathlib.CategoryTheory.LiftingProperties.Basic |
| 13 | + |
| 14 | +/-! |
| 15 | +# Strong epimorphisms |
| 16 | +
|
| 17 | +In this file, we define strong epimorphisms. A strong epimorphism is an epimorphism `f` |
| 18 | +which has the (unique) left lifting property with respect to monomorphisms. Similarly, |
| 19 | +a strong monomorphisms in a monomorphism which has the (unique) right lifting property |
| 20 | +with respect to epimorphisms. |
| 21 | +
|
| 22 | +## Main results |
| 23 | +
|
| 24 | +Besides the definition, we show that |
| 25 | +* the composition of two strong epimorphisms is a strong epimorphism, |
| 26 | +* if `f ≫ g` is a strong epimorphism, then so is `g`, |
| 27 | +* if `f` is both a strong epimorphism and a monomorphism, then it is an isomorphism |
| 28 | +
|
| 29 | +We also define classes `strong_mono_category` and `strong_epi_category` for categories in which |
| 30 | +every monomorphism or epimorphism is strong, and deduce that these categories are balanced. |
| 31 | +
|
| 32 | +## TODO |
| 33 | +
|
| 34 | +Show that the dual of a strong epimorphism is a strong monomorphism, and vice versa. |
| 35 | +
|
| 36 | +## References |
| 37 | +
|
| 38 | +* [F. Borceux, *Handbook of Categorical Algebra 1*][borceux-vol1] |
| 39 | +-/ |
| 40 | + |
| 41 | + |
| 42 | +universe v u |
| 43 | + |
| 44 | +namespace CategoryTheory |
| 45 | + |
| 46 | +variable {C : Type u} [Category.{v} C] |
| 47 | + |
| 48 | +variable {P Q : C} |
| 49 | + |
| 50 | +/-- A strong epimorphism `f` is an epimorphism which has the left lifting property |
| 51 | +with respect to monomorphisms. -/ |
| 52 | +class StrongEpi (f : P ⟶ Q) : Prop where |
| 53 | + /-- The epimorphism condition on `f` -/ |
| 54 | + epi : Epi f |
| 55 | + /-- The left lifting property with respect to all monomorphism -/ |
| 56 | + llp : ∀ ⦃X Y : C⦄ (z : X ⟶ Y) [Mono z], HasLiftingProperty f z |
| 57 | +#align category_theory.strong_epi CategoryTheory.StrongEpi |
| 58 | +#align category_theory.strong_epi.epi CategoryTheory.StrongEpi.epi |
| 59 | + |
| 60 | + |
| 61 | +theorem StrongEpi.mk' {f : P ⟶ Q} [Epi f] |
| 62 | + (hf : ∀ (X Y : C) (z : X ⟶ Y) |
| 63 | + (_ : Mono z) (u : P ⟶ X) (v : Q ⟶ Y) (sq : CommSq u f z v), sq.HasLift) : |
| 64 | + StrongEpi f := |
| 65 | + { epi := inferInstance |
| 66 | + llp := fun {X} {Y} z hz => ⟨fun {u} {v} sq => hf X Y z hz u v sq⟩ } |
| 67 | +#align category_theory.strong_epi.mk' CategoryTheory.StrongEpi.mk' |
| 68 | + |
| 69 | +/-- A strong monomorphism `f` is a monomorphism which has the right lifting property |
| 70 | +with respect to epimorphisms. -/ |
| 71 | +class StrongMono (f : P ⟶ Q) : Prop where |
| 72 | + /-- The monomorphism condition on `f` -/ |
| 73 | + mono : Mono f |
| 74 | + /-- The right lifting property with respect to all epimorphisms -/ |
| 75 | + rlp : ∀ ⦃X Y : C⦄ (z : X ⟶ Y) [Epi z], HasLiftingProperty z f |
| 76 | +#align category_theory.strong_mono CategoryTheory.StrongMono |
| 77 | + |
| 78 | +theorem StrongMono.mk' {f : P ⟶ Q} [Mono f] |
| 79 | + (hf : ∀ (X Y : C) (z : X ⟶ Y) (_ : Epi z) (u : X ⟶ P) |
| 80 | + (v : Y ⟶ Q) (sq : CommSq u z f v), sq.HasLift) : StrongMono f where |
| 81 | + mono := inferInstance |
| 82 | + rlp := fun {X} {Y} z hz => ⟨fun {u} {v} sq => hf X Y z hz u v sq⟩ |
| 83 | +#align category_theory.strong_mono.mk' CategoryTheory.StrongMono.mk' |
| 84 | + |
| 85 | +attribute [instance] StrongEpi.llp |
| 86 | + |
| 87 | +attribute [instance] StrongMono.rlp |
| 88 | + |
| 89 | +instance (priority := 100) epi_of_strongEpi (f : P ⟶ Q) [StrongEpi f] : Epi f := |
| 90 | + StrongEpi.epi |
| 91 | +#align category_theory.epi_of_strong_epi CategoryTheory.epi_of_strongEpi |
| 92 | + |
| 93 | +instance (priority := 100) mono_of_strongMono (f : P ⟶ Q) [StrongMono f] : Mono f := |
| 94 | + StrongMono.mono |
| 95 | +#align category_theory.mono_of_strong_mono CategoryTheory.mono_of_strongMono |
| 96 | + |
| 97 | +section |
| 98 | + |
| 99 | +variable {R : C} (f : P ⟶ Q) (g : Q ⟶ R) |
| 100 | + |
| 101 | +/-- The composition of two strong epimorphisms is a strong epimorphism. -/ |
| 102 | +theorem strongEpi_comp [StrongEpi f] [StrongEpi g] : StrongEpi (f ≫ g) := |
| 103 | + { epi := epi_comp _ _ |
| 104 | + llp := by |
| 105 | + intros |
| 106 | + infer_instance } |
| 107 | +#align category_theory.strong_epi_comp CategoryTheory.strongEpi_comp |
| 108 | + |
| 109 | +/-- The composition of two strong monomorphisms is a strong monomorphism. -/ |
| 110 | +theorem strongMono_comp [StrongMono f] [StrongMono g] : StrongMono (f ≫ g) := |
| 111 | + { mono := mono_comp _ _ |
| 112 | + rlp := by |
| 113 | + intros |
| 114 | + infer_instance } |
| 115 | +#align category_theory.strong_mono_comp CategoryTheory.strongMono_comp |
| 116 | + |
| 117 | +/-- If `f ≫ g` is a strong epimorphism, then so is `g`. -/ |
| 118 | +theorem strongEpi_of_strongEpi [StrongEpi (f ≫ g)] : StrongEpi g := |
| 119 | + { epi := epi_of_epi f g |
| 120 | + llp := fun {X} {Y} z _ => by |
| 121 | + constructor |
| 122 | + intro u v sq |
| 123 | + have h₀ : (f ≫ u) ≫ z = (f ≫ g) ≫ v := by simp only [Category.assoc, sq.w] |
| 124 | + exact |
| 125 | + CommSq.HasLift.mk' |
| 126 | + ⟨(CommSq.mk h₀).lift, by |
| 127 | + simp only [← cancel_mono z, Category.assoc, CommSq.fac_right, sq.w], by simp⟩ } |
| 128 | +#align category_theory.strong_epi_of_strong_epi CategoryTheory.strongEpi_of_strongEpi |
| 129 | + |
| 130 | +/-- If `f ≫ g` is a strong monomorphism, then so is `f`. -/ |
| 131 | +theorem strongMono_of_strongMono [StrongMono (f ≫ g)] : StrongMono f := |
| 132 | + { mono := mono_of_mono f g |
| 133 | + rlp := fun {X} {Y} z => by |
| 134 | + intros |
| 135 | + constructor |
| 136 | + intro u v sq |
| 137 | + have h₀ : u ≫ f ≫ g = z ≫ v ≫ g := by |
| 138 | + rw [←Category.assoc, eq_whisker sq.w, Category.assoc] |
| 139 | + exact CommSq.HasLift.mk' ⟨(CommSq.mk h₀).lift, by simp, by simp [← cancel_epi z, sq.w]⟩ } |
| 140 | +#align category_theory.strong_mono_of_strong_mono CategoryTheory.strongMono_of_strongMono |
| 141 | + |
| 142 | +/-- An isomorphism is in particular a strong epimorphism. -/ |
| 143 | +instance (priority := 100) strongEpi_of_isIso [IsIso f] : StrongEpi f |
| 144 | + where |
| 145 | + epi := by infer_instance |
| 146 | + llp {X} {Y} z := HasLiftingProperty.of_left_iso _ _ |
| 147 | +#align category_theory.strong_epi_of_is_iso CategoryTheory.strongEpi_of_isIso |
| 148 | + |
| 149 | +/-- An isomorphism is in particular a strong monomorphism. -/ |
| 150 | +instance (priority := 100) strongMono_of_isIso [IsIso f] : StrongMono f |
| 151 | + where |
| 152 | + mono := by infer_instance |
| 153 | + rlp {X} {Y} z := HasLiftingProperty.of_right_iso _ _ |
| 154 | +#align category_theory.strong_mono_of_is_iso CategoryTheory.strongMono_of_isIso |
| 155 | + |
| 156 | +theorem StrongEpi.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} |
| 157 | + (e : Arrow.mk f ≅ Arrow.mk g) [h : StrongEpi f] : StrongEpi g := |
| 158 | + { epi := by |
| 159 | + rw [Arrow.iso_w' e] |
| 160 | + haveI := epi_comp f e.hom.right |
| 161 | + apply epi_comp |
| 162 | + llp := fun {X} {Y} z => by |
| 163 | + intro |
| 164 | + apply HasLiftingProperty.of_arrow_iso_left e z } |
| 165 | +#align category_theory.strong_epi.of_arrow_iso CategoryTheory.StrongEpi.of_arrow_iso |
| 166 | + |
| 167 | +theorem StrongMono.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} |
| 168 | + (e : Arrow.mk f ≅ Arrow.mk g) [h : StrongMono f] : StrongMono g := |
| 169 | + { mono := by |
| 170 | + rw [Arrow.iso_w' e] |
| 171 | + haveI := mono_comp f e.hom.right |
| 172 | + apply mono_comp |
| 173 | + rlp := fun {X} {Y} z => by |
| 174 | + intro |
| 175 | + apply HasLiftingProperty.of_arrow_iso_right z e } |
| 176 | +#align category_theory.strong_mono.of_arrow_iso CategoryTheory.StrongMono.of_arrow_iso |
| 177 | + |
| 178 | +theorem StrongEpi.iff_of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} |
| 179 | + (e : Arrow.mk f ≅ Arrow.mk g) : StrongEpi f ↔ StrongEpi g := by |
| 180 | + constructor <;> intro |
| 181 | + exacts[StrongEpi.of_arrow_iso e, StrongEpi.of_arrow_iso e.symm] |
| 182 | +#align category_theory.strong_epi.iff_of_arrow_iso CategoryTheory.StrongEpi.iff_of_arrow_iso |
| 183 | + |
| 184 | +theorem StrongMono.iff_of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} |
| 185 | + (e : Arrow.mk f ≅ Arrow.mk g) : StrongMono f ↔ StrongMono g := by |
| 186 | + constructor <;> intro |
| 187 | + exacts[StrongMono.of_arrow_iso e, StrongMono.of_arrow_iso e.symm] |
| 188 | +#align category_theory.strong_mono.iff_of_arrow_iso CategoryTheory.StrongMono.iff_of_arrow_iso |
| 189 | + |
| 190 | +end |
| 191 | + |
| 192 | +/-- A strong epimorphism that is a monomorphism is an isomorphism. -/ |
| 193 | +theorem isIso_of_mono_of_strongEpi (f : P ⟶ Q) [Mono f] [StrongEpi f] : IsIso f := |
| 194 | + ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by aesop_cat⟩⟩ |
| 195 | +#align category_theory.is_iso_of_mono_of_strong_epi CategoryTheory.isIso_of_mono_of_strongEpi |
| 196 | + |
| 197 | +/-- A strong monomorphism that is an epimorphism is an isomorphism. -/ |
| 198 | +theorem isIso_of_epi_of_strongMono (f : P ⟶ Q) [Epi f] [StrongMono f] : IsIso f := |
| 199 | + ⟨⟨(CommSq.mk (show 𝟙 P ≫ f = f ≫ 𝟙 Q by simp)).lift, by aesop_cat⟩⟩ |
| 200 | +#align category_theory.is_iso_of_epi_of_strong_mono CategoryTheory.isIso_of_epi_of_strongMono |
| 201 | + |
| 202 | +section |
| 203 | + |
| 204 | +variable (C) |
| 205 | + |
| 206 | +/-- A strong epi category is a category in which every epimorphism is strong. -/ |
| 207 | +class StrongEpiCategory : Prop where |
| 208 | + /-- A strong epi category is a category in which every epimorphism is strong. -/ |
| 209 | + strongEpi_of_epi : ∀ {X Y : C} (f : X ⟶ Y) [Epi f], StrongEpi f |
| 210 | +#align category_theory.strong_epi_category CategoryTheory.StrongEpiCategory |
| 211 | +#align category_theory.strong_epi_category.strong_epi_of_epi CategoryTheory.StrongEpiCategory.strongEpi_of_epi |
| 212 | + |
| 213 | +/-- A strong mono category is a category in which every monomorphism is strong. -/ |
| 214 | +class StrongMonoCategory : Prop where |
| 215 | + /-- A strong mono category is a category in which every monomorphism is strong. -/ |
| 216 | + strongMono_of_mono : ∀ {X Y : C} (f : X ⟶ Y) [Mono f], StrongMono f |
| 217 | +#align category_theory.strong_mono_category CategoryTheory.StrongMonoCategory |
| 218 | +#align category_theory.strong_mono_category.strong_mono_of_mono CategoryTheory.StrongMonoCategory.strongMono_of_mono |
| 219 | + |
| 220 | +end |
| 221 | + |
| 222 | +theorem strongEpi_of_epi [StrongEpiCategory C] (f : P ⟶ Q) [Epi f] : StrongEpi f := |
| 223 | + StrongEpiCategory.strongEpi_of_epi _ |
| 224 | +#align category_theory.strong_epi_of_epi CategoryTheory.strongEpi_of_epi |
| 225 | + |
| 226 | +theorem strongMono_of_mono [StrongMonoCategory C] (f : P ⟶ Q) [Mono f] : StrongMono f := |
| 227 | + StrongMonoCategory.strongMono_of_mono _ |
| 228 | +#align category_theory.strong_mono_of_mono CategoryTheory.strongMono_of_mono |
| 229 | + |
| 230 | +section |
| 231 | + |
| 232 | +attribute [local instance] strongEpi_of_epi |
| 233 | + |
| 234 | +instance (priority := 100) balanced_of_strongEpiCategory [StrongEpiCategory C] : Balanced C |
| 235 | + where isIso_of_mono_of_epi _ _ _ := isIso_of_mono_of_strongEpi _ |
| 236 | +#align category_theory.balanced_of_strong_epi_category CategoryTheory.balanced_of_strongEpiCategory |
| 237 | + |
| 238 | +end |
| 239 | + |
| 240 | +section |
| 241 | + |
| 242 | +attribute [local instance] strongMono_of_mono |
| 243 | + |
| 244 | +instance (priority := 100) balanced_of_strongMonoCategory [StrongMonoCategory C] : Balanced C |
| 245 | + where isIso_of_mono_of_epi _ _ _ := isIso_of_epi_of_strongMono _ |
| 246 | +#align category_theory.balanced_of_strong_mono_category CategoryTheory.balanced_of_strongMonoCategory |
| 247 | + |
| 248 | +end |
| 249 | + |
| 250 | +end CategoryTheory |
| 251 | +#lint |
0 commit comments