Skip to content

Commit f81eaba

Browse files
committed
refactor: Move the data fields of MonoidalCategory into a Struct class (#7279)
This matches the approach for `CategoryStruct`, and allows us to use the notation within `MonoidalCategory`. It also makes it easier to induce the lawful structure along a faithful functor, as again it means by the time we are providing the proof fields, the notation is already available. This also eliminates `tensorUnit` vs `tensorUnit'`, adding a custom pretty-printer to provide the unprimed version with appropriate notation.
1 parent f04a8ee commit f81eaba

File tree

18 files changed

+310
-307
lines changed

18 files changed

+310
-307
lines changed

Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean

Lines changed: 20 additions & 27 deletions
Original file line numberDiff line numberDiff line change
@@ -14,6 +14,7 @@ import Mathlib.RingTheory.TensorProduct
1414
-/
1515

1616
open CategoryTheory
17+
open scoped MonoidalCategory
1718

1819
universe v u
1920

@@ -39,58 +40,49 @@ noncomputable abbrev tensorHom {W X Y Z : AlgebraCat.{u} R} (f : W ⟶ X) (g : Y
3940
tensorObj W Y ⟶ tensorObj X Z :=
4041
Algebra.TensorProduct.map f g
4142

42-
/-- Auxiliary definition used to fight a timeout when building
43-
`AlgebraCat.instMonoidalCategory`. -/
44-
@[simps!]
45-
abbrev tensorUnit : AlgebraCat.{u} R := of R R
43+
open MonoidalCategory
4644

47-
/-- Auxiliary definition used to fight a timeout when building
48-
`AlgebraCat.instMonoidalCategory`. -/
49-
noncomputable abbrev associator (X Y Z : AlgebraCat.{u} R) :
50-
tensorObj (tensorObj X Y) Z ≅ tensorObj X (tensorObj Y Z) :=
51-
(Algebra.TensorProduct.assoc R X Y Z).toAlgebraIso
45+
end instMonoidalCategory
5246

53-
open MonoidalCategory
47+
open instMonoidalCategory
48+
49+
instance : MonoidalCategoryStruct (AlgebraCat.{u} R) where
50+
tensorObj := instMonoidalCategory.tensorObj
51+
whiskerLeft X _ _ f := tensorHom (𝟙 X) f
52+
whiskerRight {X₁ X₂} (f : X₁ ⟶ X₂) Y := tensorHom f (𝟙 Y)
53+
tensorHom := tensorHom
54+
tensorUnit := of R R
55+
associator X Y Z := (Algebra.TensorProduct.assoc R X Y Z).toAlgebraIso
56+
leftUnitor X := (Algebra.TensorProduct.lid R X).toAlgebraIso
57+
rightUnitor X := (Algebra.TensorProduct.rid R R X).toAlgebraIso
5458

5559
theorem forget₂_map_associator_hom (X Y Z : AlgebraCat.{u} R) :
56-
(forget₂ (AlgebraCat R) (ModuleCat R)).map (associator X Y Z).hom =
60+
(forget₂ (AlgebraCat R) (ModuleCat R)).map (α_ X Y Z).hom =
5761
(α_
5862
(forget₂ _ (ModuleCat R) |>.obj X)
5963
(forget₂ _ (ModuleCat R) |>.obj Y)
6064
(forget₂ _ (ModuleCat R) |>.obj Z)).hom := by
6165
rfl
6266

6367
theorem forget₂_map_associator_inv (X Y Z : AlgebraCat.{u} R) :
64-
(forget₂ (AlgebraCat R) (ModuleCat R)).map (associator X Y Z).inv =
68+
(forget₂ (AlgebraCat R) (ModuleCat R)).map (α_ X Y Z).inv =
6569
(α_
6670
(forget₂ _ (ModuleCat R) |>.obj X)
6771
(forget₂ _ (ModuleCat R) |>.obj Y)
6872
(forget₂ _ (ModuleCat R) |>.obj Z)).inv := by
6973
rfl
7074

71-
end instMonoidalCategory
72-
73-
open instMonoidalCategory
74-
7575
set_option maxHeartbeats 800000 in
7676
noncomputable instance instMonoidalCategory : MonoidalCategory (AlgebraCat.{u} R) :=
7777
Monoidal.induced
7878
(forget₂ (AlgebraCat R) (ModuleCat R))
79-
{ tensorObj := instMonoidalCategory.tensorObj
80-
μIsoSymm := fun X Y => Iso.refl _
81-
whiskerLeft := fun X _ _ f => tensorHom (𝟙 _) f
82-
whiskerRight := @fun X₁ X₂ (f : X₁ ⟶ X₂) Y => tensorHom f (𝟙 _)
83-
tensorHom := tensorHom
84-
tensorUnit' := tensorUnit
79+
{ μIsoSymm := fun X Y => Iso.refl _
8580
εIsoSymm := Iso.refl _
86-
associator := associator
8781
associator_eq := fun X Y Z => by
8882
dsimp only [forget₂_module_obj, forget₂_map_associator_hom]
8983
simp only [eqToIso_refl, Iso.refl_trans, Iso.refl_symm, Iso.trans_hom, tensorIso_hom,
9084
Iso.refl_hom, MonoidalCategory.tensor_id]
9185
erw [Category.id_comp, Category.comp_id, MonoidalCategory.tensor_id, Category.comp_id]
92-
leftUnitor := fun X => (Algebra.TensorProduct.lid R X).toAlgebraIso
93-
rightUnitor := fun X => (Algebra.TensorProduct.rid R R X).toAlgebraIso
9486
rightUnitor_eq := fun X => by
9587
dsimp
9688
erw [Category.id_comp, MonoidalCategory.tensor_id, Category.id_comp]
@@ -99,8 +91,9 @@ noncomputable instance instMonoidalCategory : MonoidalCategory (AlgebraCat.{u} R
9991

10092
variable (R) in
10193
/-- `forget₂ (AlgebraCat R) (ModuleCat R)` as a monoidal functor. -/
102-
def toModuleCatMonoidalFunctor : MonoidalFunctor (AlgebraCat.{u} R) (ModuleCat.{u} R) :=
103-
Monoidal.fromInduced (forget₂ (AlgebraCat R) (ModuleCat R)) _
94+
def toModuleCatMonoidalFunctor : MonoidalFunctor (AlgebraCat.{u} R) (ModuleCat.{u} R) := by
95+
unfold instMonoidalCategory
96+
exact Monoidal.fromInduced (forget₂ (AlgebraCat R) (ModuleCat R)) _
10497

10598
instance : Faithful (toModuleCatMonoidalFunctor R).toFunctor :=
10699
forget₂_faithful _ _

Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean

Lines changed: 21 additions & 20 deletions
Original file line numberDiff line numberDiff line change
@@ -92,6 +92,26 @@ def associator (M : ModuleCat.{v} R) (N : ModuleCat.{w} R) (K : ModuleCat.{x} R)
9292
(TensorProduct.assoc R M N K).toModuleIso
9393
#align Module.monoidal_category.associator ModuleCat.MonoidalCategory.associator
9494

95+
/-- (implementation) the left unitor for R-modules -/
96+
def leftUnitor (M : ModuleCat.{u} R) : ModuleCat.of R (R ⊗[R] M) ≅ M :=
97+
(LinearEquiv.toModuleIso (TensorProduct.lid R M) : of R (R ⊗ M) ≅ of R M).trans (ofSelfIso M)
98+
#align Module.monoidal_category.left_unitor ModuleCat.MonoidalCategory.leftUnitor
99+
100+
/-- (implementation) the right unitor for R-modules -/
101+
def rightUnitor (M : ModuleCat.{u} R) : ModuleCat.of R (M ⊗[R] R) ≅ M :=
102+
(LinearEquiv.toModuleIso (TensorProduct.rid R M) : of R (M ⊗ R) ≅ of R M).trans (ofSelfIso M)
103+
#align Module.monoidal_category.right_unitor ModuleCat.MonoidalCategory.rightUnitor
104+
105+
instance : MonoidalCategoryStruct (ModuleCat.{u} R) where
106+
tensorObj := tensorObj
107+
whiskerLeft := whiskerLeft
108+
whiskerRight := whiskerRight
109+
tensorHom f g := TensorProduct.map f g
110+
tensorUnit := ModuleCat.of R R
111+
associator := associator
112+
leftUnitor := leftUnitor
113+
rightUnitor := rightUnitor
114+
95115
section
96116

97117
/-! The `associator_naturality` and `pentagon` lemmas below are very slow to elaborate.
@@ -143,11 +163,6 @@ theorem pentagon (W X Y Z : ModuleCat R) :
143163
convert pentagon_aux R W X Y Z using 1
144164
#align Module.monoidal_category.pentagon ModuleCat.MonoidalCategory.pentagon
145165

146-
/-- (implementation) the left unitor for R-modules -/
147-
def leftUnitor (M : ModuleCat.{u} R) : ModuleCat.of R (R ⊗[R] M) ≅ M :=
148-
(LinearEquiv.toModuleIso (TensorProduct.lid R M) : of R (R ⊗ M) ≅ of R M).trans (ofSelfIso M)
149-
#align Module.monoidal_category.left_unitor ModuleCat.MonoidalCategory.leftUnitor
150-
151166
theorem leftUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) :
152167
tensorHom (𝟙 (ModuleCat.of R R)) f ≫ (leftUnitor N).hom = (leftUnitor M).hom ≫ f := by
153168
-- Porting note: broken ext
@@ -162,11 +177,6 @@ theorem leftUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) :
162177
rfl
163178
#align Module.monoidal_category.left_unitor_naturality ModuleCat.MonoidalCategory.leftUnitor_naturality
164179

165-
/-- (implementation) the right unitor for R-modules -/
166-
def rightUnitor (M : ModuleCat.{u} R) : ModuleCat.of R (M ⊗[R] R) ≅ M :=
167-
(LinearEquiv.toModuleIso (TensorProduct.rid R M) : of R (M ⊗ R) ≅ of R M).trans (ofSelfIso M)
168-
#align Module.monoidal_category.right_unitor ModuleCat.MonoidalCategory.rightUnitor
169-
170180
theorem rightUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) :
171181
tensorHom f (𝟙 (ModuleCat.of R R)) ≫ (rightUnitor N).hom = (rightUnitor M).hom ≫ f := by
172182
-- Porting note: broken ext
@@ -197,17 +207,8 @@ end MonoidalCategory
197207

198208
open MonoidalCategory
199209

210+
set_option maxHeartbeats 400000 in
200211
instance monoidalCategory : MonoidalCategory (ModuleCat.{u} R) := MonoidalCategory.ofTensorHom
201-
-- data
202-
(tensorObj := MonoidalCategory.tensorObj)
203-
(tensorHom := @tensorHom _ _)
204-
(whiskerLeft := @whiskerLeft _ _)
205-
(whiskerRight := @whiskerRight _ _)
206-
(tensorUnit' := ModuleCat.of R R)
207-
(associator := associator)
208-
(leftUnitor := leftUnitor)
209-
(rightUnitor := rightUnitor)
210-
-- properties
211212
(tensor_id := fun M N ↦ tensor_id M N)
212213
(tensor_comp := fun f g h ↦ MonoidalCategory.tensor_comp f g h)
213214
(associator_naturality := fun f g h ↦ MonoidalCategory.associator_naturality f g h)

Mathlib/CategoryTheory/Bicategory/End.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -41,7 +41,7 @@ instance (X : C) : MonoidalCategory (EndMonoidal X) where
4141
tensorObj f g := f ≫ g
4242
whiskerLeft {f g h} η := f ◁ η
4343
whiskerRight {f g} η h := η ▷ h
44-
tensorUnit' := 𝟙 _
44+
tensorUnit := 𝟙 _
4545
associator f g h := α_ f g h
4646
leftUnitor f := λ_ f
4747
rightUnitor f := ρ_ f

0 commit comments

Comments
 (0)