@@ -8,16 +8,16 @@ import category_theory.limits.shapes.images
8
8
import category_theory.limits.shapes.kernels
9
9
10
10
/-!
11
- # Non-functorial cohomology groups for cochain complexes
11
+ # Cohomology groups for cochain complexes
12
12
13
13
We setup that part of the theory of cohomology groups which works in
14
14
any category with kernels and images.
15
15
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.
19
20
20
- We'll do this with stronger assumptions, later.
21
21
-/
22
22
23
23
universes v u
@@ -30,6 +30,8 @@ open category_theory.limits
30
30
variables {V : Type u} [𝒱 : category.{v} V] [has_zero_morphisms.{v} V]
31
31
include 𝒱
32
32
33
+ section
34
+
33
35
variable [has_kernels.{v} V]
34
36
/-- The map induced by a chain map between the kernels of the differentials. -/
35
37
def kernel_map {C C' : cochain_complex V} (f : C ⟶ C') (i : ℤ) :
@@ -39,10 +41,10 @@ begin
39
41
rw [category.assoc, ←comm_at f, ←category.assoc, kernel.condition, has_zero_morphisms.zero_comp],
40
42
end
41
43
42
- @[simp]
44
+ @[simp, reassoc ]
43
45
lemma kernel_map_condition {C C' : cochain_complex V} (f : C ⟶ C') (i : ℤ) :
44
46
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 ]
46
48
47
49
@[simp]
48
50
lemma kernel_map_id (C : cochain_complex.{v} V) (i : ℤ) :
@@ -53,83 +55,92 @@ lemma kernel_map_id (C : cochain_complex.{v} V) (i : ℤ) :
53
55
lemma kernel_map_comp {C C' C'' : cochain_complex.{v} V} (f : C ⟶ C')
54
56
(g : C' ⟶ C'') (i : ℤ) :
55
57
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
60
59
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.
63
60
/-- The kernels of the differentials of a cochain complex form a ℤ-graded object. -/
64
61
def kernel_functor : cochain_complex.{v} V ⥤ graded_object ℤ V :=
65
62
{ obj := λ C i, kernel (C.d i),
66
63
map := λ X Y f i, kernel_map f i }
67
64
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
+
68
83
/-!
69
84
At this point we assume that we have all images, and all equalizers.
70
85
We need to assume all equalizers, not just kernels, so that
71
86
`factor_thru_image` is an epimorphism.
72
87
-/
73
- variables [has_images.{v} V] [has_equalizers.{v} V]
88
+ variables [has_kernels.{v} V] [ has_images.{v} V] [has_equalizers.{v} V]
74
89
75
90
/--
76
91
The connecting morphism from the image of `d i` to the kernel of `d (i+1)`.
77
92
-/
78
93
def image_to_kernel_map (C : cochain_complex V) (i : ℤ) :
79
94
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]
85
101
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 }
106
107
107
108
variables [has_cokernels.{v} V]
108
109
109
110
/-- The `i`-th cohomology group of the cochain complex `C`. -/
110
111
def cohomology (C : cochain_complex V) (i : ℤ) : V :=
111
112
cokernel (image_to_kernel_map C (i-1 ))
112
113
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 }
134
145
135
146
end cochain_complex
0 commit comments