@@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Kim Morrison, Joël Riou
5
5
-/
6
6
import Mathlib.Algebra.Homology.Homotopy
7
+ import Mathlib.Algebra.Homology.ShortComplex.Retract
8
+ import Mathlib.CategoryTheory.MorphismProperty.Composition
7
9
8
10
/-!
9
11
# Quasi-isomorphisms
@@ -51,6 +53,15 @@ lemma quasiIsoAt_iff' (f : K ⟶ L) (i j k : ι) (hi : c.prev j = i) (hk : c.nex
51
53
exact ShortComplex.quasiIso_iff_of_arrow_mk_iso _ _
52
54
(Arrow.isoOfNatIso (natIsoSc' C c i j k hi hk) (Arrow.mk f))
53
55
56
+ lemma quasiIsoAt_of_retract {f : K ⟶ L} {f' : K' ⟶ L'}
57
+ (h : RetractArrow f f') (i : ι) [K.HasHomology i] [L.HasHomology i]
58
+ [K'.HasHomology i] [L'.HasHomology i] [hf' : QuasiIsoAt f' i] :
59
+ QuasiIsoAt f i := by
60
+ rw [quasiIsoAt_iff] at hf' ⊢
61
+ have : RetractArrow ((shortComplexFunctor C c i).map f)
62
+ ((shortComplexFunctor C c i).map f') := h.map (shortComplexFunctor C c i).mapArrow
63
+ exact ShortComplex.quasiIso_of_retract this
64
+
54
65
lemma quasiIsoAt_iff_isIso_homologyMap (f : K ⟶ L) (i : ι)
55
66
[K.HasHomology i] [L.HasHomology i] :
56
67
QuasiIsoAt f i ↔ IsIso (homologyMap f i) := by
@@ -223,6 +234,12 @@ lemma quasiIso_of_arrow_mk_iso (φ : K ⟶ L) (φ' : K' ⟶ L') (e : Arrow.mk φ
223
234
[hφ : QuasiIso φ] : QuasiIso φ' := by
224
235
simpa only [← quasiIso_iff_of_arrow_mk_iso φ φ' e]
225
236
237
+ lemma quasiIso_of_retractArrow {f : K ⟶ L} {f' : K' ⟶ L'}
238
+ (h : RetractArrow f f') [∀ i, K.HasHomology i] [∀ i, L.HasHomology i]
239
+ [∀ i, K'.HasHomology i] [∀ i, L'.HasHomology i] [QuasiIso f'] :
240
+ QuasiIso f where
241
+ quasiIsoAt i := quasiIsoAt_of_retract h i
242
+
226
243
namespace HomologicalComplex
227
244
228
245
section PreservesHomology
@@ -274,10 +291,35 @@ variable (C c)
274
291
def quasiIso [CategoryWithHomology C] :
275
292
MorphismProperty (HomologicalComplex C c) := fun _ _ f => QuasiIso f
276
293
277
- variable {C c}
294
+ variable {C c} [CategoryWithHomology C]
278
295
279
296
@[simp]
280
- lemma mem_quasiIso_iff [CategoryWithHomology C] (f : K ⟶ L) : quasiIso C c f ↔ QuasiIso f := by rfl
297
+ lemma mem_quasiIso_iff (f : K ⟶ L) : quasiIso C c f ↔ QuasiIso f := by rfl
298
+
299
+ instance : (quasiIso C c).IsMultiplicative where
300
+ id_mem _ := by
301
+ rw [mem_quasiIso_iff]
302
+ infer_instance
303
+ comp_mem _ _ hf hg := by
304
+ rw [mem_quasiIso_iff] at hf hg ⊢
305
+ infer_instance
306
+
307
+ instance : (quasiIso C c).HasTwoOutOfThreeProperty where
308
+ of_postcomp f g hg hfg := by
309
+ rw [mem_quasiIso_iff] at hg hfg ⊢
310
+ rwa [← quasiIso_iff_comp_right f g]
311
+ of_precomp f g hf hfg := by
312
+ rw [mem_quasiIso_iff] at hf hfg ⊢
313
+ rwa [← quasiIso_iff_comp_left f g]
314
+
315
+ instance : (quasiIso C c).IsStableUnderRetracts where
316
+ of_retract h hg := by
317
+ rw [mem_quasiIso_iff] at hg ⊢
318
+ exact quasiIso_of_retractArrow h
319
+
320
+ instance : (quasiIso C c).RespectsIso :=
321
+ MorphismProperty.respectsIso_of_isStableUnderComposition
322
+ (fun _ _ _ (_ : IsIso _) ↦ by rw [mem_quasiIso_iff]; infer_instance)
281
323
282
324
end HomologicalComplex
283
325
0 commit comments