@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Joël Riou
5
5
-/
6
6
import Mathlib.CategoryTheory.GradedObject.Associator
7
- import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor
7
+ import Mathlib.CategoryTheory.Linear.LinearFunctor
8
8
import Mathlib.Algebra.Homology.Bifunctor
9
9
10
10
/-!
@@ -187,6 +187,212 @@ lemma ι_mapBifunctor₁₂Desc (i₁ : ι₁) (i₂ : ι₂) (i₃ : ι₃)
187
187
188
188
end
189
189
190
+ variable (F₁₂ G)
191
+
192
+ /-- The first differential on a summand
193
+ of `mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄`. -/
194
+ noncomputable def d₁ (i₁ : ι₁) (i₂ : ι₂) (i₃ : ι₃) (j : ι₄) :
195
+ (G.obj ((F₁₂.obj (K₁.X i₁)).obj (K₂.X i₂))).obj (K₃.X i₃) ⟶
196
+ (mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄).X j :=
197
+ (ComplexShape.ε₁ c₁₂ c₃ c₄ (ComplexShape.π c₁ c₂ c₁₂ ⟨i₁, i₂⟩, i₃) *
198
+ ComplexShape.ε₁ c₁ c₂ c₁₂ (i₁, i₂)) •
199
+ (G.map ((F₁₂.map (K₁.d i₁ (c₁.next i₁))).app (K₂.X i₂))).app (K₃.X i₃) ≫
200
+ ιOrZero F₁₂ G K₁ K₂ K₃ c₁₂ c₄ _ i₂ i₃ j
201
+
202
+ lemma d₁_eq_zero (i₁ : ι₁) (i₂ : ι₂) (i₃ : ι₃) (j : ι₄) (h : ¬ c₁.Rel i₁ (c₁.next i₁)) :
203
+ d₁ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j = 0 := by
204
+ dsimp [d₁]
205
+ rw [shape _ _ _ h, Functor.map_zero, zero_app, Functor.map_zero, zero_app, zero_comp, smul_zero]
206
+
207
+ lemma d₁_eq {i₁ i₁' : ι₁} (h₁ : c₁.Rel i₁ i₁') (i₂ : ι₂) (i₃ : ι₃) (j : ι₄) :
208
+ d₁ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j =
209
+ (ComplexShape.ε₁ c₁₂ c₃ c₄ (ComplexShape.π c₁ c₂ c₁₂ ⟨i₁, i₂⟩, i₃) *
210
+ ComplexShape.ε₁ c₁ c₂ c₁₂ (i₁, i₂) ) •
211
+ (G.map ((F₁₂.map (K₁.d i₁ i₁')).app (K₂.X i₂))).app (K₃.X i₃) ≫
212
+ ιOrZero F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁' i₂ i₃ j := by
213
+ obtain rfl := c₁.next_eq' h₁
214
+ rfl
215
+
216
+ /-- The second differential on a summand
217
+ of `mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄`. -/
218
+ noncomputable def d₂ (i₁ : ι₁) (i₂ : ι₂) (i₃ : ι₃) (j : ι₄) :
219
+ (G.obj ((F₁₂.obj (K₁.X i₁)).obj (K₂.X i₂))).obj (K₃.X i₃) ⟶
220
+ (mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄).X j :=
221
+ (c₁₂.ε₁ c₃ c₄ (ComplexShape.π c₁ c₂ c₁₂ ⟨i₁, i₂⟩, i₃) * c₁.ε₂ c₂ c₁₂ (i₁, i₂)) •
222
+ (G.map ((F₁₂.obj (K₁.X i₁)).map (K₂.d i₂ (c₂.next i₂)))).app (K₃.X i₃) ≫
223
+ ιOrZero F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ _ i₃ j
224
+
225
+ lemma d₂_eq_zero (i₁ : ι₁) (i₂ : ι₂) (i₃ : ι₃) (j : ι₄) (h : ¬ c₂.Rel i₂ (c₂.next i₂)) :
226
+ d₂ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j = 0 := by
227
+ dsimp [d₂]
228
+ rw [shape _ _ _ h, Functor.map_zero, Functor.map_zero, zero_app, zero_comp, smul_zero]
229
+
230
+ lemma d₂_eq (i₁ : ι₁) {i₂ i₂' : ι₂} (h₂ : c₂.Rel i₂ i₂') (i₃ : ι₃) (j : ι₄) :
231
+ d₂ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j =
232
+ (c₁₂.ε₁ c₃ c₄ (ComplexShape.π c₁ c₂ c₁₂ ⟨i₁, i₂⟩, i₃) * c₁.ε₂ c₂ c₁₂ (i₁, i₂)) •
233
+ (G.map ((F₁₂.obj (K₁.X i₁)).map (K₂.d i₂ i₂'))).app (K₃.X i₃) ≫
234
+ ιOrZero F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ _ i₃ j := by
235
+ obtain rfl := c₂.next_eq' h₂
236
+ rfl
237
+
238
+ /-- The third differential on a summand
239
+ of `mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄`. -/
240
+ noncomputable def d₃ (i₁ : ι₁) (i₂ : ι₂) (i₃ : ι₃) (j : ι₄) :
241
+ (G.obj ((F₁₂.obj (K₁.X i₁)).obj (K₂.X i₂))).obj (K₃.X i₃) ⟶
242
+ (mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄).X j :=
243
+ (ComplexShape.ε₂ c₁₂ c₃ c₄ (c₁.π c₂ c₁₂ (i₁, i₂), i₃)) •
244
+ (G.obj ((F₁₂.obj (K₁.X i₁)).obj (K₂.X i₂))).map (K₃.d i₃ (c₃.next i₃)) ≫
245
+ ιOrZero F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ _ j
246
+
247
+ lemma d₃_eq_zero (i₁ : ι₁) (i₂ : ι₂) (i₃ : ι₃) (j : ι₄) (h : ¬ c₃.Rel i₃ (c₃.next i₃)) :
248
+ d₃ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j = 0 := by
249
+ dsimp [d₃]
250
+ rw [shape _ _ _ h, Functor.map_zero, zero_comp, smul_zero]
251
+
252
+ lemma d₃_eq (i₁ : ι₁) (i₂ : ι₂) {i₃ i₃' : ι₃} (h₃ : c₃.Rel i₃ i₃') (j : ι₄) :
253
+ d₃ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j =
254
+ (ComplexShape.ε₂ c₁₂ c₃ c₄ (c₁.π c₂ c₁₂ (i₁, i₂), i₃)) •
255
+ (G.obj ((F₁₂.obj (K₁.X i₁)).obj (K₂.X i₂))).map (K₃.d i₃ i₃') ≫
256
+ ιOrZero F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ _ j := by
257
+ obtain rfl := c₃.next_eq' h₃
258
+ rfl
259
+
260
+
261
+ section
262
+
263
+ variable [HasGoodTrifunctor₁₂Obj F₁₂ G K₁ K₂ K₃ c₁₂ c₄]
264
+ variable (j j' : ι₄)
265
+
266
+ /-- The first differential on `mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄`. -/
267
+ noncomputable def D₁ :
268
+ (mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄).X j ⟶
269
+ (mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄).X j' :=
270
+ mapBifunctor₁₂Desc (fun i₁ i₂ i₃ _ ↦ d₁ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j')
271
+
272
+ /-- The second differential on `mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄`. -/
273
+ noncomputable def D₂ :
274
+ (mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄).X j ⟶
275
+ (mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄).X j' :=
276
+ mapBifunctor₁₂Desc (fun i₁ i₂ i₃ _ ↦ d₂ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j')
277
+
278
+ /-- The third differential on `mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄`. -/
279
+ noncomputable def D₃ :
280
+ (mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄).X j ⟶
281
+ (mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄).X j' :=
282
+ mapBifunctor.D₂ _ _ _ _ _ _
283
+
284
+ end
285
+
286
+ section
287
+
288
+ variable (i₁ : ι₁) (i₂ : ι₂) (i₃ : ι₃) (j j' : ι₄)
289
+ (h : ComplexShape.r c₁ c₂ c₃ c₁₂ c₄ (i₁, i₂, i₃) = j)
290
+
291
+ @[reassoc (attr := simp)]
292
+ lemma ι_D₁ [HasGoodTrifunctor₁₂Obj F₁₂ G K₁ K₂ K₃ c₁₂ c₄] :
293
+ ι F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j h ≫ D₁ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ j j' =
294
+ d₁ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j' := by
295
+ simp [D₁]
296
+
297
+ @[reassoc (attr := simp)]
298
+ lemma ι_D₂ [HasGoodTrifunctor₁₂Obj F₁₂ G K₁ K₂ K₃ c₁₂ c₄] :
299
+ ι F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j h ≫ D₂ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ j j' =
300
+ d₂ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j' := by
301
+ simp [D₂]
302
+
303
+ @[reassoc (attr := simp)]
304
+ lemma ι_D₃ :
305
+ ι F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j h ≫ D₃ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ j j' =
306
+ d₃ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ i₁ i₂ i₃ j' := by
307
+ simp only [ι_eq _ _ _ _ _ _ _ _ _ _ _ _ rfl h, D₃, assoc, mapBifunctor.ι_D₂]
308
+ by_cases h₁ : c₃.Rel i₃ (c₃.next i₃)
309
+ · rw [d₃_eq _ _ _ _ _ _ _ _ _ h₁]
310
+ by_cases h₂ : ComplexShape.π c₁₂ c₃ c₄ (c₁.π c₂ c₁₂ (i₁, i₂), c₃.next i₃) = j'
311
+ · rw [mapBifunctor.d₂_eq _ _ _ _ _ h₁ _ h₂,
312
+ ιOrZero_eq _ _ _ _ _ _ _ _ _ _ _ h₂,
313
+ Linear.comp_units_smul, smul_left_cancel_iff,
314
+ ι_eq _ _ _ _ _ _ _ _ _ _ _ _ rfl h₂,
315
+ NatTrans.naturality_assoc]
316
+ · rw [mapBifunctor.d₂_eq_zero' _ _ _ _ _ h₁ _ h₂, comp_zero,
317
+ ιOrZero_eq_zero _ _ _ _ _ _ _ _ _ _ _ h₂, comp_zero, smul_zero]
318
+ · rw [mapBifunctor.d₂_eq_zero _ _ _ _ _ _ _ h₁, comp_zero,
319
+ d₃_eq_zero _ _ _ _ _ _ _ _ _ _ _ h₁]
320
+
321
+ end
322
+
323
+ lemma d_eq (j j' : ι₄) [HasGoodTrifunctor₁₂Obj F₁₂ G K₁ K₂ K₃ c₁₂ c₄] :
324
+ (mapBifunctor (mapBifunctor K₁ K₂ F₁₂ c₁₂) K₃ G c₄).d j j' =
325
+ D₁ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ j j' + D₂ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ j j' +
326
+ D₃ F₁₂ G K₁ K₂ K₃ c₁₂ c₄ j j' := by
327
+ rw [mapBifunctor.d_eq]
328
+ congr 1
329
+ ext i₁ i₂ i₃ h
330
+ simp only [Preadditive.comp_add, ι_D₁, ι_D₂]
331
+ rw [ι_eq _ _ _ _ _ _ _ _ _ _ _ _ rfl h, assoc, mapBifunctor.ι_D₁]
332
+ set i₁₂ := ComplexShape.π c₁ c₂ c₁₂ ⟨i₁, i₂⟩
333
+ by_cases h₁ : c₁₂.Rel i₁₂ (c₁₂.next i₁₂)
334
+ · by_cases h₂ : ComplexShape.π c₁₂ c₃ c₄ (c₁₂.next i₁₂, i₃) = j'
335
+ · rw [mapBifunctor.d₁_eq _ _ _ _ h₁ _ _ h₂]
336
+ simp only [mapBifunctor.d_eq, Functor.map_add, NatTrans.app_add, Preadditive.add_comp,
337
+ smul_add, Preadditive.comp_add, Linear.comp_units_smul]
338
+ congr 1
339
+ · rw [← NatTrans.comp_app_assoc, ← Functor.map_comp,
340
+ mapBifunctor.ι_D₁]
341
+ by_cases h₃ : c₁.Rel i₁ (c₁.next i₁)
342
+ · have h₄ := (ComplexShape.next_π₁ c₂ c₁₂ h₃ i₂).symm
343
+ rw [mapBifunctor.d₁_eq _ _ _ _ h₃ _ _ h₄,
344
+ d₁_eq _ _ _ _ _ _ _ h₃,
345
+ ιOrZero_eq _ _ _ _ _ _ _ _ _ _ _ (by rw [← h₂, ← h₄]; rfl),
346
+ ι_eq _ _ _ _ _ _ _ _ _ _ (c₁₂.next i₁₂) _ h₄ h₂,
347
+ Functor.map_units_smul, Functor.map_comp, NatTrans.app_units_zsmul,
348
+ NatTrans.comp_app, Linear.units_smul_comp, assoc, smul_smul]
349
+ · rw [d₁_eq_zero _ _ _ _ _ _ _ _ _ _ _ h₃,
350
+ mapBifunctor.d₁_eq_zero _ _ _ _ _ _ _ h₃,
351
+ Functor.map_zero, zero_app, zero_comp, smul_zero]
352
+ · rw [← NatTrans.comp_app_assoc, ← Functor.map_comp,
353
+ mapBifunctor.ι_D₂]
354
+ by_cases h₃ : c₂.Rel i₂ (c₂.next i₂)
355
+ · have h₄ := (ComplexShape.next_π₂ c₁ c₁₂ i₁ h₃).symm
356
+ rw [mapBifunctor.d₂_eq _ _ _ _ _ h₃ _ h₄,
357
+ d₂_eq _ _ _ _ _ _ _ _ h₃,
358
+ ιOrZero_eq _ _ _ _ _ _ _ _ _ _ _ (by rw [← h₂, ← h₄]; rfl),
359
+ ι_eq _ _ _ _ _ _ _ _ _ _ (c₁₂.next i₁₂) _ h₄ h₂,
360
+ Functor.map_units_smul, Functor.map_comp, NatTrans.app_units_zsmul,
361
+ NatTrans.comp_app, Linear.units_smul_comp, assoc, smul_smul]
362
+ · rw [d₂_eq_zero _ _ _ _ _ _ _ _ _ _ _ h₃,
363
+ mapBifunctor.d₂_eq_zero _ _ _ _ _ _ _ h₃,
364
+ Functor.map_zero, zero_app, zero_comp, smul_zero]
365
+ · rw [mapBifunctor.d₁_eq_zero' _ _ _ _ h₁ _ _ h₂, comp_zero]
366
+ trans 0 + 0
367
+ · simp
368
+ · congr 1
369
+ · by_cases h₃ : c₁.Rel i₁ (c₁.next i₁)
370
+ · rw [d₁_eq _ _ _ _ _ _ _ h₃, ιOrZero_eq_zero, comp_zero, smul_zero]
371
+ dsimp [ComplexShape.r]
372
+ intro h₄
373
+ apply h₂
374
+ rw [← h₄, ComplexShape.next_π₁ c₂ c₁₂ h₃ i₂]
375
+ · rw [d₁_eq_zero _ _ _ _ _ _ _ _ _ _ _ h₃]
376
+ · by_cases h₃ : c₂.Rel i₂ (c₂.next i₂)
377
+ · rw [d₂_eq _ _ _ _ _ _ _ _ h₃, ιOrZero_eq_zero, comp_zero, smul_zero]
378
+ dsimp [ComplexShape.r]
379
+ intro h₄
380
+ apply h₂
381
+ rw [← h₄, ComplexShape.next_π₂ c₁ c₁₂ i₁ h₃]
382
+ · rw [d₂_eq_zero _ _ _ _ _ _ _ _ _ _ _ h₃]
383
+ · rw [mapBifunctor.d₁_eq_zero _ _ _ _ _ _ _ h₁, comp_zero,
384
+ d₁_eq_zero, d₂_eq_zero, zero_add]
385
+ · intro h₂
386
+ apply h₁
387
+ have := ComplexShape.rel_π₂ c₁ c₁₂ i₁ h₂
388
+ rw [c₁₂.next_eq' this]
389
+ exact this
390
+ · intro h₂
391
+ apply h₁
392
+ have := ComplexShape.rel_π₁ c₂ c₁₂ h₂ i₂
393
+ rw [c₁₂.next_eq' this]
394
+ exact this
395
+
190
396
end mapBifunctor₁₂
191
397
192
398
end HomologicalComplex
0 commit comments