Skip to content

Commit f9a3323

Browse files
committed
feat(CategoryTheory): monomorphisms in Type are stable under coproducts (#32515)
We develop the API regarding colimit cofans in the category of types, and we apply it in order to show that monomorphisms are stable under coproducts in the category of types. From https://github.com/joelriou/topcat-model-category
1 parent c56c733 commit f9a3323

File tree

4 files changed

+196
-3
lines changed

4 files changed

+196
-3
lines changed

Mathlib/CategoryTheory/Limits/Shapes/Products.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -322,6 +322,10 @@ from a family of morphisms between the factors.
322322
abbrev Pi.map {f g : β → C} [HasProduct f] [HasProduct g] (p : ∀ b, f b ⟶ g b) : ∏ᶜ f ⟶ ∏ᶜ g :=
323323
limMap (Discrete.natTrans fun X => p X.as)
324324

325+
@[reassoc (attr := simp high)]
326+
lemma Pi.map_π {f g : β → C} [HasProduct f] [HasProduct g] (p : ∀ b, f b ⟶ g b) (b : β) :
327+
Pi.map p ≫ Pi.π g b = Pi.π f b ≫ p b := by simp
328+
325329
@[simp]
326330
lemma Pi.map_id {f : α → C} [HasProduct f] : Pi.map (fun a => 𝟙 (f a)) = 𝟙 (∏ᶜ f) := by
327331
ext; simp
@@ -437,6 +441,10 @@ abbrev Sigma.map {f g : β → C} [HasCoproduct f] [HasCoproduct g] (p : ∀ b,
437441
∐ f ⟶ ∐ g :=
438442
colimMap (Discrete.natTrans fun X => p X.as)
439443

444+
@[reassoc (attr := simp high)]
445+
lemma Sigma.ι_map {f g : β → C} [HasCoproduct f] [HasCoproduct g] (p : ∀ b, f b ⟶ g b) (b : β) :
446+
Sigma.ι f b ≫ Sigma.map p = p b ≫ Sigma.ι g b:= by simp
447+
440448
@[simp]
441449
lemma Sigma.map_id {f : α → C} [HasCoproduct f] : Sigma.map (fun a => 𝟙 (f a)) = 𝟙 (∐ f) := by
442450
ext; simp

Mathlib/CategoryTheory/Limits/Types/ColimitType.lean

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -224,6 +224,17 @@ lemma of_equiv {c' : CoconeTypes.{w₂} F} (e : c.pt ≃ c'.pt)
224224
obtain ⟨j, x, rfl⟩ := F.ιColimitType_jointly_surjective y
225225
simp_all
226226

227+
lemma iff_bijective {c' : CoconeTypes.{w₂} F}
228+
(f : c.pt → c'.pt) (hf : ∀ j x, c'.ι j x = f (c.ι j x)) :
229+
c'.IsColimit ↔ Function.Bijective f := by
230+
refine ⟨fun hc' ↦ ?_, fun h ↦ hc.of_equiv (Equiv.ofBijective _ h) hf⟩
231+
have h₁ := hc.bijective
232+
rw [← Function.Bijective.of_comp_iff _ hc.bijective]
233+
convert hc'.bijective
234+
ext x
235+
obtain ⟨j, x, rfl⟩ := F.ιColimitType_jointly_surjective x
236+
simp [hf]
237+
227238
end IsColimit
228239

229240
/-- Structure which expresses that `c : F.CoconeTypes`

Mathlib/CategoryTheory/Limits/Types/Coproducts.lean

Lines changed: 154 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -22,11 +22,162 @@ Similarly, the binary coproduct of two types `X` and `Y` identifies to
2222

2323
@[expose] public section
2424

25-
universe v u
25+
universe w v u
2626

27-
open CategoryTheory Limits
27+
namespace CategoryTheory
2828

29-
namespace CategoryTheory.Limits.Types
29+
namespace Limits
30+
31+
variable {C : Type u} (F : C → Type v)
32+
33+
/-- Given a functor `F : Discrete C ⥤ Type v`, this is a "cofan" for `F`,
34+
but we allow the point to be in `Type w` for an arbitrary universe `w`. -/
35+
abbrev CofanTypes := Functor.CoconeTypes.{w} (Discrete.functor F)
36+
37+
variable {F}
38+
39+
namespace CofanTypes
40+
41+
/-- The injection map for a cofan of a functor to types. -/
42+
abbrev inj (c : CofanTypes.{w} F) (i : C) : F i → c.pt := c.ι ⟨i⟩
43+
44+
variable (F) in
45+
/-- The cofan given by a sigma type. -/
46+
@[simps]
47+
def sigma : CofanTypes F where
48+
pt := Σ (i : C), F i
49+
ι := fun ⟨i⟩ x ↦ ⟨i, x⟩
50+
ι_naturality := by
51+
rintro ⟨i⟩ ⟨j⟩ f
52+
obtain rfl : i = j := by simpa using Discrete.eq_of_hom f
53+
rfl
54+
55+
@[simp]
56+
lemma sigma_inj (i : C) (x : F i) :
57+
(sigma F).inj i x = ⟨i, x⟩ := rfl
58+
59+
lemma isColimit_mk (c : CofanTypes.{w} F)
60+
(h₁ : ∀ (x : c.pt), ∃ (i : C) (y : F i), c.inj i y = x)
61+
(h₂ : ∀ (i : C), Function.Injective (c.inj i))
62+
(h₃ : ∀ (i j : C) (x : F i) (y : F j), c.inj i x = c.inj j y → i = j) :
63+
Functor.CoconeTypes.IsColimit c where
64+
bijective := by
65+
constructor
66+
· intro x y h
67+
obtain ⟨⟨i⟩, x, rfl⟩ := (Discrete.functor F).ιColimitType_jointly_surjective x
68+
obtain ⟨⟨j⟩, y, rfl⟩ := (Discrete.functor F).ιColimitType_jointly_surjective y
69+
obtain rfl := h₃ _ _ _ _ h
70+
obtain rfl := h₂ _ h
71+
rfl
72+
· intro x
73+
obtain ⟨i, y, rfl⟩ := h₁ x
74+
exact ⟨(Discrete.functor F).ιColimitType ⟨i⟩ y, rfl⟩
75+
76+
variable (F) in
77+
lemma isColimit_sigma : Functor.CoconeTypes.IsColimit (sigma F) :=
78+
isColimit_mk _ (by aesop)
79+
(fun _ _ _ h ↦ by rw [Sigma.ext_iff] at h; simpa using h)
80+
(fun _ _ _ _ h ↦ congr_arg Sigma.fst h)
81+
82+
variable (F) in
83+
/-- Given a cofan of a functor to types, this is a canonical map
84+
from the Sigma type to the point of the cofan. -/
85+
@[simp]
86+
def fromSigma (c : CofanTypes.{w} F) (x : Σ (i : C), F i) : c.pt :=
87+
c.inj x.1 x.2
88+
89+
lemma isColimit_iff_bijective_fromSigma (c : CofanTypes.{w} F) :
90+
c.IsColimit ↔ Function.Bijective c.fromSigma := by
91+
rw [(isColimit_sigma F).iff_bijective]
92+
aesop
93+
94+
section
95+
96+
variable {c : CofanTypes.{w} F} (hc : Functor.CoconeTypes.IsColimit c)
97+
98+
include hc
99+
100+
lemma bijective_fromSigma_of_isColimit :
101+
Function.Bijective c.fromSigma := by
102+
rwa [← isColimit_iff_bijective_fromSigma]
103+
104+
/-- The bijection from the sigma type to the point of a colimit cofan
105+
of a functor to types. -/
106+
noncomputable def equivOfIsColimit :
107+
(Σ (i : C), F i) ≃ c.pt :=
108+
Equiv.ofBijective _ (bijective_fromSigma_of_isColimit hc)
109+
110+
@[simp]
111+
lemma equivOfIsColimit_apply (i : C) (x : F i) :
112+
equivOfIsColimit hc ⟨i, x⟩ = c.inj i x := rfl
113+
114+
@[simp]
115+
lemma equivOfIsColimit_symm_apply (i : C) (x : F i) :
116+
(equivOfIsColimit hc).symm (c.inj i x) = ⟨i, x⟩ :=
117+
(equivOfIsColimit hc).injective (by simp)
118+
119+
lemma inj_jointly_surjective_of_isColimit (x : c.pt) :
120+
∃ (i : C) (y : F i), c.inj i y = x := by
121+
obtain ⟨⟨i⟩, y, rfl⟩ := hc.ι_jointly_surjective x
122+
exact ⟨i, y, rfl⟩
123+
124+
lemma inj_injective_of_isColimit (i : C) :
125+
Function.Injective (c.inj i) := by
126+
intro y₁ y₂ h
127+
simpa using (equivOfIsColimit hc).injective (a₁ := ⟨i, y₁⟩) (a₂ := ⟨i, y₂⟩) h
128+
129+
lemma eq_of_inj_apply_eq_of_isColimit
130+
{i₁ i₂ : C} (y₁ : F i₁) (y₂ : F i₂) (h : c.inj i₁ y₁ = c.inj i₂ y₂) :
131+
i₁ = i₂ :=
132+
congr_arg Sigma.fst ((equivOfIsColimit hc).injective (a₁ := ⟨i₁, y₁⟩) (a₂ := ⟨i₂, y₂⟩) h)
133+
134+
end
135+
136+
end CofanTypes
137+
138+
namespace Cofan
139+
140+
variable {C : Type u} {F : C → Type v} (c : Cofan F)
141+
142+
/-- If `F : C → Type v`, then the data of a "type-theoretic" cofan of `F`
143+
with a point in `Type v` is the same as the data of a cocone (in a categorical sense). -/
144+
@[simps]
145+
def cofanTypes :
146+
CofanTypes.{v} F where
147+
pt := c.pt
148+
ι := fun ⟨j⟩ ↦ c.inj j
149+
ι_naturality := by
150+
rintro ⟨i⟩ ⟨j⟩ f
151+
obtain rfl : i = j := by simpa using Discrete.eq_of_hom f
152+
rfl
153+
154+
lemma isColimit_cofanTypes_iff :
155+
c.cofanTypes.IsColimit ↔ Nonempty (IsColimit c) :=
156+
Functor.CoconeTypes.isColimit_iff _
157+
158+
lemma nonempty_isColimit_iff_bijective_fromSigma :
159+
Nonempty (IsColimit c) ↔ Function.Bijective c.cofanTypes.fromSigma := by
160+
rw [← isColimit_cofanTypes_iff, CofanTypes.isColimit_iff_bijective_fromSigma]
161+
162+
variable {c}
163+
164+
lemma inj_jointly_surjective_of_isColimit (hc : IsColimit c) (x : c.pt) :
165+
∃ (i : C) (y : F i), c.inj i y = x :=
166+
CofanTypes.inj_jointly_surjective_of_isColimit
167+
((isColimit_cofanTypes_iff c).2 ⟨hc⟩) x
168+
169+
lemma inj_injective_of_isColimit (hc : IsColimit c) (i : C) :
170+
Function.Injective (c.inj i) :=
171+
CofanTypes.inj_injective_of_isColimit ((isColimit_cofanTypes_iff c).2 ⟨hc⟩) i
172+
173+
lemma eq_of_inj_apply_eq_of_isColimit (hc : IsColimit c)
174+
{i₁ i₂ : C} (y₁ : F i₁) (y₂ : F i₂) (h : c.inj i₁ y₁ = c.inj i₂ y₂) :
175+
i₁ = i₂ :=
176+
CofanTypes.eq_of_inj_apply_eq_of_isColimit ((isColimit_cofanTypes_iff c).2 ⟨hc⟩) _ _ h
177+
178+
end Cofan
179+
180+
namespace Types
30181

31182
/-- The category of types has `PEmpty` as an initial object. -/
32183
def initialColimitCocone : Limits.ColimitCocone (Functor.empty (Type u)) where

Mathlib/CategoryTheory/Types/Monomorphisms.lean

Lines changed: 23 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ module
88
public import Mathlib.CategoryTheory.Limits.Connected
99
public import Mathlib.CategoryTheory.Limits.Types.Filtered
1010
public import Mathlib.CategoryTheory.Limits.Types.Pushouts
11+
public import Mathlib.CategoryTheory.Limits.Types.Coproducts
1112
public import Mathlib.CategoryTheory.MorphismProperty.Limits
1213

1314
/-!
@@ -54,4 +55,26 @@ instance : MorphismProperty.IsStableUnderFilteredColimits.{v', u'} (monomorphism
5455
simp only [← FunctorToTypes.naturality] at hk
5556
rw [← c₁.w α, types_comp_apply, types_comp_apply, hf _ hk]⟩
5657

58+
instance (T : Type u') : MorphismProperty.IsStableUnderCoproductsOfShape
59+
(monomorphisms (Type u)) T :=
60+
IsStableUnderCoproductsOfShape.mk _ _ (by
61+
intro X₁ X₂ _ _ f h
62+
simp only [monomorphisms.iff, mono_iff_injective] at h ⊢
63+
intro x₁ x₂ hx
64+
obtain ⟨i₁, x₁, rfl⟩ := Cofan.inj_jointly_surjective_of_isColimit
65+
(coproductIsCoproduct X₁) x₁
66+
obtain ⟨i₂, x₂, rfl⟩ := Cofan.inj_jointly_surjective_of_isColimit
67+
(coproductIsCoproduct X₁) x₂
68+
simp only [cofan_mk_inj] at hx ⊢
69+
replace hx : Sigma.ι X₂ i₁ (f i₁ x₁) = Sigma.ι X₂ i₂ (f i₂ x₂) := by
70+
have h₁ := congr_fun (Sigma.ι_map f i₁) x₁
71+
have h₂ := congr_fun (Sigma.ι_map f i₂) x₂
72+
dsimp at h₁ h₂
73+
rw [← h₁, ← h₂, hx]
74+
obtain rfl := Cofan.eq_of_inj_apply_eq_of_isColimit (coproductIsCoproduct X₂) _ _ hx
75+
obtain rfl := h _ (Cofan.inj_injective_of_isColimit (coproductIsCoproduct X₂) i₁ hx)
76+
rfl)
77+
78+
instance : MorphismProperty.IsStableUnderCoproducts (monomorphisms (Type u)) where
79+
5780
end CategoryTheory.Types

0 commit comments

Comments
 (0)