Skip to content

Commit 2e59248

Browse files
committed
refactor(CategoryTheory/Monoidal): split the naturality condition of monoidal functors (#9988)
Extracted from #6307. We replace `μ_natural` with `μ_natural_left` and `μ_natural_right` since we prefer to use the whiskerings to the tensor of morphisms in the refactor #6307.
1 parent 8c661e5 commit 2e59248

File tree

12 files changed

+229
-90
lines changed

12 files changed

+229
-90
lines changed

Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean

Lines changed: 7 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -186,15 +186,15 @@ theorem associativity (X Y Z : Type u) :
186186
-- In fact, it's strong monoidal, but we don't yet have a typeclass for that.
187187
/-- The free R-module functor is lax monoidal. -/
188188
@[simps]
189-
instance : LaxMonoidal.{u} (free R).obj where
189+
instance : LaxMonoidal.{u} (free R).obj := .ofTensorHom
190190
-- Send `R` to `PUnit →₀ R`
191-
ε := ε R
191+
(ε := ε R)
192192
-- Send `(α →₀ R) ⊗ (β →₀ R)` to `α × β →₀ R`
193-
μ X Y := (μ R X Y).hom
194-
μ_natural {_} {_} {_} {_} f g := μ_natural R f g
195-
left_unitality := left_unitality R
196-
right_unitality := right_unitality R
197-
associativity := associativity R
193+
(μ := fun X Y => (μ R X Y).hom)
194+
(μ_natural := fun {_} {_} {_} {_} f g μ_natural R f g)
195+
(left_unitality := left_unitality R)
196+
(right_unitality := right_unitality R)
197+
(associativity := associativity R)
198198

199199
instance : IsIso (@LaxMonoidal.ε _ _ _ _ _ _ (free R).obj _ _) := by
200200
refine' ⟨⟨Finsupp.lapply PUnit.unit, ⟨_, _⟩⟩⟩

Mathlib/CategoryTheory/Bicategory/SingleObj.lean

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -86,7 +86,12 @@ def endMonoidalStarFunctor : MonoidalFunctor (EndMonoidal (MonoidalSingleObj.sta
8686
map f := f
8787
ε := 𝟙 _
8888
μ X Y := 𝟙 _
89-
μ_natural f g := by
89+
-- The proof will be automated after merging #6307.
90+
μ_natural_left f g := by
91+
simp_rw [Category.id_comp, Category.comp_id]
92+
-- Should we provide further simp lemmas so this goal becomes visible?
93+
exact (tensor_id_comp_id_tensor _ _).symm
94+
μ_natural_right f g := by
9095
simp_rw [Category.id_comp, Category.comp_id]
9196
-- Should we provide further simp lemmas so this goal becomes visible?
9297
exact (tensor_id_comp_id_tensor _ _).symm

Mathlib/CategoryTheory/Monoidal/Braided.lean

Lines changed: 15 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -421,6 +421,18 @@ theorem tensor_μ_natural {X₁ X₂ Y₁ Y₂ U₁ U₂ V₁ V₂ : C} (f₁ :
421421
simp only [assoc]
422422
#align category_theory.tensor_μ_natural CategoryTheory.tensor_μ_natural
423423

424+
@[reassoc]
425+
theorem tensor_μ_natural_left {X₁ X₂ Y₁ Y₂ : C} (f₁: X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) (Z₁ Z₂ : C) :
426+
((f₁ ⊗ f₂) ⊗ 𝟙 (Z₁ ⊗ Z₂)) ≫ tensor_μ C (Y₁, Y₂) (Z₁, Z₂) =
427+
tensor_μ C (X₁, X₂) (Z₁, Z₂) ≫ ((f₁ ⊗ 𝟙 Z₁) ⊗ (f₂ ⊗ 𝟙 Z₂)) := by
428+
convert tensor_μ_natural C f₁ f₂ (𝟙 Z₁) (𝟙 Z₂) using 1; simp
429+
430+
@[reassoc]
431+
theorem tensor_μ_natural_right (Z₁ Z₂ : C) {X₁ X₂ Y₁ Y₂ : C} (f₁ : X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) :
432+
(𝟙 (Z₁ ⊗ Z₂) ⊗ (f₁ ⊗ f₂)) ≫ tensor_μ C (Z₁, Z₂) (Y₁, Y₂) =
433+
tensor_μ C (Z₁, Z₂) (X₁, X₂) ≫ ((𝟙 Z₁ ⊗ f₁) ⊗ (𝟙 Z₂ ⊗ f₂)) := by
434+
convert tensor_μ_natural C (𝟙 Z₁) (𝟙 Z₂) f₁ f₂ using 1; simp
435+
424436
theorem tensor_left_unitality (X₁ X₂ : C) :
425437
(λ_ (X₁ ⊗ X₂)).hom =
426438
((λ_ (𝟙_ C)).inv ⊗ 𝟙 (X₁ ⊗ X₂)) ≫
@@ -566,8 +578,9 @@ theorem tensor_associativity (X₁ X₂ Y₁ Y₂ Z₁ Z₂ : C) :
566578
def tensorMonoidal : MonoidalFunctor (C × C) C :=
567579
{ tensor C with
568580
ε := (λ_ (𝟙_ C)).inv
569-
μ := fun X Y => tensor_μ C X Y
570-
μ_natural := fun f g => tensor_μ_natural C f.1 f.2 g.1 g.2
581+
μ := tensor_μ C
582+
μ_natural_left := fun f Z => tensor_μ_natural_left C f.1 f.2 Z.1 Z.2
583+
μ_natural_right := fun Z f => tensor_μ_natural_right C Z.1 Z.2 f.1 f.2
571584
associativity := fun X Y Z => tensor_associativity C X.1 X.2 Y.1 Y.2 Z.1 Z.2
572585
left_unitality := fun ⟨X₁, X₂⟩ => tensor_left_unitality C X₁ X₂
573586
right_unitality := fun ⟨X₁, X₂⟩ => tensor_right_unitality C X₁ X₂

Mathlib/CategoryTheory/Monoidal/End.lean

Lines changed: 10 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -97,11 +97,19 @@ def tensoringRightMonoidal [MonoidalCategory.{v} C] : MonoidalFunctor C (C ⥤ C
9797
{ tensoringRight C with
9898
ε := (rightUnitorNatIso C).inv
9999
μ := fun X Y => { app := fun Z => (α_ Z X Y).hom }
100-
μ_natural := fun f g => by
100+
-- The proof will be automated after merging #6307.
101+
μ_natural_left := fun f X => by
101102
ext Z
102103
dsimp
103-
simp only [← id_tensor_comp_tensor_id g f, id_tensor_comp, ← tensor_id, Category.assoc,
104+
simp only [← id_tensor_comp_tensor_id f (𝟙 X), id_tensor_comp, ← tensor_id, Category.assoc,
104105
associator_naturality, associator_naturality_assoc]
106+
simp only [tensor_id, Category.id_comp]
107+
μ_natural_right := fun X g => by
108+
ext Z
109+
dsimp
110+
simp only [← id_tensor_comp_tensor_id (𝟙 X) g, id_tensor_comp, ← tensor_id, Category.assoc,
111+
associator_naturality, associator_naturality_assoc]
112+
simp only [tensor_id, Category.comp_id]
105113
associativity := fun X Y Z => by
106114
ext W
107115
simp [pentagon]

Mathlib/CategoryTheory/Monoidal/Free/Basic.lean

Lines changed: 10 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -341,13 +341,17 @@ def project : MonoidalFunctor (F C) D where
341341
map_comp := by rintro _ _ _ ⟨_⟩ ⟨_⟩; rfl
342342
ε := 𝟙 _
343343
μ X Y := 𝟙 _
344-
μ_natural := @fun _ _ _ _ f g => by
344+
μ_natural_left := fun f _ => by
345345
induction' f using Quotient.recOn
346-
· induction' g using Quotient.recOn
347-
· dsimp
348-
simp
349-
rfl
350-
· rfl
346+
· dsimp
347+
simp
348+
rfl
349+
· rfl
350+
μ_natural_right := fun _ f => by
351+
induction' f using Quotient.recOn
352+
· dsimp
353+
simp
354+
rfl
351355
· rfl
352356
#align category_theory.free_monoidal_category.project CategoryTheory.FreeMonoidalCategory.project
353357

Mathlib/CategoryTheory/Monoidal/Functor.lean

Lines changed: 82 additions & 21 deletions
Original file line numberDiff line numberDiff line change
@@ -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
9397
initialize_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

98103
attribute [simp] LaxMonoidalFunctor.left_unitality
99104

@@ -109,6 +114,59 @@ section
109114

110115
variable {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)]
114172
theorem 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
483543
structure as well.
484544
-/
485-
@[simps]
545+
@[simp]
486546
noncomputable 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. -/

Mathlib/CategoryTheory/Monoidal/Functorial.lean

Lines changed: 41 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -55,11 +55,14 @@ class LaxMonoidal (F : C → D) [Functorial.{v₁, v₂} F] where
5555
ε : 𝟙_ D ⟶ F (𝟙_ C)
5656
/-- tensorator -/
5757
μ : ∀ X Y : C, F X ⊗ F Y ⟶ F (X ⊗ Y)
58-
/-- naturality -/
59-
μ_natural :
60-
∀ {X Y X' Y' : C} (f : X ⟶ Y) (g : X' ⟶ Y'),
61-
(map F f ⊗ map F g) ≫ μ Y Y' = μ X X' ≫ map F (f ⊗ g) := by
62-
aesop_cat
58+
μ_natural_left :
59+
∀ {X Y : C} (f : X ⟶ Y) (X' : C),
60+
(map F f ⊗ 𝟙 (F X')) ≫ μ Y X' = μ X X' ≫ map F (f ⊗ 𝟙 X') := by
61+
aesop_cat
62+
μ_natural_right :
63+
∀ {X Y : C} (X' : C) (f : X ⟶ Y) ,
64+
(𝟙 (F X') ⊗ map F f) ≫ μ X' Y = μ X' X ≫ map F (𝟙 X' ⊗ f) := by
65+
aesop_cat
6366
/-- associativity of the tensorator -/
6467
associativity :
6568
∀ X Y Z : C,
@@ -74,9 +77,39 @@ class LaxMonoidal (F : C → D) [Functorial.{v₁, v₂} F] where
7477
aesop_cat
7578
#align category_theory.lax_monoidal CategoryTheory.LaxMonoidal
7679

77-
attribute [simp] LaxMonoidal.μ_natural
78-
79-
attribute [simp] LaxMonoidal.μ_natural
80+
/-- An unbundled description of lax monoidal functors. -/
81+
abbrev LaxMonoidal.ofTensorHom (F : C → D) [Functorial.{v₁, v₂} F]
82+
/- unit morphism -/
83+
(ε : 𝟙_ D ⟶ F (𝟙_ C))
84+
/- tensorator -/
85+
(μ : ∀ X Y : C, F X ⊗ F Y ⟶ F (X ⊗ Y))
86+
/- naturality -/
87+
(μ_natural :
88+
∀ {X Y X' Y' : C} (f : X ⟶ Y) (g : X' ⟶ Y'),
89+
(map F f ⊗ map F g) ≫ μ Y Y' = μ X X' ≫ map F (f ⊗ g) := by
90+
aesop_cat)
91+
/- associativity of the tensorator -/
92+
(associativity :
93+
∀ X Y Z : C,
94+
(μ X Y ⊗ 𝟙 (F Z)) ≫ μ (X ⊗ Y) Z ≫ map F (α_ X Y Z).hom =
95+
(α_ (F X) (F Y) (F Z)).hom ≫ (𝟙 (F X) ⊗ μ Y Z) ≫ μ X (Y ⊗ Z) := by
96+
aesop_cat)
97+
/- left unitality -/
98+
(left_unitality : ∀ X : C, (λ_ (F X)).hom = (ε ⊗ 𝟙 (F X)) ≫ μ (𝟙_ C) X ≫ map F (λ_ X).hom := by
99+
aesop_cat)
100+
/- right unitality -/
101+
(right_unitality : ∀ X : C, (ρ_ (F X)).hom = (𝟙 (F X) ⊗ ε) ≫ μ X (𝟙_ C) ≫ map F (ρ_ X).hom := by
102+
aesop_cat) :
103+
LaxMonoidal.{v₁, v₂} F where
104+
ε := ε
105+
μ := μ
106+
μ_natural_left f X := by intros; simpa using μ_natural f (𝟙 X)
107+
μ_natural_right X f := by intros; simpa using μ_natural (𝟙 X) f
108+
associativity X Y Z := by intros; simpa using associativity X Y Z
109+
left_unitality X := by intros; simpa using left_unitality X
110+
right_unitality X := by intros; simpa using right_unitality X
111+
112+
attribute [simp, nolint simpNF] LaxMonoidal.μ_natural_left LaxMonoidal.μ_natural_right
80113

81114
-- The unitality axioms cannot be used as simp lemmas because they require
82115
-- higher-order matching to figure out the `F` and `X` from `F X`.

Mathlib/CategoryTheory/Monoidal/Limits.lean

Lines changed: 13 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -47,24 +47,24 @@ theorem limitFunctorial_map {F G : J ⥤ C} (α : F ⟶ G) :
4747
variable [MonoidalCategory.{v} C]
4848

4949
@[simps]
50-
instance limitLaxMonoidal : LaxMonoidal fun F : J ⥤ C => limit F where
51-
ε :=
50+
instance limitLaxMonoidal : LaxMonoidal fun F : J ⥤ C => limit F := .ofTensorHom
51+
(ε :=
5252
limit.lift _
5353
{ pt := _
54-
π := { app := fun j => 𝟙 _ } }
55-
μ F G :=
54+
π := { app := fun j => 𝟙 _ } })
55+
(μ := fun F G =>
5656
limit.lift (F ⊗ G)
5757
{ pt := limit F ⊗ limit G
5858
π :=
5959
{ app := fun j => limit.π F j ⊗ limit.π G j
6060
naturality := fun j j' f => by
6161
dsimp
62-
simp only [Category.id_comp, ← tensor_comp, limit.w] } }
63-
μ_natural f g := by
62+
simp only [Category.id_comp, ← tensor_comp, limit.w] } })
63+
(μ_natural:= fun f g => by
6464
ext; dsimp
6565
simp only [limit.lift_π, Cones.postcompose_obj_π, Monoidal.tensorHom_app, limit.lift_map,
66-
NatTrans.comp_app, Category.assoc, ← tensor_comp, limMap_π]
67-
associativity X Y Z := by
66+
NatTrans.comp_app, Category.assoc, ← tensor_comp, limMap_π])
67+
(associativity := fun X Y Z => by
6868
ext j; dsimp
6969
simp only [limit.lift_π, Cones.postcompose_obj_π, Monoidal.associator_hom_app, limit.lift_map,
7070
NatTrans.comp_app, Category.assoc]
@@ -78,8 +78,8 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J ⥤ C => limit F where
7878
slice_rhs 2 3 =>
7979
rw [← id_tensor_comp, limit.lift_π]
8080
dsimp
81-
dsimp; simp
82-
left_unitality X := by
81+
dsimp; simp)
82+
(left_unitality := fun X => by
8383
ext j; dsimp
8484
simp only [limit.lift_map, Category.assoc, limit.lift_π, Cones.postcompose_obj_pt,
8585
Cones.postcompose_obj_π, NatTrans.comp_app, Functor.const_obj_obj, Monoidal.tensorObj_obj,
@@ -90,8 +90,8 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J ⥤ C => limit F where
9090
erw [limit.lift_π]
9191
dsimp
9292
slice_rhs 2 3 => rw [leftUnitor_naturality]
93-
simp
94-
right_unitality X := by
93+
simp)
94+
(right_unitality := fun X => by
9595
ext j; dsimp
9696
simp only [limit.lift_map, Category.assoc, limit.lift_π, Cones.postcompose_obj_pt,
9797
Cones.postcompose_obj_π, NatTrans.comp_app, Functor.const_obj_obj, Monoidal.tensorObj_obj,
@@ -102,7 +102,7 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J ⥤ C => limit F where
102102
erw [limit.lift_π]
103103
dsimp
104104
slice_rhs 2 3 => rw [rightUnitor_naturality]
105-
simp
105+
simp)
106106
#align category_theory.limits.limit_lax_monoidal CategoryTheory.Limits.limitLaxMonoidal
107107

108108
/-- The limit functor `F ↦ limit F` bundled as a lax monoidal functor. -/

0 commit comments

Comments
 (0)