Skip to content

Commit d82bc8c

Browse files
committed
feat(Algebra/Homology): quasi-isomorphisms are stable under retracts and have the 2/3 property (#20221)
1 parent f4e2c6d commit d82bc8c

File tree

3 files changed

+84
-2
lines changed

3 files changed

+84
-2
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -492,6 +492,7 @@ import Mathlib.Algebra.Homology.ShortComplex.ModuleCat
492492
import Mathlib.Algebra.Homology.ShortComplex.Preadditive
493493
import Mathlib.Algebra.Homology.ShortComplex.PreservesHomology
494494
import Mathlib.Algebra.Homology.ShortComplex.QuasiIso
495+
import Mathlib.Algebra.Homology.ShortComplex.Retract
495496
import Mathlib.Algebra.Homology.ShortComplex.RightHomology
496497
import Mathlib.Algebra.Homology.ShortComplex.ShortExact
497498
import Mathlib.Algebra.Homology.ShortComplex.SnakeLemma

Mathlib/Algebra/Homology/QuasiIso.lean

Lines changed: 44 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Kim Morrison, Joël Riou
55
-/
66
import Mathlib.Algebra.Homology.Homotopy
7+
import Mathlib.Algebra.Homology.ShortComplex.Retract
8+
import Mathlib.CategoryTheory.MorphismProperty.Composition
79

810
/-!
911
# Quasi-isomorphisms
@@ -51,6 +53,15 @@ lemma quasiIsoAt_iff' (f : K ⟶ L) (i j k : ι) (hi : c.prev j = i) (hk : c.nex
5153
exact ShortComplex.quasiIso_iff_of_arrow_mk_iso _ _
5254
(Arrow.isoOfNatIso (natIsoSc' C c i j k hi hk) (Arrow.mk f))
5355

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+
5465
lemma quasiIsoAt_iff_isIso_homologyMap (f : K ⟶ L) (i : ι)
5566
[K.HasHomology i] [L.HasHomology i] :
5667
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 φ
223234
[hφ : QuasiIso φ] : QuasiIso φ' := by
224235
simpa only [← quasiIso_iff_of_arrow_mk_iso φ φ' e]
225236

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+
226243
namespace HomologicalComplex
227244

228245
section PreservesHomology
@@ -274,10 +291,35 @@ variable (C c)
274291
def quasiIso [CategoryWithHomology C] :
275292
MorphismProperty (HomologicalComplex C c) := fun _ _ f => QuasiIso f
276293

277-
variable {C c}
294+
variable {C c} [CategoryWithHomology C]
278295

279296
@[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)
281323

282324
end HomologicalComplex
283325

Lines changed: 39 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,39 @@
1+
/-
2+
Copyright (c) 2024 Joël Riou. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Joël Riou
5+
-/
6+
7+
import Mathlib.Algebra.Homology.ShortComplex.QuasiIso
8+
import Mathlib.CategoryTheory.MorphismProperty.Retract
9+
10+
/-!
11+
# Quasi-isomorphisms of short complexes are stable under retracts
12+
13+
-/
14+
15+
namespace CategoryTheory
16+
17+
open Limits
18+
19+
namespace ShortComplex
20+
21+
variable {C : Type*} [Category C] [HasZeroMorphisms C]
22+
{S₁ T₁ S₂ T₂ : ShortComplex C}
23+
[S₁.HasHomology] [T₁.HasHomology] [S₂.HasHomology] [T₂.HasHomology]
24+
{f₁ : S₁ ⟶ T₁} {f₂ : S₂ ⟶ T₂}
25+
26+
lemma quasiIso_of_retract (h : RetractArrow f₁ f₂) [hf₂ : QuasiIso f₂] :
27+
QuasiIso f₁ := by
28+
rw [quasiIso_iff] at hf₂ ⊢
29+
have h : RetractArrow (homologyMap f₁) (homologyMap f₂) :=
30+
{ i := Arrow.homMk (u := homologyMap (show S₁ ⟶ S₂ from h.i.left))
31+
(v := homologyMap (show T₁ ⟶ T₂ from h.i.right)) (by simp [← homologyMap_comp])
32+
r := Arrow.homMk (u := homologyMap (show S₂ ⟶ S₁ from h.r.left))
33+
(v := homologyMap (show T₂ ⟶ T₁ from h.r.right)) (by simp [← homologyMap_comp])
34+
retract := by ext <;> simp [← homologyMap_comp] }
35+
exact (MorphismProperty.isomorphisms C).of_retract h hf₂
36+
37+
end ShortComplex
38+
39+
end CategoryTheory

0 commit comments

Comments
 (0)