@@ -14,6 +14,7 @@ import Mathlib.RingTheory.TensorProduct
14
14
-/
15
15
16
16
open CategoryTheory
17
+ open scoped MonoidalCategory
17
18
18
19
universe v u
19
20
@@ -39,58 +40,49 @@ noncomputable abbrev tensorHom {W X Y Z : AlgebraCat.{u} R} (f : W ⟶ X) (g : Y
39
40
tensorObj W Y ⟶ tensorObj X Z :=
40
41
Algebra.TensorProduct.map f g
41
42
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
46
44
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
52
46
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
54
58
55
59
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 =
57
61
(α_
58
62
(forget₂ _ (ModuleCat R) |>.obj X)
59
63
(forget₂ _ (ModuleCat R) |>.obj Y)
60
64
(forget₂ _ (ModuleCat R) |>.obj Z)).hom := by
61
65
rfl
62
66
63
67
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 =
65
69
(α_
66
70
(forget₂ _ (ModuleCat R) |>.obj X)
67
71
(forget₂ _ (ModuleCat R) |>.obj Y)
68
72
(forget₂ _ (ModuleCat R) |>.obj Z)).inv := by
69
73
rfl
70
74
71
- end instMonoidalCategory
72
-
73
- open instMonoidalCategory
74
-
75
75
set_option maxHeartbeats 800000 in
76
76
noncomputable instance instMonoidalCategory : MonoidalCategory (AlgebraCat.{u} R) :=
77
77
Monoidal.induced
78
78
(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 _
85
80
εIsoSymm := Iso.refl _
86
- associator := associator
87
81
associator_eq := fun X Y Z => by
88
82
dsimp only [forget₂_module_obj, forget₂_map_associator_hom]
89
83
simp only [eqToIso_refl, Iso.refl_trans, Iso.refl_symm, Iso.trans_hom, tensorIso_hom,
90
84
Iso.refl_hom, MonoidalCategory.tensor_id]
91
85
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
94
86
rightUnitor_eq := fun X => by
95
87
dsimp
96
88
erw [Category.id_comp, MonoidalCategory.tensor_id, Category.id_comp]
@@ -99,8 +91,9 @@ noncomputable instance instMonoidalCategory : MonoidalCategory (AlgebraCat.{u} R
99
91
100
92
variable (R) in
101
93
/-- `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)) _
104
97
105
98
instance : Faithful (toModuleCatMonoidalFunctor R).toFunctor :=
106
99
forget₂_faithful _ _
0 commit comments