@@ -36,7 +36,7 @@ a morphism `d : X ⟶ X⟦1⟧`, such that `d^2 = 0`.
36
36
structure differential_object :=
37
37
(X : C)
38
38
(d : X ⟶ X⟦1 ⟧)
39
- (d_squared' : d ≫ d⟦1 ⟧' = 0 . obviously)
39
+ (d_squared' : d ≫ d⟦( 1 :ℤ) ⟧' = 0 . obviously)
40
40
41
41
restate_axiom differential_object.d_squared'
42
42
attribute [simp] differential_object.d_squared
@@ -83,6 +83,11 @@ lemma comp_f {X Y Z : differential_object C} (f : X ⟶ Y) (g : Y ⟶ Z) :
83
83
(f ≫ g).f = f.f ≫ g.f :=
84
84
rfl
85
85
86
+ @[simp]
87
+ lemma eq_to_hom_f {X Y : differential_object C} (h : X = Y) :
88
+ hom.f (eq_to_hom h) = eq_to_hom (congr_arg _ h) :=
89
+ by { subst h, rw [eq_to_hom_refl, eq_to_hom_refl], refl }
90
+
86
91
variables (C)
87
92
88
93
/-- The forgetful functor taking a differential object to its underlying object. -/
@@ -95,7 +100,7 @@ instance forget_faithful : faithful (forget C) :=
95
100
96
101
instance has_zero_morphisms : has_zero_morphisms (differential_object C) :=
97
102
{ has_zero := λ X Y,
98
- ⟨{ f := 0 , }⟩}
103
+ ⟨{ f := 0 }⟩}
99
104
100
105
variables {C}
101
106
@@ -115,6 +120,16 @@ An isomorphism of differential objects gives an isomorphism of the underlying ob
115
120
@[simp] lemma iso_app_trans {X Y Z : differential_object C} (f : X ≅ Y) (g : Y ≅ Z) :
116
121
iso_app (f ≪≫ g) = iso_app f ≪≫ iso_app g := rfl
117
122
123
+ /-- An isomorphism of differential objects can be constructed
124
+ from an isomorphism of the underlying objects that commutes with the differentials. -/
125
+ @[simps] def mk_iso {X Y : differential_object C}
126
+ (f : X.X ≅ Y.X) (hf : X.d ≫ f.hom⟦1 ⟧' = f.hom ≫ Y.d) : X ≅ Y :=
127
+ { hom := ⟨f.hom, hf⟩,
128
+ inv := ⟨f.inv, by { dsimp, rw [← functor.map_iso_inv, iso.comp_inv_eq, category.assoc,
129
+ iso.eq_inv_comp, functor.map_iso_hom, ← hf], congr }⟩,
130
+ hom_inv_id' := by { ext1, dsimp, exact f.hom_inv_id },
131
+ inv_hom_id' := by { ext1, dsimp, exact f.inv_hom_id } }
132
+
118
133
end differential_object
119
134
120
135
namespace functor
0 commit comments