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

Commit 92c8d93

Browse files
committed
feat(algebra/homology): the cohomology functor (#2374)
This is the second in a series of most likely three PRs about the cohomology functor. As such, this PR depends on #2373. In the project laid out in `homology.lean`, @semorrison asks what the minimal assumptions are that are needed to get induced maps on images. In this PR, I offer a tautologial answer to this question: We get induced maps on images when there are induced maps on images. In this way, we can let type class resolution answer the question whether cohomology is functorial. In particular, the third PR will contain the fact that if our images are strong epi-mono factorizations, then we get induced maps on images. Since the regular coimage construction in regular categories is a strong epi-mono factorization, the approach in this PR generalizes the previous suggestion of requiring `V` to be regular. A quick remark about cohomology and dependent types: As you can see, at one point Lean forces us to write `i - 1 + 1` instead of `i` because these two things are not definitionally equal. I am afraid, as we do more with cohomology, there will be many cases of this issue, and to compose morphisms whose types contain different incarnations of the same integer, we will have to insert some `eq_to_hom`-esque glue and pray that we will be able to rewrite them all away in the proofs without getting the beloved `motive is not type correct` error. Maybe there is some better way to solve this problem? (Or am I overthinking this and it is not actually going to be an issue at all?)
1 parent ca98659 commit 92c8d93

File tree

1 file changed

+71
-60
lines changed

1 file changed

+71
-60
lines changed

src/algebra/homology/homology.lean

Lines changed: 71 additions & 60 deletions
Original file line numberDiff line numberDiff line change
@@ -8,16 +8,16 @@ import category_theory.limits.shapes.images
88
import category_theory.limits.shapes.kernels
99

1010
/-!
11-
# Non-functorial cohomology groups for cochain complexes
11+
# Cohomology groups for cochain complexes
1212
1313
We setup that part of the theory of cohomology groups which works in
1414
any category with kernels and images.
1515
16-
We define the cohomology groups themselves, and while we can show that
17-
chain maps induce maps on the kernels, at this level of generality
18-
chain maps do not induce maps on the images, and so not on the cohomology groups.
16+
We define the cohomology groups themselves, and show that they induce maps on kernels.
17+
18+
Under the additional assumption that our category has equalizers and functorial images, we construct
19+
induced morphisms on images and functorial induced morphisms in cohomology.
1920
20-
We'll do this with stronger assumptions, later.
2121
-/
2222

2323
universes v u
@@ -30,6 +30,8 @@ open category_theory.limits
3030
variables {V : Type u} [𝒱 : category.{v} V] [has_zero_morphisms.{v} V]
3131
include 𝒱
3232

33+
section
34+
3335
variable [has_kernels.{v} V]
3436
/-- The map induced by a chain map between the kernels of the differentials. -/
3537
def kernel_map {C C' : cochain_complex V} (f : C ⟶ C') (i : ℤ) :
@@ -39,10 +41,10 @@ begin
3941
rw [category.assoc, ←comm_at f, ←category.assoc, kernel.condition, has_zero_morphisms.zero_comp],
4042
end
4143

42-
@[simp]
44+
@[simp, reassoc]
4345
lemma kernel_map_condition {C C' : cochain_complex V} (f : C ⟶ C') (i : ℤ) :
4446
kernel_map f i ≫ kernel.ι (C'.d i) = kernel.ι (C.d i) ≫ f.f i :=
45-
by erw [limit.lift_π, fork.of_ι_app_zero]
47+
by simp [kernel_map]
4648

4749
@[simp]
4850
lemma kernel_map_id (C : cochain_complex.{v} V) (i : ℤ) :
@@ -53,83 +55,92 @@ lemma kernel_map_id (C : cochain_complex.{v} V) (i : ℤ) :
5355
lemma kernel_map_comp {C C' C'' : cochain_complex.{v} V} (f : C ⟶ C')
5456
(g : C' ⟶ C'') (i : ℤ) :
5557
kernel_map (f ≫ g) i = kernel_map f i ≫ kernel_map g i :=
56-
(cancel_mono (kernel.ι (C''.d i))).1 $
57-
by rw [kernel_map_condition, category.assoc, kernel_map_condition,
58-
←category.assoc, kernel_map_condition, category.assoc, differential_object.comp_f,
59-
graded_object.comp_apply]
58+
(cancel_mono (kernel.ι (C''.d i))).1 $ by simp
6059

61-
-- TODO: Actually, this is a functor `cochain_complex V ⥤ cochain_complex V`, but to state this
62-
-- properly we will need `has_shift` on `differential_object` first.
6360
/-- The kernels of the differentials of a cochain complex form a ℤ-graded object. -/
6461
def kernel_functor : cochain_complex.{v} V ⥤ graded_object ℤ V :=
6562
{ obj := λ C i, kernel (C.d i),
6663
map := λ X Y f i, kernel_map f i }
6764

65+
end
66+
67+
section
68+
variables [has_images.{v} V] [has_image_maps.{v} V]
69+
70+
/-- A morphism of cochain complexes induces a morphism on the images of the differentials in every
71+
degree. -/
72+
abbreviation image_map {C C' : cochain_complex.{v} V} (f : C ⟶ C') (i : ℤ) :
73+
image (C.d i) ⟶ image (C'.d i) :=
74+
image.map (arrow.hom_mk' (cochain_complex.comm_at f i).symm)
75+
76+
@[simp]
77+
lemma image_map_ι {C C' : cochain_complex.{v} V} (f : C ⟶ C') (i : ℤ) :
78+
image_map f i ≫ image.ι (C'.d i) = image.ι (C.d i) ≫ f.f (i + 1) :=
79+
image.map_hom_mk'_ι (cochain_complex.comm_at f i).symm
80+
81+
end
82+
6883
/-!
6984
At this point we assume that we have all images, and all equalizers.
7085
We need to assume all equalizers, not just kernels, so that
7186
`factor_thru_image` is an epimorphism.
7287
-/
73-
variables [has_images.{v} V] [has_equalizers.{v} V]
88+
variables [has_kernels.{v} V] [has_images.{v} V] [has_equalizers.{v} V]
7489

7590
/--
7691
The connecting morphism from the image of `d i` to the kernel of `d (i+1)`.
7792
-/
7893
def image_to_kernel_map (C : cochain_complex V) (i : ℤ) :
7994
image (C.d i) ⟶ kernel (C.d (i+1)) :=
80-
kernel.lift _ (image.ι (C.d i))
81-
begin
82-
rw ←cancel_epi (factor_thru_image (C.d i)),
83-
rw [has_zero_morphisms.comp_zero, image.fac_assoc, d_squared],
84-
end
95+
kernel.lift _ (image.ι (C.d i)) $ (cancel_epi (factor_thru_image (C.d i))).1 $ by simp
96+
97+
@[simp, reassoc]
98+
lemma image_to_kernel_map_condition (C : cochain_complex V) (i : ℤ) :
99+
image_to_kernel_map C i ≫ kernel.ι (C.d (i + 1)) = image.ι (C.d i) :=
100+
by simp [image_to_kernel_map]
85101

86-
-- TODO (a good project!):
87-
-- At this level of generality, it's just not true that a chain map
88-
-- induces maps on boundaries
89-
--
90-
-- Let's add these later, with appropriate (but hopefully fairly minimal)
91-
-- assumptions: perhaps that the category is regular?
92-
-- I think in that case we can compute `image` as the regular coimage,
93-
-- i.e. the coequalizer of the kernel pair,
94-
-- and that image has the appropriate mapping property.
95-
96-
-- def image_map {C C' : cochain_complex.{v} V} (f : C ⟶ C') (i : ℤ) :
97-
-- image (C.d i) ⟶ image (C'.d i) :=
98-
-- sorry
99-
100-
-- -- I'm not certain what the minimal assumptions required to prove the following
101-
-- -- lemma are:
102-
-- lemma induced_maps_commute {C C' : cochain_complex.{v} V} (f : C ⟶ C') (i : ℤ) :
103-
-- image_to_kernel_map C i ≫ kernel_map f (i+1) =
104-
-- image_map f i ≫ image_to_kernel_map C' i :=
105-
-- sorry
102+
@[reassoc]
103+
lemma induced_maps_commute [has_image_maps.{v} V] {C C' : cochain_complex.{v} V} (f : C ⟶ C')
104+
(i : ℤ) :
105+
image_to_kernel_map C i ≫ kernel_map f (i + 1) = image_map f i ≫ image_to_kernel_map C' i :=
106+
by { ext, simp }
106107

107108
variables [has_cokernels.{v} V]
108109

109110
/-- The `i`-th cohomology group of the cochain complex `C`. -/
110111
def cohomology (C : cochain_complex V) (i : ℤ) : V :=
111112
cokernel (image_to_kernel_map C (i-1))
112113

113-
-- TODO:
114-
115-
-- As noted above, as we don't get induced maps on boundaries with this generality,
116-
-- we can't assemble the cohomology groups into a functor. Hopefully, however,
117-
-- the commented out code below will work
118-
-- (with whatever added assumptions are needed above.)
119-
120-
-- def cohomology_map {C C' : cochain_complex.{v} V} (f : C ⟶ C') (i : ℤ) :
121-
-- C.cohomology i ⟶ C'.cohomology i :=
122-
-- cokernel.desc _ (kernel_map f (i-1) ≫ cokernel.π _)
123-
-- begin
124-
-- rw [←category.assoc, induced_maps_commute, category.assoc, cokernel.condition],
125-
-- erw [has_zero_morphisms.comp_zero],
126-
-- end
127-
128-
-- /-- The cohomology functor from chain complexes to `ℤ` graded objects in `V`. -/
129-
-- def cohomology_functor : cochain_complex.{v} V ⥤ graded_object ℤ V :=
130-
-- { obj := λ C i, cohomology C i,
131-
-- map := λ C C' f i, cohomology_map f i,
132-
-- map_id' := sorry,
133-
-- map_comp' := sorry, }
114+
variables [has_image_maps.{v} V]
115+
116+
/-- A morphism of cochain complexes induces a morphism in cohomology at every degree. -/
117+
def cohomology_map {C C' : cochain_complex.{v} V} (f : C ⟶ C') (i : ℤ) :
118+
C.cohomology i ⟶ C'.cohomology i :=
119+
cokernel.desc _ (kernel_map f (i - 1 + 1) ≫ cokernel.π _) $ by simp [induced_maps_commute_assoc]
120+
121+
@[simp, reassoc]
122+
lemma cohomology_map_condition {C C' : cochain_complex.{v} V} (f : C ⟶ C') (i : ℤ) :
123+
cokernel.π (image_to_kernel_map C (i - 1)) ≫ cohomology_map f i =
124+
kernel_map f (i - 1 + 1) ≫ cokernel.π _ :=
125+
by simp [cohomology_map]
126+
127+
@[simp]
128+
lemma cohomology_map_id (C : cochain_complex.{v} V) (i : ℤ) :
129+
cohomology_map (𝟙 C) i = 𝟙 (cohomology C i) :=
130+
begin
131+
ext,
132+
simp only [cohomology_map_condition, kernel_map_id, category.id_comp],
133+
erw [category.comp_id]
134+
end
135+
136+
@[simp]
137+
lemma cohomology_map_comp {C C' C'' : cochain_complex.{v} V} (f : C ⟶ C') (g : C' ⟶ C'') (i : ℤ) :
138+
cohomology_map (f ≫ g) i = cohomology_map f i ≫ cohomology_map g i :=
139+
by { ext, simp }
140+
141+
/-- The cohomology functor from cochain complexes to `ℤ` graded objects in `V`. -/
142+
def cohomology_functor : cochain_complex.{v} V ⥤ graded_object ℤ V :=
143+
{ obj := λ C i, cohomology C i,
144+
map := λ C C' f i, cohomology_map f i }
134145

135146
end cochain_complex

0 commit comments

Comments
 (0)