@@ -35,7 +35,6 @@ Then, for `C := Sheaf X.etale AddCommGroupCat.{u}`, we will have
35
35
sheaves over `X` shall be in `Type u`.
36
36
37
37
## TODO
38
- * construct the additive structure on `Ext`
39
38
* compute `Ext X Y 0`
40
39
* define the class in `Ext S.X₃ S.X₁ 1` of a short exact short complex `S`
41
40
* construct the long exact sequences of `Ext`.
@@ -48,7 +47,7 @@ namespace CategoryTheory
48
47
49
48
variable (C : Type u) [Category.{v} C] [Abelian C]
50
49
51
- open Localization
50
+ open Localization Limits ZeroObject
52
51
53
52
/-- The property that morphisms between single complexes in arbitrary degrees are `w`-small
54
53
in the derived category. -/
@@ -130,6 +129,178 @@ lemma comp_hom {a b : ℕ} (α : Ext X Y a) (β : Ext Y Z b) {c : ℕ} (h : a +
130
129
lemma ext {n : ℕ} {α β : Ext X Y n} (h : α.hom = β.hom) : α = β :=
131
130
homEquiv.injective h
132
131
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
+
133
304
end
134
305
135
306
end Ext
0 commit comments