Skip to content

Commit f972435

Browse files
committed
feat: the abelian group structure on Ext-groups in abelian categories (#14496)
In this PR, we construct the abelian group structure on `Ext`-groups. Co-authored-by: Joël Riou <[email protected]>
1 parent dccac13 commit f972435

File tree

3 files changed

+229
-2
lines changed

3 files changed

+229
-2
lines changed

Mathlib/Algebra/Homology/DerivedCategory/Ext.lean

+173-2
Original file line numberDiff line numberDiff line change
@@ -35,7 +35,6 @@ Then, for `C := Sheaf X.etale AddCommGroupCat.{u}`, we will have
3535
sheaves over `X` shall be in `Type u`.
3636
3737
## TODO
38-
* construct the additive structure on `Ext`
3938
* compute `Ext X Y 0`
4039
* define the class in `Ext S.X₃ S.X₁ 1` of a short exact short complex `S`
4140
* construct the long exact sequences of `Ext`.
@@ -48,7 +47,7 @@ namespace CategoryTheory
4847

4948
variable (C : Type u) [Category.{v} C] [Abelian C]
5049

51-
open Localization
50+
open Localization Limits ZeroObject
5251

5352
/-- The property that morphisms between single complexes in arbitrary degrees are `w`-small
5453
in the derived category. -/
@@ -130,6 +129,178 @@ lemma comp_hom {a b : ℕ} (α : Ext X Y a) (β : Ext Y Z b) {c : ℕ} (h : a +
130129
lemma ext {n : ℕ} {α β : Ext X Y n} (h : α.hom = β.hom) : α = β :=
131130
homEquiv.injective h
132131

132+
lemma ext_iff {n : ℕ} {α β : Ext X Y n} : α = β ↔ α.hom = β.hom :=
133+
fun h ↦ by rw [h], ext⟩
134+
135+
end
136+
137+
/-- The canonical map `(X ⟶ Y) → Ext X Y 0`. -/
138+
noncomputable def mk₀ (f : X ⟶ Y) : Ext X Y 0 := SmallShiftedHom.mk₀ _ _ (by simp)
139+
((CochainComplex.singleFunctor C 0).map f)
140+
141+
@[simp]
142+
lemma mk₀_hom [HasDerivedCategory.{w'} C] (f : X ⟶ Y) :
143+
(mk₀ f).hom = ShiftedHom.mk₀ _ (by simp) ((DerivedCategory.singleFunctor C 0).map f) := by
144+
apply SmallShiftedHom.equiv_mk₀
145+
146+
@[simp 1100]
147+
lemma mk₀_comp_mk₀ (f : X ⟶ Y) (g : Y ⟶ Z) :
148+
(mk₀ f).comp (mk₀ g) (zero_add 0) = mk₀ (f ≫ g) := by
149+
letI := HasDerivedCategory.standard C; ext; simp
150+
151+
@[simp 1100]
152+
lemma mk₀_comp_mk₀_assoc (f : X ⟶ Y) (g : Y ⟶ Z) {n : ℕ} (α : Ext Z T n) :
153+
(mk₀ f).comp ((mk₀ g).comp α (zero_add n)) (zero_add n) =
154+
(mk₀ (f ≫ g)).comp α (zero_add n) := by
155+
rw [← mk₀_comp_mk₀, comp_assoc]
156+
omega
157+
158+
variable {n : ℕ}
159+
160+
/-! The abelian group structure on `Ext X Y n` is defined by transporting the
161+
abelian group structure on the constructed derived category
162+
(given by `HasDerivedCategory.standard`). This constructed derived category
163+
is used in order to obtain most of the compatibilities satisfied by
164+
this abelian group structure. It is then shown that the bijection
165+
`homEquiv` between `Ext X Y n` and Hom-types in the derived category
166+
can be promoted to an additive equivalence for any `[HasDerivedCategory C]` instance. -/
167+
168+
noncomputable instance : AddCommGroup (Ext X Y n) :=
169+
letI := HasDerivedCategory.standard C
170+
homEquiv.addCommGroup
171+
172+
/-- The map from `Ext X Y n` to a `ShiftedHom` type in the *constructed* derived
173+
category given by `HasDerivedCategory.standard`: this definition is introduced
174+
only in order to prove properties of the abelian group structure on `Ext`-groups.
175+
Do not use this definition: use the more general `hom` instead. -/
176+
noncomputable abbrev hom' (α : Ext X Y n) :
177+
letI := HasDerivedCategory.standard C
178+
ShiftedHom ((DerivedCategory.singleFunctor C 0).obj X)
179+
((DerivedCategory.singleFunctor C 0).obj Y) (n : ℤ) :=
180+
letI := HasDerivedCategory.standard C
181+
α.hom
182+
183+
private lemma add_hom' (α β : Ext X Y n) : (α + β).hom' = α.hom' + β.hom' :=
184+
letI := HasDerivedCategory.standard C
185+
homEquiv.symm.injective (Equiv.symm_apply_apply _ _)
186+
187+
private lemma neg_hom' (α : Ext X Y n) : (-α).hom' = -α.hom' :=
188+
letI := HasDerivedCategory.standard C
189+
homEquiv.symm.injective (Equiv.symm_apply_apply _ _)
190+
191+
variable (X Y n) in
192+
private lemma zero_hom' : (0 : Ext X Y n).hom' = 0 :=
193+
letI := HasDerivedCategory.standard C
194+
homEquiv.symm.injective (Equiv.symm_apply_apply _ _)
195+
196+
@[simp]
197+
lemma add_comp (α₁ α₂ : Ext X Y n) {m : ℕ} (β : Ext Y Z m) {p : ℕ} (h : n + m = p) :
198+
(α₁ + α₂).comp β h = α₁.comp β h + α₂.comp β h := by
199+
letI := HasDerivedCategory.standard C; ext; simp [add_hom']
200+
201+
@[simp]
202+
lemma comp_add (α : Ext X Y n) {m : ℕ} (β₁ β₂ : Ext Y Z m) {p : ℕ} (h : n + m = p) :
203+
α.comp (β₁ + β₂) h = α.comp β₁ h + α.comp β₂ h := by
204+
letI := HasDerivedCategory.standard C; ext; simp [add_hom']
205+
206+
@[simp]
207+
lemma neg_comp (α : Ext X Y n) {m : ℕ} (β : Ext Y Z m) {p : ℕ} (h : n + m = p) :
208+
(-α).comp β h = -α.comp β h := by
209+
letI := HasDerivedCategory.standard C; ext; simp [neg_hom']
210+
211+
@[simp]
212+
lemma comp_neg (α : Ext X Y n) {m : ℕ} (β : Ext Y Z m) {p : ℕ} (h : n + m = p) :
213+
α.comp (-β) h = -α.comp β h := by
214+
letI := HasDerivedCategory.standard C; ext; simp [neg_hom']
215+
216+
variable (X n) in
217+
@[simp]
218+
lemma zero_comp {m : ℕ} (β : Ext Y Z m) (p : ℕ) (h : n + m = p) :
219+
(0 : Ext X Y n).comp β h = 0 := by
220+
letI := HasDerivedCategory.standard C; ext; simp [zero_hom']
221+
222+
@[simp]
223+
lemma comp_zero (α : Ext X Y n) (Z : C) (m : ℕ) (p : ℕ) (h : n + m = p) :
224+
α.comp (0 : Ext Y Z m) h = 0 := by
225+
letI := HasDerivedCategory.standard C; ext; simp [zero_hom']
226+
227+
@[simp]
228+
lemma mk₀_id_comp (α : Ext X Y n) :
229+
(mk₀ (𝟙 X)).comp α (zero_add n) = α := by
230+
letI := HasDerivedCategory.standard C; ext; simp
231+
232+
@[simp]
233+
lemma comp_mk₀_id (α : Ext X Y n) :
234+
α.comp (mk₀ (𝟙 Y)) (add_zero n) = α := by
235+
letI := HasDerivedCategory.standard C; ext; simp
236+
237+
variable (X Y) in
238+
@[simp]
239+
lemma mk₀_zero : mk₀ (0 : X ⟶ Y) = 0 := by
240+
letI := HasDerivedCategory.standard C; ext; simp [zero_hom']
241+
242+
section
243+
244+
variable [HasDerivedCategory.{w'} C]
245+
246+
variable (X Y n) in
247+
@[simp]
248+
lemma zero_hom : (0 : Ext X Y n).hom = 0 := by
249+
let β : Ext 0 Y n := 0
250+
have hβ : β.hom = 0 := by apply (Functor.map_isZero _ (isZero_zero C)).eq_of_src
251+
have : (0 : Ext X Y n) = (0 : Ext X 0 0).comp β (zero_add n) := by simp [β]
252+
rw [this, comp_hom, hβ, ShiftedHom.comp_zero]
253+
254+
attribute [local instance] preservesBinaryBiproductsOfPreservesBiproducts
255+
256+
lemma biprod_ext {X₁ X₂ : C} {α β : Ext (X₁ ⊞ X₂) Y n}
257+
(h₁ : (mk₀ biprod.inl).comp α (zero_add n) = (mk₀ biprod.inl).comp β (zero_add n))
258+
(h₂ : (mk₀ biprod.inr).comp α (zero_add n) = (mk₀ biprod.inr).comp β (zero_add n)) :
259+
α = β := by
260+
letI := HasDerivedCategory.standard C
261+
rw [ext_iff] at h₁ h₂ ⊢
262+
simp only [comp_hom, mk₀_hom, ShiftedHom.mk₀_comp] at h₁ h₂
263+
apply BinaryCofan.IsColimit.hom_ext
264+
(isBinaryBilimitOfPreserves (DerivedCategory.singleFunctor C 0)
265+
(BinaryBiproduct.isBilimit X₁ X₂)).isColimit
266+
all_goals assumption
267+
268+
@[simp]
269+
lemma add_hom (α β : Ext X Y n) : (α + β).hom = α.hom + β.hom := by
270+
let α' : Ext (X ⊞ X) Y n := (mk₀ biprod.fst).comp α (zero_add n)
271+
let β' : Ext (X ⊞ X) Y n := (mk₀ biprod.snd).comp β (zero_add n)
272+
have eq₁ : α + β = (mk₀ (biprod.lift (𝟙 X) (𝟙 X))).comp (α' + β') (zero_add n) := by
273+
simp [α', β']
274+
have eq₂ : α' + β' = homEquiv.symm (α'.hom + β'.hom) := by
275+
apply biprod_ext
276+
all_goals ext; simp [α', β', ← Functor.map_comp]
277+
simp only [eq₁, eq₂, comp_hom, Equiv.apply_symm_apply, ShiftedHom.comp_add]
278+
congr
279+
· dsimp [α']
280+
rw [comp_hom, mk₀_hom, mk₀_hom]
281+
dsimp
282+
rw [ShiftedHom.mk₀_comp_mk₀_assoc, ← Functor.map_comp,
283+
biprod.lift_fst, Functor.map_id, ShiftedHom.mk₀_id_comp]
284+
· dsimp [β']
285+
rw [comp_hom, mk₀_hom, mk₀_hom]
286+
dsimp
287+
rw [ShiftedHom.mk₀_comp_mk₀_assoc, ← Functor.map_comp,
288+
biprod.lift_snd, Functor.map_id, ShiftedHom.mk₀_id_comp]
289+
290+
lemma neg_hom (α : Ext X Y n) : (-α).hom = -α.hom := by
291+
rw [← add_right_inj α.hom, ← add_hom, add_right_neg, add_right_neg, zero_hom]
292+
293+
/-- When an instance of `[HasDerivedCategory.{w'} C]` is available, this is the additive
294+
bijection between `Ext.{w} X Y n` and a type of morphisms in the derived category. -/
295+
noncomputable def homAddEquiv {n : ℕ} :
296+
Ext.{w} X Y n ≃+ ShiftedHom ((DerivedCategory.singleFunctor C 0).obj X)
297+
((DerivedCategory.singleFunctor C 0).obj Y) (n : ℤ) where
298+
toEquiv := homEquiv
299+
map_add' := by simp
300+
301+
@[simp]
302+
lemma homAddEquiv_apply (α : Ext X Y n) : homAddEquiv α = α.hom := rfl
303+
133304
end
134305

135306
end Ext

Mathlib/CategoryTheory/Localization/SmallShiftedHom.lean

+20
Original file line numberDiff line numberDiff line change
@@ -130,6 +130,14 @@ noncomputable def comp {a b c : M} [HasSmallLocalizedShiftedHom.{w} W M X Y]
130130
SmallShiftedHom.{w} W X Z c :=
131131
SmallHom.comp f (g.shift a c h)
132132

133+
variable (W) in
134+
/-- The canonical map `(X ⟶ Y) → SmallShiftedHom.{w} W X Y m₀` when `m₀ = 0` and
135+
`[HasSmallLocalizedShiftedHom.{w} W M X Y]` holds. -/
136+
noncomputable def mk₀ [HasSmallLocalizedShiftedHom.{w} W M X Y]
137+
(m₀ : M) (hm₀ : m₀ = 0) (f : X ⟶ Y) :
138+
SmallShiftedHom.{w} W X Y m₀ :=
139+
SmallHom.mk W (f ≫ (shiftFunctorZero' C m₀ hm₀).inv.app Y)
140+
133141
end
134142

135143
section
@@ -176,6 +184,18 @@ lemma equiv_comp [HasSmallLocalizedShiftedHom.{w} W M X Y]
176184
comp_id, Functor.map_comp]
177185
rfl
178186

187+
@[simp]
188+
lemma equiv_mk₀ [HasSmallLocalizedShiftedHom.{w} W M X Y]
189+
(m₀ : M) (hm₀ : m₀ = 0) (f : X ⟶ Y) :
190+
equiv W L (SmallShiftedHom.mk₀ W m₀ hm₀ f) =
191+
ShiftedHom.mk₀ m₀ hm₀ (L.map f) := by
192+
subst hm₀
193+
dsimp [equiv, mk₀]
194+
erw [SmallHom.equiv_mk, Iso.homToEquiv_apply, Functor.map_comp]
195+
dsimp [equiv, mk₀, ShiftedHom.mk₀, shiftFunctorZero']
196+
simp only [comp_id, L.commShiftIso_zero, Functor.CommShift.isoZero_hom_app, assoc,
197+
← Functor.map_comp_assoc, Iso.inv_hom_id_app, Functor.id_obj, Functor.map_id, id_comp]
198+
179199
end
180200

181201
lemma comp_assoc {X Y Z T : C} {a₁ a₂ a₃ a₁₂ a₂₃ a : M}

Mathlib/CategoryTheory/Shift/ShiftedHom.lean

+36
Original file line numberDiff line numberDiff line change
@@ -92,10 +92,32 @@ lemma comp_mk₀_id {a : M} (f : ShiftedHom X Y a) (m₀ : M) (hm₀ : m₀ = 0)
9292
f.comp (mk₀ m₀ hm₀ (𝟙 Y)) (by rw [hm₀, zero_add]) = f := by
9393
simp [comp_mk₀]
9494

95+
@[simp 1100]
96+
lemma mk₀_comp_mk₀ (f : X ⟶ Y) (g : Y ⟶ Z) {a b c : M} (h : b + a = c)
97+
(ha : a = 0) (hb : b = 0) :
98+
(mk₀ a ha f).comp (mk₀ b hb g) h = mk₀ c (by rw [← h, ha, hb, add_zero]) (f ≫ g) := by
99+
subst ha hb
100+
obtain rfl : c = 0 := by rw [← h, zero_add]
101+
rw [mk₀_comp, mk₀, mk₀, assoc]
102+
103+
@[simp]
104+
lemma mk₀_comp_mk₀_assoc (f : X ⟶ Y) (g : Y ⟶ Z) {a : M}
105+
(ha : a = 0) {d : M} (h : ShiftedHom Z T d) :
106+
(mk₀ a ha f).comp ((mk₀ a ha g).comp h
107+
(show _ = d by rw [ha, add_zero])) (show _ = d by rw [ha, add_zero]) =
108+
(mk₀ a ha (f ≫ g)).comp h (by rw [ha, add_zero]) := by
109+
subst ha
110+
rw [← comp_assoc, mk₀_comp_mk₀]
111+
all_goals simp
112+
95113
section Preadditive
96114

97115
variable [Preadditive C]
98116

117+
variable (X Y) in
118+
@[simp]
119+
lemma mk₀_zero (m₀ : M) (hm₀ : m₀ = 0) : mk₀ m₀ hm₀ (0 : X ⟶ Y) = 0 := by simp [mk₀]
120+
99121
@[simp]
100122
lemma comp_add [∀ (a : M), (shiftFunctor C a).Additive]
101123
{a b c : M} (α : ShiftedHom X Y a) (β₁ β₂ : ShiftedHom Y Z b) (h : b + a = c) :
@@ -108,13 +130,27 @@ lemma add_comp
108130
(α₁ + α₂).comp β h = α₁.comp β h + α₂.comp β h := by
109131
rw [comp, comp, comp, Preadditive.add_comp]
110132

133+
@[simp]
134+
lemma comp_neg [∀ (a : M), (shiftFunctor C a).Additive]
135+
{a b c : M} (α : ShiftedHom X Y a) (β : ShiftedHom Y Z b) (h : b + a = c) :
136+
α.comp (-β) h = -α.comp β h := by
137+
rw [comp, comp, Functor.map_neg, Preadditive.neg_comp, Preadditive.comp_neg]
138+
139+
@[simp]
140+
lemma neg_comp
141+
{a b c : M} (α : ShiftedHom X Y a) (β : ShiftedHom Y Z b) (h : b + a = c) :
142+
(-α).comp β h = -α.comp β h := by
143+
rw [comp, comp, Preadditive.neg_comp]
144+
111145
variable (Z) in
146+
@[simp]
112147
lemma comp_zero [∀ (a : M), (shiftFunctor C a).PreservesZeroMorphisms]
113148
{a : M} (β : ShiftedHom X Y a) {b c : M} (h : b + a = c) :
114149
β.comp (0 : ShiftedHom Y Z b) h = 0 := by
115150
rw [comp, Functor.map_zero, Limits.zero_comp, Limits.comp_zero]
116151

117152
variable (X) in
153+
@[simp]
118154
lemma zero_comp (a : M) {b c : M} (β : ShiftedHom Y Z b) (h : b + a = c) :
119155
(0 : ShiftedHom X Y a).comp β h = 0 := by
120156
rw [comp, Limits.zero_comp]

0 commit comments

Comments
 (0)