@@ -69,9 +69,13 @@ structure LaxMonoidalFunctor extends C ⥤ D where
6969 ε : 𝟙_ D ⟶ obj (𝟙_ C)
7070 /-- tensorator -/
7171 μ : ∀ X Y : C, obj X ⊗ obj Y ⟶ obj (X ⊗ Y)
72- μ_natural :
73- ∀ {X Y X' Y' : C} (f : X ⟶ Y) (g : X' ⟶ Y'),
74- (map f ⊗ map g) ≫ μ Y Y' = μ X X' ≫ map (f ⊗ g) := by
72+ μ_natural_left :
73+ ∀ {X Y : C} (f : X ⟶ Y) (X' : C),
74+ (map f ⊗ 𝟙 (obj X')) ≫ μ Y X' = μ X X' ≫ map (f ⊗ 𝟙 X') := by
75+ aesop_cat
76+ μ_natural_right :
77+ ∀ {X Y : C} (X' : C) (f : X ⟶ Y) ,
78+ (𝟙 (obj X') ⊗ map f) ≫ μ X' Y = μ X' X ≫ map (𝟙 X' ⊗ f) := by
7579 aesop_cat
7680 /-- associativity of the tensorator -/
7781 associativity :
@@ -93,7 +97,8 @@ structure LaxMonoidalFunctor extends C ⥤ D where
9397initialize_simps_projections LaxMonoidalFunctor (+toFunctor, -obj, -map)
9498
9599--Porting note: was `[simp, reassoc.1]`
96- attribute [reassoc (attr := simp)] LaxMonoidalFunctor.μ_natural
100+ attribute [reassoc (attr := simp)] LaxMonoidalFunctor.μ_natural_left
101+ attribute [reassoc (attr := simp)] LaxMonoidalFunctor.μ_natural_right
97102
98103attribute [simp] LaxMonoidalFunctor.left_unitality
99104
@@ -109,6 +114,59 @@ section
109114
110115variable {C D}
111116
117+ @[reassoc (attr := simp)]
118+ theorem LaxMonoidalFunctor.μ_natural (F : LaxMonoidalFunctor C D) {X Y X' Y' : C}
119+ (f : X ⟶ Y) (g : X' ⟶ Y') :
120+ (F.map f ⊗ F.map g) ≫ F.μ Y Y' = F.μ X X' ≫ F.map (f ⊗ g) := by
121+ rw [← tensor_id_comp_id_tensor_assoc]
122+ rw [F.μ_natural_right, F.μ_natural_left_assoc]
123+ rw [← F.map_comp, tensor_id_comp_id_tensor]
124+
125+ /--
126+ A constructor for lax monoidal functors whose axioms are described by `tensorHom` instead of
127+ `whiskerLeft` and `whiskerRight`.
128+ -/
129+ @[simps]
130+ def LaxMonoidalFunctor.ofTensorHom (F : C ⥤ D)
131+ /- unit morphism -/
132+ (ε : 𝟙_ D ⟶ F.obj (𝟙_ C))
133+ /- tensorator -/
134+ (μ : ∀ X Y : C, F.obj X ⊗ F.obj Y ⟶ F.obj (X ⊗ Y))
135+ (μ_natural :
136+ ∀ {X Y X' Y' : C} (f : X ⟶ Y) (g : X' ⟶ Y'),
137+ (F.map f ⊗ F.map g) ≫ μ Y Y' = μ X X' ≫ F.map (f ⊗ g) := by
138+ aesop_cat)
139+ /- associativity of the tensorator -/
140+ (associativity :
141+ ∀ X Y Z : C,
142+ (μ X Y ⊗ 𝟙 (F.obj Z)) ≫ μ (X ⊗ Y) Z ≫ F.map (α_ X Y Z).hom =
143+ (α_ (F.obj X) (F.obj Y) (F.obj Z)).hom ≫ (𝟙 (F.obj X) ⊗ μ Y Z) ≫ μ X (Y ⊗ Z) := by
144+ aesop_cat)
145+ /- unitality -/
146+ (left_unitality :
147+ ∀ X : C, (λ_ (F.obj X)).hom = (ε ⊗ 𝟙 (F.obj X)) ≫ μ (𝟙_ C) X ≫ F.map (λ_ X).hom :=
148+ by aesop_cat)
149+ (right_unitality :
150+ ∀ X : C, (ρ_ (F.obj X)).hom = (𝟙 (F.obj X) ⊗ ε) ≫ μ X (𝟙_ C) ≫ F.map (ρ_ X).hom :=
151+ by aesop_cat) :
152+ LaxMonoidalFunctor C D where
153+ obj := F.obj
154+ map := F.map
155+ map_id := F.map_id
156+ map_comp := F.map_comp
157+ ε := ε
158+ μ := μ
159+ μ_natural_left := fun f X' => by
160+ simp_rw [← F.map_id, μ_natural]
161+ μ_natural_right := fun X' f => by
162+ simp_rw [← F.map_id, μ_natural]
163+ associativity := fun X Y Z => by
164+ simp_rw [associativity]
165+ left_unitality := fun X => by
166+ simp_rw [left_unitality]
167+ right_unitality := fun X => by
168+ simp_rw [right_unitality]
169+
112170--Porting note: was `[simp, reassoc.1]`
113171@[reassoc (attr := simp)]
114172theorem LaxMonoidalFunctor.left_unitality_inv (F : LaxMonoidalFunctor C D) (X : C) :
@@ -320,10 +378,12 @@ def comp : LaxMonoidalFunctor.{v₁, v₃} C E :=
320378 { F.toFunctor ⋙ G.toFunctor with
321379 ε := G.ε ≫ G.map F.ε
322380 μ := fun X Y => G.μ (F.obj X) (F.obj Y) ≫ G.map (F.μ X Y)
323- μ_natural := @fun _ _ _ _ f g => by
324- simp only [Functor.comp_map, assoc]
325- rw [← Category.assoc, LaxMonoidalFunctor.μ_natural, Category.assoc, ← map_comp, ← map_comp,
326- ← LaxMonoidalFunctor.μ_natural]
381+ μ_natural_left := by
382+ intro X Y f X'
383+ simp_rw [comp_obj, F.comp_map, μ_natural_left_assoc, assoc, ← G.map_comp, μ_natural_left]
384+ μ_natural_right := by
385+ intro X Y f X'
386+ simp_rw [comp_obj, F.comp_map, μ_natural_right_assoc, assoc, ← G.map_comp, μ_natural_right]
327387 associativity := fun X Y Z => by
328388 dsimp
329389 rw [id_tensor_comp]
@@ -482,17 +542,18 @@ end MonoidalFunctor
482542/-- If we have a right adjoint functor `G` to a monoidal functor `F`, then `G` has a lax monoidal
483543structure as well.
484544-/
485- @[simps ]
545+ @[simp ]
486546noncomputable def monoidalAdjoint (F : MonoidalFunctor C D) {G : D ⥤ C} (h : F.toFunctor ⊣ G) :
487- LaxMonoidalFunctor D C where
488- toFunctor := G
489- ε := h.homEquiv _ _ (inv F.ε)
490- μ X Y := h.homEquiv _ (X ⊗ Y) (inv (F.μ (G.obj X) (G.obj Y)) ≫ (h.counit.app X ⊗ h.counit.app Y))
491- μ_natural := @fun X Y X' Y' f g => by
547+ LaxMonoidalFunctor D C := LaxMonoidalFunctor.ofTensorHom
548+ (F := G)
549+ (ε := h.homEquiv _ _ (inv F.ε))
550+ (μ := fun X Y ↦
551+ h.homEquiv _ (X ⊗ Y) (inv (F.μ (G.obj X) (G.obj Y)) ≫ (h.counit.app X ⊗ h.counit.app Y)))
552+ (μ_natural := @fun X Y X' Y' f g => by
492553 rw [← h.homEquiv_naturality_left, ← h.homEquiv_naturality_right, Equiv.apply_eq_iff_eq, assoc,
493554 IsIso.eq_inv_comp, ← F.toLaxMonoidalFunctor.μ_natural_assoc, IsIso.hom_inv_id_assoc, ←
494- tensor_comp, Adjunction.counit_naturality, Adjunction.counit_naturality, tensor_comp]
495- associativity X Y Z := by
555+ tensor_comp, Adjunction.counit_naturality, Adjunction.counit_naturality, tensor_comp])
556+ ( associativity := fun X Y Z ↦ by
496557 dsimp only
497558 rw [← h.homEquiv_naturality_right, ← h.homEquiv_naturality_left, ←
498559 h.homEquiv_naturality_left, ← h.homEquiv_naturality_left, Equiv.apply_eq_iff_eq, ←
@@ -505,21 +566,21 @@ noncomputable def monoidalAdjoint (F : MonoidalFunctor C D) {G : D ⥤ C} (h : F
505566 tensor_comp_assoc, id_comp, id_comp, h.homEquiv_unit, h.homEquiv_unit, Functor.map_comp,
506567 assoc, assoc, h.counit_naturality, h.left_triangle_components_assoc, Functor.map_comp,
507568 assoc, h.counit_naturality, h.left_triangle_components_assoc]
508- simp
509- left_unitality X := by
569+ simp)
570+ ( left_unitality := fun X ↦ by
510571 rw [← h.homEquiv_naturality_right, ← h.homEquiv_naturality_left, ← Equiv.symm_apply_eq,
511572 h.homEquiv_counit, F.map_leftUnitor, h.homEquiv_unit, assoc, assoc, assoc, F.map_tensor,
512573 assoc, assoc, IsIso.hom_inv_id_assoc, ← tensor_comp_assoc, Functor.map_id, id_comp,
513574 Functor.map_comp, assoc, h.counit_naturality, h.left_triangle_components_assoc, ←
514575 leftUnitor_naturality, ← tensor_comp_assoc, id_comp, comp_id]
515- simp
516- right_unitality X := by
576+ simp)
577+ ( right_unitality := fun X ↦ by
517578 rw [← h.homEquiv_naturality_right, ← h.homEquiv_naturality_left, ← Equiv.symm_apply_eq,
518579 h.homEquiv_counit, F.map_rightUnitor, assoc, assoc, ← rightUnitor_naturality, ←
519580 tensor_comp_assoc, comp_id, id_comp, h.homEquiv_unit, F.map_tensor, assoc, assoc, assoc,
520581 IsIso.hom_inv_id_assoc, Functor.map_comp, Functor.map_id, ← tensor_comp_assoc, assoc,
521582 h.counit_naturality, h.left_triangle_components_assoc, id_comp]
522- simp
583+ simp)
523584#align category_theory.monoidal_adjoint CategoryTheory.monoidalAdjoint
524585
525586/-- If a monoidal functor `F` is an equivalence of categories then its inverse is also monoidal. -/
0 commit comments