56
56
exact preadditive.comp_add _ _ _ _ _ _,
57
57
end
58
58
59
+ @[simp]
59
60
lemma d_next_eq_d_from_from_next [has_zero_object V] (f : Π i j, C.X i ⟶ D.X j) (i : ι) :
60
61
d_next i f = C.d_from i ≫ from_next i f :=
61
62
begin
@@ -121,6 +122,7 @@ begin
121
122
exact preadditive.add_comp _ _ _ _ _ _,
122
123
end
123
124
125
+ @[simp]
124
126
lemma prev_d_eq_to_prev_d_to [has_zero_object V] (f : Π i j, C.X i ⟶ D.X j) (j : ι) :
125
127
prev_d j f = to_prev j f ≫ D.d_to j :=
126
128
begin
@@ -147,7 +149,7 @@ begin
147
149
simp, },
148
150
end
149
151
150
- @[simp] lemma to_prev'_comp_right (f : Π i j, C.X i ⟶ D.X j) (g : D ⟶ E) (j : ι) :
152
+ @[simp] lemma prev_d_comp_right (f : Π i j, C.X i ⟶ D.X j) (g : D ⟶ E) (j : ι) :
151
153
prev_d j (λ i j, f i j ≫ g.f j) = prev_d j f ≫ g.f j :=
152
154
begin
153
155
dsimp [prev_d],
@@ -240,7 +242,7 @@ def trans {e f g : C ⟶ D} (h : homotopy e f) (k : homotopy f g) : homotopy e g
240
242
241
243
/-- the sum of two homotopies is a homotopy between the sum of the respective morphisms. -/
242
244
@[simps]
243
- def add {f₁ g₁ f₂ g₂: C ⟶ D}
245
+ def add {f₁ g₁ f₂ g₂ : C ⟶ D}
244
246
(h₁ : homotopy f₁ g₁) (h₂ : homotopy f₂ g₂) : homotopy (f₁+f₂) (g₁+g₂) :=
245
247
{ hom := h₁.hom + h₂.hom,
246
248
zero' := λ i j hij, by
@@ -256,7 +258,7 @@ def comp_right {e f : C ⟶ D} (h : homotopy e f) (g : D ⟶ E) : homotopy (e
256
258
{ hom := λ i j, h.hom i j ≫ g.f j,
257
259
zero' := λ i j w, by rw [h.zero i j w, zero_comp],
258
260
comm := λ i, by simp only [h.comm i, d_next_comp_right, preadditive.add_comp,
259
- to_prev'_comp_right , comp_f], }
261
+ prev_d_comp_right , comp_f], }
260
262
261
263
/-- homotopy is closed under composition (on the left) -/
262
264
@[simps]
@@ -321,7 +323,7 @@ null_homotopic_map hom ≫ g = null_homotopic_map (λ i j, hom i j ≫ g.f j) :=
321
323
begin
322
324
ext n,
323
325
dsimp [null_homotopic_map],
324
- simp only [preadditive.add_comp, d_next_comp_right, to_prev'_comp_right ],
326
+ simp only [preadditive.add_comp, d_next_comp_right, prev_d_comp_right ],
325
327
end
326
328
327
329
/-- Compatibility of `null_homotopic_map'` with the postcomposition by a morphism
@@ -827,9 +829,7 @@ instance : inhabited (homotopy_equiv C C) := ⟨refl C⟩
827
829
828
830
end homotopy_equiv
829
831
830
- variables [has_equalizers V] [has_cokernels V] [has_images V] [has_image_maps V]
831
-
832
- variable [has_zero_object V]
832
+ variables [has_equalizers V] [has_cokernels V] [has_images V] [has_image_maps V] [has_zero_object V]
833
833
834
834
/--
835
835
Homotopic maps induce the same map on homology.
0 commit comments