Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit b535c2d

Browse files
committed
feat(algebra/homology): three lemmas on homological complexes (#12742)
1 parent 1084cee commit b535c2d

File tree

3 files changed

+47
-4
lines changed

3 files changed

+47
-4
lines changed

src/algebra/homology/additive.lean

Lines changed: 23 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -21,7 +21,7 @@ universes v u
2121
open_locale classical
2222
noncomputable theory
2323

24-
open category_theory category_theory.limits homological_complex
24+
open category_theory category_theory.category category_theory.limits homological_complex
2525

2626
variables {ι : Type*}
2727
variables {V : Type u} [category.{v} V] [preadditive V]
@@ -140,6 +140,28 @@ by tidy
140140

141141
end category_theory
142142

143+
namespace chain_complex
144+
145+
variables {W : Type*} [category W] [preadditive W]
146+
variables {α : Type*} [add_right_cancel_semigroup α] [has_one α] [decidable_eq α]
147+
148+
lemma map_chain_complex_of (F : V ⥤ W) [F.additive] (X : α → V) (d : Π n, X (n+1) ⟶ X n)
149+
(sq : ∀ n, d (n+1) ≫ d n = 0) :
150+
(F.map_homological_complex _).obj (chain_complex.of X d sq) =
151+
chain_complex.of (λ n, F.obj (X n))
152+
(λ n, F.map (d n)) (λ n, by rw [ ← F.map_comp, sq n, functor.map_zero]) :=
153+
begin
154+
apply homological_complex.ext,
155+
intros i j hij,
156+
{ have h : j+1=i := hij,
157+
subst h,
158+
simp only [category_theory.functor.map_homological_complex_obj_d, of_d,
159+
eq_to_hom_refl, comp_id, id_comp], },
160+
{ refl, }
161+
end
162+
163+
end chain_complex
164+
143165
variables [has_zero_object V] {W : Type*} [category W] [preadditive W] [has_zero_object W]
144166

145167
namespace homological_complex

src/algebra/homology/differential_object.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -47,7 +47,7 @@ by { cases h, dsimp, simp }
4747
X.d x y ≫ eq_to_hom (congr_arg X.X h) = X.d x z :=
4848
by { cases h, simp }
4949

50-
@[simp, reassoc] lemma eq_to_hom_f {X Y : differential_object (graded_object_with_shift b V)}
50+
@[simp, reassoc] lemma eq_to_hom_f' {X Y : differential_object (graded_object_with_shift b V)}
5151
(f : X ⟶ Y) {x y : β} (h : x = y) :
5252
X.X_eq_to_hom h ≫ f.f y = f.f x ≫ Y.X_eq_to_hom h :=
5353
by { cases h, simp }
@@ -81,7 +81,7 @@ def dgo_to_homological_complex :
8181
subst h,
8282
have : f.f i ≫ Y.d i = X.d i ≫ f.f (i + 1 • b) := (congr_fun f.comm i).symm,
8383
reassoc! this,
84-
simp only [category.comp_id, eq_to_hom_refl, dif_pos rfl, this, category.assoc, eq_to_hom_f]
84+
simp only [category.comp_id, eq_to_hom_refl, dif_pos rfl, this, category.assoc, eq_to_hom_f']
8585
end, } }
8686

8787
/--

src/algebra/homology/homological_complex.lean

Lines changed: 22 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -35,7 +35,7 @@ Defined in terms of these we have `C.d_from i : C.X i ⟶ C.X_next i` and
3535

3636
universes v u
3737

38-
open category_theory category_theory.limits
38+
open category_theory category_theory.category category_theory.limits
3939

4040
variables {ι : Type*}
4141
variables (V : Type u) [category.{v} V] [has_zero_morphisms V]
@@ -74,6 +74,21 @@ begin
7474
{ rw [C.shape i j hij, zero_comp] }
7575
end
7676

77+
lemma ext {C₁ C₂ : homological_complex V c} (h_X : C₁.X = C₂.X)
78+
(h_d : ∀ (i j : ι), c.rel i j → C₁.d i j ≫ eq_to_hom (congr_fun h_X j) =
79+
eq_to_hom (congr_fun h_X i) ≫ C₂.d i j) : C₁ = C₂ :=
80+
begin
81+
cases C₁,
82+
cases C₂,
83+
dsimp at h_X,
84+
subst h_X,
85+
simp only [true_and, eq_self_iff_true, heq_iff_eq],
86+
ext i j,
87+
by_cases hij : c.rel i j,
88+
{ simpa only [id_comp, eq_to_hom_refl, comp_id] using h_d i j hij, },
89+
{ rw [C₁_shape' i j hij, C₂_shape' i j hij], }
90+
end
91+
7792
end homological_complex
7893

7994
/--
@@ -176,6 +191,12 @@ end
176191
(f ≫ g).f i = f.f i ≫ g.f i :=
177192
rfl
178193

194+
@[simp]
195+
lemma eq_to_hom_f {C₁ C₂ : homological_complex V c} (h : C₁ = C₂) (n : ι) :
196+
homological_complex.hom.f (eq_to_hom h) n =
197+
eq_to_hom (congr_fun (congr_arg homological_complex.X h) n) :=
198+
by { subst h, refl, }
199+
179200
-- We'll use this later to show that `homological_complex V c` is preadditive when `V` is.
180201
lemma hom_f_injective {C₁ C₂ : homological_complex V c} :
181202
function.injective (λ f : hom C₁ C₂, f.f) :=

0 commit comments

Comments
 (0)