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

Commit c5bb320

Browse files
committed
refactor(*): Random lemmas and modifications from the shifting refactor. (#10940)
1 parent ee25d58 commit c5bb320

File tree

9 files changed

+89
-14
lines changed

9 files changed

+89
-14
lines changed

src/algebra/homology/differential_object.lean

Lines changed: 29 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -24,8 +24,35 @@ noncomputable theory
2424

2525
namespace homological_complex
2626

27-
variables {β : Type*} [add_comm_group β] (b : β)
28-
variables (V : Type*) [category V] [has_zero_morphisms V]
27+
variables {β : Type*} [add_comm_group β] {b : β}
28+
variables {V : Type*} [category V] [has_zero_morphisms V]
29+
30+
/-- Since `eq_to_hom` only preserves the fact that `X.X i = X.X j` but not `i = j`, this definition
31+
is used to aid the simplifier. -/
32+
abbreviation _root_.category_theory.differential_object.X_eq_to_hom
33+
(X : differential_object (graded_object_with_shift b V))
34+
{i j : β} (h : i = j) : X.X i ⟶ X.X j := eq_to_hom (congr_arg X.X h)
35+
36+
@[simp] lemma _root_.category_theory.differential_object.X_eq_to_hom_refl
37+
(X : differential_object (graded_object_with_shift b V)) (i : β) :
38+
X.X_eq_to_hom (refl i) = 𝟙 _ := rfl
39+
40+
@[simp, reassoc] lemma eq_to_hom_d (X : differential_object (graded_object_with_shift b V))
41+
{x y : β} (h : x = y) :
42+
X.X_eq_to_hom h ≫ X.d y = X.d x ≫ X.X_eq_to_hom (by { cases h, refl }) :=
43+
by { cases h, dsimp, simp }
44+
45+
@[simp, reassoc] lemma d_eq_to_hom (X : homological_complex V (complex_shape.up' b))
46+
{x y z : β} (h : y = z) :
47+
X.d x y ≫ eq_to_hom (congr_arg X.X h) = X.d x z :=
48+
by { cases h, simp }
49+
50+
@[simp, reassoc] lemma eq_to_hom_f {X Y : differential_object (graded_object_with_shift b V)}
51+
(f : X ⟶ Y) {x y : β} (h : x = y) :
52+
X.X_eq_to_hom h ≫ f.f y = f.f x ≫ Y.X_eq_to_hom h :=
53+
by { cases h, simp }
54+
55+
variables (b V)
2956

3057
/--
3158
The functor from differential graded objects to homological complexes.

src/category_theory/differential_object.lean

Lines changed: 17 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -36,7 +36,7 @@ a morphism `d : X ⟶ X⟦1⟧`, such that `d^2 = 0`.
3636
structure differential_object :=
3737
(X : C)
3838
(d : X ⟶ X⟦1⟧)
39-
(d_squared' : d ≫ d⟦1⟧' = 0 . obviously)
39+
(d_squared' : d ≫ d⟦(1:ℤ)⟧' = 0 . obviously)
4040

4141
restate_axiom differential_object.d_squared'
4242
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) :
8383
(f ≫ g).f = f.f ≫ g.f :=
8484
rfl
8585

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+
8691
variables (C)
8792

8893
/-- The forgetful functor taking a differential object to its underlying object. -/
@@ -95,7 +100,7 @@ instance forget_faithful : faithful (forget C) :=
95100

96101
instance has_zero_morphisms : has_zero_morphisms (differential_object C) :=
97102
{ has_zero := λ X Y,
98-
⟨{ f := 0, }⟩}
103+
⟨{ f := 0 }⟩}
99104

100105
variables {C}
101106

@@ -115,6 +120,16 @@ An isomorphism of differential objects gives an isomorphism of the underlying ob
115120
@[simp] lemma iso_app_trans {X Y Z : differential_object C} (f : X ≅ Y) (g : Y ≅ Z) :
116121
iso_app (f ≪≫ g) = iso_app f ≪≫ iso_app g := rfl
117122

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+
118133
end differential_object
119134

120135
namespace functor

src/category_theory/equivalence.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -392,6 +392,8 @@ mk' ::
392392

393393
restate_axiom is_equivalence.functor_unit_iso_comp'
394394

395+
attribute [simp, reassoc] is_equivalence.functor_unit_iso_comp
396+
395397
namespace is_equivalence
396398

397399
instance of_equivalence (F : C ≌ D) : is_equivalence F.functor :=
@@ -433,6 +435,12 @@ is_equivalence.of_equivalence F.as_equivalence.symm
433435
@[simp] lemma as_equivalence_inverse (F : C ⥤ D) [is_equivalence F] :
434436
F.as_equivalence.inverse = inv F := rfl
435437

438+
@[simp] lemma as_equivalence_unit {F : C ⥤ D} [h : is_equivalence F] :
439+
F.as_equivalence.unit_iso = @@is_equivalence.unit_iso _ _ h := rfl
440+
441+
@[simp] lemma as_equivalence_counit {F : C ⥤ D} [is_equivalence F] :
442+
F.as_equivalence.counit_iso = is_equivalence.counit_iso := rfl
443+
436444
@[simp] lemma inv_inv (F : C ⥤ D) [is_equivalence F] :
437445
inv (inv F) = F := rfl
438446

src/category_theory/graded_object.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -82,6 +82,10 @@ begin
8282
simp,
8383
end
8484

85+
@[simp] lemma eq_to_hom_apply {β : Type w} {X Y : Π b : β, C} (h : X = Y) (b : β) :
86+
(eq_to_hom h : X ⟶ Y) b = eq_to_hom (by subst h) :=
87+
by { subst h, refl }
88+
8589
/--
8690
The equivalence between β-graded objects and γ-graded objects,
8791
given an equivalence between β and γ.

src/category_theory/limits/has_limits.lean

Lines changed: 2 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -491,8 +491,7 @@ lemma has_limits_of_size_shrink [has_limits_of_size.{(max v₁ v₂) (max u₁ u
491491
⟨λ J hJ, by exactI has_limits_of_shape_of_equivalence
492492
(ulift_hom_ulift_category.equiv.{v₂ u₂} J).symm⟩
493493

494-
@[priority 100]
495-
instance has_smallest_limits_of_has_limits [has_limits C] :
494+
lemma has_smallest_limits_of_has_limits [has_limits C] :
496495
has_limits_of_size.{0 0} C := has_limits_of_size_shrink.{0 0} C
497496

498497
end limit
@@ -958,8 +957,7 @@ lemma has_colimits_of_size_shrink [has_colimits_of_size.{(max v₁ v₂) (max u
958957
⟨λ J hJ, by exactI has_colimits_of_shape_of_equivalence
959958
(ulift_hom_ulift_category.equiv.{v₂ u₂} J).symm⟩
960959

961-
@[priority 100]
962-
instance has_smallest_colimits_of_has_colimits [has_colimits C] :
960+
lemma has_smallest_colimits_of_has_colimits [has_colimits C] :
963961
has_colimits_of_size.{0 0} C := has_colimits_of_size_shrink.{0 0} C
964962

965963
end colimit

src/category_theory/monoidal/End.lean

Lines changed: 19 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -121,7 +121,7 @@ end
121121

122122
@[simp, reassoc]
123123
lemma μ_naturalityₗ {m n m' : M} (f : m ⟶ m') (X : C) :
124-
(F.obj n).map((F.map f).app X) ≫ (F.μ m' n).app X =
124+
(F.obj n).map ((F.map f).app X) ≫ (F.μ m' n).app X =
125125
(F.μ m n).app X ≫ (F.map (f ⊗ 𝟙 n)).app X :=
126126
begin
127127
rw ← μ_naturality₂ F f (𝟙 n) X,
@@ -137,6 +137,24 @@ begin
137137
simp,
138138
end
139139

140+
@[simp, reassoc]
141+
lemma μ_inv_naturalityₗ {m n m' : M} (f : m ⟶ m') (X : C) :
142+
(F.μ_iso m n).inv.app X ≫ (F.obj n).map ((F.map f).app X) =
143+
(F.map (f ⊗ 𝟙 n)).app X ≫ (F.μ_iso m' n).inv.app X :=
144+
begin
145+
rw [← is_iso.comp_inv_eq, category.assoc, ← is_iso.eq_inv_comp],
146+
simp,
147+
end
148+
149+
@[simp, reassoc]
150+
lemma μ_inv_naturalityᵣ {m n n' : M} (g : n ⟶ n') (X : C) :
151+
(F.μ_iso m n).inv.app X ≫ (F.map g).app ((F.obj m).obj X) =
152+
(F.map (𝟙 m ⊗ g)).app X ≫ (F.μ_iso m n').inv.app X :=
153+
begin
154+
rw [← is_iso.comp_inv_eq, category.assoc, ← is_iso.eq_inv_comp],
155+
simp,
156+
end
157+
140158
@[reassoc]
141159
lemma left_unitality_app (n : M) (X : C) :
142160
(F.obj n).map (F.ε.app X) ≫ (F.μ (𝟙_M) n).app X

src/category_theory/monoidal/discrete.lean

Lines changed: 7 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -23,9 +23,11 @@ variables (M : Type u) [monoid M]
2323

2424
namespace category_theory
2525

26+
@[to_additive]
2627
instance monoid_discrete : monoid (discrete M) := by { dsimp [discrete], apply_instance }
2728

28-
instance : monoidal_category (discrete M) :=
29+
@[to_additive discrete.add_monoidal]
30+
instance discrete.monoidal : monoidal_category (discrete M) :=
2931
{ tensor_unit := 1,
3032
tensor_obj := λ X Y, X * Y,
3133
tensor_hom := λ W X Y Z f g, eq_to_hom (by rw [eq_of_hom f, eq_of_hom g]),
@@ -39,7 +41,8 @@ variables {M} {N : Type u} [monoid N]
3941
A multiplicative morphism between monoids gives a monoidal functor between the corresponding
4042
discrete monoidal categories.
4143
-/
42-
@[simps]
44+
@[to_additive dicrete.add_monoidal_functor "An additive morphism between add_monoids gives a
45+
monoidal functor between the corresponding discrete monoidal categories.", simps]
4346
def discrete.monoidal_functor (F : M →* N) : monoidal_functor (discrete M) (discrete N) :=
4447
{ obj := F,
4548
map := λ X Y f, eq_to_hom (F.congr_arg (eq_of_hom f)),
@@ -51,6 +54,8 @@ variables {K : Type u} [monoid K]
5154
/--
5255
The monoidal natural isomorphism corresponding to composing two multiplicative morphisms.
5356
-/
57+
@[to_additive dicrete.add_monoidal_functor_comp "The monoidal natural isomorphism corresponding to
58+
composing two additive morphisms."]
5459
def discrete.monoidal_functor_comp (F : M →* N) (G : N →* K) :
5560
discrete.monoidal_functor F ⊗⋙ discrete.monoidal_functor G ≅
5661
discrete.monoidal_functor (G.comp F) :=

src/category_theory/natural_isomorphism.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -124,7 +124,7 @@ by simp only [←category.assoc, cancel_mono]
124124
by simp only [←category.assoc, cancel_mono]
125125

126126
@[simp] lemma inv_inv_app {F G : C ⥤ D} (e : F ≅ G) (X : C) :
127-
inv (e.inv.app X) = e.hom.app X := by { ext, simp }
127+
inv (e.inv.app X) = e.hom.app X := by { ext, simp }
128128

129129
end
130130

src/category_theory/triangulated/basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -40,14 +40,14 @@ structure triangle := mk' ::
4040
(obj₃ : C)
4141
(mor₁ : obj₁ ⟶ obj₂)
4242
(mor₂ : obj₂ ⟶ obj₃)
43-
(mor₃ : obj₃ ⟶ obj₁⟦1⟧)
43+
(mor₃ : obj₃ ⟶ obj₁⟦(1:ℤ)⟧)
4444

4545
/--
4646
A triangle `(X,Y,Z,f,g,h)` in `C` is defined by the morphisms `f : X ⟶ Y`, `g : Y ⟶ Z`
4747
and `h : Z ⟶ X⟦1⟧`.
4848
-/
4949
@[simps]
50-
def triangle.mk {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) (h : Z ⟶ X⟦1⟧) : triangle C :=
50+
def triangle.mk {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) (h : Z ⟶ X⟦(1:ℤ)⟧) : triangle C :=
5151
{ obj₁ := X,
5252
obj₂ := Y,
5353
obj₃ := Z,

0 commit comments

Comments
 (0)