From a6c7eebdbb42162d30ab1447d0397a2cb3010bea Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 3 Aug 2023 00:00:54 +0900 Subject: [PATCH 01/62] monoidal structure on the Drinfeld center --- .../Algebra/Category/FGModuleCat/Basic.lean | 6 +- .../Category/ModuleCat/Adjunctions.lean | 21 +- .../Category/ModuleCat/Monoidal/Basic.lean | 94 ++- Mathlib/CategoryTheory/Bicategory/End.lean | 9 +- .../CategoryTheory/Bicategory/SingleObj.lean | 16 +- Mathlib/CategoryTheory/Closed/Monoidal.lean | 29 +- Mathlib/CategoryTheory/Monoidal/Braided.lean | 101 +-- Mathlib/CategoryTheory/Monoidal/Category.lean | 644 +++++++++++++++--- Mathlib/CategoryTheory/Monoidal/Center.lean | 166 +++-- .../Monoidal/CoherenceLemmas.lean | 14 +- Mathlib/CategoryTheory/Monoidal/Discrete.lean | 4 +- Mathlib/CategoryTheory/Monoidal/End.lean | 40 +- .../CategoryTheory/Monoidal/Free/Basic.lean | 192 ++++-- .../Monoidal/Free/Coherence.lean | 229 ++++--- Mathlib/CategoryTheory/Monoidal/Functor.lean | 201 ++++-- .../CategoryTheory/Monoidal/Functorial.lean | 74 +- Mathlib/CategoryTheory/Monoidal/Linear.lean | 22 +- .../Monoidal/NaturalTransformation.lean | 39 +- .../OfChosenFiniteProducts/Basic.lean | 28 +- .../Monoidal/OfHasFiniteProducts.lean | 124 ++-- .../CategoryTheory/Monoidal/Preadditive.lean | 93 ++- .../CategoryTheory/Monoidal/Rigid/Basic.lean | 191 +++--- .../CategoryTheory/Monoidal/Subcategory.lean | 50 +- .../CategoryTheory/Monoidal/Transport.lean | 199 ++++-- .../Monoidal/Types/Coyoneda.lean | 28 +- Mathlib/CategoryTheory/Shift/Basic.lean | 9 +- Mathlib/Tactic/CategoryTheory/Coherence.lean | 44 +- 27 files changed, 1792 insertions(+), 875 deletions(-) diff --git a/Mathlib/Algebra/Category/FGModuleCat/Basic.lean b/Mathlib/Algebra/Category/FGModuleCat/Basic.lean index a88d302649f8f..18aed4837bcc0 100644 --- a/Mathlib/Algebra/Category/FGModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/FGModuleCat/Basic.lean @@ -268,15 +268,15 @@ theorem FGModuleCatEvaluation_apply (f : FGModuleCatDual K V) (x : V) : set_option maxHeartbeats 1600000 in private theorem coevaluation_evaluation : letI V' : FGModuleCat K := FGModuleCatDual K V - (πŸ™ V' βŠ— FGModuleCatCoevaluation K V) ≫ (Ξ±_ V' V V').inv ≫ (FGModuleCatEvaluation K V βŠ— πŸ™ V') = + (V' ◁ FGModuleCatCoevaluation K V) ≫ (Ξ±_ V' V V').inv ≫ (FGModuleCatEvaluation K V β–· V') = (ρ_ V').hom ≫ (Ξ»_ V').inv := by apply contractLeft_assoc_coevaluation K V -- Porting note: extremely slow, was fast in mathlib3. set_option maxHeartbeats 1600000 in private theorem evaluation_coevaluation : - (FGModuleCatCoevaluation K V βŠ— πŸ™ V) ≫ - (Ξ±_ V (FGModuleCatDual K V) V).hom ≫ (πŸ™ V βŠ— FGModuleCatEvaluation K V) = + (FGModuleCatCoevaluation K V β–· V) ≫ + (Ξ±_ V (FGModuleCatDual K V) V).hom ≫ (V ◁ FGModuleCatEvaluation K V) = (Ξ»_ V).hom ≫ (ρ_ V).inv := by apply contractLeft_assoc_coevaluation' K V diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index b6ef518db20b7..fd61c102674c4 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -179,20 +179,23 @@ theorem associativity (X Y Z : Type u) : CategoryTheory.associator_hom_apply] #align Module.free.associativity ModuleCat.Free.associativity --- In fact, it's strong monoidal, but we don't yet have a typeclass for that. -/-- The free R-module functor is lax monoidal. -/ +/-- The free R-module functor is lax monoidal. The structure part. -/ @[simps] -instance : LaxMonoidal.{u} (free R).obj where +instance : LaxMonoidalStruct.{u} (free R).obj where -- Send `R` to `PUnit β†’β‚€ R` Ξ΅ := Ξ΅ R -- Send `(Ξ± β†’β‚€ R) βŠ— (Ξ² β†’β‚€ R)` to `Ξ± Γ— Ξ² β†’β‚€ R` ΞΌ X Y := (ΞΌ R X Y).hom - ΞΌ_natural {_} {_} {_} {_} f g := ΞΌ_natural R f g - left_unitality := left_unitality R - right_unitality := right_unitality R - associativity := associativity R -instance : IsIso (@LaxMonoidal.Ξ΅ _ _ _ _ _ _ (free R).obj _ _) := by +-- In fact, it's strong monoidal, but we don't yet have a typeclass for that. +/-- The free R-module functor is lax monoidal. The property part. -/ +instance : LaxMonoidal.{u} (free R).obj := .ofTensorHom + (ΞΌ_natural := fun {_} {_} {_} {_} f g ↦ ΞΌ_natural R f g) + (left_unitality := left_unitality R) + (right_unitality := right_unitality R) + (associativity := associativity R) + +instance : IsIso (@LaxMonoidalStruct.Ξ΅ _ _ _ _ _ _ (free R).obj _ _) := by refine' ⟨⟨Finsupp.lapply PUnit.unit, ⟨_, _⟩⟩⟩ Β· -- Porting note: broken ext apply LinearMap.ext_ring @@ -221,7 +224,7 @@ variable [CommRing R] def monoidalFree : MonoidalFunctor (Type u) (ModuleCat.{u} R) := { LaxMonoidalFunctor.of (free R).obj with -- Porting note: used to be dsimp - Ξ΅_isIso := (by infer_instance : IsIso (@LaxMonoidal.Ξ΅ _ _ _ _ _ _ (free R).obj _ _)) + Ξ΅_isIso := inferInstanceAs <| IsIso LaxMonoidalStruct.Ξ΅ ΞΌ_isIso := fun X Y => by dsimp; infer_instance } #align Module.monoidal_free ModuleCat.monoidalFree diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean index 133a748199f31..3eee418fe0149 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean @@ -61,6 +61,22 @@ def tensorHom {M N M' N' : ModuleCat R} (f : M ⟢ N) (g : M' ⟢ N') : TensorProduct.map f g #align Module.monoidal_category.tensor_hom ModuleCat.MonoidalCategory.tensorHom +/-- (implementation) left whiskering for R-modules -/ +def whiskerLeft (M : ModuleCat R) {N₁ Nβ‚‚ : ModuleCat R} (f : N₁ ⟢ Nβ‚‚) : + tensorObj M N₁ ⟢ tensorObj M Nβ‚‚ := + f.lTensor M + +/-- (implementation) right whiskering for R-modules -/ +def whiskerRight {M₁ Mβ‚‚ : ModuleCat R} (f : M₁ ⟢ Mβ‚‚) (N : ModuleCat R) : + tensorObj M₁ N ⟢ tensorObj Mβ‚‚ N := + f.rTensor N + +-- theorem whiskerLeft_id (M N : ModuleCat R) : tensorHom (πŸ™ M) (πŸ™ N) = πŸ™ (ModuleCat.of R (M βŠ— N)) := by +-- -- Porting note: even with high priority ext fails to find this +-- apply TensorProduct.ext +-- rfl + + theorem tensor_id (M N : ModuleCat R) : tensorHom (πŸ™ M) (πŸ™ N) = πŸ™ (ModuleCat.of R (M βŠ— N)) := by -- Porting note: even with high priority ext fails to find this apply TensorProduct.ext @@ -105,9 +121,9 @@ variable (R) private theorem pentagon_aux (W X Y Z : Type _) [AddCommMonoid W] [AddCommMonoid X] [AddCommMonoid Y] [AddCommMonoid Z] [Module R W] [Module R X] [Module R Y] [Module R Z] : - ((map (1 : W β†’β‚—[R] W) (assoc R X Y Z).toLinearMap).comp + (((assoc R X Y Z).toLinearMap.lTensor W).comp (assoc R W (X βŠ—[R] Y) Z).toLinearMap).comp - (map ↑(assoc R W X Y) (1 : Z β†’β‚—[R] Z)) = + ((assoc R W X Y).rTensor Z) = (assoc R W X (Y βŠ—[R] Z)).toLinearMap.comp (assoc R (W βŠ—[R] X) Y Z).toLinearMap := by apply TensorProduct.ext_fourfold intro w x y z @@ -127,10 +143,11 @@ theorem associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : ModuleCat R} (f -- Porting note: very slow! set_option maxHeartbeats 1600000 in theorem pentagon (W X Y Z : ModuleCat R) : - tensorHom (associator W X Y).hom (πŸ™ Z) ≫ - (associator W (tensorObj X Y) Z).hom ≫ tensorHom (πŸ™ W) (associator X Y Z).hom = + whiskerRight (associator W X Y).hom Z ≫ + (associator W (tensorObj X Y) Z).hom ≫ whiskerLeft W (associator X Y Z).hom = (associator (tensorObj W X) Y Z).hom ≫ (associator W X (tensorObj Y Z)).hom := by - convert pentagon_aux R W X Y Z using 1 + sorry + -- convert pentagon_aux R W X Y Z using 1 #align Module.monoidal_category.pentagon ModuleCat.MonoidalCategory.pentagon /-- (implementation) the left unitor for R-modules -/ @@ -183,26 +200,47 @@ theorem triangle (M N : ModuleCat.{u} R) : exact (TensorProduct.smul_tmul _ _ _).symm #align Module.monoidal_category.triangle ModuleCat.MonoidalCategory.triangle +-- theorem id_whiskerLeft_aux {M N : ModuleCat R} (f : M ⟢ N) : +-- whiskerLeft (of R R) f = TensorProduct.lift (LinearMap.complβ‚‚ (TensorProduct.mk _ _ _) f) := by +-- sorry + + +-- theorem id_whiskerLeft {M N : ModuleCat R} (f : M ⟢ N) : +-- whiskerLeft (of R R) f = (leftUnitor M).hom ≫ f ≫ (leftUnitor N).inv := by +-- -- rw [id_whiskerLeft_aux] +-- apply TensorProduct.ext +-- apply LinearMap.ext_ring +-- apply LinearMap.ext; intro x +-- dsimp +-- -- dsimp [LinearMap.comprβ‚‚, whiskerLeft] +-- change ((leftUnitor N).hom) ((tensorHom (πŸ™ (of R R)) f) ((1 : R) βŠ—β‚œ[R] x)) = +-- f (((leftUnitor M).hom) (1 βŠ—β‚œ[R] x)) +-- change _ = _ ≫ f _ ≫ _ +-- change ((leftUnitor N).hom) _ = f _ +-- erw [TensorProduct.lid_tmul, TensorProduct.lid_tmul] +-- rw [LinearMap.map_smul] +-- rfl + end MonoidalCategory open MonoidalCategory -instance monoidalCategory : MonoidalCategory (ModuleCat.{u} R) where +instance monoidalCategory : MonoidalCategory (ModuleCat.{u} R) := MonoidalCategory.ofTensorHom -- data - tensorObj := tensorObj - tensorHom := @tensorHom _ _ - tensorUnit' := ModuleCat.of R R - associator := associator - leftUnitor := leftUnitor - rightUnitor := rightUnitor + (tensorObj := MonoidalCategory.tensorObj) + (tensorHom := @tensorHom _ _) + (tensorUnit' := ModuleCat.of R R) + (associator := associator) + (leftUnitor := leftUnitor) + (rightUnitor := rightUnitor) -- properties - tensor_id M N := tensor_id M N - tensor_comp f g h := MonoidalCategory.tensor_comp f g h - associator_naturality f g h := MonoidalCategory.associator_naturality f g h - leftUnitor_naturality f := MonoidalCategory.leftUnitor_naturality f - rightUnitor_naturality f := rightUnitor_naturality f - pentagon M N K L := pentagon M N K L - triangle M N := triangle M N + (tensor_id := fun M N ↦ tensor_id M N) + (tensor_comp := fun f g h ↦ MonoidalCategory.tensor_comp f g h) + (associator_naturality := fun f g h ↦ MonoidalCategory.associator_naturality f g h) + (leftUnitor_naturality := fun f ↦ MonoidalCategory.leftUnitor_naturality f) + (rightUnitor_naturality := fun f ↦ rightUnitor_naturality f) + (pentagon := fun M N K L ↦ pentagon M N K L) + (triangle := fun M N ↦ triangle M N) #align Module.monoidal_category ModuleCat.monoidalCategory /-- Remind ourselves that the monoidal unit, being just `R`, is still a commutative ring. -/ @@ -263,24 +301,24 @@ instance : MonoidalPreadditive (ModuleCat.{u} R) := by Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] - rw [LinearMap.zero_apply, MonoidalCategory.hom_apply, LinearMap.zero_apply, + rw [LinearMap.zero_apply, ← id_tensorHom, MonoidalCategory.hom_apply, LinearMap.zero_apply, TensorProduct.tmul_zero] Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] - rw [LinearMap.zero_apply, MonoidalCategory.hom_apply, LinearMap.zero_apply, + rw [LinearMap.zero_apply, ← tensorHom_id, MonoidalCategory.hom_apply, LinearMap.zero_apply, TensorProduct.zero_tmul] Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] - rw [LinearMap.add_apply, MonoidalCategory.hom_apply, MonoidalCategory.hom_apply] - erw [MonoidalCategory.hom_apply] + rw [LinearMap.add_apply] + repeat rw [← id_tensorHom, MonoidalCategory.hom_apply] rw [LinearMap.add_apply, TensorProduct.tmul_add] Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] - rw [LinearMap.add_apply, MonoidalCategory.hom_apply, MonoidalCategory.hom_apply] - erw [MonoidalCategory.hom_apply] + rw [LinearMap.add_apply] + repeat rw [← tensorHom_id, MonoidalCategory.hom_apply] rw [LinearMap.add_apply, TensorProduct.add_tmul] -- Porting note: simp wasn't firing but rw was, annoying @@ -289,12 +327,14 @@ instance : MonoidalLinear R (ModuleCat.{u} R) := by Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] - rw [LinearMap.smul_apply, MonoidalCategory.hom_apply, MonoidalCategory.hom_apply, + rw [LinearMap.smul_apply, ← id_tensorHom, MonoidalCategory.hom_apply, + ← id_tensorHom, MonoidalCategory.hom_apply, LinearMap.smul_apply, TensorProduct.tmul_smul] Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] - rw [LinearMap.smul_apply, MonoidalCategory.hom_apply, MonoidalCategory.hom_apply, + rw [LinearMap.smul_apply, ← tensorHom_id, MonoidalCategory.hom_apply, + ← tensorHom_id, MonoidalCategory.hom_apply, LinearMap.smul_apply, TensorProduct.smul_tmul, TensorProduct.tmul_smul] end ModuleCat diff --git a/Mathlib/CategoryTheory/Bicategory/End.lean b/Mathlib/CategoryTheory/Bicategory/End.lean index a27ab4ccb101a..45f73bf95a8ea 100644 --- a/Mathlib/CategoryTheory/Bicategory/End.lean +++ b/Mathlib/CategoryTheory/Bicategory/End.lean @@ -39,15 +39,12 @@ open Bicategory attribute [local simp] EndMonoidal in instance (X : C) : MonoidalCategory (EndMonoidal X) where tensorObj f g := f ≫ g - tensorHom {f g} h i Ξ· ΞΈ := Ξ· β–· h ≫ g ◁ ΞΈ + whiskerLeft {f g h} Ξ· := f ◁ Ξ· + whiskerRight {f g} Ξ· h := Ξ· β–· h tensorUnit' := πŸ™ _ associator f g h := Ξ±_ f g h leftUnitor f := Ξ»_ f rightUnitor f := ρ_ f - tensor_comp := by - intros - dsimp - rw [Bicategory.whiskerLeft_comp, Bicategory.comp_whiskerRight, Category.assoc, Category.assoc, - Bicategory.whisker_exchange_assoc] + whisker_exchange := whisker_exchange end CategoryTheory diff --git a/Mathlib/CategoryTheory/Bicategory/SingleObj.lean b/Mathlib/CategoryTheory/Bicategory/SingleObj.lean index de9f6a18a608e..10cf6749f614b 100644 --- a/Mathlib/CategoryTheory/Bicategory/SingleObj.lean +++ b/Mathlib/CategoryTheory/Bicategory/SingleObj.lean @@ -53,18 +53,12 @@ instance : Bicategory (MonoidalSingleObj C) where Hom _ _ := C id _ := πŸ™_ C comp X Y := tensorObj X Y - whiskerLeft X Y Z f := tensorHom (πŸ™ X) f - whiskerRight f Z := tensorHom f (πŸ™ Z) + whiskerLeft X Y Z f := X ◁ f + whiskerRight f Z := f β–· Z associator X Y Z := Ξ±_ X Y Z leftUnitor X := Ξ»_ X rightUnitor X := ρ_ X - comp_whiskerLeft _ _ _ _ _ := by - simp_rw [associator_inv_naturality, Iso.hom_inv_id_assoc, tensor_id] - whisker_assoc _ _ _ _ _ := by simp_rw [associator_inv_naturality, Iso.hom_inv_id_assoc] - whiskerRight_comp _ _ _ := by simp_rw [← tensor_id, associator_naturality, Iso.inv_hom_id_assoc] - id_whiskerLeft _ := by simp_rw [leftUnitor_inv_naturality, Iso.hom_inv_id_assoc] - whiskerRight_id _ := by simp_rw [rightUnitor_inv_naturality, Iso.hom_inv_id_assoc] - pentagon _ _ _ _ := by simp_rw [pentagon] + whisker_exchange := whisker_exchange namespace MonoidalSingleObj @@ -86,10 +80,6 @@ def endMonoidalStarFunctor : MonoidalFunctor (EndMonoidal (MonoidalSingleObj.sta map f := f Ξ΅ := πŸ™ _ ΞΌ X Y := πŸ™ _ - ΞΌ_natural f g := by - simp_rw [Category.id_comp, Category.comp_id] - -- Should we provide further simp lemmas so this goal becomes visible? - exact (tensor_id_comp_id_tensor _ _).symm #align category_theory.monoidal_single_obj.End_monoidal_star_functor CategoryTheory.MonoidalSingleObj.endMonoidalStarFunctor /-- The equivalence between the endomorphisms of the single object diff --git a/Mathlib/CategoryTheory/Closed/Monoidal.lean b/Mathlib/CategoryTheory/Closed/Monoidal.lean index de6b16de08a6c..f5e6b1c9e3bfe 100644 --- a/Mathlib/CategoryTheory/Closed/Monoidal.lean +++ b/Mathlib/CategoryTheory/Closed/Monoidal.lean @@ -70,10 +70,7 @@ def unitClosed : Closed (πŸ™_ C) where { toFun := fun a => (leftUnitor X).inv ≫ a invFun := fun a => (leftUnitor X).hom ≫ a left_inv := by aesop_cat - right_inv := by aesop_cat } - homEquiv_naturality_left_symm := fun f g => by - dsimp - rw [leftUnitor_naturality_assoc] } } + right_inv := by aesop_cat } } } #align category_theory.unit_closed CategoryTheory.unitClosed variable (A B : C) {X X' Y Y' Z : C} @@ -115,13 +112,13 @@ theorem ihom_adjunction_unit : (ihom.adjunction A).unit = coev A := @[reassoc (attr := simp)] theorem ev_naturality {X Y : C} (f : X ⟢ Y) : - (πŸ™ A βŠ— (ihom A).map f) ≫ (ev A).app Y = (ev A).app X ≫ f := + A ◁ (ihom A).map f ≫ (ev A).app Y = (ev A).app X ≫ f := (ev A).naturality f #align category_theory.ihom.ev_naturality CategoryTheory.ihom.ev_naturality @[reassoc (attr := simp)] theorem coev_naturality {X Y : C} (f : X ⟢ Y) : - f ≫ (coev A).app Y = (coev A).app X ≫ (ihom A).map (πŸ™ A βŠ— f) := + f ≫ (coev A).app Y = (coev A).app X ≫ (ihom A).map (A ◁ f) := (coev A).naturality f #align category_theory.ihom.coev_naturality CategoryTheory.ihom.coev_naturality @@ -130,8 +127,8 @@ set_option quotPrecheck false in notation A " ⟢[" C "] " B:10 => (@ihom C _ _ A _).obj B @[reassoc (attr := simp)] -theorem ev_coev : (πŸ™ A βŠ— (coev A).app B) ≫ (ev A).app (A βŠ— B) = πŸ™ (A βŠ— B) := - Adjunction.left_triangle_components (ihom.adjunction A) +theorem ev_coev : (A ◁ (coev A).app B) ≫ (ev A).app (A βŠ— B) = πŸ™ (A βŠ— B) := + (ihom.adjunction A).left_triangle_components #align category_theory.ihom.ev_coev CategoryTheory.ihom.ev_coev @[reassoc (attr := simp)] @@ -173,7 +170,7 @@ theorem homEquiv_symm_apply_eq (f : Y ⟢ A ⟢[C] X) : #align category_theory.monoidal_closed.hom_equiv_symm_apply_eq CategoryTheory.MonoidalClosed.homEquiv_symm_apply_eq @[reassoc] -theorem curry_natural_left (f : X ⟢ X') (g : A βŠ— X' ⟢ Y) : curry ((πŸ™ _ βŠ— f) ≫ g) = f ≫ curry g := +theorem curry_natural_left (f : X ⟢ X') (g : A βŠ— X' ⟢ Y) : curry (_ ◁ f ≫ g) = f ≫ curry g := Adjunction.homEquiv_naturality_left _ _ _ #align category_theory.monoidal_closed.curry_natural_left CategoryTheory.MonoidalClosed.curry_natural_left @@ -191,7 +188,7 @@ theorem uncurry_natural_right (f : X ⟢ A ⟢[C] Y) (g : Y ⟢ Y') : @[reassoc] theorem uncurry_natural_left (f : X ⟢ X') (g : X' ⟢ A ⟢[C] Y) : - uncurry (f ≫ g) = (πŸ™ _ βŠ— f) ≫ uncurry g := + uncurry (f ≫ g) = _ ◁ f ≫ uncurry g := Adjunction.homEquiv_naturality_left_symm _ _ _ #align category_theory.monoidal_closed.uncurry_natural_left CategoryTheory.MonoidalClosed.uncurry_natural_left @@ -214,7 +211,7 @@ theorem eq_curry_iff (f : A βŠ— Y ⟢ X) (g : Y ⟢ A ⟢[C] X) : g = curry f #align category_theory.monoidal_closed.eq_curry_iff CategoryTheory.MonoidalClosed.eq_curry_iff -- I don't think these two should be simp. -theorem uncurry_eq (g : Y ⟢ A ⟢[C] X) : uncurry g = (πŸ™ A βŠ— g) ≫ (ihom.ev A).app X := +theorem uncurry_eq (g : Y ⟢ A ⟢[C] X) : uncurry g = (A ◁ g) ≫ (ihom.ev A).app X := Adjunction.homEquiv_counit _ #align category_theory.monoidal_closed.uncurry_eq CategoryTheory.MonoidalClosed.uncurry_eq @@ -233,7 +230,7 @@ theorem uncurry_injective : Function.Injective (uncurry : (Y ⟢ A ⟢[C] X) β†’ variable (A X) theorem uncurry_id_eq_ev : uncurry (πŸ™ (A ⟢[C] X)) = (ihom.ev A).app X := by - rw [uncurry_eq, tensor_id, id_comp] + simp [uncurry_eq] #align category_theory.monoidal_closed.uncurry_id_eq_ev CategoryTheory.MonoidalClosed.uncurry_id_eq_ev theorem curry_id_eq_coev : curry (πŸ™ _) = (ihom.coev A).app X := by @@ -252,19 +249,19 @@ def pre (f : B ⟢ A) : ihom A ⟢ ihom B := @[reassoc (attr := simp)] theorem id_tensor_pre_app_comp_ev (f : B ⟢ A) (X : C) : - (πŸ™ B βŠ— (pre f).app X) ≫ (ihom.ev B).app X = (f βŠ— πŸ™ (A ⟢[C] X)) ≫ (ihom.ev A).app X := + B ◁ (pre f).app X ≫ (ihom.ev B).app X = f β–· (A ⟢[C] X) ≫ (ihom.ev A).app X := transferNatTransSelf_counit _ _ ((tensoringLeft C).map f) X #align category_theory.monoidal_closed.id_tensor_pre_app_comp_ev CategoryTheory.MonoidalClosed.id_tensor_pre_app_comp_ev @[simp] theorem uncurry_pre (f : B ⟢ A) (X : C) : - MonoidalClosed.uncurry ((pre f).app X) = (f βŠ— πŸ™ _) ≫ (ihom.ev A).app X := by - rw [uncurry_eq, id_tensor_pre_app_comp_ev] + MonoidalClosed.uncurry ((pre f).app X) = f β–· _ ≫ (ihom.ev A).app X := by + simp [uncurry_eq] #align category_theory.monoidal_closed.uncurry_pre CategoryTheory.MonoidalClosed.uncurry_pre @[reassoc (attr := simp)] theorem coev_app_comp_pre_app (f : B ⟢ A) : - (ihom.coev A).app X ≫ (pre f).app (A βŠ— X) = (ihom.coev B).app X ≫ (ihom B).map (f βŠ— πŸ™ _) := + (ihom.coev A).app X ≫ (pre f).app (A βŠ— X) = (ihom.coev B).app X ≫ (ihom B).map (f β–· _) := unit_transferNatTransSelf _ _ ((tensoringLeft C).map f) X #align category_theory.monoidal_closed.coev_app_comp_pre_app CategoryTheory.MonoidalClosed.coev_app_comp_pre_app diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index f255f1285a826..60e24c5a631a8 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -93,10 +93,10 @@ def braidedCategoryOfFaithful {C D : Type _} [Category C] [Category D] [Monoidal refine (cancel_epi (F.ΞΌ _ _ βŠ— πŸ™ _)).1 ?_ rw [Functor.map_comp, Functor.map_comp, Functor.map_comp, Functor.map_comp, ← LaxMonoidalFunctor.ΞΌ_natural_assoc, Functor.map_id, ← comp_tensor_id_assoc, w, - comp_tensor_id, assoc, LaxMonoidalFunctor.associativity_assoc, - LaxMonoidalFunctor.associativity_assoc, ← LaxMonoidalFunctor.ΞΌ_natural, Functor.map_id, ← + comp_tensor_id, assoc, LaxMonoidalFunctor.associativity'_assoc, + LaxMonoidalFunctor.associativity'_assoc, ← LaxMonoidalFunctor.ΞΌ_natural, Functor.map_id, ← id_tensor_comp_assoc, w, id_tensor_comp_assoc, reassoc_of% w, braiding_naturality_assoc, - LaxMonoidalFunctor.associativity, hexagon_forward_assoc] + LaxMonoidalFunctor.associativity', hexagon_forward_assoc] hexagon_reverse := by intros apply F.toFunctor.map_injective @@ -104,10 +104,10 @@ def braidedCategoryOfFaithful {C D : Type _} [Category C] [Category D] [Monoidal refine (cancel_epi (πŸ™ _ βŠ— F.ΞΌ _ _)).1 ?_ rw [Functor.map_comp, Functor.map_comp, Functor.map_comp, Functor.map_comp, ← LaxMonoidalFunctor.ΞΌ_natural_assoc, Functor.map_id, ← id_tensor_comp_assoc, w, - id_tensor_comp_assoc, LaxMonoidalFunctor.associativity_inv_assoc, - LaxMonoidalFunctor.associativity_inv_assoc, ← LaxMonoidalFunctor.ΞΌ_natural, + id_tensor_comp_assoc, LaxMonoidalFunctor.associativity_inv'_assoc, + LaxMonoidalFunctor.associativity_inv'_assoc, ← LaxMonoidalFunctor.ΞΌ_natural, Functor.map_id, ← comp_tensor_id_assoc, w, comp_tensor_id_assoc, reassoc_of% w, - braiding_naturality_assoc, LaxMonoidalFunctor.associativity_inv, hexagon_reverse_assoc] + braiding_naturality_assoc, LaxMonoidalFunctor.associativity_inv', hexagon_reverse_assoc] #align category_theory.braided_category_of_faithful CategoryTheory.braidedCategoryOfFaithful /-- Pull back a braiding along a fully faithful monoidal functor. -/ @@ -140,8 +140,8 @@ variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory C] [BraidedCatego theorem braiding_leftUnitor_aux₁ (X : C) : (Ξ±_ (πŸ™_ C) (πŸ™_ C) X).hom ≫ (πŸ™ (πŸ™_ C) βŠ— (Ξ²_ X (πŸ™_ C)).inv) ≫ (Ξ±_ _ X _).inv ≫ ((Ξ»_ X).hom βŠ— πŸ™ _) = - ((Ξ»_ _).hom βŠ— πŸ™ X) ≫ (Ξ²_ X (πŸ™_ C)).inv := - by rw [← leftUnitor_tensor, leftUnitor_naturality]; simp + ((Ξ»_ _).hom βŠ— πŸ™ X) ≫ (Ξ²_ X (πŸ™_ C)).inv := by + coherence #align category_theory.braiding_left_unitor_aux₁ CategoryTheory.braiding_leftUnitor_aux₁ theorem braiding_leftUnitor_auxβ‚‚ (X : C) : @@ -161,7 +161,7 @@ theorem braiding_leftUnitor_auxβ‚‚ (X : C) : _ = (Ξ±_ _ _ _).hom ≫ (πŸ™ _ βŠ— (Ξ»_ _).hom) ≫ (Ξ²_ _ _).hom ≫ (Ξ²_ X _).inv := by (slice_lhs 2 3 => rw [← braiding_naturality]); simp only [assoc] _ = (Ξ±_ _ _ _).hom ≫ (πŸ™ _ βŠ— (Ξ»_ _).hom) := by rw [Iso.hom_inv_id, comp_id] - _ = (ρ_ X).hom βŠ— πŸ™ (πŸ™_ C) := by rw [triangle] + _ = (ρ_ X).hom βŠ— πŸ™ (πŸ™_ C) := by rw [triangle'] #align category_theory.braiding_left_unitor_auxβ‚‚ CategoryTheory.braiding_leftUnitor_auxβ‚‚ @@ -173,8 +173,8 @@ theorem braiding_leftUnitor (X : C) : (Ξ²_ X (πŸ™_ C)).hom ≫ (Ξ»_ X).hom = ( theorem braiding_rightUnitor_aux₁ (X : C) : (Ξ±_ X (πŸ™_ C) (πŸ™_ C)).inv ≫ ((Ξ²_ (πŸ™_ C) X).inv βŠ— πŸ™ (πŸ™_ C)) ≫ (Ξ±_ _ X _).hom ≫ (πŸ™ _ βŠ— (ρ_ X).hom) = - (πŸ™ X βŠ— (ρ_ _).hom) ≫ (Ξ²_ (πŸ™_ C) X).inv := - by rw [← rightUnitor_tensor, rightUnitor_naturality]; simp + (πŸ™ X βŠ— (ρ_ _).hom) ≫ (Ξ²_ (πŸ™_ C) X).inv := by + coherence #align category_theory.braiding_right_unitor_aux₁ CategoryTheory.braiding_rightUnitor_aux₁ theorem braiding_rightUnitor_auxβ‚‚ (X : C) : @@ -194,7 +194,7 @@ theorem braiding_rightUnitor_auxβ‚‚ (X : C) : _ = (Ξ±_ _ _ _).inv ≫ ((ρ_ _).hom βŠ— πŸ™ _) ≫ (Ξ²_ _ X).hom ≫ (Ξ²_ _ _).inv := by (slice_lhs 2 3 => rw [← braiding_naturality]); simp only [assoc] _ = (Ξ±_ _ _ _).inv ≫ ((ρ_ _).hom βŠ— πŸ™ _) := by rw [Iso.hom_inv_id, comp_id] - _ = πŸ™ (πŸ™_ C) βŠ— (Ξ»_ X).hom := by rw [triangle_assoc_comp_right] + _ = πŸ™ (πŸ™_ C) βŠ— (Ξ»_ X).hom := by rw [triangle_assoc_comp_right'] #align category_theory.braiding_right_unitor_auxβ‚‚ CategoryTheory.braiding_rightUnitor_auxβ‚‚ @@ -410,7 +410,8 @@ theorem tensor_ΞΌ_natural {X₁ Xβ‚‚ Y₁ Yβ‚‚ U₁ Uβ‚‚ V₁ Vβ‚‚ : C} (f₁ : (gβ‚‚ : Uβ‚‚ ⟢ Vβ‚‚) : ((f₁ βŠ— fβ‚‚) βŠ— g₁ βŠ— gβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (V₁, Vβ‚‚) = tensor_ΞΌ C (X₁, Xβ‚‚) (U₁, Uβ‚‚) ≫ ((f₁ βŠ— g₁) βŠ— fβ‚‚ βŠ— gβ‚‚) := by - dsimp [tensor_ΞΌ] + dsimp only [tensor_ΞΌ] + simp_rw [← id_tensorHom, ← tensorHom_id] slice_lhs 1 2 => rw [associator_naturality] slice_lhs 2 3 => rw [← tensor_comp, comp_id f₁, ← id_comp f₁, associator_inv_naturality, tensor_comp] @@ -426,34 +427,32 @@ theorem tensor_left_unitality (X₁ Xβ‚‚ : C) : (Ξ»_ (X₁ βŠ— Xβ‚‚)).hom = ((Ξ»_ (πŸ™_ C)).inv βŠ— πŸ™ (X₁ βŠ— Xβ‚‚)) ≫ tensor_ΞΌ C (πŸ™_ C, πŸ™_ C) (X₁, Xβ‚‚) ≫ ((Ξ»_ X₁).hom βŠ— (Ξ»_ Xβ‚‚).hom) := by - dsimp [tensor_ΞΌ] + dsimp only [tensor_ΞΌ] have : ((Ξ»_ (πŸ™_ C)).inv βŠ— πŸ™ (X₁ βŠ— Xβ‚‚)) ≫ (Ξ±_ (πŸ™_ C) (πŸ™_ C) (X₁ βŠ— Xβ‚‚)).hom ≫ (πŸ™ (πŸ™_ C) βŠ— (Ξ±_ (πŸ™_ C) X₁ Xβ‚‚).inv) = πŸ™ (πŸ™_ C) βŠ— (Ξ»_ X₁).inv βŠ— πŸ™ Xβ‚‚ := - by pure_coherence + by coherence slice_rhs 1 3 => rw [this] clear this slice_rhs 1 2 => rw [← tensor_comp, ← tensor_comp, comp_id, comp_id, leftUnitor_inv_braiding] - simp only [assoc] - coherence + simp [tensorHom_id, id_tensorHom, tensorHom_def] #align category_theory.tensor_left_unitality CategoryTheory.tensor_left_unitality theorem tensor_right_unitality (X₁ Xβ‚‚ : C) : (ρ_ (X₁ βŠ— Xβ‚‚)).hom = (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— (Ξ»_ (πŸ™_ C)).inv) ≫ tensor_ΞΌ C (X₁, Xβ‚‚) (πŸ™_ C, πŸ™_ C) ≫ ((ρ_ X₁).hom βŠ— (ρ_ Xβ‚‚).hom) := by - dsimp [tensor_ΞΌ] + dsimp only [tensor_ΞΌ] have : (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— (Ξ»_ (πŸ™_ C)).inv) ≫ (Ξ±_ X₁ Xβ‚‚ (πŸ™_ C βŠ— πŸ™_ C)).hom ≫ (πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ (πŸ™_ C) (πŸ™_ C)).inv) = (Ξ±_ X₁ Xβ‚‚ (πŸ™_ C)).hom ≫ (πŸ™ X₁ βŠ— (ρ_ Xβ‚‚).inv βŠ— πŸ™ (πŸ™_ C)) := - by pure_coherence + by coherence slice_rhs 1 3 => rw [this] clear this slice_rhs 2 3 => rw [← tensor_comp, ← tensor_comp, comp_id, comp_id, rightUnitor_inv_braiding] - simp only [assoc] - coherence + simp [tensorHom_id, id_tensorHom, tensorHom_def] #align category_theory.tensor_right_unitality CategoryTheory.tensor_right_unitality /- @@ -496,7 +495,7 @@ theorem tensor_associativity (X₁ Xβ‚‚ Y₁ Yβ‚‚ Z₁ Zβ‚‚ : C) : (πŸ™ X₁ βŠ— (Ξ±_ ((Y₁ βŠ— Z₁) βŠ— Xβ‚‚) Yβ‚‚ Zβ‚‚).hom) ≫ (πŸ™ X₁ βŠ— (Ξ±_ (Y₁ βŠ— Z₁) Xβ‚‚ (Yβ‚‚ βŠ— Zβ‚‚)).hom) ≫ (Ξ±_ X₁ (Y₁ βŠ— Z₁) (Xβ‚‚ βŠ— Yβ‚‚ βŠ— Zβ‚‚)).inv := - by pure_coherence + by monoidal_simps; pure_coherence rw [this]; clear this slice_lhs 2 4 => rw [tensor_ΞΌ_def₁] slice_lhs 4 5 => rw [← tensor_id, associator_naturality] @@ -514,7 +513,7 @@ theorem tensor_associativity (X₁ Xβ‚‚ Y₁ Yβ‚‚ Z₁ Zβ‚‚ : C) : ((πŸ™ X₁ βŠ— (Ξ±_ (Y₁ βŠ— Xβ‚‚) Yβ‚‚ Z₁).hom) βŠ— πŸ™ Zβ‚‚) ≫ ((πŸ™ X₁ βŠ— (Ξ±_ Y₁ Xβ‚‚ (Yβ‚‚ βŠ— Z₁)).hom) βŠ— πŸ™ Zβ‚‚) ≫ ((πŸ™ X₁ βŠ— πŸ™ Y₁ βŠ— (Ξ±_ Xβ‚‚ Yβ‚‚ Z₁).inv) βŠ— πŸ™ Zβ‚‚) := - by pure_coherence + by monoidal_simps; pure_coherence slice_lhs 2 6 => rw [this] clear this slice_lhs 1 3 => rw [← tensor_comp, ← tensor_comp, tensor_ΞΌ_def₁, tensor_comp, tensor_comp] @@ -530,10 +529,11 @@ theorem tensor_associativity (X₁ Xβ‚‚ Y₁ Yβ‚‚ Z₁ Zβ‚‚ : C) : tensor_comp, tensor_comp, tensor_comp] slice_lhs 11 12 => rw [← tensor_comp, ← tensor_comp, Iso.hom_inv_id] - simp + rw [comp_id, id_comp, id_tensorHom, tensorHom_id, MonoidalCategory.whiskerLeft_id, id_whiskerRight] simp only [assoc, id_comp] slice_lhs 10 11 => rw [← tensor_comp, ← tensor_comp, ← tensor_comp, Iso.hom_inv_id] + simp only [id_tensorHom, tensorHom_id, comp_id, id_comp] simp simp only [assoc, id_comp] slice_lhs 9 10 => rw [associator_naturality] @@ -550,7 +550,7 @@ theorem tensor_associativity (X₁ Xβ‚‚ Y₁ Yβ‚‚ Z₁ Zβ‚‚ : C) : (πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ Y₁ ((Z₁ βŠ— Yβ‚‚) βŠ— Zβ‚‚)).hom) ≫ (Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— (Z₁ βŠ— Yβ‚‚) βŠ— Zβ‚‚)).inv ≫ (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— πŸ™ Y₁ βŠ— (Ξ±_ Z₁ Yβ‚‚ Zβ‚‚).hom) ≫ (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— (Ξ±_ Y₁ Z₁ (Yβ‚‚ βŠ— Zβ‚‚)).inv) := - by pure_coherence + by monoidal_simps; pure_coherence slice_lhs 7 12 => rw [this] clear this slice_lhs 6 7 => rw [associator_naturality] @@ -558,52 +558,57 @@ theorem tensor_associativity (X₁ Xβ‚‚ Y₁ Yβ‚‚ Z₁ Zβ‚‚ : C) : slice_lhs 8 9 => rw [← tensor_comp, associator_naturality, tensor_comp] slice_lhs 9 10 => rw [associator_inv_naturality] slice_lhs 10 12 => rw [← tensor_comp, ← tensor_comp, ← tensor_ΞΌ_defβ‚‚, tensor_comp, tensor_comp] - dsimp + dsimp only [tensor_obj, prodMonoidal_tensorObj] + simp_rw [← Category.assoc] + congr 2 coherence #align category_theory.tensor_associativity CategoryTheory.tensor_associativity +@[simp] +def tensorLaxMonoidal : LaxMonoidalFunctor (C Γ— C) C := LaxMonoidalFunctor.ofTensorHom + (F := tensor C) + (Ξ΅ := (Ξ»_ (πŸ™_ C)).inv) + (ΞΌ := fun X Y => tensor_ΞΌ C X Y) + (ΞΌ_natural := fun f g => tensor_ΞΌ_natural C f.1 f.2 g.1 g.2) + (associativity := fun X Y Z => tensor_associativity C X.1 X.2 Y.1 Y.2 Z.1 Z.2) + (left_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_left_unitality C X₁ Xβ‚‚) + (right_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_right_unitality C X₁ Xβ‚‚) + /-- The tensor product functor from `C Γ— C` to `C` as a monoidal functor. -/ @[simps!] def tensorMonoidal : MonoidalFunctor (C Γ— C) C := - { tensor C with - Ξ΅ := (Ξ»_ (πŸ™_ C)).inv - ΞΌ := fun X Y => tensor_ΞΌ C X Y - ΞΌ_natural := fun f g => tensor_ΞΌ_natural C f.1 f.2 g.1 g.2 - associativity := fun X Y Z => tensor_associativity C X.1 X.2 Y.1 Y.2 Z.1 Z.2 - left_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_left_unitality C X₁ Xβ‚‚ - right_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_right_unitality C X₁ Xβ‚‚ - ΞΌ_isIso := by dsimp [tensor_ΞΌ]; infer_instance } + { tensorLaxMonoidal C with + ΞΌ_isIso := by dsimp [tensor_ΞΌ]; infer_instance + Ξ΅_isIso := by dsimp; infer_instance } #align category_theory.tensor_monoidal CategoryTheory.tensorMonoidal theorem leftUnitor_monoidal (X₁ Xβ‚‚ : C) : (Ξ»_ X₁).hom βŠ— (Ξ»_ Xβ‚‚).hom = tensor_ΞΌ C (πŸ™_ C, X₁) (πŸ™_ C, Xβ‚‚) ≫ ((Ξ»_ (πŸ™_ C)).hom βŠ— πŸ™ (X₁ βŠ— Xβ‚‚)) ≫ (Ξ»_ (X₁ βŠ— Xβ‚‚)).hom := by - dsimp [tensor_ΞΌ] + dsimp only [tensor_ΞΌ] have : (Ξ»_ X₁).hom βŠ— (Ξ»_ Xβ‚‚).hom = (Ξ±_ (πŸ™_ C) X₁ (πŸ™_ C βŠ— Xβ‚‚)).hom ≫ (πŸ™ (πŸ™_ C) βŠ— (Ξ±_ X₁ (πŸ™_ C) Xβ‚‚).inv) ≫ (Ξ»_ ((X₁ βŠ— πŸ™_ C) βŠ— Xβ‚‚)).hom ≫ ((ρ_ X₁).hom βŠ— πŸ™ Xβ‚‚) := - by pure_coherence + by coherence rw [this]; clear this rw [← braiding_leftUnitor] - slice_lhs 3 4 => rw [← id_comp (πŸ™ Xβ‚‚), tensor_comp] - slice_lhs 3 4 => rw [← leftUnitor_naturality] + dsimp only [tensor_obj, prodMonoidal_tensorObj] coherence #align category_theory.left_unitor_monoidal CategoryTheory.leftUnitor_monoidal theorem rightUnitor_monoidal (X₁ Xβ‚‚ : C) : (ρ_ X₁).hom βŠ— (ρ_ Xβ‚‚).hom = tensor_ΞΌ C (X₁, πŸ™_ C) (Xβ‚‚, πŸ™_ C) ≫ (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— (Ξ»_ (πŸ™_ C)).hom) ≫ (ρ_ (X₁ βŠ— Xβ‚‚)).hom := by - dsimp [tensor_ΞΌ] + dsimp only [tensor_ΞΌ] have : (ρ_ X₁).hom βŠ— (ρ_ Xβ‚‚).hom = (Ξ±_ X₁ (πŸ™_ C) (Xβ‚‚ βŠ— πŸ™_ C)).hom ≫ (πŸ™ X₁ βŠ— (Ξ±_ (πŸ™_ C) Xβ‚‚ (πŸ™_ C)).inv) ≫ (πŸ™ X₁ βŠ— (ρ_ (πŸ™_ C βŠ— Xβ‚‚)).hom) ≫ (πŸ™ X₁ βŠ— (Ξ»_ Xβ‚‚).hom) := - by pure_coherence + by coherence rw [this]; clear this rw [← braiding_rightUnitor] - slice_lhs 3 4 => rw [← id_comp (πŸ™ X₁), tensor_comp, id_comp] - slice_lhs 3 4 => rw [← tensor_comp, ← rightUnitor_naturality, tensor_comp] + dsimp only [tensor_obj, prodMonoidal_tensorObj] coherence #align category_theory.right_unitor_monoidal CategoryTheory.rightUnitor_monoidal @@ -614,7 +619,7 @@ theorem associator_monoidal_aux (W X Y Z : C) : (Ξ±_ (W βŠ— X) Y Z).inv ≫ ((Ξ²_ (W βŠ— X) Y).hom βŠ— πŸ™ Z) ≫ ((Ξ±_ Y W X).inv βŠ— πŸ™ Z) ≫ (Ξ±_ (Y βŠ— W) X Z).hom ≫ (πŸ™ (Y βŠ— W) βŠ— (Ξ²_ X Z).hom) := by - slice_rhs 1 2 => rw [← pentagon_inv] + slice_rhs 1 2 => rw [← pentagon_inv'] slice_rhs 3 5 => rw [← tensor_comp, ← tensor_comp, hexagon_reverse, tensor_comp, tensor_comp] slice_rhs 5 6 => rw [associator_naturality] slice_rhs 6 7 => rw [tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] @@ -644,7 +649,7 @@ theorem associator_monoidal (X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C) : (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— (Yβ‚‚ βŠ— X₃) βŠ— Y₃)).inv ≫ (πŸ™ (X₁ βŠ— Y₁) βŠ— πŸ™ Xβ‚‚ βŠ— (Ξ±_ Yβ‚‚ X₃ Y₃).hom) ≫ (πŸ™ (X₁ βŠ— Y₁) βŠ— (Ξ±_ Xβ‚‚ Yβ‚‚ (X₃ βŠ— Y₃)).inv) := - by pure_coherence + by coherence rw [this]; clear this slice_lhs 2 4 => rw [← tensor_comp, ← tensor_comp, tensor_ΞΌ_def₁, tensor_comp, tensor_comp] slice_lhs 4 5 => rw [← tensor_id, associator_inv_naturality] @@ -663,7 +668,7 @@ theorem associator_monoidal (X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C) : (Ξ±_ X₁ (Xβ‚‚ βŠ— (Y₁ βŠ— Yβ‚‚) βŠ— X₃) Y₃).inv ≫ ((πŸ™ X₁ βŠ— πŸ™ Xβ‚‚ βŠ— (Ξ±_ Y₁ Yβ‚‚ X₃).hom) βŠ— πŸ™ Y₃) ≫ ((πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ Y₁ (Yβ‚‚ βŠ— X₃)).inv) βŠ— πŸ™ Y₃) := - by pure_coherence + by coherence slice_lhs 2 6 => rw [this] clear this slice_lhs 1 3 => rw [tensor_ΞΌ_def₁] @@ -694,13 +699,15 @@ theorem associator_monoidal (X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C) : (Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚ βŠ— X₃) βŠ— Yβ‚‚) Y₃).hom ≫ (πŸ™ X₁ βŠ— (Ξ±_ (Y₁ βŠ— Xβ‚‚ βŠ— X₃) Yβ‚‚ Y₃).hom) ≫ (πŸ™ X₁ βŠ— (Ξ±_ Y₁ (Xβ‚‚ βŠ— X₃) (Yβ‚‚ βŠ— Y₃)).hom) ≫ (Ξ±_ X₁ Y₁ ((Xβ‚‚ βŠ— X₃) βŠ— Yβ‚‚ βŠ— Y₃)).inv := - by pure_coherence + by coherence slice_lhs 9 16 => rw [this] clear this slice_lhs 8 9 => rw [associator_naturality] slice_lhs 9 10 => rw [← tensor_comp, associator_naturality, tensor_comp] slice_lhs 10 12 => rw [tensor_id, ← tensor_ΞΌ_defβ‚‚] - dsimp + dsimp only [tensor_obj, prodMonoidal_tensorObj] + simp_rw [← Category.assoc] + congr 2 coherence #align category_theory.associator_monoidal CategoryTheory.associator_monoidal diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index b0f37f1b380d9..59ffdc83e42a3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -60,6 +60,14 @@ open CategoryTheory.Iso namespace CategoryTheory +def tensorHomDefault {C : Type u} [Category.{v} C] + (tensorObj : C β†’ C β†’ C) + (whiskerLeft : βˆ€ (X : C) {Y₁ Yβ‚‚ : C} (_f : Y₁ ⟢ Yβ‚‚), tensorObj X Y₁ ⟢ tensorObj X Yβ‚‚) + (whiskerRight : βˆ€ {X₁ Xβ‚‚ : C} (_f : X₁ ⟢ Xβ‚‚) (Y : C), tensorObj X₁ Y ⟢ tensorObj Xβ‚‚ Y) + ⦃X₁ Y₁ Xβ‚‚ Yβ‚‚ : C⦄ (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : + (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚) := + whiskerLeft X₁ g ≫ whiskerRight f Yβ‚‚ + /-- In a monoidal category, we can take the tensor product of objects, `X βŠ— Y` and of morphisms `f βŠ— g`. Tensor product does not need to be strictly associative on objects, but there is a @@ -71,78 +79,98 @@ See . -/ -- Porting note: The Mathport did not translate the temporary notation class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] where - /-- curried tensor product of objects -/ tensorObj : C β†’ C β†’ C - /-- curried tensor product of morphisms -/ - tensorHom : βˆ€ {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C}, (X₁ ⟢ Y₁) β†’ (Xβ‚‚ ⟢ Yβ‚‚) β†’ (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚) - /-- Tensor product of identity maps is the identity: `(πŸ™ X₁ βŠ— πŸ™ Xβ‚‚) = πŸ™ (X₁ βŠ— Xβ‚‚)` -/ - tensor_id : βˆ€ X₁ Xβ‚‚ : C, tensorHom (πŸ™ X₁) (πŸ™ Xβ‚‚) = πŸ™ (tensorObj X₁ Xβ‚‚) := by aesop_cat - /-- - Composition of tensor products is tensor product of compositions: - `(f₁ βŠ— g₁) ∘ (fβ‚‚ βŠ— gβ‚‚) = (f₁ ∘ fβ‚‚) βŠ— (g₁ βŠ— gβ‚‚)` - -/ - tensor_comp : - βˆ€ {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : Y₁ ⟢ Z₁) (gβ‚‚ : Yβ‚‚ ⟢ Zβ‚‚), - tensorHom (f₁ ≫ g₁) (fβ‚‚ ≫ gβ‚‚) = tensorHom f₁ fβ‚‚ ≫ tensorHom g₁ gβ‚‚ := by - aesop_cat + whiskerLeft (X : C) {Y₁ Yβ‚‚ : C} (f : Y₁ ⟢ Yβ‚‚) : tensorObj X Y₁ ⟢ tensorObj X Yβ‚‚ + whiskerRight {X₁ Xβ‚‚ : C} (f : X₁ ⟢ Xβ‚‚) (Y : C) : tensorObj X₁ Y ⟢ tensorObj Xβ‚‚ Y + tensorHom {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚) := + whiskerLeft X₁ g ≫ whiskerRight f Yβ‚‚ + tensorHom_def {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : + tensorHom f g = whiskerLeft X₁ g ≫ whiskerRight f Yβ‚‚ := by + aesop_cat -- Porting note: Adding a prime here, so I can later define `tensorUnit` unprimed with explicit -- argument `C` /-- The tensor unity in the monoidal structure `πŸ™_C` -/ tensorUnit' : C + /-- The left unitor: `πŸ™_C βŠ— X ≃ X` -/ + leftUnitor : βˆ€ X : C, tensorObj tensorUnit' X β‰… X + /-- The right unitor: `X βŠ— πŸ™_C ≃ X` -/ + rightUnitor : βˆ€ X : C, tensorObj X tensorUnit' β‰… X /-- The associator isomorphism `(X βŠ— Y) βŠ— Z ≃ X βŠ— (Y βŠ— Z)` -/ associator : βˆ€ X Y Z : C, tensorObj (tensorObj X Y) Z β‰… tensorObj X (tensorObj Y Z) - /-- Naturality of the associator isomorphism: `(f₁ βŠ— fβ‚‚) βŠ— f₃ ≃ f₁ βŠ— (fβ‚‚ βŠ— f₃)` -/ - associator_naturality : - βˆ€ {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (f₃ : X₃ ⟢ Y₃), - tensorHom (tensorHom f₁ fβ‚‚) f₃ ≫ (associator Y₁ Yβ‚‚ Y₃).hom = - (associator X₁ Xβ‚‚ X₃).hom ≫ tensorHom f₁ (tensorHom fβ‚‚ f₃) := by + whiskerLeft_id : βˆ€ (X Y : C), whiskerLeft X (πŸ™ Y) = πŸ™ (tensorObj X Y) := by aesop_cat - /-- The left unitor: `πŸ™_C βŠ— X ≃ X` -/ - leftUnitor : βˆ€ X : C, tensorObj tensorUnit' X β‰… X - /-- - Naturality of the left unitor, commutativity of `πŸ™_C βŠ— X ⟢ πŸ™_C βŠ— Y ⟢ Y` and `πŸ™_C βŠ— X ⟢ X ⟢ Y` - -/ - leftUnitor_naturality : + whiskerLeft_comp : + βˆ€ (W : C) {X Y Z : C} (f : X ⟢ Y) (g : Y ⟢ Z), + whiskerLeft W (f ≫ g) = whiskerLeft W f ≫ whiskerLeft W g := by + aesop_cat + id_whiskerLeft : βˆ€ {X Y : C} (f : X ⟢ Y), - tensorHom (πŸ™ tensorUnit') f ≫ (leftUnitor Y).hom = (leftUnitor X).hom ≫ f := by + whiskerLeft tensorUnit' f = (leftUnitor X).hom ≫ f ≫ (leftUnitor Y).inv := by aesop_cat - /-- The right unitor: `X βŠ— πŸ™_C ≃ X` -/ - rightUnitor : βˆ€ X : C, tensorObj X tensorUnit' β‰… X - /-- - Naturality of the right unitor: commutativity of `X βŠ— πŸ™_C ⟢ Y βŠ— πŸ™_C ⟢ Y` and `X βŠ— πŸ™_C ⟢ X ⟢ Y` - -/ - rightUnitor_naturality : + tensor_whiskerLeft : + βˆ€ (X Y : C) {Z Z' : C} (f : Z ⟢ Z'), + whiskerLeft (tensorObj X Y) f = + (associator X Y Z).hom ≫ whiskerLeft X (whiskerLeft Y f) ≫ (associator X Y Z').inv := by + aesop_cat + id_whiskerRight : βˆ€ (X Y : C), whiskerRight (πŸ™ X) Y = πŸ™ (tensorObj X Y) := by + aesop_cat + comp_whiskerRight : + βˆ€ {W X Y : C} (f : W ⟢ X) (g : X ⟢ Y) (Z : C), + whiskerRight (f ≫ g) Z = whiskerRight f Z ≫ whiskerRight g Z := by + aesop_cat + whiskerRight_id : βˆ€ {X Y : C} (f : X ⟢ Y), - tensorHom f (πŸ™ tensorUnit') ≫ (rightUnitor Y).hom = (rightUnitor X).hom ≫ f := by + whiskerRight f tensorUnit' = (rightUnitor X).hom ≫ f ≫ (rightUnitor Y).inv := by + aesop_cat + whiskerRight_tensor : + βˆ€ {X X' : C} (f : X ⟢ X') (Y Z : C), + whiskerRight f (tensorObj Y Z) = + (associator X Y Z).inv ≫ whiskerRight (whiskerRight f Y) Z ≫ (associator X' Y Z).hom := by + aesop_cat + whisker_assoc : + βˆ€ (X : C) {Y Y' : C} (f : Y ⟢ Y') (Z : C), + whiskerRight (whiskerLeft X f) Z = + (associator X Y Z).hom ≫ whiskerLeft X (whiskerRight f Z) ≫ (associator X Y' Z).inv := by + aesop_cat + whisker_exchange : + βˆ€ {W X Y Z : C} (f : W ⟢ X) (g : Y ⟢ Z), + whiskerLeft W g ≫ whiskerRight f Z = whiskerRight f Y ≫ whiskerLeft X g := by aesop_cat - /-- - The pentagon identity relating the isomorphism between `X βŠ— (Y βŠ— (Z βŠ— W))` and `((X βŠ— Y) βŠ— Z) βŠ— W` - -/ pentagon : βˆ€ W X Y Z : C, - tensorHom (associator W X Y).hom (πŸ™ Z) ≫ - (associator W (tensorObj X Y) Z).hom ≫ tensorHom (πŸ™ W) (associator X Y Z).hom = + whiskerRight (associator W X Y).hom Z ≫ + (associator W (tensorObj X Y) Z).hom ≫ whiskerLeft W (associator X Y Z).hom = (associator (tensorObj W X) Y Z).hom ≫ (associator W X (tensorObj Y Z)).hom := by aesop_cat - /-- - The identity relating the isomorphisms between `X βŠ— (πŸ™_C βŠ— Y)`, `(X βŠ— πŸ™_C) βŠ— Y` and `X βŠ— Y` - -/ triangle : βˆ€ X Y : C, - (associator X tensorUnit' Y).hom ≫ tensorHom (πŸ™ X) (leftUnitor Y).hom = - tensorHom (rightUnitor X).hom (πŸ™ Y) := by + (associator X tensorUnit' Y).hom ≫ whiskerLeft X (leftUnitor Y).hom = + whiskerRight (rightUnitor X).hom Y := by aesop_cat #align category_theory.monoidal_category CategoryTheory.MonoidalCategory -attribute [simp] MonoidalCategory.tensor_id -attribute [reassoc] MonoidalCategory.tensor_comp -attribute [simp] MonoidalCategory.tensor_comp -attribute [reassoc] MonoidalCategory.associator_naturality -attribute [reassoc] MonoidalCategory.leftUnitor_naturality -attribute [reassoc] MonoidalCategory.rightUnitor_naturality -attribute [reassoc] MonoidalCategory.pentagon +-- attribute [simp] MonoidalCategory.tensor_id +-- attribute [reassoc] MonoidalCategory.tensor_comp +-- attribute [simp] MonoidalCategory.tensor_comp +-- attribute [reassoc] MonoidalCategory.associator_naturality +-- attribute [reassoc] MonoidalCategory.leftUnitor_naturality +-- attribute [reassoc] MonoidalCategory.rightUnitor_naturality +attribute [reassoc (attr := simp)] MonoidalCategory.pentagon attribute [reassoc (attr := simp)] MonoidalCategory.triangle +-- attribute [simp] MonoidalCategory.whiskerLeft_eq_tensorHom_id MonoidalCategory.whiskerRight_eq_id_tensorHom + +namespace MonoidalCategory + +attribute [reassoc] + whiskerLeft_comp id_whiskerLeft tensor_whiskerLeft comp_whiskerRight whiskerRight_id + whiskerRight_tensor whisker_assoc whisker_exchange + +attribute [simp] + whiskerLeft_id whiskerLeft_comp id_whiskerLeft tensor_whiskerLeft id_whiskerRight comp_whiskerRight + whiskerRight_id whiskerRight_tensor whisker_assoc +end MonoidalCategory + -- Porting Note: This is here to make `tensorUnit` explicitly depend on `C`, which was done in -- Lean 3 using the `[]` notation in the `tensorUnit'` field. open CategoryTheory.MonoidalCategory in @@ -153,10 +181,7 @@ abbrev MonoidalCategory.tensorUnit (C : Type u) [Category.{v} C] [MonoidalCatego namespace MonoidalCategory /-- Notation for `tensorObj`, the tensor product of objects in a monoidal category -/ -scoped infixr:70 " βŠ— " => tensorObj - -/-- Notation for `tensorHom`, the tensor product of morphisms in a monoidal category -/ -scoped infixr:70 " βŠ— " => tensorHom +scoped infixr:70 " βŠ— " => MonoidalCategory.tensorObj /-- Notation for `tensorUnit`, the two-sided identity of `βŠ—` -/ scoped notation "πŸ™_" => tensorUnit @@ -170,12 +195,105 @@ scoped notation "Ξ»_" => leftUnitor /-- Notation for the `rightUnitor`: `X βŠ— πŸ™_C ≃ X` -/ scoped notation "ρ_" => rightUnitor +scoped infixr:81 " ◁ " => whiskerLeft +scoped infixl:81 " β–· " => whiskerRight + +variable {C : Type u} [π’ž : Category.{v} C] [MonoidalCategory C] + +/-- Notation for `tensorHom`, the tensor product of morphisms in a monoidal category -/ +scoped infixr:70 " βŠ— " => MonoidalCategory.tensorHom + +attribute [reassoc] tensorHom_def +attribute [local simp] tensorHom_def + +@[reassoc] +theorem tensorHom_def' {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : + f βŠ— g = f β–· Xβ‚‚ ≫ Y₁ ◁ g := by + rw [← whisker_exchange] + apply tensorHom_def + + +/-- Tensor product of identity maps is the identity: `(πŸ™ X₁ βŠ— πŸ™ Xβ‚‚) = πŸ™ (X₁ βŠ— Xβ‚‚)` -/ +theorem tensor_id (X₁ Xβ‚‚ : C) : tensorHom (πŸ™ X₁) (πŸ™ Xβ‚‚) = πŸ™ (tensorObj X₁ Xβ‚‚) := by simp + +/-- +Composition of tensor products is tensor product of compositions: +`(f₁ βŠ— g₁) ∘ (fβ‚‚ βŠ— gβ‚‚) = (f₁ ∘ fβ‚‚) βŠ— (g₁ βŠ— gβ‚‚)` +-/ +@[reassoc, simp] +theorem tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : Y₁ ⟢ Z₁) (gβ‚‚ : Yβ‚‚ ⟢ Zβ‚‚) : + (f₁ ≫ g₁) βŠ— (fβ‚‚ ≫ gβ‚‚) = (f₁ βŠ— fβ‚‚) ≫ (g₁ βŠ— gβ‚‚) := by + simp [whisker_exchange_assoc] + +theorem id_tensorHom (X : C) {Y₁ Yβ‚‚ : C} (f : Y₁ ⟢ Yβ‚‚) : + (πŸ™ X) βŠ— f = X ◁ f := by + simp + +theorem tensorHom_id {X₁ Xβ‚‚ : C} (f : X₁ ⟢ Xβ‚‚) (Y : C) : + f βŠ— (πŸ™ Y) = f β–· Y := by + simp + end MonoidalCategory open scoped MonoidalCategory open MonoidalCategory -variable (C : Type u) [π’ž : Category.{v} C] [MonoidalCategory C] +variable {C : Type u} [π’ž : Category.{v} C] [MonoidalCategory C] + +namespace MonoidalCategory + +-- @[simp] +-- theorem tensorHom_def {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : +-- f βŠ— g = X₁ ◁ g ≫ f β–· Yβ‚‚ := by +-- rw [← whiskerLeft_eq_tensorHom_id, ← whiskerRight_eq_id_tensorHom, ← tensor_comp, id_comp, comp_id] + +@[reassoc (attr := simp)] +theorem hom_inv_whiskerLeft (X : C) {Y Z : C} (f : Y β‰… Z) : + X ◁ f.hom ≫ X ◁ f.inv = πŸ™ (X βŠ— Y) := by rw [← whiskerLeft_comp, hom_inv_id, whiskerLeft_id] + +@[reassoc (attr := simp)] +theorem hom_inv_whiskerRight {X Y : C} (f : X β‰… Y) (Z : C) : + f.hom β–· Z ≫ f.inv β–· Z = πŸ™ (X βŠ— Z) := by rw [← comp_whiskerRight, hom_inv_id, id_whiskerRight] + +@[reassoc (attr := simp)] +theorem inv_hom_whiskerLeft (X : C) {Y Z : C} (f : Y β‰… Z) : + X ◁ f.inv ≫ X ◁ f.hom = πŸ™ (X βŠ— Z) := by rw [← whiskerLeft_comp, inv_hom_id, whiskerLeft_id] + +@[reassoc (attr := simp)] +theorem inv_hom_whiskerRight {X Y : C} (f : X β‰… Y) (Z : C) : + f.inv β–· Z ≫ f.hom β–· Z = πŸ™ (Y βŠ— Z) := by rw [← comp_whiskerRight, inv_hom_id, id_whiskerRight] + +/-- The left whiskering of a 2-isomorphism is a 2-isomorphism. -/ +@[simps] +def whiskerLeftIso (X : C) {Y Z : C} (f : Y β‰… Z) : X βŠ— Y β‰… X βŠ— Z where + hom := X ◁ f.hom + inv := X ◁ f.inv + +instance whiskerLeft_isIso (X : C) {Y Z : C} (f : Y ⟢ Z) [IsIso f] : IsIso (X ◁ f) := + IsIso.of_iso (whiskerLeftIso X (asIso f)) + +@[simp] +theorem inv_whiskerLeft (X : C) {Y Z : C} (f : Y ⟢ Z) [IsIso f] : + inv (X ◁ f) = X ◁ inv f := by + aesop_cat_nonterminal + simp only [← whiskerLeft_comp, whiskerLeft_id, IsIso.hom_inv_id] + +/-- The right whiskering of a 2-isomorphism is a 2-isomorphism. -/ +@[simps!] +def whiskerRightIso {X Y : C} (f : X β‰… Y) (Z : C) : X βŠ— Z β‰… Y βŠ— Z where + hom := f.hom β–· Z + inv := f.inv β–· Z + +instance whiskerRight_isIso {X Y : C} (f : X ⟢ Y) (Z : C) [IsIso f] : IsIso (f β–· Z) := + IsIso.of_iso (whiskerRightIso (asIso f) Z) + +@[simp] +theorem inv_whiskerRight {X Y : C} (f : X ⟢ Y) (Z : C) [IsIso f] : + inv (f β–· Z) = inv f β–· Z := by + aesop_cat_nonterminal + simp only [← comp_whiskerRight, id_whiskerRight, IsIso.hom_inv_id] + +end MonoidalCategory /-- The tensor product of two isomorphisms is an isomorphism. -/ @[simps] @@ -196,6 +314,8 @@ section variable {C : Type u} [Category.{v} C] [MonoidalCategory.{v} C] +attribute [local simp] tensorHom_def + instance tensor_isIso {W X Y Z : C} (f : W ⟢ X) [IsIso f] (g : Y ⟢ Z) [IsIso g] : IsIso (f βŠ— g) := IsIso.of_iso (asIso f βŠ— asIso g) #align category_theory.monoidal_category.tensor_is_iso CategoryTheory.MonoidalCategory.tensor_isIso @@ -205,9 +325,78 @@ theorem inv_tensor {W X Y Z : C} (f : W ⟢ X) [IsIso f] (g : Y ⟢ Z) [IsIso g] inv (f βŠ— g) = inv f βŠ— inv g := by -- Porting note: Replaced `ext` with `aesop_cat_nonterminal` aesop_cat_nonterminal - simp [← tensor_comp] + simp [whisker_exchange] #align category_theory.monoidal_category.inv_tensor CategoryTheory.MonoidalCategory.inv_tensor +section pentagon + +variable (W X Y Z : C) + +@[reassoc (attr := simp)] +theorem pentagon_inv : + W ◁ (Ξ±_ X Y Z).inv ≫ (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z = + (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv := + eq_of_inv_eq_inv (by simp) + +@[reassoc] +theorem pentagon_inv' (W X Y Z : C) : + (πŸ™ W βŠ— (Ξ±_ X Y Z).inv) ≫ (Ξ±_ W (X βŠ— Y) Z).inv ≫ ((Ξ±_ W X Y).inv βŠ— πŸ™ Z) = + (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv := + CategoryTheory.eq_of_inv_eq_inv (by simp [pentagon]) +#align category_theory.monoidal_category.pentagon_inv CategoryTheory.MonoidalCategory.pentagon_inv' + +@[reassoc (attr := simp)] +theorem pentagon_inv_inv_hom_hom_inv : + (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z ≫ (Ξ±_ (W βŠ— X) Y Z).hom = + W ◁ (Ξ±_ X Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).inv := by + rw [← cancel_epi (W ◁ (Ξ±_ X Y Z).inv), ← cancel_mono (Ξ±_ (W βŠ— X) Y Z).inv] + simp + +@[reassoc (attr := simp)] +theorem pentagon_inv_hom_hom_hom_inv : + (Ξ±_ (W βŠ— X) Y Z).inv ≫ (Ξ±_ W X Y).hom β–· Z ≫ (Ξ±_ W (X βŠ— Y) Z).hom = + (Ξ±_ W X (Y βŠ— Z)).hom ≫ W ◁ (Ξ±_ X Y Z).inv := + eq_of_inv_eq_inv (by simp) + +@[reassoc (attr := simp)] +theorem pentagon_hom_inv_inv_inv_inv : + W ◁ (Ξ±_ X Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv = + (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z := + by simp [← cancel_epi (W ◁ (Ξ±_ X Y Z).inv)] + +@[reassoc (attr := simp)] +theorem pentagon_hom_hom_inv_hom_hom : + (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom ≫ W ◁ (Ξ±_ X Y Z).inv = + (Ξ±_ W X Y).hom β–· Z ≫ (Ξ±_ W (X βŠ— Y) Z).hom := + eq_of_inv_eq_inv (by simp) + +@[reassoc (attr := simp)] +theorem pentagon_hom_inv_inv_inv_hom : + (Ξ±_ W X (Y βŠ— Z)).hom ≫ W ◁ (Ξ±_ X Y Z).inv ≫ (Ξ±_ W (X βŠ— Y) Z).inv = + (Ξ±_ (W βŠ— X) Y Z).inv ≫ (Ξ±_ W X Y).hom β–· Z := by + rw [← cancel_epi (Ξ±_ W X (Y βŠ— Z)).inv, ← cancel_mono ((Ξ±_ W X Y).inv β–· Z)] + simp + +@[reassoc (attr := simp)] +theorem pentagon_hom_hom_inv_inv_hom : + (Ξ±_ W (X βŠ— Y) Z).hom ≫ W ◁ (Ξ±_ X Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).inv = + (Ξ±_ W X Y).inv β–· Z ≫ (Ξ±_ (W βŠ— X) Y Z).hom := + eq_of_inv_eq_inv (by simp) + +@[reassoc (attr := simp)] +theorem pentagon_inv_hom_hom_hom_hom : + (Ξ±_ W X Y).inv β–· Z ≫ (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom = + (Ξ±_ W (X βŠ— Y) Z).hom ≫ W ◁ (Ξ±_ X Y Z).hom := + by simp [← cancel_epi ((Ξ±_ W X Y).hom β–· Z)] + +@[reassoc (attr := simp)] +theorem pentagon_inv_inv_hom_inv_inv : + (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv ≫ (Ξ±_ W X Y).hom β–· Z = + W ◁ (Ξ±_ X Y Z).inv ≫ (Ξ±_ W (X βŠ— Y) Z).inv := + eq_of_inv_eq_inv (by simp) + +end pentagon + variable {U V W X Y Z : C} theorem tensor_dite {P : Prop} [Decidable P] {W X Y Z : C} (f : W ⟢ X) (g : P β†’ (Y ⟢ Z)) @@ -244,26 +433,16 @@ theorem tensor_id_comp_id_tensor (f : W ⟢ X) (g : Y ⟢ Z) : (g βŠ— πŸ™ W) simp #align category_theory.monoidal_category.tensor_id_comp_id_tensor CategoryTheory.MonoidalCategory.tensor_id_comp_id_tensor -@[simp] -theorem rightUnitor_conjugation {X Y : C} (f : X ⟢ Y) : - f βŠ— πŸ™ (πŸ™_ C) = (ρ_ X).hom ≫ f ≫ (ρ_ Y).inv := by - rw [← rightUnitor_naturality_assoc, Iso.hom_inv_id, Category.comp_id] -#align category_theory.monoidal_category.right_unitor_conjugation CategoryTheory.MonoidalCategory.rightUnitor_conjugation - -@[simp] -theorem leftUnitor_conjugation {X Y : C} (f : X ⟢ Y) : πŸ™ (πŸ™_ C) βŠ— f = (Ξ»_ X).hom ≫ f ≫ (Ξ»_ Y).inv - := by rw [← leftUnitor_naturality_assoc, Iso.hom_inv_id, Category.comp_id] -#align category_theory.monoidal_category.left_unitor_conjugation CategoryTheory.MonoidalCategory.leftUnitor_conjugation +-- @[simp] +-- theorem rightUnitor_conjugation {X Y : C} (f : X ⟢ Y) : +-- f βŠ— πŸ™ (πŸ™_ C) = (ρ_ X).hom ≫ f ≫ (ρ_ Y).inv := by +-- rw [← rightUnitor_naturality_assoc, Iso.hom_inv_id, Category.comp_id] +-- #align category_theory.monoidal_category.right_unitor_conjugation CategoryTheory.MonoidalCategory.rightUnitor_conjugation -@[reassoc] -theorem leftUnitor_inv_naturality {X X' : C} (f : X ⟢ X') : - f ≫ (Ξ»_ X').inv = (Ξ»_ X).inv ≫ (πŸ™ _ βŠ— f) := by simp -#align category_theory.monoidal_category.left_unitor_inv_naturality CategoryTheory.MonoidalCategory.leftUnitor_inv_naturality - -@[reassoc] -theorem rightUnitor_inv_naturality {X X' : C} (f : X ⟢ X') : - f ≫ (ρ_ X').inv = (ρ_ X).inv ≫ (f βŠ— πŸ™ _) := by simp -#align category_theory.monoidal_category.right_unitor_inv_naturality CategoryTheory.MonoidalCategory.rightUnitor_inv_naturality +-- @[simp] +-- theorem leftUnitor_conjugation {X Y : C} (f : X ⟢ Y) : πŸ™ (πŸ™_ C) βŠ— f = (Ξ»_ X).hom ≫ f ≫ (Ξ»_ Y).inv +-- := by rw [← leftUnitor_naturality_assoc, Iso.hom_inv_id, Category.comp_id] +-- #align category_theory.monoidal_category.left_unitor_conjugation CategoryTheory.MonoidalCategory.leftUnitor_conjugation theorem tensor_left_iff {X Y : C} (f g : X ⟢ Y) : πŸ™ (πŸ™_ C) βŠ— f = πŸ™ (πŸ™_ C) βŠ— g ↔ f = g := by simp #align category_theory.monoidal_category.tensor_left_iff CategoryTheory.MonoidalCategory.tensor_left_iff @@ -277,47 +456,205 @@ but we prove them directly as they are used in proving the coherence theorem. -/ section +@[reassoc (attr := simp)] +theorem triangle_assoc_comp_right (X Y : C) : + (Ξ±_ X (πŸ™_ C) Y).inv ≫ ((ρ_ X).hom β–· Y) = X ◁ (Ξ»_ Y).hom := by + rw [← triangle, Iso.inv_hom_id_assoc] + +@[reassoc (attr := simp)] +theorem triangle_assoc_comp_right_inv (X Y : C) : + (ρ_ X).inv β–· Y ≫ (Ξ±_ X (πŸ™_ C) Y).hom = X ◁ (Ξ»_ Y).inv := by + simp [← cancel_mono (X ◁ (Ξ»_ Y).hom)] + +@[reassoc (attr := simp)] +theorem triangle_assoc_comp_left_inv (X Y : C) : + (X ◁ (Ξ»_ Y).inv) ≫ (Ξ±_ X (πŸ™_ C) Y).inv = (ρ_ X).inv β–· Y := by + simp [← cancel_mono ((ρ_ X).hom β–· Y)] + @[reassoc] -theorem pentagon_inv (W X Y Z : C) : - (πŸ™ W βŠ— (Ξ±_ X Y Z).inv) ≫ (Ξ±_ W (X βŠ— Y) Z).inv ≫ ((Ξ±_ W X Y).inv βŠ— πŸ™ Z) = - (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv := - CategoryTheory.eq_of_inv_eq_inv (by simp [pentagon]) -#align category_theory.monoidal_category.pentagon_inv CategoryTheory.MonoidalCategory.pentagon_inv +theorem associator_naturality_left {X X' : C} (f : X ⟢ X') (Y Z : C) : + f β–· Y β–· Z ≫ (Ξ±_ X' Y Z).hom = (Ξ±_ X Y Z).hom ≫ f β–· (Y βŠ— Z) := by simp + +@[reassoc] +theorem associator_inv_naturality_left {X X' : C} (f : X ⟢ X') (Y Z : C) : + f β–· (Y βŠ— Z) ≫ (Ξ±_ X' Y Z).inv = (Ξ±_ X Y Z).inv ≫ f β–· Y β–· Z := by simp + +@[reassoc] +theorem whiskerRight_tensor_symm {X X' : C} (f : X ⟢ X') (Y Z : C) : + f β–· Y β–· Z = (Ξ±_ X Y Z).hom ≫ f β–· (Y βŠ— Z) ≫ (Ξ±_ X' Y Z).inv := by simp + +@[reassoc] +theorem associator_naturality_middle (X : C) {Y Y' : C} (f : Y ⟢ Y') (Z : C) : + (X ◁ f) β–· Z ≫ (Ξ±_ X Y' Z).hom = (Ξ±_ X Y Z).hom ≫ X ◁ f β–· Z := by simp + +@[reassoc] +theorem associator_inv_naturality_middle (X : C) {Y Y' : C} (f : Y ⟢ Y') (Z : C) : + X ◁ f β–· Z ≫ (Ξ±_ X Y' Z).inv = (Ξ±_ X Y Z).inv ≫ (X ◁ f) β–· Z := by simp + +@[reassoc] +theorem whisker_assoc_symm (X : C) {Y Y' : C} (f : Y ⟢ Y') (Z : C) : + X ◁ f β–· Z = (Ξ±_ X Y Z).inv ≫ (X ◁ f) β–· Z ≫ (Ξ±_ X Y' Z).hom := by simp + +@[reassoc] +theorem associator_naturality_right (X Y : C) {Z Z' : C} (f : Z ⟢ Z') : + (X βŠ— Y) ◁ f ≫ (Ξ±_ X Y Z').hom = (Ξ±_ X Y Z).hom ≫ X ◁ Y ◁ f := by simp + +@[reassoc] +theorem associator_inv_naturality_right (X Y : C) {Z Z' : C} (f : Z ⟢ Z') : + X ◁ Y ◁ f ≫ (Ξ±_ X Y Z').inv = (Ξ±_ X Y Z).inv ≫ (X βŠ— Y) ◁ f := by simp + +@[reassoc] +theorem tensor_whiskerLeft_symm (X Y : C) {Z Z' : C} (f : Z ⟢ Z') : + X ◁ Y ◁ f = (Ξ±_ X Y Z).inv ≫ (X βŠ— Y) ◁ f ≫ (Ξ±_ X Y Z').hom := by simp + +@[reassoc] +theorem leftUnitor_naturality {X Y : C} (f : X ⟢ Y) : + πŸ™_ C ◁ f ≫ (Ξ»_ Y).hom = (Ξ»_ X).hom ≫ f := + by simp + +@[reassoc] +theorem leftUnitor_inv_naturality {X Y : C} (f : X ⟢ Y) : + f ≫ (Ξ»_ Y).inv = (Ξ»_ X).inv ≫ (_ ◁ f) := by simp + +theorem id_whiskerLeft_symm {X X' : C} (f : X ⟢ X') : + f = (Ξ»_ X).inv ≫ πŸ™_ C ◁ f ≫ (Ξ»_ X').hom := by + simp + +@[reassoc] +theorem rightUnitor_naturality {X Y : C} (f : X ⟢ Y) : + f β–· πŸ™_ C ≫ (ρ_ Y).hom = (ρ_ X).hom ≫ f := by + simp + +@[reassoc] +theorem rightUnitor_inv_naturality {X X' : C} (f : X ⟢ X') : + f ≫ (ρ_ X').inv = (ρ_ X).inv ≫ (f β–· _) := by simp + +theorem whiskerRight_id_symm {X Y : C} (f : X ⟢ Y) : + f = (ρ_ X).inv ≫ f β–· πŸ™_ C ≫ (ρ_ Y).hom := by + simp + +theorem whiskerLeft_iff {X Y : C} (f g : X ⟢ Y) : πŸ™_ C ◁ f = πŸ™_ C ◁ g ↔ f = g := by simp + +theorem whiskerRight_iff {X Y : C} (f g : X ⟢ Y) : f β–· πŸ™_ C = g β–· πŸ™_ C ↔ f = g := by simp + +/-- We state it as a simp lemma, which is regarded as an involved version of +`id_whiskerRight X Y : πŸ™ X β–· Y = πŸ™ (X βŠ— Y)`. +-/ @[reassoc, simp] -theorem rightUnitor_tensor (X Y : C) : - (ρ_ (X βŠ— Y)).hom = (Ξ±_ X Y (πŸ™_ C)).hom ≫ (πŸ™ X βŠ— (ρ_ Y).hom) := by - rw [← tensor_right_iff, comp_tensor_id, ← cancel_mono (Ξ±_ X Y (πŸ™_ C)).hom, assoc, - associator_naturality, ← triangle_assoc, ← triangle, id_tensor_comp, pentagon_assoc, ← - associator_naturality, tensor_id] -#align category_theory.monoidal_category.right_unitor_tensor CategoryTheory.MonoidalCategory.rightUnitor_tensor +theorem leftUnitor_whiskerRight (X Y : C) : + (Ξ»_ X).hom β–· Y = (Ξ±_ (πŸ™_ C) X Y).hom ≫ (Ξ»_ (X βŠ— Y)).hom := by + rw [← whiskerLeft_iff, whiskerLeft_comp, ← cancel_epi (Ξ±_ _ _ _).hom, ← + cancel_epi ((Ξ±_ _ _ _).hom β–· _), pentagon_assoc, triangle, ← associator_naturality_middle, ← + comp_whiskerRight_assoc, triangle, associator_naturality_left] @[reassoc, simp] +theorem leftUnitor_inv_whiskerRight (X Y : C) : + (Ξ»_ X).inv β–· Y = (Ξ»_ (X βŠ— Y)).inv ≫ (Ξ±_ (πŸ™_ C) X Y).inv := + eq_of_inv_eq_inv (by simp) + +@[reassoc, simp] +theorem whiskerLeft_rightUnitor (X Y : C) : + X ◁ (ρ_ Y).hom = (Ξ±_ X Y (πŸ™_ C)).inv ≫ (ρ_ (X βŠ— Y)).hom := by + rw [← whiskerRight_iff, comp_whiskerRight, ← cancel_epi (Ξ±_ _ _ _).inv, ← + cancel_epi (X ◁ (Ξ±_ _ _ _).inv), pentagon_inv_assoc, triangle_assoc_comp_right, ← + associator_inv_naturality_middle, ← whiskerLeft_comp_assoc, triangle_assoc_comp_right, + associator_inv_naturality_right] + +@[reassoc, simp] +theorem whiskerLeft_rightUnitor_inv (X Y : C) : + X ◁ (ρ_ Y).inv = (ρ_ (X βŠ— Y)).inv ≫ (Ξ±_ X Y (πŸ™_ C)).hom := + eq_of_inv_eq_inv (by simp) + +@[reassoc] +theorem leftUnitor_tensor (X Y : C) : + (Ξ»_ (X βŠ— Y)).hom = (Ξ±_ (πŸ™_ C) X Y).inv ≫ (Ξ»_ X).hom β–· Y := by simp + +@[reassoc] +theorem leftUnitor_tensor_inv (X Y : C) : + (Ξ»_ (X βŠ— Y)).inv = (Ξ»_ X).inv β–· Y ≫ (Ξ±_ (πŸ™_ C) X Y).hom := by simp + +@[reassoc] +theorem rightUnitor_tensor (X Y : C) : + (ρ_ (X βŠ— Y)).hom = (Ξ±_ X Y (πŸ™_ C)).hom ≫ X ◁ (ρ_ Y).hom := by simp + +@[reassoc] theorem rightUnitor_tensor_inv (X Y : C) : + (ρ_ (X βŠ— Y)).inv = X ◁ (ρ_ Y).inv ≫ (Ξ±_ X Y (πŸ™_ C)).inv := by simp + +@[reassoc] +theorem associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (f₃ : X₃ ⟢ Y₃) : + tensorHom (tensorHom f₁ fβ‚‚) f₃ ≫ (associator Y₁ Yβ‚‚ Y₃).hom = + (associator X₁ Xβ‚‚ X₃).hom ≫ tensorHom f₁ (tensorHom fβ‚‚ f₃) := by + simp + +@[reassoc] +theorem triangle' (X Y : C) : + (Ξ±_ X (πŸ™_ C) Y).hom ≫ (πŸ™ X βŠ— (Ξ»_ Y).hom) = (ρ_ X).hom βŠ— πŸ™ Y := by + simp +#align category_theory.monoidal_category.triangle CategoryTheory.MonoidalCategory.triangle' + +@[reassoc] +theorem pentagon' (W X Y Z : C) : + ((Ξ±_ W X Y).hom βŠ— πŸ™ Z) ≫ (Ξ±_ W (X βŠ— Y) Z).hom ≫ (πŸ™ W βŠ— (Ξ±_ X Y Z).hom) = (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom := by + simp +#align category_theory.monoidal_category.pentagon CategoryTheory.MonoidalCategory.pentagon' + +@[reassoc] +theorem rightUnitor_tensor' (X Y : C) : + (ρ_ (X βŠ— Y)).hom = (Ξ±_ X Y (πŸ™_ C)).hom ≫ (πŸ™ X βŠ— (ρ_ Y).hom) := by + simp +#align category_theory.monoidal_category.right_unitor_tensor CategoryTheory.MonoidalCategory.rightUnitor_tensor' + +@[reassoc] +theorem rightUnitor_tensor_inv' (X Y : C) : (ρ_ (X βŠ— Y)).inv = (πŸ™ X βŠ— (ρ_ Y).inv) ≫ (Ξ±_ X Y (πŸ™_ C)).inv := eq_of_inv_eq_inv (by simp) -#align category_theory.monoidal_category.right_unitor_tensor_inv CategoryTheory.MonoidalCategory.rightUnitor_tensor_inv +#align category_theory.monoidal_category.right_unitor_tensor_inv CategoryTheory.MonoidalCategory.rightUnitor_tensor_inv' -@[reassoc (attr := simp)] -theorem triangle_assoc_comp_right (X Y : C) : +@[reassoc] +theorem triangle_assoc_comp_right' (X Y : C) : (Ξ±_ X (πŸ™_ C) Y).inv ≫ ((ρ_ X).hom βŠ— πŸ™ Y) = πŸ™ X βŠ— (Ξ»_ Y).hom := by - rw [← triangle, Iso.inv_hom_id_assoc] -#align category_theory.monoidal_category.triangle_assoc_comp_right CategoryTheory.MonoidalCategory.triangle_assoc_comp_right + simp +#align category_theory.monoidal_category.triangle_assoc_comp_right CategoryTheory.MonoidalCategory.triangle_assoc_comp_right' -@[reassoc (attr := simp)] -theorem triangle_assoc_comp_left_inv (X Y : C) : +@[reassoc] +theorem triangle_assoc_comp_left_inv' (X Y : C) : (πŸ™ X βŠ— (Ξ»_ Y).inv) ≫ (Ξ±_ X (πŸ™_ C) Y).inv = (ρ_ X).inv βŠ— πŸ™ Y := by - apply (cancel_mono ((ρ_ X).hom βŠ— πŸ™ Y)).1 - simp only [triangle_assoc_comp_right, assoc] - rw [← id_tensor_comp, Iso.inv_hom_id, ← comp_tensor_id, Iso.inv_hom_id] -#align category_theory.monoidal_category.triangle_assoc_comp_left_inv CategoryTheory.MonoidalCategory.triangle_assoc_comp_left_inv + simp +#align category_theory.monoidal_category.triangle_assoc_comp_left_inv CategoryTheory.MonoidalCategory.triangle_assoc_comp_left_inv' + +theorem rightUnitor_conjugation {X Y : C} (f : X ⟢ Y) : + f βŠ— πŸ™ (πŸ™_ C) = (ρ_ X).hom ≫ f ≫ (ρ_ Y).inv := by + simp +#align category_theory.monoidal_category.right_unitor_conjugation CategoryTheory.MonoidalCategory.rightUnitor_conjugation + +theorem leftUnitor_conjugation {X Y : C} (f : X ⟢ Y) : πŸ™ (πŸ™_ C) βŠ— f = (Ξ»_ X).hom ≫ f ≫ (Ξ»_ Y).inv := by simp +#align category_theory.monoidal_category.left_unitor_conjugation CategoryTheory.MonoidalCategory.leftUnitor_conjugation + +@[reassoc] +theorem leftUnitor_naturality' {X Y : C} (f : X ⟢ Y) : + (πŸ™ (πŸ™_ C) βŠ— f) ≫ (Ξ»_ Y).hom = (Ξ»_ X).hom ≫ f := + by simp + +@[reassoc] +theorem rightUnitor_naturality' {X Y : C} (f : X ⟢ Y) : + (f βŠ— πŸ™ (πŸ™_ C)) ≫ (ρ_ Y).hom = (ρ_ X).hom ≫ f := by + simp + +@[reassoc] +theorem leftUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : f ≫ (Ξ»_ X').inv = (Ξ»_ X).inv ≫ (πŸ™ _ βŠ— f) := by simp +#align category_theory.monoidal_category.left_unitor_inv_naturality CategoryTheory.MonoidalCategory.leftUnitor_inv_naturality' + +@[reassoc] +theorem rightUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : f ≫ (ρ_ X').inv = (ρ_ X).inv ≫ (f βŠ— πŸ™ _) := by simp +#align category_theory.monoidal_category.right_unitor_inv_naturality CategoryTheory.MonoidalCategory.rightUnitor_inv_naturality' end @[reassoc] theorem associator_inv_naturality {X Y Z X' Y' Z' : C} (f : X ⟢ X') (g : Y ⟢ Y') (h : Z ⟢ Z') : (f βŠ— g βŠ— h) ≫ (Ξ±_ X' Y' Z').inv = (Ξ±_ X Y Z).inv ≫ ((f βŠ— g) βŠ— h) := by - rw [comp_inv_eq, assoc, associator_naturality] simp #align category_theory.monoidal_category.associator_inv_naturality CategoryTheory.MonoidalCategory.associator_inv_naturality @@ -395,17 +732,82 @@ theorem tensor_inv_hom_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) rw [← tensor_comp, IsIso.inv_hom_id, comp_tensor_id] #align category_theory.monoidal_category.tensor_inv_hom_id' CategoryTheory.MonoidalCategory.tensor_inv_hom_id' +def ofTensorHom + (tensorObj : C β†’ C β†’ C) + (tensorHom : βˆ€ {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C}, (X₁ ⟢ Y₁) β†’ (Xβ‚‚ ⟢ Yβ‚‚) β†’ (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚)) + (tensor_id : βˆ€ X₁ Xβ‚‚ : C, tensorHom (πŸ™ X₁) (πŸ™ Xβ‚‚) = πŸ™ (tensorObj X₁ Xβ‚‚) := by + aesop_cat) + (tensor_comp : + βˆ€ {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : Y₁ ⟢ Z₁) (gβ‚‚ : Yβ‚‚ ⟢ Zβ‚‚), + tensorHom (f₁ ≫ g₁) (fβ‚‚ ≫ gβ‚‚) = tensorHom f₁ fβ‚‚ ≫ tensorHom g₁ gβ‚‚ := by + aesop_cat) + (tensorUnit' : C) + (associator : βˆ€ X Y Z : C, tensorObj (tensorObj X Y) Z β‰… tensorObj X (tensorObj Y Z)) + (associator_naturality : + βˆ€ {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (f₃ : X₃ ⟢ Y₃), + tensorHom (tensorHom f₁ fβ‚‚) f₃ ≫ (associator Y₁ Yβ‚‚ Y₃).hom = + (associator X₁ Xβ‚‚ X₃).hom ≫ tensorHom f₁ (tensorHom fβ‚‚ f₃) := by + aesop_cat) + (leftUnitor : βˆ€ X : C, tensorObj tensorUnit' X β‰… X) + (leftUnitor_naturality : + βˆ€ {X Y : C} (f : X ⟢ Y), + tensorHom (πŸ™ tensorUnit') f ≫ (leftUnitor Y).hom = (leftUnitor X).hom ≫ f := by + aesop_cat) + (rightUnitor : βˆ€ X : C, tensorObj X tensorUnit' β‰… X) + (rightUnitor_naturality : + βˆ€ {X Y : C} (f : X ⟢ Y), + tensorHom f (πŸ™ tensorUnit') ≫ (rightUnitor Y).hom = (rightUnitor X).hom ≫ f := by + aesop_cat) + (pentagon : + βˆ€ W X Y Z : C, + tensorHom (associator W X Y).hom (πŸ™ Z) ≫ + (associator W (tensorObj X Y) Z).hom ≫ tensorHom (πŸ™ W) (associator X Y Z).hom = + (associator (tensorObj W X) Y Z).hom ≫ (associator W X (tensorObj Y Z)).hom := by + aesop_cat) + (triangle : + βˆ€ X Y : C, + (associator X tensorUnit' Y).hom ≫ tensorHom (πŸ™ X) (leftUnitor Y).hom = + tensorHom (rightUnitor X).hom (πŸ™ Y) := by + aesop_cat) : + MonoidalCategory C where + tensorObj := tensorObj + tensorHom := tensorHom + whiskerLeft X _ _ f := tensorHom (πŸ™ X) f + whiskerRight f X := tensorHom f (πŸ™ X) + tensorHom_def f g := by rw [← tensor_comp, id_comp, comp_id] + tensorUnit' := tensorUnit' + leftUnitor := leftUnitor + rightUnitor := rightUnitor + associator := associator + whiskerLeft_id := by intros; simp [tensor_id] + whiskerLeft_comp := by intros; simp [← tensor_comp] + id_whiskerLeft := by intros; rw [← assoc, ← leftUnitor_naturality]; simp + tensor_whiskerLeft := by intros; rw [← assoc, ← associator_naturality]; simp [tensor_id] + id_whiskerRight := by intros; simp [tensor_id] + comp_whiskerRight := by intros; simp [← tensor_comp] + whiskerRight_id := by intros; rw [← assoc, ← rightUnitor_naturality]; simp + whiskerRight_tensor := by intros; rw [associator_naturality]; simp [tensor_id] + whisker_assoc := by intros; rw [← assoc, ← associator_naturality]; simp + whisker_exchange := by intros; dsimp; rw [← tensor_comp, ← tensor_comp]; simp + pentagon := pentagon + triangle := triangle + end section variable (C : Type u) [Category.{v} C] [MonoidalCategory.{v} C] +attribute [local simp] tensorHom_def + /-- The tensor product expressed as a functor. -/ @[simps] def tensor : C Γ— C β₯€ C where obj X := X.1 βŠ— X.2 map {X Y : C Γ— C} (f : X ⟢ Y) := f.1 βŠ— f.2 + map_comp := by + intro X Y Z f g + simp [whisker_exchange_assoc] #align category_theory.monoidal_category.tensor CategoryTheory.MonoidalCategory.tensor /-- The left-associated triple tensor product as a functor. -/ @@ -413,6 +815,10 @@ def leftAssocTensor : C Γ— C Γ— C β₯€ C where obj X := (X.1 βŠ— X.2.1) βŠ— X.2.2 map {X Y : C Γ— C Γ— C} (f : X ⟢ Y) := (f.1 βŠ— f.2.1) βŠ— f.2.2 + map_comp := by + intro X Y Z f g + simp [-tensorHom_def, tensor_comp] + #align category_theory.monoidal_category.left_assoc_tensor CategoryTheory.MonoidalCategory.leftAssocTensor @[simp] @@ -430,6 +836,9 @@ def rightAssocTensor : C Γ— C Γ— C β₯€ C where obj X := X.1 βŠ— X.2.1 βŠ— X.2.2 map {X Y : C Γ— C Γ— C} (f : X ⟢ Y) := f.1 βŠ— f.2.1 βŠ— f.2.2 + map_comp := by + intro X Y Z f g + simp [-tensorHom_def, tensor_comp] #align category_theory.monoidal_category.right_assoc_tensor CategoryTheory.MonoidalCategory.rightAssocTensor @[simp] @@ -445,13 +854,13 @@ theorem rightAssocTensor_map {X Y} (f : X ⟢ Y) : (rightAssocTensor C).map f = /-- The functor `fun X ↦ πŸ™_ C βŠ— X`. -/ def tensorUnitLeft : C β₯€ C where obj X := πŸ™_ C βŠ— X - map {X Y : C} (f : X ⟢ Y) := πŸ™ (πŸ™_ C) βŠ— f + map {X Y : C} (f : X ⟢ Y) := πŸ™_ C ◁ f #align category_theory.monoidal_category.tensor_unit_left CategoryTheory.MonoidalCategory.tensorUnitLeft /-- The functor `fun X ↦ X βŠ— πŸ™_ C`. -/ def tensorUnitRight : C β₯€ C where obj X := X βŠ— πŸ™_ C - map {X Y : C} (f : X ⟢ Y) := f βŠ— πŸ™ (πŸ™_ C) + map {X Y : C} (f : X ⟢ Y) := f β–· πŸ™_ C #align category_theory.monoidal_category.tensor_unit_right CategoryTheory.MonoidalCategory.tensorUnitRight -- We can express the associator and the unitors, given componentwise above, @@ -504,17 +913,14 @@ variable {C : Type u} [Category.{v} C] [MonoidalCategory.{v} C] @[simps] def tensorLeft (X : C) : C β₯€ C where obj Y := X βŠ— Y - map {Y} {Y'} f := πŸ™ X βŠ— f + map {Y} {Y'} f := X ◁ f #align category_theory.monoidal_category.tensor_left CategoryTheory.MonoidalCategory.tensorLeft /-- Tensoring on the left with `X βŠ— Y` is naturally isomorphic to tensoring on the left with `Y`, and then again with `X`. -/ def tensorLeftTensor (X Y : C) : tensorLeft (X βŠ— Y) β‰… tensorLeft Y β‹™ tensorLeft X := - NatIso.ofComponents (associator _ _) fun {Z} {Z'} f => by - dsimp - rw [← tensor_id] - apply associator_naturality + NatIso.ofComponents (associator _ _) fun {Z} {Z'} f => by simp #align category_theory.monoidal_category.tensor_left_tensor CategoryTheory.MonoidalCategory.tensorLeftTensor @[simp] @@ -532,7 +938,7 @@ theorem tensorLeftTensor_inv_app (X Y Z : C) : @[simps] def tensorRight (X : C) : C β₯€ C where obj Y := Y βŠ— X - map {Y} {Y'} f := f βŠ— πŸ™ X + map {Y} {Y'} f := f β–· X #align category_theory.monoidal_category.tensor_right CategoryTheory.MonoidalCategory.tensorRight -- Porting Note: This used to be `variable (C)` but it seems like Lean 4 parses that differently @@ -545,7 +951,11 @@ TODO: show this is an op-monoidal functor. @[simps] def tensoringLeft : C β₯€ C β₯€ C where obj := tensorLeft - map {X} {Y} f := { app := fun Z => f βŠ— πŸ™ Z } + map {X} {Y} f := { + app := fun Z => f β–· Z + naturality := by + intros + simp [whisker_exchange] } #align category_theory.monoidal_category.tensoring_left CategoryTheory.MonoidalCategory.tensoringLeft instance : Faithful (tensoringLeft C) where @@ -561,7 +971,11 @@ We later show this is a monoidal functor. @[simps] def tensoringRight : C β₯€ C β₯€ C where obj := tensorRight - map {X} {Y} f := { app := fun Z => πŸ™ Z βŠ— f } + map {X} {Y} f := { + app := fun Z => Z ◁ f + naturality := by + intros + simp [whisker_exchange] } #align category_theory.monoidal_category.tensoring_right CategoryTheory.MonoidalCategory.tensoringRight instance : Faithful (tensoringRight C) where @@ -577,10 +991,7 @@ variable {C : Type u} [Category.{v} C] [MonoidalCategory.{v} C] tensoring on the right with `X`, and then again with `Y`. -/ def tensorRightTensor (X Y : C) : tensorRight (X βŠ— Y) β‰… tensorRight X β‹™ tensorRight Y := - NatIso.ofComponents (fun Z => (associator Z X Y).symm) fun {Z} {Z'} f => by - dsimp - rw [← tensor_id] - apply associator_inv_naturality + NatIso.ofComponents (fun Z => (associator Z X Y).symm) fun {Z} {Z'} f => by simp #align category_theory.monoidal_category.tensor_right_tensor CategoryTheory.MonoidalCategory.tensorRightTensor @[simp] @@ -607,17 +1018,26 @@ variable (C₁ : Type u₁) [Category.{v₁} C₁] [MonoidalCategory.{v₁} C₁ variable (Cβ‚‚ : Type uβ‚‚) [Category.{vβ‚‚} Cβ‚‚] [MonoidalCategory.{vβ‚‚} Cβ‚‚] attribute [local simp] associator_naturality leftUnitor_naturality rightUnitor_naturality pentagon +attribute [local simp] tensorHom_def -@[simps! tensorObj tensorHom tensorUnit' associator] +@[simps! tensorObj tensorUnit' whiskerLeft whiskerRight associator] instance prodMonoidal : MonoidalCategory (C₁ Γ— Cβ‚‚) where tensorObj X Y := (X.1 βŠ— Y.1, X.2 βŠ— Y.2) tensorHom f g := (f.1 βŠ— g.1, f.2 βŠ— g.2) + whiskerLeft X _ _ f := (whiskerLeft X.1 f.1, whiskerLeft X.2 f.2) + whiskerRight f X := (whiskerRight f.1 X.1, whiskerRight f.2 X.2) + whisker_exchange := by simp [whisker_exchange] tensorUnit' := (πŸ™_ C₁, πŸ™_ Cβ‚‚) associator X Y Z := (Ξ±_ X.1 Y.1 Z.1).prod (Ξ±_ X.2 Y.2 Z.2) leftUnitor := fun ⟨X₁, Xβ‚‚βŸ© => (Ξ»_ X₁).prod (Ξ»_ Xβ‚‚) rightUnitor := fun ⟨X₁, Xβ‚‚βŸ© => (ρ_ X₁).prod (ρ_ Xβ‚‚) #align category_theory.monoidal_category.prod_monoidal CategoryTheory.MonoidalCategory.prodMonoidal +@[simp] +theorem prodMonoidal_tensorHom {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C₁ Γ— Cβ‚‚} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : + f βŠ— g = (f.1 βŠ— g.1, f.2 βŠ— g.2) := + rfl + @[simp] theorem prodMonoidal_leftUnitor_hom_fst (X : C₁ Γ— Cβ‚‚) : ((Ξ»_ X).hom : πŸ™_ _ βŠ— X ⟢ X).1 = (Ξ»_ X.1).hom := by diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index cf98e943b68a9..f5ac70d24356e 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -56,9 +56,9 @@ structure HalfBraiding (X : C) where Ξ² : βˆ€ U, X βŠ— U β‰… U βŠ— X monoidal : βˆ€ U U', (Ξ² (U βŠ— U')).hom = (Ξ±_ _ _ _).inv ≫ - ((Ξ² U).hom βŠ— πŸ™ U') ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ U βŠ— (Ξ² U').hom) ≫ (Ξ±_ _ _ _).inv := by + ((Ξ² U).hom β–· U') ≫ (Ξ±_ _ _ _).hom ≫ (U ◁ (Ξ² U').hom) ≫ (Ξ±_ _ _ _).inv := by aesop_cat - naturality : βˆ€ {U U'} (f : U ⟢ U'), (πŸ™ X βŠ— f) ≫ (Ξ² U').hom = (Ξ² U).hom ≫ (f βŠ— πŸ™ X) := by + naturality : βˆ€ {U U'} (f : U ⟢ U'), (X ◁ f) ≫ (Ξ² U').hom = (Ξ² U).hom ≫ (f β–· X) := by aesop_cat #align category_theory.half_braiding CategoryTheory.HalfBraiding @@ -84,7 +84,7 @@ variable {C} @[ext] -- @[nolint has_nonempty_instance] -- Porting note: This linter does not exist yet. structure Hom (X Y : Center C) where f : X.1 ⟢ Y.1 - comm : βˆ€ U, (f βŠ— πŸ™ U) ≫ (Y.2.Ξ² U).hom = (X.2.Ξ² U).hom ≫ (πŸ™ U βŠ— f) := by aesop_cat + comm : βˆ€ U, (f β–· U) ≫ (Y.2.Ξ² U).hom = (X.2.Ξ² U).hom ≫ (U ◁ f) := by aesop_cat #align category_theory.center.hom CategoryTheory.Center.Hom attribute [reassoc (attr := simp)] Hom.comm @@ -118,7 +118,8 @@ a morphism whose underlying morphism is an isomorphism. def isoMk {X Y : Center C} (f : X ⟢ Y) [IsIso f.f] : X β‰… Y where hom := f inv := ⟨inv f.f, - fun U => by simp [← cancel_epi (f.f βŠ— πŸ™ U), ← comp_tensor_id_assoc, ← id_tensor_comp]⟩ + fun U => by simp [← cancel_epi (f.f β–· U), ← comp_whiskerRight_assoc, + ← MonoidalCategory.whiskerLeft_comp] ⟩ #align category_theory.center.iso_mk CategoryTheory.Center.isoMk instance isIso_of_f_isIso {X Y : Center C} (f : X ⟢ Y) [IsIso f.f] : IsIso f := by @@ -126,38 +127,43 @@ instance isIso_of_f_isIso {X Y : Center C} (f : X ⟢ Y) [IsIso f.f] : IsIso f : infer_instance #align category_theory.center.is_iso_of_f_is_iso CategoryTheory.Center.isIso_of_f_isIso +attribute [local simp] id_tensorHom tensorHom_id + /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ @[simps] def tensorObj (X Y : Center C) : Center C := ⟨X.1 βŠ— Y.1, { Ξ² := fun U => Ξ±_ _ _ _ β‰ͺ≫ - (Iso.refl X.1 βŠ— Y.2.Ξ² U) β‰ͺ≫ (Ξ±_ _ _ _).symm β‰ͺ≫ (X.2.Ξ² U βŠ— Iso.refl Y.1) β‰ͺ≫ Ξ±_ _ _ _ + (whiskerLeftIso X.1 (Y.2.Ξ² U)) β‰ͺ≫ (Ξ±_ _ _ _).symm β‰ͺ≫ + (whiskerRightIso (X.2.Ξ² U) Y.1) β‰ͺ≫ Ξ±_ _ _ _ monoidal := fun U U' => by - dsimp - simp only [comp_tensor_id, id_tensor_comp, Category.assoc, HalfBraiding.monoidal] - -- On the RHS, we'd like to commute `((X.snd.Ξ² U).hom βŠ— πŸ™ Y.fst) βŠ— πŸ™ U'` - -- and `πŸ™ U βŠ— πŸ™ X.fst βŠ— (Y.snd.Ξ² U').hom` past each other, - -- but there are some associators we need to get out of the way first. - slice_rhs 6 8 => rw [pentagon] - slice_rhs 5 6 => rw [associator_naturality] - slice_rhs 7 8 => rw [← associator_naturality] - slice_rhs 6 7 => - rw [tensor_id, tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id, - ← tensor_id, ← tensor_id] - -- Now insert associators as needed to make the four half-braidings look identical - slice_rhs 10 10 => rw [associator_inv_conjugation] - slice_rhs 7 7 => rw [associator_inv_conjugation] - slice_rhs 6 6 => rw [associator_conjugation] - slice_rhs 3 3 => rw [associator_conjugation] - -- Finish with an application of the coherence theorem. - coherence - naturality := fun f => by - dsimp - rw [Category.assoc, Category.assoc, Category.assoc, Category.assoc, - id_tensor_associator_naturality_assoc, ← id_tensor_comp_assoc, HalfBraiding.naturality, - id_tensor_comp_assoc, associator_inv_naturality_assoc, ← comp_tensor_id_assoc, - HalfBraiding.naturality, comp_tensor_id_assoc, associator_naturality, ← tensor_id] }⟩ + dsimp only [Iso.trans_hom, whiskerLeftIso_hom, Iso.symm_hom, whiskerRightIso_hom] + simp only [HalfBraiding.monoidal] + -- We'd like to commute `X.1 ◁ U ◁ (HalfBraiding.Ξ² Y.2 U').hom` + -- and `((HalfBraiding.Ξ² X.2 U).hom β–· U' β–· Y.1)` past each other. + -- We do this with the help of the monoidal composition `βŠ—β‰«`. + calc + _ = πŸ™ _ βŠ—β‰« + X.1 ◁ (HalfBraiding.Ξ² Y.2 U).hom β–· U' βŠ—β‰« + (_ ◁ (HalfBraiding.Ξ² Y.2 U').hom ≫ (HalfBraiding.Ξ² X.2 U).hom β–· _) βŠ—β‰« + U ◁ (HalfBraiding.Ξ² X.2 U').hom β–· Y.1 βŠ—β‰« πŸ™ _ := ?eq1 + _ = _ := ?eq2 + case eq1 => coherence + case eq2 => rw [whisker_exchange]; coherence + naturality := fun {U U'} f => by + dsimp only [Iso.trans_hom, whiskerLeftIso_hom, Iso.symm_hom, whiskerRightIso_hom] + calc + _ = πŸ™ _ βŠ—β‰« + (X.1 ◁ (Y.1 ◁ f ≫ (HalfBraiding.Ξ² Y.2 U').hom)) βŠ—β‰« + (HalfBraiding.Ξ² X.2 U').hom β–· Y.1 βŠ—β‰« πŸ™ _ := ?eq1 + _ = πŸ™ _ βŠ—β‰« + X.1 ◁ (HalfBraiding.Ξ² Y.2 U).hom βŠ—β‰« + (X.1 ◁ f ≫ (HalfBraiding.Ξ² X.2 U').hom) β–· Y.1 βŠ—β‰« πŸ™ _ := ?eq2 + _ = _ := ?eq3 + case eq1 => coherence + case eq2 => rw [HalfBraiding.naturality]; coherence + case eq3 => rw [HalfBraiding.naturality]; coherence }⟩ #align category_theory.center.tensor_obj CategoryTheory.Center.tensorObj /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ @@ -166,55 +172,66 @@ def tensorHom {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Center C} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚ where f := f.f βŠ— g.f comm U := by - dsimp - rw [Category.assoc, Category.assoc, Category.assoc, Category.assoc, associator_naturality_assoc, - ← tensor_id_comp_id_tensor, Category.assoc, ← id_tensor_comp_assoc, g.comm, - id_tensor_comp_assoc, tensor_id_comp_id_tensor_assoc, ← id_tensor_comp_tensor_id, - Category.assoc, associator_inv_naturality_assoc, id_tensor_associator_inv_naturality_assoc, - tensor_id, id_tensor_comp_tensor_id_assoc, ← tensor_id_comp_id_tensor g.f, Category.assoc, - ← comp_tensor_id_assoc, f.comm, comp_tensor_id_assoc, id_tensor_associator_naturality, - associator_naturality_assoc, ← id_tensor_comp, tensor_id_comp_id_tensor] + sorry #align category_theory.center.tensor_hom CategoryTheory.Center.tensorHom +/-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ +def whiskerLeft (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) : + tensorObj X Y₁ ⟢ tensorObj X Yβ‚‚ where + f := X.1 ◁ f.f + comm U := by + dsimp only [tensorObj_fst, tensorObj_snd_Ξ², Iso.trans_hom, whiskerLeftIso_hom, + Iso.symm_hom, whiskerRightIso_hom] + calc + _ = πŸ™ _ βŠ—β‰« + X.fst ◁ (f.f β–· U ≫ (HalfBraiding.Ξ² Yβ‚‚.snd U).hom) βŠ—β‰« + (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst βŠ—β‰« πŸ™ _ := ?_ + _ = πŸ™ _ βŠ—β‰« + X.fst ◁ (HalfBraiding.Ξ² Y₁.snd U).hom βŠ—β‰« + ((X.fst βŠ— U) ◁ f.f ≫ (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst) βŠ—β‰« πŸ™ _ := ?eq1 + _ = _ := ?eq2 + case eq1 => rw [f.comm]; coherence + case eq2 => rw [whisker_exchange]; coherence + any_goals coherence + +/-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ +def whiskerRight {X₁ Xβ‚‚ : Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) : + tensorObj X₁ Y ⟢ tensorObj Xβ‚‚ Y where + f := f.f β–· Y.1 + comm U := by + dsimp only [tensorObj_fst, tensorObj_snd_Ξ², Iso.trans_hom, whiskerLeftIso_hom, + Iso.symm_hom, whiskerRightIso_hom] + calc + _ = πŸ™ _ βŠ—β‰« + (f.f β–· (Y.fst βŠ— U) ≫ Xβ‚‚.fst ◁ (HalfBraiding.Ξ² Y.snd U).hom) βŠ—β‰« + (HalfBraiding.Ξ² Xβ‚‚.snd U).hom β–· Y.fst βŠ—β‰« πŸ™ _ := ?_ + _ = πŸ™ _ βŠ—β‰« + X₁.fst ◁ (HalfBraiding.Ξ² Y.snd U).hom βŠ—β‰« + (f.f β–· U ≫ (HalfBraiding.Ξ² Xβ‚‚.snd U).hom) β–· Y.fst βŠ—β‰« πŸ™ _ := ?eq1 + _ = _ := ?eq2 + case eq1 => rw [← whisker_exchange]; coherence + case eq2 => rw [f.comm]; coherence + any_goals coherence + /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ @[simps] def tensorUnit : Center C := - βŸ¨πŸ™_ C, - { Ξ² := fun U => Ξ»_ U β‰ͺ≫ (ρ_ U).symm - monoidal := fun U U' => by simp - naturality := fun f => by - dsimp - rw [leftUnitor_naturality_assoc, rightUnitor_inv_naturality, Category.assoc] }⟩ + βŸ¨πŸ™_ C, { Ξ² := fun U => Ξ»_ U β‰ͺ≫ (ρ_ U).symm }⟩ #align category_theory.center.tensor_unit CategoryTheory.Center.tensorUnit /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ def associator (X Y Z : Center C) : tensorObj (tensorObj X Y) Z β‰… tensorObj X (tensorObj Y Z) := - isoMk - ⟨(Ξ±_ X.1 Y.1 Z.1).hom, fun U => by - dsimp - simp only [comp_tensor_id, id_tensor_comp, ← tensor_id, associator_conjugation] - coherence⟩ + isoMk ⟨(Ξ±_ X.1 Y.1 Z.1).hom, fun U => by simp⟩ #align category_theory.center.associator CategoryTheory.Center.associator /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ def leftUnitor (X : Center C) : tensorObj tensorUnit X β‰… X := - isoMk - ⟨(Ξ»_ X.1).hom, fun U => by - dsimp - simp only [Category.comp_id, Category.assoc, tensor_inv_hom_id, comp_tensor_id, - tensor_id_comp_id_tensor, triangle_assoc_comp_right_inv] - rw [← leftUnitor_tensor, leftUnitor_naturality, leftUnitor_tensor'_assoc]⟩ + isoMk ⟨(Ξ»_ X.1).hom, fun U => by simp⟩ #align category_theory.center.left_unitor CategoryTheory.Center.leftUnitor /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ def rightUnitor (X : Center C) : tensorObj X tensorUnit β‰… X := - isoMk - ⟨(ρ_ X.1).hom, fun U => by - dsimp - simp only [tensor_id_comp_id_tensor_assoc, triangle_assoc, id_tensor_comp, Category.assoc] - rw [← tensor_id_comp_id_tensor_assoc (ρ_ U).inv, cancel_epi, ← rightUnitor_tensor_inv_assoc, - ← rightUnitor_inv_naturality_assoc] - simp⟩ + isoMk ⟨(ρ_ X.1).hom, fun U => by simp⟩ #align category_theory.center.right_unitor CategoryTheory.Center.rightUnitor section @@ -223,13 +240,17 @@ attribute [local simp] associator_naturality leftUnitor_naturality rightUnitor_n attribute [local simp] Center.associator Center.leftUnitor Center.rightUnitor +attribute [local simp] Center.whiskerLeft Center.whiskerRight + instance : MonoidalCategory (Center C) where tensorObj X Y := tensorObj X Y - tensorHom f g := tensorHom f g + whiskerLeft X _ _ f := whiskerLeft X f + whiskerRight f X := whiskerRight f X tensorUnit' := tensorUnit associator := associator leftUnitor := leftUnitor rightUnitor := rightUnitor + whisker_exchange := by intros; ext; simp [whisker_exchange] @[simp] theorem tensor_fst (X Y : Center C) : (X βŠ— Y).1 = X.1 βŠ— Y.1 := @@ -240,13 +261,24 @@ theorem tensor_fst (X Y : Center C) : (X βŠ— Y).1 = X.1 βŠ— Y.1 := theorem tensor_Ξ² (X Y : Center C) (U : C) : (X βŠ— Y).2.Ξ² U = Ξ±_ _ _ _ β‰ͺ≫ - (Iso.refl X.1 βŠ— Y.2.Ξ² U) β‰ͺ≫ (Ξ±_ _ _ _).symm β‰ͺ≫ (X.2.Ξ² U βŠ— Iso.refl Y.1) β‰ͺ≫ Ξ±_ _ _ _ := + (whiskerLeftIso X.1 (Y.2.Ξ² U)) β‰ͺ≫ (Ξ±_ _ _ _).symm β‰ͺ≫ + (whiskerRightIso (X.2.Ξ² U) Y.1) β‰ͺ≫ Ξ±_ _ _ _ := rfl #align category_theory.center.tensor_Ξ² CategoryTheory.Center.tensor_Ξ² @[simp] -theorem tensor_f {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Center C} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : (f βŠ— g).f = f.f βŠ— g.f := +theorem whiskerLeft_f (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) : + (X ◁ f).f = X.1 ◁ f.f:= by rfl + +@[simp] +theorem whiskerRight_f {X₁ Xβ‚‚ : Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) : + (f β–· Y).f = f.f β–· Y.1 := by + rfl + +@[simp] +theorem tensor_f {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Center C} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : (f βŠ— g).f = f.f βŠ— g.f := by + rw [tensorHom_def, tensorHom_def, comp_f, whiskerLeft_f, whiskerRight_f] #align category_theory.center.tensor_f CategoryTheory.Center.tensor_f @[simp] @@ -323,9 +355,9 @@ instance braidedCategoryCenter : BraidedCategory (Center C) where braiding := braiding braiding_naturality f g := by ext - dsimp - rw [← tensor_id_comp_id_tensor, Category.assoc, HalfBraiding.naturality, f.comm_assoc, - id_tensor_comp_tensor_id] + dsimp only [comp_f, braiding_hom_f] + erw [tensorHom_def', Category.assoc, HalfBraiding.naturality, f.comm_assoc, ← tensorHom_def, + tensor_f] #align category_theory.center.braided_category_center CategoryTheory.Center.braidedCategoryCenter -- `aesop_cat` handles the hexagon axioms diff --git a/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean b/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean index bb57c869710a8..6092a2dddbcba 100644 --- a/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean +++ b/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean @@ -24,21 +24,23 @@ namespace CategoryTheory.MonoidalCategory variable {C : Type _} [Category C] [MonoidalCategory C] +attribute [local simp] id_tensorHom tensorHom_id + -- See Proposition 2.2.4 of @[reassoc] -theorem leftUnitor_tensor' (X Y : C) : +theorem leftUnitor_tensor'' (X Y : C) : (Ξ±_ (πŸ™_ C) X Y).hom ≫ (Ξ»_ (X βŠ— Y)).hom = (Ξ»_ X).hom βŠ— πŸ™ Y := by coherence -#align category_theory.monoidal_category.left_unitor_tensor' CategoryTheory.MonoidalCategory.leftUnitor_tensor' +#align category_theory.monoidal_category.left_unitor_tensor' CategoryTheory.MonoidalCategory.leftUnitor_tensor'' -@[reassoc, simp] -theorem leftUnitor_tensor (X Y : C) : +@[reassoc] +theorem leftUnitor_tensor' (X Y : C) : (Ξ»_ (X βŠ— Y)).hom = (Ξ±_ (πŸ™_ C) X Y).inv ≫ ((Ξ»_ X).hom βŠ— πŸ™ Y) := by coherence #align category_theory.monoidal_category.left_unitor_tensor CategoryTheory.MonoidalCategory.leftUnitor_tensor @[reassoc] -theorem leftUnitor_tensor_inv (X Y : C) : +theorem leftUnitor_tensor_inv' (X Y : C) : (Ξ»_ (X βŠ— Y)).inv = ((Ξ»_ X).inv βŠ— πŸ™ Y) ≫ (Ξ±_ (πŸ™_ C) X Y).hom := by coherence #align category_theory.monoidal_category.left_unitor_tensor_inv CategoryTheory.MonoidalCategory.leftUnitor_tensor_inv @@ -60,7 +62,7 @@ theorem pentagon_inv_inv_hom (W X Y Z : C) : #align category_theory.monoidal_category.pentagon_inv_inv_hom CategoryTheory.MonoidalCategory.pentagon_inv_inv_hom @[reassoc (attr := simp)] -theorem triangle_assoc_comp_right_inv (X Y : C) : +theorem triangle_assoc_comp_right_inv' (X Y : C) : ((ρ_ X).inv βŠ— πŸ™ Y) ≫ (Ξ±_ X (πŸ™_ C) Y).hom = πŸ™ X βŠ— (Ξ»_ Y).inv := by coherence #align category_theory.monoidal_category.triangle_assoc_comp_right_inv CategoryTheory.MonoidalCategory.triangle_assoc_comp_right_inv diff --git a/Mathlib/CategoryTheory/Monoidal/Discrete.lean b/Mathlib/CategoryTheory/Monoidal/Discrete.lean index 6f7c69ce4bdf3..f303296b04ede 100644 --- a/Mathlib/CategoryTheory/Monoidal/Discrete.lean +++ b/Mathlib/CategoryTheory/Monoidal/Discrete.lean @@ -30,7 +30,9 @@ instance Discrete.monoidal : MonoidalCategory (Discrete M) where tensorUnit' := Discrete.mk 1 tensorObj X Y := Discrete.mk (X.as * Y.as) - tensorHom f g := eqToHom (by dsimp; rw [eq_of_hom f, eq_of_hom g]) + -- tensorHom f g := eqToHom (by dsimp; rw [eq_of_hom f, eq_of_hom g]) + whiskerLeft X _ _ f := eqToHom (by dsimp; rw [eq_of_hom f]) + whiskerRight f X := eqToHom (by dsimp; rw [eq_of_hom f]) leftUnitor X := Discrete.eqToIso (one_mul X.as) rightUnitor X := Discrete.eqToIso (mul_one X.as) associator X Y Z := Discrete.eqToIso (mul_assoc _ _ _) diff --git a/Mathlib/CategoryTheory/Monoidal/End.lean b/Mathlib/CategoryTheory/Monoidal/End.lean index c4d28e456bb3d..445f6ba7b89f8 100644 --- a/Mathlib/CategoryTheory/Monoidal/End.lean +++ b/Mathlib/CategoryTheory/Monoidal/End.lean @@ -32,7 +32,9 @@ with tensor product given by composition of functors -/ def endofunctorMonoidalCategory : MonoidalCategory (C β₯€ C) where tensorObj F G := F β‹™ G - tensorHom Ξ± Ξ² := Ξ± β—« Ξ² + whiskerLeft X _ _ F := whiskerLeft X F + whiskerRight F X := whiskerRight F X + -- tensorHom Ξ± Ξ² := Ξ± β—« Ξ² tensorUnit' := 𝟭 C associator F G H := Functor.associator F G H leftUnitor F := Functor.leftUnitor F @@ -59,6 +61,14 @@ attribute [local instance] endofunctorMonoidalCategory {F G H K : C β₯€ C} {Ξ± : F ⟢ G} {Ξ² : H ⟢ K} (X : C) : (Ξ± βŠ— Ξ²).app X = Ξ².app (F.obj X) ≫ K.map (Ξ±.app X) := rfl +@[simp] theorem endofunctorMonoidalCategory_whiskerLeft_app + {F G H : C β₯€ C} {Ξ± : F ⟢ G} (X : C) : + (H ◁ Ξ±).app X = Ξ±.app (H.obj X) := rfl + +@[simp] theorem endofunctorMonoidalCategory_whiskerRight_app + {F G H : C β₯€ C} {Ξ± : F ⟢ G} (X : C) : + (Ξ± β–· H).app X = H.map (Ξ±.app X) := rfl + @[simp] theorem endofunctorMonoidalCategory_associator_hom_app (F G H : C β₯€ C) (X : C) : (Ξ±_ F G H).hom.app X = πŸ™ _ := rfl @@ -87,22 +97,10 @@ def tensoringRightMonoidal [MonoidalCategory.{v} C] : MonoidalFunctor C (C β₯€ C { tensoringRight C with Ξ΅ := (rightUnitorNatIso C).inv ΞΌ := fun X Y => { app := fun Z => (Ξ±_ Z X Y).hom } - ΞΌ_natural := fun f g => by - ext Z - dsimp - simp only [← id_tensor_comp_tensor_id g f, id_tensor_comp, ← tensor_id, Category.assoc, - associator_naturality, associator_naturality_assoc] - associativity := fun X Y Z => by - ext W - simp [pentagon] ΞΌ_isIso := fun X Y => -- We could avoid needing to do this explicitly by -- constructing a partially applied analogue of `associatorNatIso`. - ⟨⟨{ app := fun Z => (Ξ±_ Z X Y).inv - naturality := fun Z Z' f => by - dsimp - rw [← associator_inv_naturality] - simp }, + ⟨⟨{ app := fun Z => (Ξ±_ Z X Y).inv }, by aesop_cat⟩⟩ } #align category_theory.tensoring_right_monoidal CategoryTheory.tensoringRightMonoidal @@ -168,23 +166,23 @@ theorem ΞΌ_naturalityβ‚‚ {m n m' n' : M} (f : m ⟢ m') (g : n ⟢ n') (X : C) : @[reassoc (attr := simp)] theorem ΞΌ_naturalityβ‚— {m n m' : M} (f : m ⟢ m') (X : C) : (F.obj n).map ((F.map f).app X) ≫ (F.ΞΌ m' n).app X = - (F.ΞΌ m n).app X ≫ (F.map (f βŠ— πŸ™ n)).app X := by - rw [← ΞΌ_naturalityβ‚‚ F f (πŸ™ n) X] + (F.ΞΌ m n).app X ≫ (F.map (f β–· n)).app X := by + rw [← tensorHom_id, ← ΞΌ_naturalityβ‚‚ F f (πŸ™ n) X] simp #align category_theory.ΞΌ_naturalityβ‚— CategoryTheory.ΞΌ_naturalityβ‚— @[reassoc (attr := simp)] theorem ΞΌ_naturalityα΅£ {m n n' : M} (g : n ⟢ n') (X : C) : (F.map g).app ((F.obj m).obj X) ≫ (F.ΞΌ m n').app X = - (F.ΞΌ m n).app X ≫ (F.map (πŸ™ m βŠ— g)).app X := by - rw [← ΞΌ_naturalityβ‚‚ F (πŸ™ m) g X] + (F.ΞΌ m n).app X ≫ (F.map (m ◁ g)).app X := by + rw [← id_tensorHom, ← ΞΌ_naturalityβ‚‚ F (πŸ™ m) g X] simp #align category_theory.ΞΌ_naturalityα΅£ CategoryTheory.ΞΌ_naturalityα΅£ @[reassoc (attr := simp)] theorem ΞΌ_inv_naturalityβ‚— {m n m' : M} (f : m ⟢ m') (X : C) : (F.ΞΌIso m n).inv.app X ≫ (F.obj n).map ((F.map f).app X) = - (F.map (f βŠ— πŸ™ n)).app X ≫ (F.ΞΌIso m' n).inv.app X := by + (F.map (f β–· n)).app X ≫ (F.ΞΌIso m' n).inv.app X := by rw [← IsIso.comp_inv_eq, Category.assoc, ← IsIso.eq_inv_comp] simp #align category_theory.ΞΌ_inv_naturalityβ‚— CategoryTheory.ΞΌ_inv_naturalityβ‚— @@ -192,7 +190,7 @@ theorem ΞΌ_inv_naturalityβ‚— {m n m' : M} (f : m ⟢ m') (X : C) : @[reassoc (attr := simp)] theorem ΞΌ_inv_naturalityα΅£ {m n n' : M} (g : n ⟢ n') (X : C) : (F.ΞΌIso m n).inv.app X ≫ (F.map g).app ((F.obj m).obj X) = - (F.map (πŸ™ m βŠ— g)).app X ≫ (F.ΞΌIso m n').inv.app X := by + (F.map (m ◁ g)).app X ≫ (F.ΞΌIso m n').inv.app X := by rw [← IsIso.comp_inv_eq, Category.assoc, ← IsIso.eq_inv_comp] simp #align category_theory.ΞΌ_inv_naturalityα΅£ CategoryTheory.ΞΌ_inv_naturalityα΅£ @@ -317,7 +315,7 @@ noncomputable def unitOfTensorIsoUnit (m n : M) (h : m βŠ— n β‰… πŸ™_ M) : F.ob then `F.obj m` and `F.obj n` forms a self-equivalence of `C`. -/ @[simps] noncomputable def equivOfTensorIsoUnit (m n : M) (h₁ : m βŠ— n β‰… πŸ™_ M) (hβ‚‚ : n βŠ— m β‰… πŸ™_ M) - (H : (h₁.hom βŠ— πŸ™ m) ≫ (Ξ»_ m).hom = (Ξ±_ m n m).hom ≫ (πŸ™ m βŠ— hβ‚‚.hom) ≫ (ρ_ m).hom) : C β‰Œ C + (H : (h₁.hom β–· m) ≫ (Ξ»_ m).hom = (Ξ±_ m n m).hom ≫ (m ◁ hβ‚‚.hom) ≫ (ρ_ m).hom) : C β‰Œ C where functor := F.obj m inverse := F.obj n diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean index 2083f7832bfe5..6df3588710fc9 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean @@ -71,7 +71,11 @@ inductive Hom : F C β†’ F C β†’ Type u | ρ_hom (X : F C) : Hom (X.tensor Unit) X | ρ_inv (X : F C) : Hom X (X.tensor Unit) | comp {X Y Z} (f : Hom X Y) (g : Hom Y Z) : Hom X Z - | tensor {W X Y Z} (f : Hom W Y) (g : Hom X Z) : Hom (W.tensor X) (Y.tensor Z) + -- | tensor {W X Y Z} (f : Hom W Y) (g : Hom X Z) : Hom (W.tensor X) (Y.tensor Z) + | whiskerLeft (X : F C) {Y Z : F C} (f : Hom Y Z) : + Hom (X.tensor Y) (X.tensor Z) + -- `f` cannot be earlier than `Z` since it is a recursive argument. + | whiskerRight {X Y : F C} (f : Hom X Y) (Z : F C) : Hom (X.tensor Z) (Y.tensor Z) #align category_theory.free_monoidal_category.hom CategoryTheory.FreeMonoidalCategory.Hom local infixr:10 " ⟢ᡐ " => Hom @@ -84,37 +88,64 @@ inductive HomEquiv : βˆ€ {X Y : F C}, (X ⟢ᡐ Y) β†’ (X ⟢ᡐ Y) β†’ Prop | trans {X Y} {f g h : X ⟢ᡐ Y} : HomEquiv f g β†’ HomEquiv g h β†’ HomEquiv f h | comp {X Y Z} {f f' : X ⟢ᡐ Y} {g g' : Y ⟢ᡐ Z} : HomEquiv f f' β†’ HomEquiv g g' β†’ HomEquiv (f.comp g) (f'.comp g') - | tensor {W X Y Z} {f f' : W ⟢ᡐ X} {g g' : Y ⟢ᡐ Z} : - HomEquiv f f' β†’ HomEquiv g g' β†’ HomEquiv (f.tensor g) (f'.tensor g') + -- | tensor {W X Y Z} {f f' : W ⟢ᡐ X} {g g' : Y ⟢ᡐ Z} : + -- HomEquiv f f' β†’ HomEquiv g g' β†’ HomEquiv (f.tensor g) (f'.tensor g') + | whisker_left (X) {Y Z} (f f' : Y ⟢ᡐ Z) : + HomEquiv f f' β†’ HomEquiv (f.whiskerLeft X) (f'.whiskerLeft X) + | whisker_right {Y Z} (f f' : Y ⟢ᡐ Z) (X) : + HomEquiv f f' β†’ HomEquiv (f.whiskerRight X) (f'.whiskerRight X) | comp_id {X Y} (f : X ⟢ᡐ Y) : HomEquiv (f.comp (Hom.id _)) f | id_comp {X Y} (f : X ⟢ᡐ Y) : HomEquiv ((Hom.id _).comp f) f | assoc {X Y U V : F C} (f : X ⟢ᡐ U) (g : U ⟢ᡐ V) (h : V ⟢ᡐ Y) : HomEquiv ((f.comp g).comp h) (f.comp (g.comp h)) - | tensor_id {X Y} : HomEquiv ((Hom.id X).tensor (Hom.id Y)) (Hom.id _) - | tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : F C} (f₁ : X₁ ⟢ᡐ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ᡐ Yβ‚‚) (g₁ : Y₁ ⟢ᡐ Z₁) - (gβ‚‚ : Yβ‚‚ ⟢ᡐ Zβ‚‚) : - HomEquiv ((f₁.comp g₁).tensor (fβ‚‚.comp gβ‚‚)) ((f₁.tensor fβ‚‚).comp (g₁.tensor gβ‚‚)) + | whisker_left_id (X Y) : HomEquiv ((Hom.id Y).whiskerLeft X) (Hom.id (X.tensor Y)) + | whisker_left_comp (W) {X Y Z} (f : X ⟢ᡐ Y) (g : Y ⟢ᡐ Z) : + HomEquiv ((f.comp g).whiskerLeft W) ((f.whiskerLeft W).comp (g.whiskerLeft W)) + | id_whisker_left {X Y} (f : X ⟢ᡐ Y) : + HomEquiv (f.whiskerLeft Unit) ((Hom.l_hom X).comp <| f.comp (Hom.l_inv Y)) + | tensor_whisker_left (X Y) {Z Z'} (f : Z ⟢ᡐ Z') : + HomEquiv (f.whiskerLeft (X.tensor Y)) + ((Hom.Ξ±_hom X Y Z).comp <| ((f.whiskerLeft Y).whiskerLeft X).comp <| Hom.Ξ±_inv X Y Z') + + | id_whisker_right (X Y) : HomEquiv ((Hom.id X).whiskerRight Y) (Hom.id (X.tensor Y)) + | comp_whisker_right {X Y Z} (W) (f : X ⟢ᡐ Y) (g : Y ⟢ᡐ Z) : + HomEquiv ((f.comp g).whiskerRight W) ((f.whiskerRight W).comp <| g.whiskerRight W) + | whisker_right_id {X Y} (f : X ⟢ᡐ Y) : + HomEquiv (f.whiskerRight Unit) ((Hom.ρ_hom X).comp <| f.comp <| Hom.ρ_inv Y) + | whisker_right_tensor {X X'} (f : X ⟢ᡐ X') (Y Z) : + HomEquiv (f.whiskerRight <| Y.tensor Z) + ((Hom.Ξ±_inv X Y Z).comp <| ((f.whiskerRight Y).whiskerRight Z).comp <| Hom.Ξ±_hom X' Y Z) + | whisker_assoc (X) {Y Y'} (f : Y ⟢ᡐ Y') (Z) : + HomEquiv ((f.whiskerLeft X).whiskerRight Z) + ((Hom.Ξ±_hom X Y Z).comp <| ((f.whiskerRight Z).whiskerLeft X).comp <| Hom.Ξ±_inv X Y' Z) + | whisker_exchange {W X Y Z} (f : W ⟢ᡐ X) (g : Y ⟢ᡐ Z) : + HomEquiv ((g.whiskerLeft W).comp <| f.whiskerRight Z) ((f.whiskerRight Y).comp <| g.whiskerLeft X) + + -- | tensor_id {X Y} : HomEquiv ((Hom.id X).tensor (Hom.id Y)) (Hom.id _) + -- | tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : F C} (f₁ : X₁ ⟢ᡐ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ᡐ Yβ‚‚) (g₁ : Y₁ ⟢ᡐ Z₁) + -- (gβ‚‚ : Yβ‚‚ ⟢ᡐ Zβ‚‚) : + -- HomEquiv ((f₁.comp g₁).tensor (fβ‚‚.comp gβ‚‚)) ((f₁.tensor fβ‚‚).comp (g₁.tensor gβ‚‚)) | Ξ±_hom_inv {X Y Z} : HomEquiv ((Hom.Ξ±_hom X Y Z).comp (Hom.Ξ±_inv X Y Z)) (Hom.id _) | Ξ±_inv_hom {X Y Z} : HomEquiv ((Hom.Ξ±_inv X Y Z).comp (Hom.Ξ±_hom X Y Z)) (Hom.id _) - | associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃} (f₁ : X₁ ⟢ᡐ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ᡐ Yβ‚‚) (f₃ : X₃ ⟢ᡐ Y₃) : - HomEquiv (((f₁.tensor fβ‚‚).tensor f₃).comp (Hom.Ξ±_hom Y₁ Yβ‚‚ Y₃)) - ((Hom.Ξ±_hom X₁ Xβ‚‚ X₃).comp (f₁.tensor (fβ‚‚.tensor f₃))) + -- | associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃} (f₁ : X₁ ⟢ᡐ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ᡐ Yβ‚‚) (f₃ : X₃ ⟢ᡐ Y₃) : + -- HomEquiv (((f₁.tensor fβ‚‚).tensor f₃).comp (Hom.Ξ±_hom Y₁ Yβ‚‚ Y₃)) + -- ((Hom.Ξ±_hom X₁ Xβ‚‚ X₃).comp (f₁.tensor (fβ‚‚.tensor f₃))) | ρ_hom_inv {X} : HomEquiv ((Hom.ρ_hom X).comp (Hom.ρ_inv X)) (Hom.id _) | ρ_inv_hom {X} : HomEquiv ((Hom.ρ_inv X).comp (Hom.ρ_hom X)) (Hom.id _) - | ρ_naturality {X Y} (f : X ⟢ᡐ Y) : - HomEquiv ((f.tensor (Hom.id Unit)).comp (Hom.ρ_hom Y)) ((Hom.ρ_hom X).comp f) + -- | ρ_naturality {X Y} (f : X ⟢ᡐ Y) : + -- HomEquiv ((f.tensor (Hom.id Unit)).comp (Hom.ρ_hom Y)) ((Hom.ρ_hom X).comp f) | l_hom_inv {X} : HomEquiv ((Hom.l_hom X).comp (Hom.l_inv X)) (Hom.id _) | l_inv_hom {X} : HomEquiv ((Hom.l_inv X).comp (Hom.l_hom X)) (Hom.id _) - | l_naturality {X Y} (f : X ⟢ᡐ Y) : - HomEquiv (((Hom.id Unit).tensor f).comp (Hom.l_hom Y)) ((Hom.l_hom X).comp f) + -- | l_naturality {X Y} (f : X ⟢ᡐ Y) : + -- HomEquiv (((Hom.id Unit).tensor f).comp (Hom.l_hom Y)) ((Hom.l_hom X).comp f) | pentagon {W X Y Z} : HomEquiv - (((Hom.Ξ±_hom W X Y).tensor (Hom.id Z)).comp - ((Hom.Ξ±_hom W (X.tensor Y) Z).comp ((Hom.id W).tensor (Hom.Ξ±_hom X Y Z)))) + (((Hom.Ξ±_hom W X Y).whiskerRight Z).comp + ((Hom.Ξ±_hom W (X.tensor Y) Z).comp ((Hom.Ξ±_hom X Y Z).whiskerLeft W))) ((Hom.Ξ±_hom (W.tensor X) Y Z).comp (Hom.Ξ±_hom W X (Y.tensor Z))) | triangle {X Y} : - HomEquiv ((Hom.Ξ±_hom X Unit Y).comp ((Hom.id X).tensor (Hom.l_hom Y))) - ((Hom.ρ_hom X).tensor (Hom.id Y)) + HomEquiv ((Hom.Ξ±_hom X Unit Y).comp ((Hom.l_hom Y).whiskerLeft X)) + ((Hom.ρ_hom X).whiskerRight Y) set_option linter.uppercaseLean3 false #align category_theory.free_monoidal_category.HomEquiv CategoryTheory.FreeMonoidalCategory.HomEquiv @@ -155,29 +186,58 @@ instance categoryFreeMonoidalCategory : Category.{u} (F C) where instance : MonoidalCategory (F C) where tensorObj X Y := FreeMonoidalCategory.tensor X Y - tensorHom := @fun X₁ Y₁ Xβ‚‚ Yβ‚‚ => - Quotient.mapβ‚‚ Hom.tensor <| by - intro _ _ h _ _ h' - exact HomEquiv.tensor h h' - tensor_id X Y := Quotient.sound tensor_id - tensor_comp := @fun X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ => by - rintro ⟨fβ‚βŸ© ⟨fβ‚‚βŸ© ⟨gβ‚βŸ© ⟨gβ‚‚βŸ© - exact Quotient.sound (tensor_comp _ _ _ _) + -- tensorHom := @fun X₁ Y₁ Xβ‚‚ Yβ‚‚ => + -- Quotient.mapβ‚‚ Hom.tensor <| by + -- intro _ _ h _ _ h' + -- exact HomEquiv.tensor h h' + whiskerLeft := fun X _ _ f => Quotient.map (Hom.whiskerLeft X) (HomEquiv.whisker_left X) f + whiskerRight := fun f Y => Quotient.map (fun _f => Hom.whiskerRight _f Y) (fun _ _ h => HomEquiv.whisker_right _ _ _ h) f tensorUnit' := FreeMonoidalCategory.Unit associator X Y Z := ⟨⟦Hom.Ξ±_hom X Y Z⟧, ⟦Hom.Ξ±_inv X Y Z⟧, Quotient.sound Ξ±_hom_inv, Quotient.sound Ξ±_inv_hom⟩ - associator_naturality := @fun X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ => by - rintro ⟨fβ‚βŸ© ⟨fβ‚‚βŸ© ⟨fβ‚ƒβŸ© - exact Quotient.sound (associator_naturality _ _ _) + -- associator_naturality := @fun X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ => by + -- rintro ⟨fβ‚βŸ© ⟨fβ‚‚βŸ© ⟨fβ‚ƒβŸ© + -- exact Quotient.sound (associator_naturality _ _ _) leftUnitor X := ⟨⟦Hom.l_hom X⟧, ⟦Hom.l_inv X⟧, Quotient.sound l_hom_inv, Quotient.sound l_inv_hom⟩ - leftUnitor_naturality := @fun X Y => by - rintro ⟨f⟩ - exact Quotient.sound (l_naturality _) rightUnitor X := ⟨⟦Hom.ρ_hom X⟧, ⟦Hom.ρ_inv X⟧, Quotient.sound ρ_hom_inv, Quotient.sound ρ_inv_hom⟩ - rightUnitor_naturality := @fun X Y => by - rintro ⟨f⟩ - exact Quotient.sound (ρ_naturality _) + whisker_exchange := by + rintro W X Y Z ⟨f⟩ ⟨g⟩ + apply Quotient.sound + apply HomEquiv.whisker_exchange + whiskerLeft_id := fun X Y ↦ Quotient.sound (HomEquiv.whisker_left_id X Y) + whiskerLeft_comp := by + rintro W X Y Z ⟨f⟩ ⟨g⟩ + apply Quotient.sound + apply HomEquiv.whisker_left_comp + id_whiskerLeft := by + rintro X Y ⟨f⟩ + apply Quotient.sound + apply HomEquiv.id_whisker_left + tensor_whiskerLeft := by + rintro X Y Z Z' ⟨f⟩ + apply Quotient.sound + apply HomEquiv.tensor_whisker_left + id_whiskerRight := by + intro X Y + apply Quotient.sound + apply HomEquiv.id_whisker_right + comp_whiskerRight := by + rintro W X Y ⟨f⟩ ⟨g⟩ Z + apply Quotient.sound + apply HomEquiv.comp_whisker_right + whiskerRight_id := by + rintro X Y ⟨f⟩ + apply Quotient.sound + apply HomEquiv.whisker_right_id + whiskerRight_tensor := by + rintro X X' ⟨f⟩ Y Z + apply Quotient.sound + apply HomEquiv.whisker_right_tensor + whisker_assoc := by + rintro X Y Y' ⟨f⟩ Z + apply Quotient.sound + apply HomEquiv.whisker_assoc pentagon W X Y Z := Quotient.sound pentagon triangle X Y := Quotient.sound triangle @@ -187,11 +247,21 @@ theorem mk_comp {X Y Z : F C} (f : X ⟢ᡐ Y) (g : Y ⟢ᡐ Z) : rfl #align category_theory.free_monoidal_category.mk_comp CategoryTheory.FreeMonoidalCategory.mk_comp +-- @[simp] +-- theorem mk_tensor {X₁ Y₁ Xβ‚‚ Yβ‚‚ : F C} (f : X₁ ⟢ᡐ Y₁) (g : Xβ‚‚ ⟢ᡐ Yβ‚‚) : +-- ⟦f.tensor g⟧ = @MonoidalCategory.tensorHom (F C) _ _ _ _ _ _ ⟦f⟧ ⟦g⟧ := +-- rfl +-- #align category_theory.free_monoidal_category.mk_tensor CategoryTheory.FreeMonoidalCategory.mk_tensor + +@[simp] +theorem mk_whiskerLeft (X : F C) {Y₁ Yβ‚‚ : F C} (f : Y₁ ⟢ᡐ Yβ‚‚) : + ⟦f.whiskerLeft X⟧ = MonoidalCategory.whiskerLeft (C := F C) (X := X) (f := ⟦f⟧) := + rfl + @[simp] -theorem mk_tensor {X₁ Y₁ Xβ‚‚ Yβ‚‚ : F C} (f : X₁ ⟢ᡐ Y₁) (g : Xβ‚‚ ⟢ᡐ Yβ‚‚) : - ⟦f.tensor g⟧ = @MonoidalCategory.tensorHom (F C) _ _ _ _ _ _ ⟦f⟧ ⟦g⟧ := +theorem mk_whiskerRight {X₁ Xβ‚‚ : F C} (f : X₁ ⟢ᡐ Xβ‚‚) (Y : F C) : + ⟦f.whiskerRight Y⟧ = MonoidalCategory.whiskerRight (C := F C) (f := ⟦f⟧) (Y := Y) := rfl -#align category_theory.free_monoidal_category.mk_tensor CategoryTheory.FreeMonoidalCategory.mk_tensor @[simp] theorem mk_id {X : F C} : ⟦Hom.id X⟧ = πŸ™ X := @@ -266,7 +336,9 @@ def projectMapAux : βˆ€ {X Y : F C}, (X ⟢ᡐ Y) β†’ (projectObj f X ⟢ projec | _, _, ρ_hom _ => (ρ_ _).hom | _, _, ρ_inv _ => (ρ_ _).inv | _, _, Hom.comp f g => projectMapAux f ≫ projectMapAux g - | _, _, Hom.tensor f g => projectMapAux f βŠ— projectMapAux g + | _, _, Hom.whiskerLeft X p => projectObj f X ◁ projectMapAux p + | _, _, Hom.whiskerRight p X => projectMapAux p β–· projectObj f X + -- | _, _, Hom.tensor f g => projectMapAux f βŠ— projectMapAux g #align category_theory.free_monoidal_category.project_map_aux CategoryTheory.FreeMonoidalCategory.projectMapAux -- Porting note: this declaration generates the same panic. @@ -279,31 +351,37 @@ def projectMap (X Y : F C) : (X ⟢ Y) β†’ (projectObj f X ⟢ projectObj f Y) : | symm _ _ _ hfg' => exact hfg'.symm | trans _ _ hfg hgh => exact hfg.trans hgh | comp _ _ hf hg => dsimp only [projectMapAux]; rw [hf, hg] - | tensor _ _ hfg hfg' => dsimp only [projectMapAux]; rw [hfg, hfg'] + -- | tensor _ _ hfg hfg' => dsimp only [projectMapAux]; rw [hfg, hfg'] + | whisker_left _ _ _ _ h => dsimp only [projectMapAux]; rw [h] + | whisker_right _ _ _ _ h => dsimp only [projectMapAux]; rw [h] | comp_id => dsimp only [projectMapAux]; rw [Category.comp_id] | id_comp => dsimp only [projectMapAux]; rw [Category.id_comp] | assoc => dsimp only [projectMapAux]; rw [Category.assoc] - | tensor_id => dsimp only [projectMapAux]; rw [MonoidalCategory.tensor_id]; rfl - | tensor_comp => dsimp only [projectMapAux]; rw [MonoidalCategory.tensor_comp] + -- | tensor_id => dsimp only [projectMapAux]; rw [MonoidalCategory.tensor_id]; rfl + -- | tensor_comp => dsimp only [projectMapAux]; rw [MonoidalCategory.tensor_comp] | Ξ±_hom_inv => dsimp only [projectMapAux]; rw [Iso.hom_inv_id] | Ξ±_inv_hom => dsimp only [projectMapAux]; rw [Iso.inv_hom_id] - | associator_naturality => - dsimp only [projectMapAux]; rw [MonoidalCategory.associator_naturality] + -- | associator_naturality => + -- dsimp only [projectMapAux]; rw [MonoidalCategory.associator_naturality] | ρ_hom_inv => dsimp only [projectMapAux]; rw [Iso.hom_inv_id] | ρ_inv_hom => dsimp only [projectMapAux]; rw [Iso.inv_hom_id] - | ρ_naturality => - dsimp only [projectMapAux, projectObj]; rw [MonoidalCategory.rightUnitor_naturality] + -- | ρ_naturality => + -- dsimp only [projectMapAux, projectObj]; rw [MonoidalCategory.rightUnitor_naturality] | l_hom_inv => dsimp only [projectMapAux]; rw [Iso.hom_inv_id] | l_inv_hom => dsimp only [projectMapAux]; rw [Iso.inv_hom_id] - | l_naturality => - dsimp only [projectMapAux, projectObj] - exact MonoidalCategory.leftUnitor_naturality _ + -- | l_naturality => + -- dsimp only [projectMapAux, projectObj] + -- exact MonoidalCategory.leftUnitor_naturality _ | pentagon => dsimp only [projectMapAux] exact MonoidalCategory.pentagon _ _ _ _ | triangle => dsimp only [projectMapAux] exact MonoidalCategory.triangle _ _ + | whisker_exchange => + dsimp only [projectMapAux, projectObj]; simp [MonoidalCategory.whisker_exchange] + | _ => + dsimp only [projectMapAux, projectObj]; simp #align category_theory.free_monoidal_category.project_map CategoryTheory.FreeMonoidalCategory.projectMap end @@ -319,13 +397,17 @@ def project : MonoidalFunctor (F C) D where map_comp := by rintro _ _ _ ⟨_⟩ ⟨_⟩; rfl Ξ΅ := πŸ™ _ ΞΌ X Y := πŸ™ _ - ΞΌ_natural := @fun _ _ _ _ f g => by + ΞΌ_natural_left := fun f _ => by + induction' f using Quotient.recOn + Β· dsimp + simp + rfl + Β· rfl + ΞΌ_natural_right := fun _ f => by induction' f using Quotient.recOn - Β· induction' g using Quotient.recOn - Β· dsimp - simp - rfl - Β· rfl + Β· dsimp + simp + rfl Β· rfl #align category_theory.free_monoidal_category.project CategoryTheory.FreeMonoidalCategory.project diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean index afba63f365e0d..bae8738e28fb3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean @@ -41,6 +41,24 @@ universe u namespace CategoryTheory +namespace MonoidalCategory + +variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] + +@[simp] +theorem whiskerLeft_eqToHom (X : C) {Y Z : C} (f : Y = Z) : + X ◁ eqToHom f = eqToHom (congr_argβ‚‚ tensorObj rfl f) := by + cases f + simp only [whiskerLeft_id, eqToHom_refl] + +@[simp] +theorem eqToHom_whiskerRight {X Y : C} (f : X = Y) (Z : C) : + eqToHom f β–· Z = eqToHom (congr_argβ‚‚ tensorObj f rfl) := by + cases f + simp only [id_whiskerRight, eqToHom_refl] + +end MonoidalCategory + open MonoidalCategory namespace FreeMonoidalCategory @@ -78,27 +96,38 @@ def inclusionObj : NormalMonoidalObject C β†’ F C #align category_theory.free_monoidal_category.inclusion_obj CategoryTheory.FreeMonoidalCategory.inclusionObj /-- The discrete subcategory of objects in normal form includes into the free monoidal category. -/ -@[simp] def inclusion : N C β₯€ F C := Discrete.functor inclusionObj #align category_theory.free_monoidal_category.inclusion CategoryTheory.FreeMonoidalCategory.inclusion +@[simp] +theorem inclusion_obj (X : N C) : + inclusion.obj X = inclusionObj X.as := + rfl + +@[simp] +theorem inclusion_map {X Y : N C} (f : X ⟢ Y) : + inclusion.map f = eqToHom (congr_arg _ (Discrete.ext _ _ (Discrete.eq_of_hom f))) := by + rcases f with ⟨⟨⟩⟩ + cases Discrete.ext _ _ (by assumption) + apply inclusion.map_id + /-- Auxiliary definition for `normalize`. -/ @[simp] -def normalizeObj : F C β†’ NormalMonoidalObject C β†’ N C - | Unit, n => ⟨n⟩ - | of X, n => ⟨NormalMonoidalObject.tensor n X⟩ - | tensor X Y, n => normalizeObj Y (normalizeObj X n).as +def normalizeObj : F C β†’ NormalMonoidalObject C β†’ NormalMonoidalObject C + | Unit, n => n + | of X, n => NormalMonoidalObject.tensor n X + | tensor X Y, n => normalizeObj Y (normalizeObj X n) #align category_theory.free_monoidal_category.normalize_obj CategoryTheory.FreeMonoidalCategory.normalizeObj @[simp] -theorem normalizeObj_unitor (n : NormalMonoidalObject C) : normalizeObj (πŸ™_ (F C)) n = ⟨n⟩ := +theorem normalizeObj_unitor (n : NormalMonoidalObject C) : normalizeObj (πŸ™_ (F C)) n = n := rfl #align category_theory.free_monoidal_category.normalize_obj_unitor CategoryTheory.FreeMonoidalCategory.normalizeObj_unitor @[simp] theorem normalizeObj_tensor (X Y : F C) (n : NormalMonoidalObject C) : - normalizeObj (X βŠ— Y) n = normalizeObj Y (normalizeObj X n).as := + normalizeObj (X βŠ— Y) n = normalizeObj Y (normalizeObj X n) := rfl #align category_theory.free_monoidal_category.normalize_obj_tensor CategoryTheory.FreeMonoidalCategory.normalizeObj_tensor @@ -114,8 +143,8 @@ open Hom @[simp] def normalizeMapAux : βˆ€ {X Y : F C}, (X ⟢ᡐ Y) β†’ - ((Discrete.functor (normalizeObj X) : _ β₯€ N C) ⟢ Discrete.functor (normalizeObj Y)) - | _, _, Hom.id _ => πŸ™ _ + ((Discrete.functor (fun n ↦ ⟨normalizeObj X n⟩) : N C β₯€ N C) ⟢ Discrete.functor fun n ↦ ⟨normalizeObj Y n⟩) + | _, _, Hom.id _ => by dsimp; exact πŸ™ _ | _, _, Ξ±_hom X Y Z => by dsimp; exact Discrete.natTrans (fun _ => πŸ™ _) | _, _, Ξ±_inv _ _ _ => by dsimp; exact Discrete.natTrans (fun _ => πŸ™ _) | _, _, l_hom _ => by dsimp; exact Discrete.natTrans (fun _ => πŸ™ _) @@ -123,10 +152,13 @@ def normalizeMapAux : | _, _, ρ_hom _ => by dsimp; exact Discrete.natTrans (fun _ => πŸ™ _) | _, _, ρ_inv _ => by dsimp; exact Discrete.natTrans (fun _ => πŸ™ _) | _, _, (@comp _ _ _ _ f g) => normalizeMapAux f ≫ normalizeMapAux g - | _, _, (@Hom.tensor _ T _ _ W f g) => by + | _, _, (@Hom.whiskerLeft _ T _ W f) => by + dsimp + exact Discrete.natTrans (fun ⟨X⟩ => (normalizeMapAux f).app ⟨normalizeObj T X⟩) + | _, _, (@Hom.whiskerRight _ T _ f W) => by dsimp - exact Discrete.natTrans (fun ⟨X⟩ => (normalizeMapAux g).app (normalizeObj T X) ≫ - (Discrete.functor (normalizeObj W) : _ β₯€ N C).map ((normalizeMapAux f).app ⟨X⟩)) + exact Discrete.natTrans <| fun X => + (Discrete.functor fun n ↦ ⟨normalizeObj W n⟩ : N C β₯€ N C).map <| (normalizeMapAux f).app X #align category_theory.free_monoidal_category.normalize_map_aux CategoryTheory.FreeMonoidalCategory.normalizeMapAux end @@ -140,7 +172,7 @@ variable (C) `πŸ™_ C`. -/ @[simp] def normalize : F C β₯€ N C β₯€ N C where - obj X := Discrete.functor (normalizeObj X) + obj X := Discrete.functor (fun n ↦ ⟨normalizeObj X n⟩) map {X Y} := Quotient.lift normalizeMapAux (by aesop_cat) #align category_theory.free_monoidal_category.normalize CategoryTheory.FreeMonoidalCategory.normalize @@ -163,22 +195,21 @@ def fullNormalize : F C β₯€ N C where @[simp] def tensorFunc : F C β₯€ N C β₯€ F C where obj X := Discrete.functor fun n => inclusion.obj ⟨n⟩ βŠ— X - map f := Discrete.natTrans (fun n => πŸ™ _ βŠ— f) + map f := Discrete.natTrans (fun n => _ ◁ f) #align category_theory.free_monoidal_category.tensor_func CategoryTheory.FreeMonoidalCategory.tensorFunc -theorem tensorFunc_map_app {X Y : F C} (f : X ⟢ Y) (n) : ((tensorFunc C).map f).app n = πŸ™ _ βŠ— f := +theorem tensorFunc_map_app {X Y : F C} (f : X ⟢ Y) (n) : ((tensorFunc C).map f).app n = _ ◁ f := rfl #align category_theory.free_monoidal_category.tensor_func_map_app CategoryTheory.FreeMonoidalCategory.tensorFunc_map_app theorem tensorFunc_obj_map (Z : F C) {n n' : N C} (f : n ⟢ n') : - ((tensorFunc C).obj Z).map f = inclusion.map f βŠ— πŸ™ Z := by + ((tensorFunc C).obj Z).map f = inclusion.map f β–· Z := by cases n cases n' rcases f with ⟨⟨h⟩⟩ dsimp at h subst h simp - #align category_theory.free_monoidal_category.tensor_func_obj_map CategoryTheory.FreeMonoidalCategory.tensorFunc_obj_map /-- Auxiliary definition for `normalizeIso`. Here we construct the isomorphism between @@ -192,6 +223,25 @@ def normalizeIsoApp : (Ξ±_ _ _ _).symm β‰ͺ≫ tensorIso (normalizeIsoApp X n) (Iso.refl _) β‰ͺ≫ normalizeIsoApp _ _ #align category_theory.free_monoidal_category.normalize_iso_app CategoryTheory.FreeMonoidalCategory.normalizeIsoApp +/-- Almost non-definitionally equall to `normalizeIsoApp`, but has a better definitional property +in the proof of `normalize_naturality`. -/ +@[simp] +def normalizeIsoApp' : + βˆ€ (X : F C) (n : NormalMonoidalObject C), inclusionObj n βŠ— X β‰… inclusionObj (normalizeObj X n) + | of _, _ => Iso.refl _ + | Unit, _ => ρ_ _ + | tensor X _, n => + (Ξ±_ _ _ _).symm β‰ͺ≫ tensorIso (normalizeIsoApp' X n) (Iso.refl _) β‰ͺ≫ normalizeIsoApp' _ _ + +theorem normalizeIsoApp_eq : βˆ€ (X : F C) (n : N C), normalizeIsoApp C X n = normalizeIsoApp' C X n.as + | of X, _ => rfl + | Unit, _ => rfl + | tensor X Y, n => by + rw [normalizeIsoApp, normalizeIsoApp'] + rw [normalizeIsoApp_eq X n] + rw [normalizeIsoApp_eq Y ⟨normalizeObj X n.as⟩] + rfl + @[simp] theorem normalizeIsoApp_tensor (X Y : F C) (n : N C) : normalizeIsoApp C (X βŠ— Y) n = @@ -233,82 +283,78 @@ theorem discrete_functor_map_eq_id (g : X ⟢ X) : (Discrete.functor f).map g = end +section + +variable {C} + +theorem normalizeObj_congr (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ Y) : + normalizeObj X n = normalizeObj Y n := by + rcases f with ⟨f'⟩ + apply @congr_fun _ _ fun n => normalizeObj X n + clear n f + induction f' with + | comp _ _ _ _ => apply Eq.trans <;> assumption + | whiskerLeft _ _ ih => funext; apply congr_fun ih + | whiskerRight _ _ ih => + funext + apply congr_argβ‚‚ _ rfl (congr_fun ih _) + | _ => funext; rfl + +attribute [local simp] id_tensorHom tensorHom_id + +theorem normalize_naturality (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ Y) : + inclusionObj n ◁ f ≫ (normalizeIsoApp' C Y n).hom = + (normalizeIsoApp' C X n).hom ≫ + inclusion.map (eqToHom (Discrete.ext _ _ + ((normalizeObj_congr n f)))) := by + induction' f using Quotient.recOn with f'; swap; rfl + revert n + induction f' with + | id => erw [mk_id]; simp + | Ξ±_hom X Y Z => erw [mk_Ξ±_hom]; simp + | Ξ±_inv X Y Z => erw [mk_Ξ±_inv]; simp + | l_hom X => erw [mk_l_hom]; simp + | l_inv X => erw [mk_l_inv]; simp + | ρ_hom X => erw [mk_ρ_hom]; simp + | ρ_inv X => erw [mk_ρ_inv]; simp + | comp f g ihf ihg => + intro n + erw [mk_comp] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 2 3 => rw [ihg] + slice_lhs 1 2 => rw [ihf] + simp + | whiskerLeft X f ih => + intro n + erw [mk_whiskerLeft] + dsimp only [tensor_eq_tensor, normalizeObj_tensor, normalizeIsoApp', Iso.trans_hom, Iso.symm_hom, tensorIso_hom, + Iso.refl_hom, Function.comp_apply, inclusion_obj] + simp only [MonoidalCategory.whiskerLeft_id, Category.assoc, Category.id_comp, tensorHom_id] + rw [associator_inv_naturality_right_assoc, whisker_exchange_assoc, ih] + | @whiskerRight X Y h Ξ·' ih => + intro n + erw [mk_whiskerRight] + dsimp + simp only [MonoidalCategory.whiskerLeft_id, Category.assoc, Category.id_comp, tensorHom_id] + rw [associator_inv_naturality_middle_assoc, ← comp_whiskerRight_assoc] + erw [ih] + dsimp + simp + have := dcongr_arg (fun x => (normalizeIsoApp' C Ξ·' x).hom) + (normalizeObj_congr n (Quotient.mk (setoidHom X Y) h)) + dsimp at this; simp [this] + +end + /-- The isomorphism between `n βŠ— X` and `normalize X n` is natural (in both `X` and `n`, but naturality in `n` is trivial and was "proved" in `normalizeIsoAux`). This is the real heart of our proof of the coherence theorem. -/ def normalizeIso : tensorFunc C β‰… normalize' C := - NatIso.ofComponents (normalizeIsoAux C) - (by -- porting note: the proof has been mostly rewritten - rintro X Y f - induction' f using Quotient.recOn with f; swap; rfl - induction' f with _ X₁ Xβ‚‚ X₃ _ _ _ _ _ _ _ _ _ _ _ _ h₁ hβ‚‚ X₁ Xβ‚‚ Y₁ Yβ‚‚ f g h₁ hβ‚‚ - Β· simp only [mk_id, Functor.map_id, Category.comp_id, Category.id_comp] - Β· ext n - dsimp - rw [mk_Ξ±_hom, NatTrans.comp_app, NatTrans.comp_app] - dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] - simp only [comp_tensor_id, associator_conjugation, tensor_id, - Category.comp_id] - simp only [← Category.assoc] - congr 4 - simp only [Category.assoc, ← cancel_epi (πŸ™ (inclusionObj n.as) βŠ— (Ξ±_ X₁ Xβ‚‚ X₃).inv), - pentagon_inv_assoc (inclusionObj n.as) X₁ Xβ‚‚ X₃, - tensor_inv_hom_id_assoc, tensor_id, Category.id_comp, Iso.inv_hom_id, - Category.comp_id] - Β· ext n - dsimp - rw [mk_Ξ±_inv, NatTrans.comp_app, NatTrans.comp_app] - dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] - simp only [Category.assoc, comp_tensor_id, tensor_id, Category.comp_id, - pentagon_inv_assoc, ← associator_inv_naturality_assoc] - rfl - Β· ext n - dsimp [Functor.comp] - rw [mk_l_hom, NatTrans.comp_app, NatTrans.comp_app] - dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] - simp only [triangle_assoc_comp_right_assoc, Category.assoc, Category.comp_id] - rfl - Β· ext n - dsimp [Functor.comp] - rw [mk_l_inv, NatTrans.comp_app, NatTrans.comp_app] - dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] - simp only [triangle_assoc_comp_left_inv_assoc, inv_hom_id_tensor_assoc, tensor_id, - Category.id_comp, Category.comp_id] - rfl - Β· ext n - dsimp - rw [mk_ρ_hom, NatTrans.comp_app, NatTrans.comp_app] - dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] - simp only [← (Iso.inv_comp_eq _).2 (rightUnitor_tensor _ _), Category.assoc, - ← rightUnitor_naturality, Category.comp_id]; rfl - Β· ext n - dsimp - rw [mk_ρ_inv, NatTrans.comp_app, NatTrans.comp_app] - dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] - simp only [← (Iso.eq_comp_inv _).1 (rightUnitor_tensor_inv _ _), rightUnitor_conjugation, - Category.assoc, Iso.hom_inv_id_assoc, Iso.inv_hom_id_assoc, Iso.inv_hom_id, - Discrete.functor, Category.comp_id, Function.comp] - Β· rw [mk_comp, Functor.map_comp, Functor.map_comp, Category.assoc, hβ‚‚, reassoc_of% h₁] - Β· ext ⟨n⟩ - replace h₁ := NatTrans.congr_app h₁ ⟨n⟩ - replace hβ‚‚ := NatTrans.congr_app hβ‚‚ ((Discrete.functor (normalizeObj X₁)).obj ⟨n⟩) - have h₃ := (normalizeIsoAux _ Yβ‚‚).hom.naturality ((normalizeMapAux f).app ⟨n⟩) - have hβ‚„ : βˆ€ (X₃ Y₃ : N C) (Ο† : X₃ ⟢ Y₃), (Discrete.functor inclusionObj).map Ο† βŠ— πŸ™ Yβ‚‚ = - (Discrete.functor fun n ↦ inclusionObj n βŠ— Yβ‚‚).map Ο† := by - rintro ⟨Xβ‚ƒβŸ© ⟨Yβ‚ƒβŸ© Ο† - obtain rfl : X₃ = Y₃ := Ο†.1.1 - simp only [discrete_functor_map_eq_id, tensor_id] - rfl - rw [NatTrans.comp_app, NatTrans.comp_app] at h₁ hβ‚‚ ⊒ - dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, - Functor.comp, Discrete.natTrans] at h₁ hβ‚‚ h₃ ⊒ - rw [mk_tensor, associator_inv_naturality_assoc, ← tensor_comp_assoc, h₁, - Category.assoc, Category.comp_id, ← @Category.id_comp (F C) _ _ _ (@Quotient.mk _ _ g), - tensor_comp, Category.assoc, Category.assoc, Functor.map_comp] - congr 2 - erw [← reassoc_of% hβ‚‚] - rw [← h₃, ← Category.assoc, ← id_tensor_comp_tensor_id, hβ‚„] - rfl ) + NatIso.ofComponents (normalizeIsoAux C) <| by + intro X Y f + ext ⟨n⟩ + convert normalize_naturality n f using 1 + any_goals dsimp [NatIso.ofComponents]; simp; congr; apply normalizeIsoApp_eq #align category_theory.free_monoidal_category.normalize_iso CategoryTheory.FreeMonoidalCategory.normalizeIso /-- The isomorphism between an object and its normal form is natural. -/ @@ -319,7 +365,7 @@ def fullNormalizeIso : 𝟭 (F C) β‰… fullNormalize C β‹™ inclusion := intro X Y f dsimp rw [leftUnitor_inv_naturality_assoc, Category.assoc, Iso.cancel_iso_inv_left] - exact + convert congr_arg (fun f => NatTrans.app f (Discrete.mk NormalMonoidalObject.Unit)) ((normalizeIso.{u} C).hom.naturality f)) #align category_theory.free_monoidal_category.full_normalize_iso CategoryTheory.FreeMonoidalCategory.fullNormalizeIso @@ -352,7 +398,8 @@ def inverseAux : βˆ€ {X Y : F C}, (X ⟢ᡐ Y) β†’ (Y ⟢ᡐ X) | _, _, l_hom _ => l_inv _ | _, _, l_inv _ => l_hom _ | _, _, comp f g => (inverseAux g).comp (inverseAux f) - | _, _, Hom.tensor f g => (inverseAux f).tensor (inverseAux g) + | _, _, .whiskerLeft X f => (inverseAux f).whiskerLeft X + | _, _, .whiskerRight f X => (inverseAux f).whiskerRight X #align category_theory.free_monoidal_category.inverse_aux CategoryTheory.FreeMonoidalCategory.inverseAux end diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index 169b5f183f272..e13974b784759 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -69,20 +69,24 @@ structure LaxMonoidalFunctor extends C β₯€ D where Ξ΅ : πŸ™_ D ⟢ obj (πŸ™_ C) /-- tensorator -/ ΞΌ : βˆ€ X Y : C, obj X βŠ— obj Y ⟢ obj (X βŠ— Y) - ΞΌ_natural : - βˆ€ {X Y X' Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y'), - (map f βŠ— map g) ≫ ΞΌ Y Y' = ΞΌ X X' ≫ map (f βŠ— g) := by + ΞΌ_natural_left : + βˆ€ {X Y : C} (f : X ⟢ Y) (X' : C), + (map f β–· obj X') ≫ ΞΌ Y X' = ΞΌ X X' ≫ map (f β–· X') := by + aesop_cat + ΞΌ_natural_right : + βˆ€ {X Y : C} (X' : C) (f : X ⟢ Y) , + (obj X' ◁ map f) ≫ ΞΌ X' Y = ΞΌ X' X ≫ map (X' ◁ f) := by aesop_cat /-- associativity of the tensorator -/ associativity : βˆ€ X Y Z : C, - (ΞΌ X Y βŠ— πŸ™ (obj Z)) ≫ ΞΌ (X βŠ— Y) Z ≫ map (Ξ±_ X Y Z).hom = - (Ξ±_ (obj X) (obj Y) (obj Z)).hom ≫ (πŸ™ (obj X) βŠ— ΞΌ Y Z) ≫ ΞΌ X (Y βŠ— Z) := by + (ΞΌ X Y β–· obj Z) ≫ ΞΌ (X βŠ— Y) Z ≫ map (Ξ±_ X Y Z).hom = + (Ξ±_ (obj X) (obj Y) (obj Z)).hom ≫ (obj X ◁ ΞΌ Y Z) ≫ ΞΌ X (Y βŠ— Z) := by aesop_cat -- unitality - left_unitality : βˆ€ X : C, (Ξ»_ (obj X)).hom = (Ξ΅ βŠ— πŸ™ (obj X)) ≫ ΞΌ (πŸ™_ C) X ≫ map (Ξ»_ X).hom := + left_unitality : βˆ€ X : C, (Ξ»_ (obj X)).hom = (Ξ΅ β–· obj X) ≫ ΞΌ (πŸ™_ C) X ≫ map (Ξ»_ X).hom := by aesop_cat - right_unitality : βˆ€ X : C, (ρ_ (obj X)).hom = (πŸ™ (obj X) βŠ— Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ map (ρ_ X).hom := + right_unitality : βˆ€ X : C, (ρ_ (obj X)).hom = (obj X ◁ Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ map (ρ_ X).hom := by aesop_cat #align category_theory.lax_monoidal_functor CategoryTheory.LaxMonoidalFunctor @@ -93,7 +97,9 @@ structure LaxMonoidalFunctor extends C β₯€ D where initialize_simps_projections LaxMonoidalFunctor (+toFunctor, -obj, -map) --Porting note: was `[simp, reassoc.1]` -attribute [reassoc (attr := simp)] LaxMonoidalFunctor.ΞΌ_natural +-- attribute [reassoc (attr := simp)] LaxMonoidalFunctor.ΞΌ_natural +attribute [reassoc (attr := simp)] LaxMonoidalFunctor.ΞΌ_natural_left +attribute [reassoc (attr := simp)] LaxMonoidalFunctor.ΞΌ_natural_right attribute [simp] LaxMonoidalFunctor.left_unitality @@ -109,10 +115,72 @@ section variable {C D} +attribute [local simp] tensorHom_def + +@[reassoc (attr := simp)] +theorem LaxMonoidalFunctor.ΞΌ_natural (F : LaxMonoidalFunctor C D) {X Y X' Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y') : + (F.map f βŠ— F.map g) ≫ F.ΞΌ Y Y' = F.ΞΌ X X' ≫ F.map (f βŠ— g) := by + simp only [tensorHom_def, assoc, F.ΞΌ_natural_left, F.ΞΌ_natural_right_assoc, map_comp] + +@[reassoc] +theorem LaxMonoidalFunctor.associativity' (F : LaxMonoidalFunctor C D) (X Y Z : C) : + (F.ΞΌ X Y βŠ— πŸ™ (F.obj Z)) ≫ F.ΞΌ (X βŠ— Y) Z ≫ F.map (Ξ±_ X Y Z).hom = + (Ξ±_ (F.obj X) (F.obj Y) (F.obj Z)).hom ≫ (πŸ™ (F.obj X) βŠ— F.ΞΌ Y Z) ≫ F.ΞΌ X (Y βŠ— Z) := by + simp + +@[reassoc] +theorem LaxMonoidalFunctor.left_unitality' (F : LaxMonoidalFunctor C D) (X : C) : + (Ξ»_ (F.obj X)).hom = (F.Ξ΅ βŠ— πŸ™ (F.obj X)) ≫ F.ΞΌ (πŸ™_ C) X ≫ F.map (Ξ»_ X).hom := by + simp + +@[reassoc] +theorem LaxMonoidalFunctor.right_unitality' (F : LaxMonoidalFunctor C D) (X : C) : + (ρ_ (F.obj X)).hom = (πŸ™ (F.obj X) βŠ— F.Ξ΅) ≫ F.ΞΌ X (πŸ™_ C) ≫ F.map (ρ_ X).hom := by + simp + +@[simps] +def LaxMonoidalFunctor.ofTensorHom (F : C β₯€ D) + /- unit morphism -/ + (Ξ΅ : πŸ™_ D ⟢ F.obj (πŸ™_ C)) + /- tensorator -/ + (ΞΌ : βˆ€ X Y : C, F.obj X βŠ— F.obj Y ⟢ F.obj (X βŠ— Y)) + (ΞΌ_natural : + βˆ€ {X Y X' Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y'), + (F.map f βŠ— F.map g) ≫ ΞΌ Y Y' = ΞΌ X X' ≫ F.map (f βŠ— g) := by + aesop_cat) + /- associativity of the tensorator -/ + (associativity : + βˆ€ X Y Z : C, + (ΞΌ X Y βŠ— πŸ™ (F.obj Z)) ≫ ΞΌ (X βŠ— Y) Z ≫ F.map (Ξ±_ X Y Z).hom = + (Ξ±_ (F.obj X) (F.obj Y) (F.obj Z)).hom ≫ (πŸ™ (F.obj X) βŠ— ΞΌ Y Z) ≫ ΞΌ X (Y βŠ— Z) := by + aesop_cat) + /- unitality -/ + (left_unitality : βˆ€ X : C, (Ξ»_ (F.obj X)).hom = (Ξ΅ βŠ— πŸ™ (F.obj X)) ≫ ΞΌ (πŸ™_ C) X ≫ F.map (Ξ»_ X).hom := + by aesop_cat) + (right_unitality : βˆ€ X : C, (ρ_ (F.obj X)).hom = (πŸ™ (F.obj X) βŠ— Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ F.map (ρ_ X).hom := + by aesop_cat) : + LaxMonoidalFunctor C D where + obj := F.obj + map := F.map + map_id := F.map_id + map_comp := F.map_comp + Ξ΅ := Ξ΅ + ΞΌ := ΞΌ + ΞΌ_natural_left := fun f X' => by + simp_rw [← tensorHom_id, ← F.map_id, ΞΌ_natural] + ΞΌ_natural_right := fun X' f => by + simp_rw [← id_tensorHom, ← F.map_id, ΞΌ_natural] + associativity := fun X Y Z => by + simp_rw [← tensorHom_id, ← id_tensorHom, associativity] + left_unitality := fun X => by + simp_rw [← tensorHom_id, ← id_tensorHom, left_unitality] + right_unitality := fun X => by + simp_rw [← tensorHom_id, ← id_tensorHom, right_unitality] + --Porting note: was `[simp, reassoc.1]` @[reassoc (attr := simp)] theorem LaxMonoidalFunctor.left_unitality_inv (F : LaxMonoidalFunctor C D) (X : C) : - (Ξ»_ (F.obj X)).inv ≫ (F.Ξ΅ βŠ— πŸ™ (F.obj X)) ≫ F.ΞΌ (πŸ™_ C) X = F.map (Ξ»_ X).inv := by + (Ξ»_ (F.obj X)).inv ≫ (F.Ξ΅ β–· F.obj X) ≫ F.ΞΌ (πŸ™_ C) X = F.map (Ξ»_ X).inv := by rw [Iso.inv_comp_eq, F.left_unitality, Category.assoc, Category.assoc, ← F.toFunctor.map_comp, Iso.hom_inv_id, F.toFunctor.map_id, comp_id] #align category_theory.lax_monoidal_functor.left_unitality_inv CategoryTheory.LaxMonoidalFunctor.left_unitality_inv @@ -120,7 +188,7 @@ theorem LaxMonoidalFunctor.left_unitality_inv (F : LaxMonoidalFunctor C D) (X : --Porting note: was `[simp, reassoc.1]` @[reassoc (attr := simp)] theorem LaxMonoidalFunctor.right_unitality_inv (F : LaxMonoidalFunctor C D) (X : C) : - (ρ_ (F.obj X)).inv ≫ (πŸ™ (F.obj X) βŠ— F.Ξ΅) ≫ F.ΞΌ X (πŸ™_ C) = F.map (ρ_ X).inv := by + (ρ_ (F.obj X)).inv ≫ (F.obj X ◁ F.Ξ΅) ≫ F.ΞΌ X (πŸ™_ C) = F.map (ρ_ X).inv := by rw [Iso.inv_comp_eq, F.right_unitality, Category.assoc, Category.assoc, ← F.toFunctor.map_comp, Iso.hom_inv_id, F.toFunctor.map_id, comp_id] #align category_theory.lax_monoidal_functor.right_unitality_inv CategoryTheory.LaxMonoidalFunctor.right_unitality_inv @@ -128,12 +196,18 @@ theorem LaxMonoidalFunctor.right_unitality_inv (F : LaxMonoidalFunctor C D) (X : --Porting note: was `[simp, reassoc.1]` @[reassoc (attr := simp)] theorem LaxMonoidalFunctor.associativity_inv (F : LaxMonoidalFunctor C D) (X Y Z : C) : - (πŸ™ (F.obj X) βŠ— F.ΞΌ Y Z) ≫ F.ΞΌ X (Y βŠ— Z) ≫ F.map (Ξ±_ X Y Z).inv = - (Ξ±_ (F.obj X) (F.obj Y) (F.obj Z)).inv ≫ (F.ΞΌ X Y βŠ— πŸ™ (F.obj Z)) ≫ F.ΞΌ (X βŠ— Y) Z := by + (F.obj X ◁ F.ΞΌ Y Z) ≫ F.ΞΌ X (Y βŠ— Z) ≫ F.map (Ξ±_ X Y Z).inv = + (Ξ±_ (F.obj X) (F.obj Y) (F.obj Z)).inv ≫ (F.ΞΌ X Y β–· F.obj Z) ≫ F.ΞΌ (X βŠ— Y) Z := by rw [Iso.eq_inv_comp, ← F.associativity_assoc, ← F.toFunctor.map_comp, Iso.hom_inv_id, F.toFunctor.map_id, comp_id] #align category_theory.lax_monoidal_functor.associativity_inv CategoryTheory.LaxMonoidalFunctor.associativity_inv +@[reassoc] +theorem LaxMonoidalFunctor.associativity_inv' (F : LaxMonoidalFunctor C D) (X Y Z : C) : + (πŸ™ (F.obj X) βŠ— F.ΞΌ Y Z) ≫ F.ΞΌ X (Y βŠ— Z) ≫ F.map (Ξ±_ X Y Z).inv = + (Ξ±_ (F.obj X) (F.obj Y) (F.obj Z)).inv ≫ (F.ΞΌ X Y βŠ— πŸ™ (F.obj Z)) ≫ F.ΞΌ (X βŠ— Y) Z := by + simp + end /-- @@ -202,22 +276,38 @@ theorem map_tensor {X Y X' Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y') : F.map (f βŠ— g) = inv (F.ΞΌ X X') ≫ (F.map f βŠ— F.map g) ≫ F.ΞΌ Y Y' := by simp #align category_theory.monoidal_functor.map_tensor CategoryTheory.MonoidalFunctor.map_tensor +theorem map_whiskerLeft (X : C) {Y Z : C} (f : Y ⟢ Z) : + F.map (X ◁ f) = inv (F.ΞΌ X Y) ≫ (F.obj X ◁ F.map f) ≫ F.ΞΌ X Z := by simp + +theorem map_whiskerRight {X Y : C} (f : X ⟢ Y) (Z : C) : + F.map (f β–· Z) = inv (F.ΞΌ X Z) ≫ (F.map f β–· F.obj Z) ≫ F.ΞΌ Y Z := by simp + theorem map_leftUnitor (X : C) : - F.map (Ξ»_ X).hom = inv (F.ΞΌ (πŸ™_ C) X) ≫ (inv F.Ξ΅ βŠ— πŸ™ (F.obj X)) ≫ (Ξ»_ (F.obj X)).hom := by + F.map (Ξ»_ X).hom = inv (F.ΞΌ (πŸ™_ C) X) ≫ (inv F.Ξ΅ β–· F.obj X) ≫ (Ξ»_ (F.obj X)).hom := by simp only [LaxMonoidalFunctor.left_unitality] slice_rhs 2 3 => - rw [← comp_tensor_id] + rw [← comp_whiskerRight] simp simp -#align category_theory.monoidal_functor.map_left_unitor CategoryTheory.MonoidalFunctor.map_leftUnitor + +theorem map_leftUnitor' (X : C) : + F.map (Ξ»_ X).hom = inv (F.ΞΌ (πŸ™_ C) X) ≫ (inv F.Ξ΅ βŠ— πŸ™ (F.obj X)) ≫ (Ξ»_ (F.obj X)).hom := by + rw [tensorHom_id] + apply map_leftUnitor +#align category_theory.monoidal_functor.map_left_unitor CategoryTheory.MonoidalFunctor.map_leftUnitor' theorem map_rightUnitor (X : C) : - F.map (ρ_ X).hom = inv (F.ΞΌ X (πŸ™_ C)) ≫ (πŸ™ (F.obj X) βŠ— inv F.Ξ΅) ≫ (ρ_ (F.obj X)).hom := by + F.map (ρ_ X).hom = inv (F.ΞΌ X (πŸ™_ C)) ≫ (F.obj X ◁ inv F.Ξ΅) ≫ (ρ_ (F.obj X)).hom := by simp only [LaxMonoidalFunctor.right_unitality] slice_rhs 2 3 => - rw [← id_tensor_comp] + rw [← MonoidalCategory.whiskerLeft_comp] simp simp + +theorem map_rightUnitor' (X : C) : + F.map (ρ_ X).hom = inv (F.ΞΌ X (πŸ™_ C)) ≫ (πŸ™ (F.obj X) βŠ— inv F.Ξ΅) ≫ (ρ_ (F.obj X)).hom := by + rw [id_tensorHom] + apply map_rightUnitor #align category_theory.monoidal_functor.map_right_unitor CategoryTheory.MonoidalFunctor.map_rightUnitor /-- The tensorator as a natural isomorphism. -/ @@ -229,7 +319,7 @@ noncomputable def ΞΌNatIso : apply F.ΞΌIso) (by intros - apply F.toLaxMonoidalFunctor.ΞΌ_natural) + apply F.ΞΌ_natural) #align category_theory.monoidal_functor.ΞΌ_nat_iso CategoryTheory.MonoidalFunctor.ΞΌNatIso @[simp] @@ -268,18 +358,15 @@ theorem Ξ΅_hom_inv_id : F.Ξ΅ ≫ F.Ξ΅Iso.inv = πŸ™ _ := @[simps!] noncomputable def commTensorLeft (X : C) : F.toFunctor β‹™ tensorLeft (F.toFunctor.obj X) β‰… tensorLeft X β‹™ F.toFunctor := - NatIso.ofComponents (fun Y => F.ΞΌIso X Y) @fun Y Z f => by - convert F.ΞΌ_natural (πŸ™ X) f using 2 - simp + NatIso.ofComponents (fun Y => F.ΞΌIso X Y) fun f => F.ΞΌ_natural_right X f + #align category_theory.monoidal_functor.comm_tensor_left CategoryTheory.MonoidalFunctor.commTensorLeft /-- Monoidal functors commute with right tensoring up to isomorphism -/ @[simps!] noncomputable def commTensorRight (X : C) : F.toFunctor β‹™ tensorRight (F.toFunctor.obj X) β‰… tensorRight X β‹™ F.toFunctor := - NatIso.ofComponents (fun Y => F.ΞΌIso Y X) @fun Y Z f => by - convert F.ΞΌ_natural f (πŸ™ X) using 2 - simp + NatIso.ofComponents (fun Y => F.ΞΌIso Y X) fun f => F.ΞΌ_natural_left f X #align category_theory.monoidal_functor.comm_tensor_right CategoryTheory.MonoidalFunctor.commTensorRight end @@ -320,32 +407,18 @@ def comp : LaxMonoidalFunctor.{v₁, v₃} C E := { F.toFunctor β‹™ G.toFunctor with Ξ΅ := G.Ξ΅ ≫ G.map F.Ξ΅ ΞΌ := fun X Y => G.ΞΌ (F.obj X) (F.obj Y) ≫ G.map (F.ΞΌ X Y) - ΞΌ_natural := @fun _ _ _ _ f g => by - simp only [Functor.comp_map, assoc] - rw [← Category.assoc, LaxMonoidalFunctor.ΞΌ_natural, Category.assoc, ← map_comp, ← map_comp, - ← LaxMonoidalFunctor.ΞΌ_natural] + ΞΌ_natural_left := by + intro X Y f X' + simp_rw [comp_obj, F.comp_map, ΞΌ_natural_left_assoc, assoc, ← G.map_comp, ΞΌ_natural_left] + ΞΌ_natural_right := by + intro X Y f X' + simp_rw [comp_obj, F.comp_map, ΞΌ_natural_right_assoc, assoc, ← G.map_comp, ΞΌ_natural_right] associativity := fun X Y Z => by dsimp - rw [id_tensor_comp] - slice_rhs 3 4 => rw [← G.toFunctor.map_id, G.ΞΌ_natural] + simp only [comp_whiskerRight, assoc, ΞΌ_natural_left_assoc, MonoidalCategory.whiskerLeft_comp, + ΞΌ_natural_right_assoc] slice_rhs 1 3 => rw [← G.associativity] - rw [comp_tensor_id] - slice_lhs 2 3 => rw [← G.toFunctor.map_id, G.ΞΌ_natural] - rw [Category.assoc, Category.assoc, Category.assoc, Category.assoc, Category.assoc, ← - G.toFunctor.map_comp, ← G.toFunctor.map_comp, ← G.toFunctor.map_comp, ← - G.toFunctor.map_comp, F.associativity] - left_unitality := fun X => by - dsimp - rw [G.left_unitality, comp_tensor_id, Category.assoc, Category.assoc] - apply congr_arg - rw [F.left_unitality, map_comp, ← NatTrans.id_app, ← Category.assoc, ← - LaxMonoidalFunctor.ΞΌ_natural, NatTrans.id_app, map_id, ← Category.assoc, map_comp] - right_unitality := fun X => by - dsimp - rw [G.right_unitality, id_tensor_comp, Category.assoc, Category.assoc] - apply congr_arg - rw [F.right_unitality, map_comp, ← NatTrans.id_app, ← Category.assoc, ← - LaxMonoidalFunctor.ΞΌ_natural, NatTrans.id_app, map_id, ← Category.assoc, map_comp] } + simp_rw [Category.assoc, ←G.toFunctor.map_comp, F.associativity] } #align category_theory.lax_monoidal_functor.comp CategoryTheory.LaxMonoidalFunctor.comp @[inherit_doc] @@ -482,44 +555,44 @@ end MonoidalFunctor /-- If we have a right adjoint functor `G` to a monoidal functor `F`, then `G` has a lax monoidal structure as well. -/ -@[simps] +@[simp] noncomputable def monoidalAdjoint (F : MonoidalFunctor C D) {G : D β₯€ C} (h : F.toFunctor ⊣ G) : - LaxMonoidalFunctor D C where - toFunctor := G - Ξ΅ := h.homEquiv _ _ (inv F.Ξ΅) - ΞΌ X Y := h.homEquiv _ (X βŠ— Y) (inv (F.ΞΌ (G.obj X) (G.obj Y)) ≫ (h.counit.app X βŠ— h.counit.app Y)) - ΞΌ_natural := @fun X Y X' Y' f g => by + LaxMonoidalFunctor D C := LaxMonoidalFunctor.ofTensorHom + (F := G) + (Ξ΅ := h.homEquiv _ _ (inv F.Ξ΅)) + (ΞΌ := fun X Y ↦ h.homEquiv _ (X βŠ— Y) (inv (F.ΞΌ (G.obj X) (G.obj Y)) ≫ (h.counit.app X βŠ— h.counit.app Y))) + (ΞΌ_natural := @fun X Y X' Y' f g => by rw [← h.homEquiv_naturality_left, ← h.homEquiv_naturality_right, Equiv.apply_eq_iff_eq, assoc, IsIso.eq_inv_comp, ← F.toLaxMonoidalFunctor.ΞΌ_natural_assoc, IsIso.hom_inv_id_assoc, ← - tensor_comp, Adjunction.counit_naturality, Adjunction.counit_naturality, tensor_comp] - associativity X Y Z := by + tensor_comp, Adjunction.counit_naturality, Adjunction.counit_naturality, tensor_comp]) + (associativity := fun X Y Z ↦ by dsimp only rw [← h.homEquiv_naturality_right, ← h.homEquiv_naturality_left, ← h.homEquiv_naturality_left, ← h.homEquiv_naturality_left, Equiv.apply_eq_iff_eq, ← cancel_epi (F.toLaxMonoidalFunctor.ΞΌ (G.obj X βŠ— G.obj Y) (G.obj Z)), ← cancel_epi (F.toLaxMonoidalFunctor.ΞΌ (G.obj X) (G.obj Y) βŠ— πŸ™ (F.obj (G.obj Z))), - F.toLaxMonoidalFunctor.associativity_assoc (G.obj X) (G.obj Y) (G.obj Z), ← + F.toLaxMonoidalFunctor.associativity'_assoc (G.obj X) (G.obj Y) (G.obj Z), ← F.toLaxMonoidalFunctor.ΞΌ_natural_assoc, assoc, IsIso.hom_inv_id_assoc, ← F.toLaxMonoidalFunctor.ΞΌ_natural_assoc, IsIso.hom_inv_id_assoc, ← tensor_comp, ← tensor_comp, id_comp, Functor.map_id, Functor.map_id, id_comp, ← tensor_comp_assoc, ← tensor_comp_assoc, id_comp, id_comp, h.homEquiv_unit, h.homEquiv_unit, Functor.map_comp, assoc, assoc, h.counit_naturality, h.left_triangle_components_assoc, Functor.map_comp, assoc, h.counit_naturality, h.left_triangle_components_assoc] - simp - left_unitality X := by + simp) + (left_unitality := fun X ↦ by rw [← h.homEquiv_naturality_right, ← h.homEquiv_naturality_left, ← Equiv.symm_apply_eq, - h.homEquiv_counit, F.map_leftUnitor, h.homEquiv_unit, assoc, assoc, assoc, F.map_tensor, + h.homEquiv_counit, F.map_leftUnitor', h.homEquiv_unit, assoc, assoc, assoc, F.map_tensor, assoc, assoc, IsIso.hom_inv_id_assoc, ← tensor_comp_assoc, Functor.map_id, id_comp, Functor.map_comp, assoc, h.counit_naturality, h.left_triangle_components_assoc, ← - leftUnitor_naturality, ← tensor_comp_assoc, id_comp, comp_id] - simp - right_unitality X := by + leftUnitor_naturality', ← tensor_comp_assoc, id_comp, comp_id] + simp) + (right_unitality := fun X ↦ by rw [← h.homEquiv_naturality_right, ← h.homEquiv_naturality_left, ← Equiv.symm_apply_eq, - h.homEquiv_counit, F.map_rightUnitor, assoc, assoc, ← rightUnitor_naturality, ← + h.homEquiv_counit, F.map_rightUnitor', assoc, assoc, ← rightUnitor_naturality', ← tensor_comp_assoc, comp_id, id_comp, h.homEquiv_unit, F.map_tensor, assoc, assoc, assoc, IsIso.hom_inv_id_assoc, Functor.map_comp, Functor.map_id, ← tensor_comp_assoc, assoc, h.counit_naturality, h.left_triangle_components_assoc, id_comp] - simp + simp) #align category_theory.monoidal_adjoint CategoryTheory.monoidalAdjoint /-- If a monoidal functor `F` is an equivalence of categories then its inverse is also monoidal. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Functorial.lean b/Mathlib/CategoryTheory/Monoidal/Functorial.lean index 2e729969ef121..5582488929ed8 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functorial.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functorial.lean @@ -47,36 +47,68 @@ open MonoidalCategory variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] {D : Type uβ‚‚} [Category.{vβ‚‚} D] [MonoidalCategory.{vβ‚‚} D] --- Perhaps in the future we'll redefine `LaxMonoidalFunctor` in terms of this, --- but that isn't the immediate plan. -/-- An unbundled description of lax monoidal functors. -/ -class LaxMonoidal (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] where +class LaxMonoidalStruct (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] where /-- unit morphism -/ Ξ΅ : πŸ™_ D ⟢ F (πŸ™_ C) /-- tensorator -/ ΞΌ : βˆ€ X Y : C, F X βŠ— F Y ⟢ F (X βŠ— Y) - /-- naturality -/ - ΞΌ_natural : - βˆ€ {X Y X' Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y'), - (map F f βŠ— map F g) ≫ ΞΌ Y Y' = ΞΌ X X' ≫ map F (f βŠ— g) := by - aesop_cat + +-- Perhaps in the future we'll redefine `LaxMonoidalFunctor` in terms of this, +-- but that isn't the immediate plan. +/-- An unbundled description of lax monoidal functors. -/ +class LaxMonoidal (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] extends LaxMonoidalStruct F where + ΞΌ_natural_left : + βˆ€ {X Y : C} (f : X ⟢ Y) (X' : C), + (map F f β–· F X') ≫ ΞΌ Y X' = ΞΌ X X' ≫ map F (f β–· X') := by + aesop_cat + ΞΌ_natural_right : + βˆ€ {X Y : C} (X' : C) (f : X ⟢ Y) , + (F X' ◁ map F f) ≫ ΞΌ X' Y = ΞΌ X' X ≫ map F (X' ◁ f) := by + aesop_cat /-- associativity of the tensorator -/ associativity : βˆ€ X Y Z : C, - (ΞΌ X Y βŠ— πŸ™ (F Z)) ≫ ΞΌ (X βŠ— Y) Z ≫ map F (Ξ±_ X Y Z).hom = - (Ξ±_ (F X) (F Y) (F Z)).hom ≫ (πŸ™ (F X) βŠ— ΞΌ Y Z) ≫ ΞΌ X (Y βŠ— Z) := by - aesop_cat - /-- left unitality -/ - left_unitality : βˆ€ X : C, (Ξ»_ (F X)).hom = (Ξ΅ βŠ— πŸ™ (F X)) ≫ ΞΌ (πŸ™_ C) X ≫ map F (Ξ»_ X).hom := by + (ΞΌ X Y β–· F Z) ≫ ΞΌ (X βŠ— Y) Z ≫ map F (Ξ±_ X Y Z).hom = + (Ξ±_ (F X) (F Y) (F Z)).hom ≫ (F X ◁ ΞΌ Y Z) ≫ ΞΌ X (Y βŠ— Z) := by aesop_cat - /-- right unitality -/ - right_unitality : βˆ€ X : C, (ρ_ (F X)).hom = (πŸ™ (F X) βŠ— Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ map F (ρ_ X).hom := by - aesop_cat -#align category_theory.lax_monoidal CategoryTheory.LaxMonoidal - -attribute [simp] LaxMonoidal.ΞΌ_natural + -- unitality + left_unitality : βˆ€ X : C, (Ξ»_ (F X)).hom = (Ξ΅ β–· F X) ≫ ΞΌ (πŸ™_ C) X ≫ map F (Ξ»_ X).hom := + by aesop_cat + right_unitality : βˆ€ X : C, (ρ_ (F X)).hom = (F X ◁ Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ map F (ρ_ X).hom := + by aesop_cat -attribute [simp] LaxMonoidal.ΞΌ_natural +-- Perhaps in the future we'll redefine `LaxMonoidalFunctor` in terms of this, +-- but that isn't the immediate plan. +open LaxMonoidalStruct in +/-- An unbundled description of lax monoidal functors. -/ +def LaxMonoidal.ofTensorHom (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] [LaxMonoidalStruct F] + /- naturality -/ + (ΞΌ_natural : + βˆ€ {X Y X' Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y'), + (map F f βŠ— map F g) ≫ ΞΌ Y Y' = ΞΌ X X' ≫ map F (f βŠ— g) := by + aesop_cat) + /- associativity of the tensorator -/ + (associativity : + βˆ€ X Y Z : C, + (ΞΌ X Y βŠ— πŸ™ (F Z)) ≫ ΞΌ (X βŠ— Y) Z ≫ map F (Ξ±_ X Y Z).hom = + (Ξ±_ (F X) (F Y) (F Z)).hom ≫ (πŸ™ (F X) βŠ— ΞΌ Y Z) ≫ ΞΌ X (Y βŠ— Z) := by + aesop_cat) + /- left unitality -/ + (left_unitality : βˆ€ X : C, (Ξ»_ (F X)).hom = (Ξ΅ βŠ— πŸ™ (F X)) ≫ ΞΌ (πŸ™_ C) X ≫ map F (Ξ»_ X).hom := by + aesop_cat) + /- right unitality -/ + (right_unitality : βˆ€ X : C, (ρ_ (F X)).hom = (πŸ™ (F X) βŠ— Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ map F (ρ_ X).hom := by + aesop_cat) : + LaxMonoidal.{v₁, vβ‚‚} F where + Ξ΅ := Ξ΅ + ΞΌ := ΞΌ + ΞΌ_natural_left := sorry + ΞΌ_natural_right := sorry + associativity := sorry + left_unitality := sorry + right_unitality := sorry + +attribute [simp] LaxMonoidal.ΞΌ_natural_left LaxMonoidal.ΞΌ_natural_right -- The unitality axioms cannot be used as simp lemmas because they require -- higher-order matching to figure out the `F` and `X` from `F X`. diff --git a/Mathlib/CategoryTheory/Monoidal/Linear.lean b/Mathlib/CategoryTheory/Monoidal/Linear.lean index 4315144fb05b0..5c5bb91f2a6e6 100644 --- a/Mathlib/CategoryTheory/Monoidal/Linear.lean +++ b/Mathlib/CategoryTheory/Monoidal/Linear.lean @@ -32,13 +32,13 @@ variable [MonoidalCategory C] /-- A category is `MonoidalLinear R` if tensoring is `R`-linear in both factors. -/ class MonoidalLinear [MonoidalPreadditive C] : Prop where - tensor_smul : βˆ€ {W X Y Z : C} (f : W ⟢ X) (r : R) (g : Y ⟢ Z), f βŠ— r β€’ g = r β€’ (f βŠ— g) := by + whiskerLeft_smul : βˆ€ (X : C) {Y Z : C} (r : R) (f : Y ⟢ Z) , X ◁ (r β€’ f) = r β€’ (X ◁ f) := by aesop_cat - smul_tensor : βˆ€ {W X Y Z : C} (r : R) (f : W ⟢ X) (g : Y ⟢ Z), r β€’ f βŠ— g = r β€’ (f βŠ— g) := by + smul_whiskerRight : βˆ€ (r : R) {Y Z : C} (f : Y ⟢ Z) (X : C), (r β€’ f) β–· X = r β€’ (f β–· X) := by aesop_cat #align category_theory.monoidal_linear CategoryTheory.MonoidalLinear -attribute [simp] MonoidalLinear.tensor_smul MonoidalLinear.smul_tensor +attribute [simp] MonoidalLinear.whiskerLeft_smul MonoidalLinear.smul_whiskerRight variable {C} variable [MonoidalPreadditive C] [MonoidalLinear R C] @@ -60,16 +60,16 @@ ensures that the domain is linear monoidal. -/ theorem monoidalLinearOfFaithful {D : Type _} [Category D] [Preadditive D] [Linear R D] [MonoidalCategory D] [MonoidalPreadditive D] (F : MonoidalFunctor D C) [Faithful F.toFunctor] [F.toFunctor.Additive] [F.toFunctor.Linear R] : MonoidalLinear R D := - { tensor_smul := by - intros W X Y Z f r g + { whiskerLeft_smul := by + intros X Y Z r f apply F.toFunctor.map_injective - simp only [F.toFunctor.map_smul r (f βŠ— g), F.toFunctor.map_smul r g, F.map_tensor, - MonoidalLinear.tensor_smul, Linear.smul_comp, Linear.comp_smul] - smul_tensor := by - intros W X Y Z r f g + rw [F.map_whiskerLeft] + simp + smul_whiskerRight := by + intros r X Y f Z apply F.toFunctor.map_injective - simp only [F.toFunctor.map_smul r (f βŠ— g), F.toFunctor.map_smul r f, F.map_tensor, - MonoidalLinear.smul_tensor, Linear.smul_comp, Linear.comp_smul] } + rw [F.map_whiskerRight] + simp } #align category_theory.monoidal_linear_of_faithful CategoryTheory.monoidalLinearOfFaithful end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index f293f02c878e7..275eead75c874 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -33,6 +33,8 @@ namespace CategoryTheory open MonoidalCategory +attribute [local simp] tensorHom_def + variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] {D : Type uβ‚‚} [Category.{vβ‚‚} D] [MonoidalCategory.{vβ‚‚} D] @@ -76,7 +78,8 @@ instance (F : LaxMonoidalFunctor C D) : Inhabited (MonoidalNatTrans F F) := @[simps!] def vcomp {F G H : LaxMonoidalFunctor C D} (Ξ± : MonoidalNatTrans F G) (Ξ² : MonoidalNatTrans G H) : MonoidalNatTrans F H := - { NatTrans.vcomp Ξ±.toNatTrans Ξ².toNatTrans with } + { NatTrans.vcomp Ξ±.toNatTrans Ξ².toNatTrans with + tensor := fun X Y => by simp [whisker_exchange_assoc] } #align category_theory.monoidal_nat_trans.vcomp CategoryTheory.MonoidalNatTrans.vcomp instance categoryLaxMonoidalFunctor : Category (LaxMonoidalFunctor C D) where @@ -117,8 +120,13 @@ def hcomp {F G : LaxMonoidalFunctor C D} {H K : LaxMonoidalFunctor D E} (Ξ± : Mo dsimp; simp conv_lhs => rw [← K.toFunctor.map_comp, Ξ±.unit] tensor := fun X Y => by - dsimp; simp - conv_lhs => rw [← K.toFunctor.map_comp, Ξ±.tensor, K.toFunctor.map_comp] } + simp only [LaxMonoidalFunctor.comp_toFunctor, comp_obj, LaxMonoidalFunctor.comp_ΞΌ, NatTrans.hcomp_app, assoc, + NatTrans.naturality_assoc, tensor_assoc, tensorHom_def, whisker_exchange_assoc, MonoidalCategory.whiskerLeft_comp, + comp_whiskerRight, LaxMonoidalFunctor.ΞΌ_natural_left_assoc, LaxMonoidalFunctor.ΞΌ_natural_right_assoc] + simp_rw [← K.toFunctor.map_comp] + congr 4 + simp + } #align category_theory.monoidal_nat_trans.hcomp CategoryTheory.MonoidalNatTrans.hcomp section @@ -155,9 +163,9 @@ def ofComponents (app : βˆ€ X : C, F.obj X β‰… G.obj X) dsimp rw [← unit', assoc, Iso.hom_inv_id, comp_id] tensor := fun X Y => by - dsimp - rw [Iso.comp_inv_eq, assoc, tensor', ← tensor_comp_assoc, - Iso.inv_hom_id, Iso.inv_hom_id, tensor_id, id_comp] } + dsimp only + rw [Iso.comp_inv_eq, assoc, tensor'] + simp [whisker_exchange_assoc]} #align category_theory.monoidal_nat_iso.of_components CategoryTheory.MonoidalNatIso.ofComponents @[simp] @@ -193,17 +201,14 @@ def monoidalUnit (F : MonoidalFunctor C D) [IsEquivalence F.toFunctor] : simp only [Adjunction.homEquiv_unit, Adjunction.homEquiv_naturality_right, id_comp, assoc] simp only [← Functor.map_comp, assoc] - erw [e.counit_app_functor, e.counit_app_functor, - F.toLaxMonoidalFunctor.ΞΌ_natural, IsIso.inv_hom_id_assoc] - simp only [CategoryTheory.IsEquivalence.inv_fun_map] - slice_rhs 2 3 => erw [Iso.hom_inv_id_app] - dsimp - simp only [CategoryTheory.Category.id_comp] - slice_rhs 1 2 => - rw [← tensor_comp, Iso.hom_inv_id_app, Iso.hom_inv_id_app] - dsimp - rw [tensor_id] - simp } + erw [e.counit_app_functor, e.counit_app_functor] + simp only [comp_obj, asEquivalence_functor, asEquivalence_inverse, id_obj, LaxMonoidalFunctor.ΞΌ_natural_left, + LaxMonoidalFunctor.ΞΌ_natural_right_assoc, IsIso.inv_hom_id_assoc, map_comp, IsEquivalence.inv_fun_map, assoc, + Iso.hom_inv_id_app_assoc, tensorHom_def] + slice_rhs 3 4 => erw [Iso.hom_inv_id_app] + simp only [id_obj, id_comp, assoc, whisker_exchange_assoc] + simp only [← whiskerLeft_comp_assoc, ← comp_whiskerRight_assoc] + simp [- whiskerLeft_comp, - comp_whiskerRight] } #align category_theory.monoidal_unit CategoryTheory.monoidalUnit instance (F : MonoidalFunctor C D) [IsEquivalence F.toFunctor] : IsIso (monoidalUnit F) := diff --git a/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean b/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean index 9e64c577827c1..44e1b1f1c70c3 100644 --- a/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean @@ -331,20 +331,20 @@ end MonoidalOfChosenFiniteProducts open MonoidalOfChosenFiniteProducts /-- A category with a terminal object and binary products has a natural monoidal structure. -/ -def monoidalOfChosenFiniteProducts : MonoidalCategory C where - tensorUnit' := 𝒯.cone.pt - tensorObj X Y := tensorObj ℬ X Y - tensorHom f g := tensorHom ℬ f g - tensor_id := tensor_id ℬ - tensor_comp f₁ fβ‚‚ g₁ gβ‚‚ := tensor_comp ℬ f₁ fβ‚‚ g₁ gβ‚‚ - associator X Y Z := BinaryFan.associatorOfLimitCone ℬ X Y Z - leftUnitor X := BinaryFan.leftUnitor 𝒯.isLimit (ℬ 𝒯.cone.pt X).isLimit - rightUnitor X := BinaryFan.rightUnitor 𝒯.isLimit (ℬ X 𝒯.cone.pt).isLimit - pentagon := pentagon ℬ - triangle := triangle 𝒯 ℬ - leftUnitor_naturality f := leftUnitor_naturality 𝒯 ℬ f - rightUnitor_naturality f := rightUnitor_naturality 𝒯 ℬ f - associator_naturality f₁ fβ‚‚ f₃ := associator_naturality ℬ f₁ fβ‚‚ f₃ +def monoidalOfChosenFiniteProducts : MonoidalCategory C := .ofTensorHom + (tensorUnit' := 𝒯.cone.pt) + (tensorObj := tensorObj ℬ) + (tensorHom := tensorHom ℬ) + (tensor_id := tensor_id ℬ) + (tensor_comp := tensor_comp ℬ) + (associator := BinaryFan.associatorOfLimitCone ℬ) + (leftUnitor := fun X ↦ BinaryFan.leftUnitor 𝒯.isLimit (ℬ 𝒯.cone.pt X).isLimit) + (rightUnitor := fun X ↦ BinaryFan.rightUnitor 𝒯.isLimit (ℬ X 𝒯.cone.pt).isLimit) + (pentagon := pentagon ℬ) + (triangle := triangle 𝒯 ℬ) + (leftUnitor_naturality := leftUnitor_naturality 𝒯 ℬ) + (rightUnitor_naturality := rightUnitor_naturality 𝒯 ℬ) + (associator_naturality := associator_naturality ℬ) #align category_theory.monoidal_of_chosen_finite_products CategoryTheory.monoidalOfChosenFiniteProducts namespace MonoidalOfChosenFiniteProducts diff --git a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean index ba99a75e5ebfd..bfc462eeee96f 100644 --- a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean +++ b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean @@ -45,39 +45,21 @@ open CategoryTheory.Limits section /-- A category with a terminal object and binary products has a natural monoidal structure. -/ -def monoidalOfHasFiniteProducts [HasTerminal C] [HasBinaryProducts C] : MonoidalCategory C where - tensorUnit' := ⊀_ C - tensorObj X Y := X β¨― Y - tensorHom f g := Limits.prod.map f g - associator := prod.associator - leftUnitor P := prod.leftUnitor P - rightUnitor P := prod.rightUnitor P - pentagon := prod.pentagon - triangle := prod.triangle - associator_naturality := @prod.associator_naturality _ _ _ +def monoidalOfHasFiniteProducts [HasTerminal C] [HasBinaryProducts C] : MonoidalCategory C := + .ofTensorHom + (tensorUnit' := ⊀_ C) + (tensorObj := fun X Y ↦ X β¨― Y) + (tensorHom := fun f g ↦ Limits.prod.map f g) + (associator := prod.associator) + (leftUnitor := fun P ↦ prod.leftUnitor P) + (rightUnitor := fun P ↦ prod.rightUnitor P) + (pentagon := prod.pentagon) + (triangle := prod.triangle) + (associator_naturality := @prod.associator_naturality _ _ _) #align category_theory.monoidal_of_has_finite_products CategoryTheory.monoidalOfHasFiniteProducts end -section - -attribute [local instance] monoidalOfHasFiniteProducts - -open MonoidalCategory - -/-- The monoidal structure coming from finite products is symmetric. --/ -@[simps] -def symmetricOfHasFiniteProducts [HasTerminal C] [HasBinaryProducts C] : SymmetricCategory C where - braiding X Y := Limits.prod.braiding X Y - braiding_naturality f g := by dsimp [tensorHom]; simp - hexagon_forward X Y Z := by dsimp [monoidalOfHasFiniteProducts]; simp - hexagon_reverse X Y Z := by dsimp [monoidalOfHasFiniteProducts]; simp - symmetry X Y := by dsimp; simp; rfl -#align category_theory.symmetric_of_has_finite_products CategoryTheory.symmetricOfHasFiniteProducts - -end - namespace monoidalOfHasFiniteProducts variable [HasTerminal C] [HasBinaryProducts C] @@ -126,42 +108,47 @@ theorem associator_hom (X Y Z : C) : rfl #align category_theory.monoidal_of_has_finite_products.associator_hom CategoryTheory.monoidalOfHasFiniteProducts.associator_hom -end monoidalOfHasFiniteProducts - -section - -/-- A category with an initial object and binary coproducts has a natural monoidal structure. -/ -def monoidalOfHasFiniteCoproducts [HasInitial C] [HasBinaryCoproducts C] : MonoidalCategory C where - tensorUnit' := βŠ₯_ C - tensorObj X Y := X β¨Ώ Y - tensorHom f g := Limits.coprod.map f g - associator := coprod.associator - leftUnitor := coprod.leftUnitor - rightUnitor := coprod.rightUnitor - pentagon := coprod.pentagon - triangle := coprod.triangle - associator_naturality := @coprod.associator_naturality _ _ _ -#align category_theory.monoidal_of_has_finite_coproducts CategoryTheory.monoidalOfHasFiniteCoproducts +theorem associator_inv (X Y Z : C) : + (Ξ±_ X Y Z).inv = + prod.lift (prod.lift prod.fst (prod.snd ≫ prod.fst)) (prod.snd ≫ prod.snd) := + rfl -end +end monoidalOfHasFiniteProducts section -attribute [local instance] monoidalOfHasFiniteCoproducts +attribute [local instance] monoidalOfHasFiniteProducts open MonoidalCategory -/-- The monoidal structure coming from finite coproducts is symmetric. +/-- The monoidal structure coming from finite products is symmetric. -/ @[simps] -def symmetricOfHasFiniteCoproducts [HasInitial C] [HasBinaryCoproducts C] : - SymmetricCategory C where - braiding := Limits.coprod.braiding - braiding_naturality f g := by dsimp [tensorHom]; simp - hexagon_forward X Y Z := by dsimp [monoidalOfHasFiniteCoproducts]; simp - hexagon_reverse X Y Z := by dsimp [monoidalOfHasFiniteCoproducts]; simp - symmetry X Y := by dsimp; simp; rfl -#align category_theory.symmetric_of_has_finite_coproducts CategoryTheory.symmetricOfHasFiniteCoproducts +def symmetricOfHasFiniteProducts [HasTerminal C] [HasBinaryProducts C] : SymmetricCategory C where + braiding X Y := Limits.prod.braiding X Y + braiding_naturality f g := by dsimp; simp + hexagon_forward X Y Z := by dsimp [monoidalOfHasFiniteProducts.associator_hom]; simp + hexagon_reverse X Y Z := by dsimp [monoidalOfHasFiniteProducts.associator_inv]; simp + symmetry X Y := by dsimp; simp +#align category_theory.symmetric_of_has_finite_products CategoryTheory.symmetricOfHasFiniteProducts + +end + +section + +/-- A category with an initial object and binary coproducts has a natural monoidal structure. -/ +def monoidalOfHasFiniteCoproducts [HasInitial C] [HasBinaryCoproducts C] : MonoidalCategory C := + .ofTensorHom + (tensorUnit' := βŠ₯_ C) + (tensorObj := fun X Y ↦ X β¨Ώ Y) + (tensorHom := fun f g ↦ Limits.coprod.map f g) + (associator := coprod.associator) + (leftUnitor := coprod.leftUnitor) + (rightUnitor := coprod.rightUnitor) + (pentagon := coprod.pentagon) + (triangle := coprod.triangle) + (associator_naturality := @coprod.associator_naturality _ _ _) +#align category_theory.monoidal_of_has_finite_coproducts CategoryTheory.monoidalOfHasFiniteCoproducts end @@ -212,6 +199,31 @@ theorem associator_hom (X Y Z : C) : rfl #align category_theory.monoidal_of_has_finite_coproducts.associator_hom CategoryTheory.monoidalOfHasFiniteCoproducts.associator_hom +theorem associator_inv (X Y Z : C) : + (Ξ±_ X Y Z).inv = + coprod.desc (coprod.inl ≫ coprod.inl) (coprod.desc (coprod.inr ≫ coprod.inl) coprod.inr) := + rfl + end monoidalOfHasFiniteCoproducts +section + +attribute [local instance] monoidalOfHasFiniteCoproducts + +open MonoidalCategory + +/-- The monoidal structure coming from finite coproducts is symmetric. +-/ +@[simps] +def symmetricOfHasFiniteCoproducts [HasInitial C] [HasBinaryCoproducts C] : + SymmetricCategory C where + braiding := Limits.coprod.braiding + braiding_naturality f g := by dsimp [tensorHom]; simp + hexagon_forward X Y Z := by dsimp [monoidalOfHasFiniteCoproducts.associator_hom]; simp + hexagon_reverse X Y Z := by dsimp [monoidalOfHasFiniteCoproducts.associator_inv]; simp + symmetry X Y := by dsimp; simp +#align category_theory.symmetric_of_has_finite_coproducts CategoryTheory.symmetricOfHasFiniteCoproducts + +end + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index 520ba3884fa3c..73b02c9e9ff6b 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -34,24 +34,20 @@ Note we don't `extend Preadditive C` here, as `Abelian C` already extends it, and we'll need to have both typeclasses sometimes. -/ class MonoidalPreadditive : Prop where - /-- tensoring on the right with a zero morphism gives zero -/ - tensor_zero : βˆ€ {W X Y Z : C} (f : W ⟢ X), f βŠ— (0 : Y ⟢ Z) = 0 := by aesop_cat - /-- tensoring on the left with a zero morphism gives zero -/ - zero_tensor : βˆ€ {W X Y Z : C} (f : Y ⟢ Z), (0 : W ⟢ X) βŠ— f = 0 := by aesop_cat - /-- left tensoring with a morphism is compatible with addition -/ - tensor_add : βˆ€ {W X Y Z : C} (f : W ⟢ X) (g h : Y ⟢ Z), f βŠ— (g + h) = f βŠ— g + f βŠ— h := by - aesop_cat - /-- right tensoring with a morphism is compatible with addition -/ - add_tensor : βˆ€ {W X Y Z : C} (f g : W ⟢ X) (h : Y ⟢ Z), (f + g) βŠ— h = f βŠ— h + g βŠ— h := by - aesop_cat + whiskerLeft_zero : βˆ€ {X Y Z : C}, X ◁ (0 : Y ⟢ Z) = 0 := by aesop_cat + zero_whiskerRight : βˆ€ {X Y Z : C}, (0 : Y ⟢ Z) β–· X = 0 := by aesop_cat + whiskerLeft_add : βˆ€ {X Y Z : C} (f g : Y ⟢ Z), X ◁ (f + g) = X ◁ f + X ◁ g := by aesop_cat + add_whiskerRight : βˆ€ {X Y Z : C} (f g : Y ⟢ Z), (f + g) β–· X = f β–· X + g β–· X := by aesop_cat #align category_theory.monoidal_preadditive CategoryTheory.MonoidalPreadditive -attribute [simp] MonoidalPreadditive.tensor_zero MonoidalPreadditive.zero_tensor +-- attribute [simp] MonoidalPreadditive.tensor_zero MonoidalPreadditive.zero_tensor +attribute [simp] MonoidalPreadditive.whiskerLeft_zero MonoidalPreadditive.zero_whiskerRight +attribute [simp] MonoidalPreadditive.whiskerLeft_add MonoidalPreadditive.add_whiskerRight variable {C} variable [MonoidalPreadditive C] -attribute [local simp] MonoidalPreadditive.tensor_add MonoidalPreadditive.add_tensor +-- attribute [local simp] MonoidalPreadditive.tensor_add MonoidalPreadditive.add_tensor instance tensorLeft_additive (X : C) : (tensorLeft X).Additive where #align category_theory.tensor_left_additive CategoryTheory.tensorLeft_additive @@ -70,46 +66,46 @@ ensures that the domain is monoidal preadditive. -/ theorem monoidalPreadditive_of_faithful {D} [Category D] [Preadditive D] [MonoidalCategory D] (F : MonoidalFunctor D C) [Faithful F.toFunctor] [F.toFunctor.Additive] : MonoidalPreadditive D := - { tensor_zero := by + { whiskerLeft_zero := by intros apply F.toFunctor.map_injective - simp [F.map_tensor] - zero_tensor := by + simp [F.map_whiskerLeft] + zero_whiskerRight := by intros apply F.toFunctor.map_injective - simp [F.map_tensor] - tensor_add := by + simp [F.map_whiskerRight] + whiskerLeft_add := by intros apply F.toFunctor.map_injective - simp only [F.map_tensor, Functor.map_add, Preadditive.comp_add, Preadditive.add_comp, - MonoidalPreadditive.tensor_add] - add_tensor := by + simp only [F.map_whiskerLeft, Functor.map_add, Preadditive.comp_add, Preadditive.add_comp, + MonoidalPreadditive.whiskerLeft_add] + add_whiskerRight := by intros apply F.toFunctor.map_injective - simp only [F.map_tensor, Functor.map_add, Preadditive.comp_add, Preadditive.add_comp, - MonoidalPreadditive.add_tensor] } + simp only [F.map_whiskerRight, Functor.map_add, Preadditive.comp_add, Preadditive.add_comp, + MonoidalPreadditive.add_whiskerRight] } #align category_theory.monoidal_preadditive_of_faithful CategoryTheory.monoidalPreadditive_of_faithful open BigOperators +attribute [local simp] tensorHom_def + +theorem whiskerLeft_sum (P : C) {Q R : C} {J : Type _} (s : Finset J) (g : J β†’ (Q ⟢ R)) : + P ◁ βˆ‘ j in s, g j = βˆ‘ j in s, P ◁ g j := + map_sum ((tensoringLeft C).obj P).mapAddHom g s + +theorem sum_whiskerRight {Q R : C} {J : Type _} (s : Finset J) (g : J β†’ (Q ⟢ R)) (P : C) : + (βˆ‘ j in s, g j) β–· P = βˆ‘ j in s, g j β–· P := + map_sum ((tensoringRight C).obj P).mapAddHom g s + theorem tensor_sum {P Q R S : C} {J : Type _} (s : Finset J) (f : P ⟢ Q) (g : J β†’ (R ⟢ S)) : (f βŠ— βˆ‘ j in s, g j) = βˆ‘ j in s, f βŠ— g j := by - rw [← tensor_id_comp_id_tensor] - let tQ := (((tensoringLeft C).obj Q).mapAddHom : (R ⟢ S) β†’+ _) - change _ ≫ tQ _ = _ - rw [tQ.map_sum, Preadditive.comp_sum] - dsimp [Functor.mapAddHom] - simp only [tensor_id_comp_id_tensor] + simp [whiskerLeft_sum, sum_whiskerRight, Preadditive.sum_comp] #align category_theory.tensor_sum CategoryTheory.tensor_sum theorem sum_tensor {P Q R S : C} {J : Type _} (s : Finset J) (f : P ⟢ Q) (g : J β†’ (R ⟢ S)) : (βˆ‘ j in s, g j) βŠ— f = βˆ‘ j in s, g j βŠ— f := by - rw [← tensor_id_comp_id_tensor] - let tQ := (((tensoringRight C).obj P).mapAddHom : (R ⟢ S) β†’+ _) - change tQ _ ≫ _ = _ - rw [tQ.map_sum, Preadditive.sum_comp] - dsimp [Functor.mapAddHom] - simp only [tensor_id_comp_id_tensor] + simp [whiskerLeft_sum, sum_whiskerRight, Preadditive.comp_sum] #align category_theory.sum_tensor CategoryTheory.sum_tensor -- In a closed monoidal category, this would hold because @@ -120,6 +116,8 @@ instance (X : C) : PreservesFiniteBiproducts (tensorLeft X) where { preserves := fun {f} => { preserves := fun {b} i => isBilimitOfTotal _ (by dsimp + simp_rw [← whiskerLeft_comp] + simp_rw [← tensorHom_id, ← id_tensorHom] simp only [← tensor_comp, Category.comp_id, ← tensor_sum, ← tensor_id, IsBilimit.total i]) } } @@ -128,6 +126,7 @@ instance (X : C) : PreservesFiniteBiproducts (tensorRight X) where { preserves := fun {f} => { preserves := fun {b} i => isBilimitOfTotal _ (by dsimp + simp_rw [← tensorHom_id, ← id_tensorHom] simp only [← tensor_comp, Category.comp_id, ← sum_tensor, ← tensor_id, IsBilimit.total i]) } } @@ -141,7 +140,7 @@ def leftDistributor {J : Type} [Fintype J] (X : C) (f : J β†’ C) : X βŠ— ⨁ f @[simp] theorem leftDistributor_hom {J : Type} [Fintype J] (X : C) (f : J β†’ C) : (leftDistributor X f).hom = - βˆ‘ j : J, (πŸ™ X βŠ— biproduct.Ο€ f j) ≫ biproduct.ΞΉ (fun j => X βŠ— f j) j := by + βˆ‘ j : J, (X ◁ biproduct.Ο€ f j) ≫ biproduct.ΞΉ (fun j => X βŠ— f j) j := by ext dsimp [leftDistributor, Functor.mapBiproduct, Functor.mapBicone] erw [biproduct.lift_Ο€] @@ -151,7 +150,7 @@ theorem leftDistributor_hom {J : Type} [Fintype J] (X : C) (f : J β†’ C) : @[simp] theorem leftDistributor_inv {J : Type} [Fintype J] (X : C) (f : J β†’ C) : - (leftDistributor X f).inv = βˆ‘ j : J, biproduct.Ο€ _ j ≫ (πŸ™ X βŠ— biproduct.ΞΉ f j) := by + (leftDistributor X f).inv = βˆ‘ j : J, biproduct.Ο€ _ j ≫ (X ◁ biproduct.ΞΉ f j) := by ext dsimp [leftDistributor, Functor.mapBiproduct, Functor.mapBicone] simp only [Preadditive.comp_sum, biproduct.ΞΉ_Ο€_assoc, dite_comp, zero_comp, @@ -167,11 +166,10 @@ theorem leftDistributor_assoc {J : Type} [Fintype J] (X Y : C) (f : J β†’ C) : asIso_hom, comp_zero, comp_dite, Preadditive.sum_comp, Preadditive.comp_sum, tensor_sum, id_tensor_comp, tensorIso_hom, leftDistributor_hom, biproduct.mapIso_hom, biproduct.ΞΉ_map, biproduct.ΞΉ_Ο€, Finset.sum_dite_irrel, Finset.sum_dite_eq', Finset.sum_const_zero] + simp_rw [← tensorHom_id, ← id_tensorHom] simp only [← id_tensor_comp, biproduct.ΞΉ_Ο€] simp only [id_tensor_comp, tensor_dite, comp_dite] - simp only [Category.comp_id, comp_zero, MonoidalPreadditive.tensor_zero, eqToHom_refl, - tensor_id, if_true, dif_ctx_congr, Finset.sum_congr, Finset.mem_univ, Finset.sum_dite_eq'] - simp only [← tensor_id, associator_naturality, Iso.inv_hom_id_assoc] + simp #align category_theory.left_distributor_assoc CategoryTheory.leftDistributor_assoc /-- The isomorphism showing how tensor product on the right distributes over direct sums. -/ @@ -182,7 +180,7 @@ def rightDistributor {J : Type} [Fintype J] (X : C) (f : J β†’ C) : (⨁ f) βŠ— @[simp] theorem rightDistributor_hom {J : Type} [Fintype J] (X : C) (f : J β†’ C) : (rightDistributor X f).hom = - βˆ‘ j : J, (biproduct.Ο€ f j βŠ— πŸ™ X) ≫ biproduct.ΞΉ (fun j => f j βŠ— X) j := by + βˆ‘ j : J, (biproduct.Ο€ f j β–· X) ≫ biproduct.ΞΉ (fun j => f j βŠ— X) j := by ext dsimp [rightDistributor, Functor.mapBiproduct, Functor.mapBicone] erw [biproduct.lift_Ο€] @@ -192,7 +190,7 @@ theorem rightDistributor_hom {J : Type} [Fintype J] (X : C) (f : J β†’ C) : @[simp] theorem rightDistributor_inv {J : Type} [Fintype J] (X : C) (f : J β†’ C) : - (rightDistributor X f).inv = βˆ‘ j : J, biproduct.Ο€ _ j ≫ (biproduct.ΞΉ f j βŠ— πŸ™ X) := by + (rightDistributor X f).inv = βˆ‘ j : J, biproduct.Ο€ _ j ≫ (biproduct.ΞΉ f j β–· X) := by ext dsimp [rightDistributor, Functor.mapBiproduct, Functor.mapBicone] simp only [biproduct.ΞΉ_desc, Preadditive.comp_sum, ne_eq, biproduct.ΞΉ_Ο€_assoc, dite_comp, @@ -208,11 +206,9 @@ theorem rightDistributor_assoc {J : Type} [Fintype J] (X Y : C) (f : J β†’ C) : comp_tensor_id, tensorIso_hom, rightDistributor_hom, biproduct.mapIso_hom, biproduct.ΞΉ_map, biproduct.ΞΉ_Ο€, Finset.sum_dite_irrel, Finset.sum_dite_eq', Finset.sum_const_zero, Finset.mem_univ, if_true] + simp_rw [← tensorHom_id, ← id_tensorHom] simp only [← comp_tensor_id, biproduct.ΞΉ_Ο€, dite_tensor, comp_dite] - simp only [Category.comp_id, comp_tensor_id, eqToHom_refl, tensor_id, comp_zero, - MonoidalPreadditive.zero_tensor, if_true, dif_ctx_congr, Finset.mem_univ, Finset.sum_congr, - Finset.sum_dite_eq'] - simp only [← tensor_id, associator_inv_naturality, Iso.hom_inv_id_assoc] + simp #align category_theory.right_distributor_assoc CategoryTheory.rightDistributor_assoc theorem leftDistributor_rightDistributor_assoc {J : Type _} [Fintype J] (X Y : C) (f : J β†’ C) : @@ -226,13 +222,10 @@ theorem leftDistributor_rightDistributor_assoc {J : Type _} [Fintype J] (X Y : C tensor_sum, comp_tensor_id, tensorIso_hom, leftDistributor_hom, rightDistributor_hom, biproduct.mapIso_hom, biproduct.ΞΉ_map, biproduct.ΞΉ_Ο€, Finset.sum_dite_irrel, Finset.sum_dite_eq', Finset.sum_const_zero, Finset.mem_univ, if_true] + simp_rw [← tensorHom_id, ← id_tensorHom] simp only [← comp_tensor_id, ← id_tensor_comp_assoc, Category.assoc, biproduct.ΞΉ_Ο€, comp_dite, dite_comp, tensor_dite, dite_tensor] - simp only [Category.comp_id, Category.id_comp, Category.assoc, id_tensor_comp, comp_zero, - zero_comp, MonoidalPreadditive.tensor_zero, MonoidalPreadditive.zero_tensor, comp_tensor_id, - eqToHom_refl, tensor_id, if_true, dif_ctx_congr, Finset.sum_congr, Finset.mem_univ, - Finset.sum_dite_eq'] - simp only [associator_inv_naturality, Iso.hom_inv_id_assoc] + simp #align category_theory.left_distributor_right_distributor_assoc CategoryTheory.leftDistributor_rightDistributor_assoc end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean index ff16d3c8af4b1..8e3973e93919a 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean @@ -64,6 +64,8 @@ open CategoryTheory MonoidalCategory universe v v₁ vβ‚‚ v₃ u u₁ uβ‚‚ u₃ +attribute [local simp] id_tensorHom tensorHom_id + noncomputable section namespace CategoryTheory @@ -76,10 +78,10 @@ class ExactPairing (X Y : C) where coevaluation' : πŸ™_ C ⟢ X βŠ— Y evaluation' : Y βŠ— X ⟢ πŸ™_ C coevaluation_evaluation' : - (πŸ™ Y βŠ— coevaluation') ≫ (Ξ±_ _ _ _).inv ≫ (evaluation' βŠ— πŸ™ Y) = (ρ_ Y).hom ≫ (Ξ»_ Y).inv := by + (Y ◁ coevaluation') ≫ (Ξ±_ _ _ _).inv ≫ (evaluation' β–· Y) = (ρ_ Y).hom ≫ (Ξ»_ Y).inv := by aesop_cat evaluation_coevaluation' : - (coevaluation' βŠ— πŸ™ X) ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ X βŠ— evaluation') = (Ξ»_ X).hom ≫ (ρ_ X).inv := by + (coevaluation' β–· X) ≫ (Ξ±_ _ _ _).hom ≫ (X ◁ evaluation') = (Ξ»_ X).hom ≫ (ρ_ X).inv := by aesop_cat #align category_theory.exact_pairing CategoryTheory.ExactPairing @@ -100,13 +102,23 @@ notation "Ξ·_" => ExactPairing.coevaluation notation "Ξ΅_" => ExactPairing.evaluation lemma coevaluation_evaluation : - (πŸ™ Y βŠ— Ξ·_ _ _) ≫ (Ξ±_ _ _ _).inv ≫ (Ξ΅_ X _ βŠ— πŸ™ Y) = (ρ_ Y).hom ≫ (Ξ»_ Y).inv := + (Y ◁ Ξ·_ _ _) ≫ (Ξ±_ _ _ _).inv ≫ (Ξ΅_ X _ β–· Y) = (ρ_ Y).hom ≫ (Ξ»_ Y).inv := coevaluation_evaluation' lemma evaluation_coevaluation : - (Ξ·_ _ _ βŠ— πŸ™ X) ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ X βŠ— Ξ΅_ _ Y) = (Ξ»_ X).hom ≫ (ρ_ X).inv := + (Ξ·_ _ _ β–· X) ≫ (Ξ±_ _ _ _).hom ≫ (X ◁ Ξ΅_ _ Y) = (Ξ»_ X).hom ≫ (ρ_ X).inv := evaluation_coevaluation' +-- temporary lemma +lemma coevaluation_evaluation'' : + (πŸ™ Y βŠ— Ξ·_ _ _) ≫ (Ξ±_ _ _ _).inv ≫ (Ξ΅_ X _ βŠ— πŸ™ Y) = (ρ_ Y).hom ≫ (Ξ»_ Y).inv := by + simp [coevaluation_evaluation] + +-- temporary lemma +lemma evaluation_coevaluation'' : + (Ξ·_ _ _ βŠ— πŸ™ X) ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ X βŠ— Ξ΅_ _ Y) = (Ξ»_ X).hom ≫ (ρ_ X).inv := by + simp [evaluation_coevaluation] + end ExactPairing attribute [reassoc (attr := simp)] ExactPairing.coevaluation_evaluation @@ -115,8 +127,8 @@ attribute [reassoc (attr := simp)] ExactPairing.evaluation_coevaluation instance exactPairingUnit : ExactPairing (πŸ™_ C) (πŸ™_ C) where coevaluation' := (ρ_ _).inv evaluation' := (ρ_ _).hom - coevaluation_evaluation' := by coherence - evaluation_coevaluation' := by coherence + coevaluation_evaluation' := by rw [← id_tensorHom, ← tensorHom_id]; coherence + evaluation_coevaluation' := by rw [← id_tensorHom, ← tensorHom_id]; coherence #align category_theory.exact_pairing_unit CategoryTheory.exactPairingUnit /-- A class of objects which have a right dual. -/ @@ -167,12 +179,12 @@ theorem rightDual_leftDual {X : C} [HasLeftDual X] : (ᘁX)ᘁ = X := /-- The right adjoint mate `fᘁ : Xᘁ ⟢ Yᘁ` of a morphism `f : X ⟢ Y`. -/ def rightAdjointMate {X Y : C} [HasRightDual X] [HasRightDual Y] (f : X ⟢ Y) : Yᘁ ⟢ Xᘁ := - (ρ_ _).inv ≫ (πŸ™ _ βŠ— Ξ·_ _ _) ≫ (πŸ™ _ βŠ— f βŠ— πŸ™ _) ≫ (Ξ±_ _ _ _).inv ≫ (Ξ΅_ _ _ βŠ— πŸ™ _) ≫ (Ξ»_ _).hom + (ρ_ _).inv ≫ (_ ◁ Ξ·_ _ _) ≫ (_ ◁ f β–· _) ≫ (Ξ±_ _ _ _).inv ≫ (Ξ΅_ _ _ β–· _) ≫ (Ξ»_ _).hom #align category_theory.right_adjoint_mate CategoryTheory.rightAdjointMate /-- The left adjoint mate `ᘁf : ᘁY ⟢ ᘁX` of a morphism `f : X ⟢ Y`. -/ def leftAdjointMate {X Y : C} [HasLeftDual X] [HasLeftDual Y] (f : X ⟢ Y) : ᘁY ⟢ ᘁX := - (Ξ»_ _).inv ≫ (Ξ·_ (ᘁX) X βŠ— πŸ™ _) ≫ ((πŸ™ _ βŠ— f) βŠ— πŸ™ _) ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ _ βŠ— Ξ΅_ _ _) ≫ (ρ_ _).hom + (Ξ»_ _).inv ≫ (Ξ·_ (ᘁX) X β–· _) ≫ ((_ ◁ f) β–· _) ≫ (Ξ±_ _ _ _).hom ≫ (_ ◁ Ξ΅_ _ _) ≫ (ρ_ _).hom #align category_theory.left_adjoint_mate CategoryTheory.leftAdjointMate notation f "ᘁ" => rightAdjointMate f @@ -180,25 +192,24 @@ notation "ᘁ" f => leftAdjointMate f @[simp] theorem rightAdjointMate_id {X : C} [HasRightDual X] : (πŸ™ X)ᘁ = πŸ™ (Xᘁ) := by - simp only [rightAdjointMate, MonoidalCategory.tensor_id, Category.id_comp, - coevaluation_evaluation_assoc, Category.comp_id, Iso.inv_hom_id] + simp [rightAdjointMate] #align category_theory.right_adjoint_mate_id CategoryTheory.rightAdjointMate_id @[simp] theorem leftAdjointMate_id {X : C} [HasLeftDual X] : (ᘁ(πŸ™ X)) = πŸ™ (ᘁX) := by - simp only [leftAdjointMate, MonoidalCategory.tensor_id, Category.id_comp, - evaluation_coevaluation_assoc, Category.comp_id, Iso.inv_hom_id] + simp [leftAdjointMate] #align category_theory.left_adjoint_mate_id CategoryTheory.leftAdjointMate_id theorem rightAdjointMate_comp {X Y Z : C} [HasRightDual X] [HasRightDual Y] {f : X ⟢ Y} {g : Xᘁ ⟢ Z} : fᘁ ≫ g = (ρ_ (Yᘁ)).inv ≫ - (πŸ™ _ βŠ— Ξ·_ X (Xᘁ)) ≫ (πŸ™ _ βŠ— f βŠ— g) ≫ (Ξ±_ (Yᘁ) Y Z).inv ≫ (Ξ΅_ Y (Yᘁ) βŠ— πŸ™ _) ≫ (Ξ»_ Z).hom := by + (_ ◁ Ξ·_ X (Xᘁ)) ≫ (_ ◁ (f βŠ— g)) ≫ (Ξ±_ (Yᘁ) Y Z).inv ≫ Ξ΅_ Y (Yᘁ) β–· _ ≫ (Ξ»_ Z).hom := by dsimp only [rightAdjointMate] + simp_rw [← id_tensorHom, ← tensorHom_id] rw [Category.assoc, Category.assoc, associator_inv_naturality_assoc, associator_inv_naturality_assoc, ← tensor_id_comp_id_tensor g, Category.assoc, Category.assoc, - Category.assoc, Category.assoc, id_tensor_comp_tensor_id_assoc, ← leftUnitor_naturality, + Category.assoc, Category.assoc, id_tensor_comp_tensor_id_assoc, ← leftUnitor_naturality', tensor_id_comp_id_tensor_assoc] #align category_theory.right_adjoint_mate_comp CategoryTheory.rightAdjointMate_comp @@ -206,11 +217,12 @@ theorem leftAdjointMate_comp {X Y Z : C} [HasLeftDual X] [HasLeftDual Y] {f : X {g : (ᘁX) ⟢ Z} : (ᘁf) ≫ g = (Ξ»_ _).inv ≫ - (Ξ·_ (ᘁX) X βŠ— πŸ™ _) ≫ ((g βŠ— f) βŠ— πŸ™ _) ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ _ βŠ— Ξ΅_ _ _) ≫ (ρ_ _).hom := by + (Ξ·_ (ᘁX) X β–· _) ≫ ((g βŠ— f) β–· _) ≫ (Ξ±_ _ _ _).hom ≫ (_ ◁ Ξ΅_ _ _) ≫ (ρ_ _).hom := by dsimp only [leftAdjointMate] + simp_rw [← id_tensorHom, ← tensorHom_id] rw [Category.assoc, Category.assoc, associator_naturality_assoc, associator_naturality_assoc, ← id_tensor_comp_tensor_id _ g, Category.assoc, Category.assoc, Category.assoc, Category.assoc, - tensor_id_comp_id_tensor_assoc, ← rightUnitor_naturality, id_tensor_comp_tensor_id_assoc] + tensor_id_comp_id_tensor_assoc, ← rightUnitor_naturality', id_tensor_comp_tensor_id_assoc] #align category_theory.left_adjoint_mate_comp CategoryTheory.leftAdjointMate_comp /-- The composition of right adjoint mates is the adjoint mate of the composition. -/ @@ -218,6 +230,7 @@ theorem leftAdjointMate_comp {X Y Z : C} [HasLeftDual X] [HasLeftDual Y] {f : X theorem comp_rightAdjointMate {X Y Z : C} [HasRightDual X] [HasRightDual Y] [HasRightDual Z] {f : X ⟢ Y} {g : Y ⟢ Z} : (f ≫ g)ᘁ = gᘁ ≫ fᘁ := by rw [rightAdjointMate_comp] + simp_rw [rightAdjointMate, ← id_tensorHom, ← tensorHom_id] simp only [rightAdjointMate, comp_tensor_id, Iso.cancel_iso_inv_left, id_tensor_comp, Category.assoc] symm @@ -235,20 +248,16 @@ theorem comp_rightAdjointMate {X Y Z : C} [HasRightDual X] [HasRightDual Y] [Has rw [← @id_tensor_comp C] congr 1 rw [← id_tensor_comp_tensor_id (Ξ»_ (Xᘁ)).hom g, id_tensor_rightUnitor_inv, Category.assoc, - Category.assoc, rightUnitor_inv_naturality_assoc, ← associator_naturality_assoc, tensor_id, + Category.assoc, rightUnitor_inv_naturality'_assoc, ← associator_naturality_assoc, tensor_id, tensor_id_comp_id_tensor_assoc, ← associator_naturality_assoc] slice_rhs 2 3 => rw [← tensor_comp, tensor_id, Category.comp_id, ← Category.id_comp (Ξ·_ Y (Yᘁ)), tensor_comp] rw [← id_tensor_comp_tensor_id _ (Ξ·_ Y (Yᘁ)), ← tensor_id] - repeat' rw [@Category.assoc C] - rw [pentagon_hom_inv_assoc, ← associator_naturality_assoc, associator_inv_naturality_assoc] - slice_rhs 5 7 => rw [← comp_tensor_id, ← comp_tensor_id, evaluation_coevaluation, comp_tensor_id] - rw [associator_inv_naturality_assoc] - slice_rhs 4 5 => rw [← tensor_comp, leftUnitor_naturality, tensor_comp] - repeat' rw [@Category.assoc C] - rw [triangle_assoc_comp_right_inv_assoc, ← leftUnitor_tensor_assoc, leftUnitor_naturality_assoc, - unitors_equal, ← Category.assoc, ← Category.assoc] - simp + simp only [tensorHom_id, id_tensorHom, id_whiskerLeft, id_whiskerRight, whiskerRight_tensor, Category.assoc, + Iso.inv_hom_id_assoc, pentagon_hom_hom_inv_hom_hom] + rw [← associator_naturality_middle_assoc] + simp_rw [← comp_whiskerRight_assoc, evaluation_coevaluation] + coherence #align category_theory.comp_right_adjoint_mate CategoryTheory.comp_rightAdjointMate /-- The composition of left adjoint mates is the adjoint mate of the composition. -/ @@ -256,6 +265,7 @@ theorem comp_rightAdjointMate {X Y Z : C} [HasRightDual X] [HasRightDual Y] [Has theorem comp_leftAdjointMate {X Y Z : C} [HasLeftDual X] [HasLeftDual Y] [HasLeftDual Z] {f : X ⟢ Y} {g : Y ⟢ Z} : (ᘁf ≫ g) = (ᘁg) ≫ ᘁf := by rw [leftAdjointMate_comp] + simp_rw [leftAdjointMate, ← id_tensorHom, ← tensorHom_id] simp only [leftAdjointMate, id_tensor_comp, Iso.cancel_iso_inv_left, comp_tensor_id, Category.assoc] symm @@ -273,20 +283,15 @@ theorem comp_leftAdjointMate {X Y Z : C} [HasLeftDual X] [HasLeftDual Y] [HasLef rw [← @comp_tensor_id C] congr 1 rw [← tensor_id_comp_id_tensor g (ρ_ (ᘁX)).hom, leftUnitor_inv_tensor_id, Category.assoc, - Category.assoc, leftUnitor_inv_naturality_assoc, ← associator_inv_naturality_assoc, tensor_id, + Category.assoc, leftUnitor_inv_naturality'_assoc, ← associator_inv_naturality_assoc, tensor_id, id_tensor_comp_tensor_id_assoc, ← associator_inv_naturality_assoc] slice_rhs 2 3 => rw [← tensor_comp, tensor_id, Category.comp_id, ← Category.id_comp (Ξ·_ (ᘁY) Y), tensor_comp] rw [← tensor_id_comp_id_tensor (Ξ·_ (ᘁY) Y), ← tensor_id] repeat' rw [@Category.assoc C] rw [pentagon_inv_hom_assoc, ← associator_inv_naturality_assoc, associator_naturality_assoc] - slice_rhs 5 7 => rw [← id_tensor_comp, ← id_tensor_comp, coevaluation_evaluation, id_tensor_comp] - rw [associator_naturality_assoc] - slice_rhs 4 5 => rw [← tensor_comp, rightUnitor_naturality, tensor_comp] - repeat' rw [@Category.assoc C] - rw [triangle_assoc_comp_left_inv_assoc, ← rightUnitor_tensor_assoc, - rightUnitor_naturality_assoc, ← unitors_equal, ← Category.assoc, ← Category.assoc] - simp + slice_rhs 5 7 => rw [← id_tensor_comp, ← id_tensor_comp, coevaluation_evaluation''] + coherence #align category_theory.comp_left_adjoint_mate CategoryTheory.comp_leftAdjointMate /-- Given an exact pairing on `Y Y'`, @@ -299,10 +304,11 @@ This adjunction is often referred to as "Frobenius reciprocity" in the fusion categories / planar algebras / subfactors literature. -/ def tensorLeftHomEquiv (X Y Y' Z : C) [ExactPairing Y Y'] : (Y' βŠ— X ⟢ Z) ≃ (X ⟢ Y βŠ— Z) where - toFun f := (Ξ»_ _).inv ≫ (Ξ·_ _ _ βŠ— πŸ™ _) ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ _ βŠ— f) - invFun f := (πŸ™ Y' βŠ— f) ≫ (Ξ±_ _ _ _).inv ≫ (Ξ΅_ _ _ βŠ— πŸ™ _) ≫ (Ξ»_ _).hom + toFun f := (Ξ»_ _).inv ≫ (Ξ·_ _ _ β–· _) ≫ (Ξ±_ _ _ _).hom ≫ (_ ◁ f) + invFun f := (Y' ◁ f) ≫ (Ξ±_ _ _ _).inv ≫ (Ξ΅_ _ _ β–· _) ≫ (Ξ»_ _).hom left_inv f := by dsimp + simp_rw [← id_tensorHom, ← tensorHom_id] simp only [id_tensor_comp] slice_lhs 4 5 => rw [associator_inv_naturality] slice_lhs 5 6 => rw [tensor_id, id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] @@ -311,13 +317,14 @@ def tensorLeftHomEquiv (X Y Y' Z : C) [ExactPairing Y Y'] : (Y' βŠ— X ⟢ Z) ≃ (Ξ±_ Y' (Y βŠ— Y') X).hom ≫ (πŸ™ Y' βŠ— (Ξ±_ Y Y' X).hom) ≫ (Ξ±_ Y' Y (Y' βŠ— X)).inv ≫ (Ξ±_ (Y' βŠ— Y) Y' X).inv = (Ξ±_ _ _ _).inv βŠ— πŸ™ _ - pure_coherence + coherence slice_lhs 4 7 => rw [c] - slice_lhs 3 5 => rw [← comp_tensor_id, ← comp_tensor_id, coevaluation_evaluation] + slice_lhs 3 5 => rw [← comp_tensor_id, ← comp_tensor_id, coevaluation_evaluation''] simp only [leftUnitor_conjugation] coherence right_inv f := by dsimp + simp_rw [← id_tensorHom, ← tensorHom_id] simp only [id_tensor_comp] slice_lhs 3 4 => rw [← associator_naturality] slice_lhs 2 3 => rw [tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] @@ -326,9 +333,9 @@ def tensorLeftHomEquiv (X Y Y' Z : C) [ExactPairing Y Y'] : (Y' βŠ— X ⟢ Z) ≃ (Ξ±_ (Y βŠ— Y') Y Z).hom ≫ (Ξ±_ Y Y' (Y βŠ— Z)).hom ≫ (πŸ™ Y βŠ— (Ξ±_ Y' Y Z).inv) ≫ (Ξ±_ Y (Y' βŠ— Y) Z).inv = (Ξ±_ _ _ _).hom βŠ— πŸ™ Z - pure_coherence + coherence slice_lhs 5 8 => rw [c] - slice_lhs 4 6 => rw [← comp_tensor_id, ← comp_tensor_id, evaluation_coevaluation] + slice_lhs 4 6 => rw [← comp_tensor_id, ← comp_tensor_id, evaluation_coevaluation''] simp only [leftUnitor_conjugation] coherence #align category_theory.tensor_left_hom_equiv CategoryTheory.tensorLeftHomEquiv @@ -338,10 +345,11 @@ we get a bijection on hom-sets `(X βŠ— Y ⟢ Z) ≃ (X ⟢ Z βŠ— Y')` by "pulling the string on the right" up or down. -/ def tensorRightHomEquiv (X Y Y' Z : C) [ExactPairing Y Y'] : (X βŠ— Y ⟢ Z) ≃ (X ⟢ Z βŠ— Y') where - toFun f := (ρ_ _).inv ≫ (πŸ™ _ βŠ— Ξ·_ _ _) ≫ (Ξ±_ _ _ _).inv ≫ (f βŠ— πŸ™ _) - invFun f := (f βŠ— πŸ™ _) ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ _ βŠ— Ξ΅_ _ _) ≫ (ρ_ _).hom + toFun f := (ρ_ _).inv ≫ (_ ◁ Ξ·_ _ _) ≫ (Ξ±_ _ _ _).inv ≫ (f β–· _) + invFun f := (f β–· _) ≫ (Ξ±_ _ _ _).hom ≫ (_ ◁ Ξ΅_ _ _) ≫ (ρ_ _).hom left_inv f := by dsimp + simp_rw [← id_tensorHom, ← tensorHom_id] simp only [comp_tensor_id] slice_lhs 4 5 => rw [associator_naturality] slice_lhs 5 6 => rw [tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] @@ -350,13 +358,14 @@ def tensorRightHomEquiv (X Y Y' Z : C) [ExactPairing Y Y'] : (X βŠ— Y ⟢ Z) ≃ (Ξ±_ X (Y βŠ— Y') Y).inv ≫ ((Ξ±_ X Y Y').inv βŠ— πŸ™ Y) ≫ (Ξ±_ (X βŠ— Y) Y' Y).hom ≫ (Ξ±_ X Y (Y' βŠ— Y)).hom = πŸ™ _ βŠ— (Ξ±_ _ _ _).hom - pure_coherence + coherence slice_lhs 4 7 => rw [c] - slice_lhs 3 5 => rw [← id_tensor_comp, ← id_tensor_comp, evaluation_coevaluation] + slice_lhs 3 5 => rw [← id_tensor_comp, ← id_tensor_comp, evaluation_coevaluation''] simp only [rightUnitor_conjugation] coherence right_inv f := by dsimp + simp_rw [← id_tensorHom, ← tensorHom_id] simp only [comp_tensor_id] slice_lhs 3 4 => rw [← associator_inv_naturality] slice_lhs 2 3 => rw [tensor_id, id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] @@ -365,41 +374,37 @@ def tensorRightHomEquiv (X Y Y' Z : C) [ExactPairing Y Y'] : (X βŠ— Y ⟢ Z) ≃ (Ξ±_ Z Y' (Y βŠ— Y')).inv ≫ (Ξ±_ (Z βŠ— Y') Y Y').inv ≫ ((Ξ±_ Z Y' Y).hom βŠ— πŸ™ Y') ≫ (Ξ±_ Z (Y' βŠ— Y) Y').hom = πŸ™ _ βŠ— (Ξ±_ _ _ _).inv - pure_coherence + coherence slice_lhs 5 8 => rw [c] - slice_lhs 4 6 => rw [← id_tensor_comp, ← id_tensor_comp, coevaluation_evaluation] + slice_lhs 4 6 => rw [← id_tensor_comp, ← id_tensor_comp, coevaluation_evaluation''] simp only [rightUnitor_conjugation] coherence #align category_theory.tensor_right_hom_equiv CategoryTheory.tensorRightHomEquiv theorem tensorLeftHomEquiv_naturality {X Y Y' Z Z' : C} [ExactPairing Y Y'] (f : Y' βŠ— X ⟢ Z) (g : Z ⟢ Z') : - (tensorLeftHomEquiv X Y Y' Z') (f ≫ g) = (tensorLeftHomEquiv X Y Y' Z) f ≫ (πŸ™ Y βŠ— g) := by - dsimp [tensorLeftHomEquiv] - simp only [id_tensor_comp, Category.assoc] + (tensorLeftHomEquiv X Y Y' Z') (f ≫ g) = (tensorLeftHomEquiv X Y Y' Z) f ≫ (Y ◁ g) := by + simp [tensorLeftHomEquiv] #align category_theory.tensor_left_hom_equiv_naturality CategoryTheory.tensorLeftHomEquiv_naturality theorem tensorLeftHomEquiv_symm_naturality {X X' Y Y' Z : C} [ExactPairing Y Y'] (f : X ⟢ X') (g : X' ⟢ Y βŠ— Z) : (tensorLeftHomEquiv X Y Y' Z).symm (f ≫ g) = - (πŸ™ _ βŠ— f) ≫ (tensorLeftHomEquiv X' Y Y' Z).symm g := by - dsimp [tensorLeftHomEquiv] - simp only [id_tensor_comp, Category.assoc] + (_ ◁ f) ≫ (tensorLeftHomEquiv X' Y Y' Z).symm g := by + simp [tensorLeftHomEquiv] #align category_theory.tensor_left_hom_equiv_symm_naturality CategoryTheory.tensorLeftHomEquiv_symm_naturality theorem tensorRightHomEquiv_naturality {X Y Y' Z Z' : C} [ExactPairing Y Y'] (f : X βŠ— Y ⟢ Z) (g : Z ⟢ Z') : - (tensorRightHomEquiv X Y Y' Z') (f ≫ g) = (tensorRightHomEquiv X Y Y' Z) f ≫ (g βŠ— πŸ™ Y') := by - dsimp [tensorRightHomEquiv] - simp only [comp_tensor_id, Category.assoc] + (tensorRightHomEquiv X Y Y' Z') (f ≫ g) = (tensorRightHomEquiv X Y Y' Z) f ≫ (g β–· Y') := by + simp [tensorRightHomEquiv] #align category_theory.tensor_right_hom_equiv_naturality CategoryTheory.tensorRightHomEquiv_naturality theorem tensorRightHomEquiv_symm_naturality {X X' Y Y' Z : C} [ExactPairing Y Y'] (f : X ⟢ X') (g : X' ⟢ Z βŠ— Y') : (tensorRightHomEquiv X Y Y' Z).symm (f ≫ g) = - (f βŠ— πŸ™ Y) ≫ (tensorRightHomEquiv X' Y Y' Z).symm g := by - dsimp [tensorRightHomEquiv] - simp only [comp_tensor_id, Category.assoc] + (f β–· Y) ≫ (tensorRightHomEquiv X' Y Y' Z).symm g := by + simp [tensorRightHomEquiv] #align category_theory.tensor_right_hom_equiv_symm_naturality CategoryTheory.tensorRightHomEquiv_symm_naturality /-- If `Y Y'` have an exact pairing, @@ -442,12 +447,7 @@ theorem tensorLeftHomEquiv_tensor {X X' Y Y' Z Z' : C} [ExactPairing Y Y'] (f : (g : X' ⟢ Z') : (tensorLeftHomEquiv (X βŠ— X') Y Y' (Z βŠ— Z')).symm ((f βŠ— g) ≫ (Ξ±_ _ _ _).hom) = (Ξ±_ _ _ _).inv ≫ ((tensorLeftHomEquiv X Y Y' Z).symm f βŠ— g) := by - dsimp [tensorLeftHomEquiv] - simp only [id_tensor_comp] - simp only [associator_inv_conjugation] - slice_lhs 2 2 => rw [← id_tensor_comp_tensor_id] - conv_rhs => rw [← id_tensor_comp_tensor_id, comp_tensor_id, comp_tensor_id] - simp; coherence + simp [tensorLeftHomEquiv, tensorHom_def] #align category_theory.tensor_left_hom_equiv_tensor CategoryTheory.tensorLeftHomEquiv_tensor /-- `tensorRightHomEquiv` commutes with tensoring on the left -/ @@ -456,29 +456,31 @@ theorem tensorRightHomEquiv_tensor {X X' Y Y' Z Z' : C} [ExactPairing Y Y'] (f : (tensorRightHomEquiv (X' βŠ— X) Y Y' (Z' βŠ— Z)).symm ((g βŠ— f) ≫ (Ξ±_ _ _ _).inv) = (Ξ±_ _ _ _).hom ≫ (g βŠ— (tensorRightHomEquiv X Y Y' Z).symm f) := by dsimp [tensorRightHomEquiv] + simp_rw [← id_tensorHom, ← tensorHom_id] simp only [comp_tensor_id] simp only [associator_conjugation] slice_lhs 2 2 => rw [← tensor_id_comp_id_tensor] conv_rhs => rw [← tensor_id_comp_id_tensor, id_tensor_comp, id_tensor_comp] simp only [← tensor_id, associator_conjugation] - simp; coherence + simp #align category_theory.tensor_right_hom_equiv_tensor CategoryTheory.tensorRightHomEquiv_tensor @[simp] theorem tensorLeftHomEquiv_symm_coevaluation_comp_id_tensor {Y Y' Z : C} [ExactPairing Y Y'] - (f : Y' ⟢ Z) : (tensorLeftHomEquiv _ _ _ _).symm (Ξ·_ _ _ ≫ (πŸ™ Y βŠ— f)) = (ρ_ _).hom ≫ f := by + (f : Y' ⟢ Z) : (tensorLeftHomEquiv _ _ _ _).symm (Ξ·_ _ _ ≫ (Y ◁ f)) = (ρ_ _).hom ≫ f := by dsimp [tensorLeftHomEquiv] + simp_rw [← id_tensorHom, ← tensorHom_id] rw [id_tensor_comp] slice_lhs 2 3 => rw [associator_inv_naturality] slice_lhs 3 4 => rw [tensor_id, id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_lhs 1 3 => rw [coevaluation_evaluation] + slice_lhs 1 3 => rw [coevaluation_evaluation''] simp #align category_theory.tensor_left_hom_equiv_symm_coevaluation_comp_id_tensor CategoryTheory.tensorLeftHomEquiv_symm_coevaluation_comp_id_tensor @[simp] theorem tensorLeftHomEquiv_symm_coevaluation_comp_tensor_id {X Y : C} [HasRightDual X] [HasRightDual Y] (f : X ⟢ Y) : - (tensorLeftHomEquiv _ _ _ _).symm (Ξ·_ _ _ ≫ (f βŠ— πŸ™ (Xᘁ))) = (ρ_ _).hom ≫ fᘁ := by + (tensorLeftHomEquiv _ _ _ _).symm (Ξ·_ _ _ ≫ (f β–· (Xᘁ))) = (ρ_ _).hom ≫ fᘁ := by dsimp [tensorLeftHomEquiv, rightAdjointMate] simp #align category_theory.tensor_left_hom_equiv_symm_coevaluation_comp_tensor_id CategoryTheory.tensorLeftHomEquiv_symm_coevaluation_comp_tensor_id @@ -486,99 +488,104 @@ theorem tensorLeftHomEquiv_symm_coevaluation_comp_tensor_id {X Y : C} [HasRightD @[simp] theorem tensorRightHomEquiv_symm_coevaluation_comp_id_tensor {X Y : C} [HasLeftDual X] [HasLeftDual Y] (f : X ⟢ Y) : - (tensorRightHomEquiv _ (ᘁY) _ _).symm (Ξ·_ (ᘁX) X ≫ (πŸ™ (ᘁX) βŠ— f)) = (Ξ»_ _).hom ≫ ᘁf := by + (tensorRightHomEquiv _ (ᘁY) _ _).symm (Ξ·_ (ᘁX) X ≫ ((ᘁX) ◁ f)) = (Ξ»_ _).hom ≫ ᘁf := by dsimp [tensorRightHomEquiv, leftAdjointMate] simp #align category_theory.tensor_right_hom_equiv_symm_coevaluation_comp_id_tensor CategoryTheory.tensorRightHomEquiv_symm_coevaluation_comp_id_tensor @[simp] theorem tensorRightHomEquiv_symm_coevaluation_comp_tensor_id {Y Y' Z : C} [ExactPairing Y Y'] - (f : Y ⟢ Z) : (tensorRightHomEquiv _ Y _ _).symm (Ξ·_ Y Y' ≫ (f βŠ— πŸ™ Y')) = (Ξ»_ _).hom ≫ f := by + (f : Y ⟢ Z) : (tensorRightHomEquiv _ Y _ _).symm (Ξ·_ Y Y' ≫ (f β–· Y')) = (Ξ»_ _).hom ≫ f := by dsimp [tensorRightHomEquiv] + simp_rw [← id_tensorHom, ← tensorHom_id] rw [comp_tensor_id] slice_lhs 2 3 => rw [associator_naturality] slice_lhs 3 4 => rw [tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] - slice_lhs 1 3 => rw [evaluation_coevaluation] + slice_lhs 1 3 => rw [evaluation_coevaluation''] simp #align category_theory.tensor_right_hom_equiv_symm_coevaluation_comp_tensor_id CategoryTheory.tensorRightHomEquiv_symm_coevaluation_comp_tensor_id @[simp] theorem tensorLeftHomEquiv_id_tensor_comp_evaluation {Y Z : C} [HasLeftDual Z] (f : Y ⟢ ᘁZ) : - (tensorLeftHomEquiv _ _ _ _) ((πŸ™ Z βŠ— f) ≫ Ξ΅_ _ _) = f ≫ (ρ_ _).inv := by + (tensorLeftHomEquiv _ _ _ _) ((Z ◁ f) ≫ Ξ΅_ _ _) = f ≫ (ρ_ _).inv := by dsimp [tensorLeftHomEquiv] + simp_rw [← id_tensorHom, ← tensorHom_id] rw [id_tensor_comp] slice_lhs 3 4 => rw [← associator_naturality] slice_lhs 2 3 => rw [tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] - slice_lhs 3 5 => rw [evaluation_coevaluation] + slice_lhs 3 5 => rw [evaluation_coevaluation''] simp #align category_theory.tensor_left_hom_equiv_id_tensor_comp_evaluation CategoryTheory.tensorLeftHomEquiv_id_tensor_comp_evaluation @[simp] theorem tensorLeftHomEquiv_tensor_id_comp_evaluation {X Y : C} [HasLeftDual X] [HasLeftDual Y] - (f : X ⟢ Y) : (tensorLeftHomEquiv _ _ _ _) ((f βŠ— πŸ™ _) ≫ Ξ΅_ _ _) = (ᘁf) ≫ (ρ_ _).inv := by + (f : X ⟢ Y) : (tensorLeftHomEquiv _ _ _ _) ((f β–· _) ≫ Ξ΅_ _ _) = (ᘁf) ≫ (ρ_ _).inv := by dsimp [tensorLeftHomEquiv, leftAdjointMate] simp #align category_theory.tensor_left_hom_equiv_tensor_id_comp_evaluation CategoryTheory.tensorLeftHomEquiv_tensor_id_comp_evaluation @[simp] theorem tensorRightHomEquiv_id_tensor_comp_evaluation {X Y : C} [HasRightDual X] [HasRightDual Y] - (f : X ⟢ Y) : (tensorRightHomEquiv _ _ _ _) ((πŸ™ (Yᘁ) βŠ— f) ≫ Ξ΅_ _ _) = fᘁ ≫ (Ξ»_ _).inv := by + (f : X ⟢ Y) : (tensorRightHomEquiv _ _ _ _) (((Yᘁ) ◁ f) ≫ Ξ΅_ _ _) = fᘁ ≫ (Ξ»_ _).inv := by dsimp [tensorRightHomEquiv, rightAdjointMate] simp #align category_theory.tensor_right_hom_equiv_id_tensor_comp_evaluation CategoryTheory.tensorRightHomEquiv_id_tensor_comp_evaluation @[simp] theorem tensorRightHomEquiv_tensor_id_comp_evaluation {X Y : C} [HasRightDual X] (f : Y ⟢ Xᘁ) : - (tensorRightHomEquiv _ _ _ _) ((f βŠ— πŸ™ X) ≫ Ξ΅_ X (Xᘁ)) = f ≫ (Ξ»_ _).inv := by + (tensorRightHomEquiv _ _ _ _) ((f β–· X) ≫ Ξ΅_ X (Xᘁ)) = f ≫ (Ξ»_ _).inv := by dsimp [tensorRightHomEquiv] + simp_rw [← id_tensorHom, ← tensorHom_id] rw [comp_tensor_id] slice_lhs 3 4 => rw [← associator_inv_naturality] slice_lhs 2 3 => rw [tensor_id, id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_lhs 3 5 => rw [coevaluation_evaluation] + slice_lhs 3 5 => rw [coevaluation_evaluation''] simp #align category_theory.tensor_right_hom_equiv_tensor_id_comp_evaluation CategoryTheory.tensorRightHomEquiv_tensor_id_comp_evaluation -- Next four lemmas passing `fᘁ` or `ᘁf` through (co)evaluations. theorem coevaluation_comp_rightAdjointMate {X Y : C} [HasRightDual X] [HasRightDual Y] (f : X ⟢ Y) : - Ξ·_ Y (Yᘁ) ≫ (πŸ™ _ βŠ— fᘁ) = Ξ·_ _ _ ≫ (f βŠ— πŸ™ _) := by + Ξ·_ Y (Yᘁ) ≫ (_ ◁ fᘁ) = Ξ·_ _ _ ≫ (f β–· _) := by apply_fun (tensorLeftHomEquiv _ Y (Yᘁ) _).symm simp #align category_theory.coevaluation_comp_right_adjoint_mate CategoryTheory.coevaluation_comp_rightAdjointMate theorem leftAdjointMate_comp_evaluation {X Y : C} [HasLeftDual X] [HasLeftDual Y] (f : X ⟢ Y) : - (πŸ™ X βŠ— ᘁf) ≫ Ξ΅_ _ _ = (f βŠ— πŸ™ _) ≫ Ξ΅_ _ _ := by + (X ◁ ᘁf) ≫ Ξ΅_ _ _ = (f β–· _) ≫ Ξ΅_ _ _ := by apply_fun tensorLeftHomEquiv _ (ᘁX) X _ simp #align category_theory.left_adjoint_mate_comp_evaluation CategoryTheory.leftAdjointMate_comp_evaluation theorem coevaluation_comp_leftAdjointMate {X Y : C} [HasLeftDual X] [HasLeftDual Y] (f : X ⟢ Y) : - Ξ·_ (ᘁY) Y ≫ ((ᘁf) βŠ— πŸ™ Y) = Ξ·_ (ᘁX) X ≫ (πŸ™ (ᘁX) βŠ— f) := by + Ξ·_ (ᘁY) Y ≫ ((ᘁf) β–· Y) = Ξ·_ (ᘁX) X ≫ ((ᘁX) ◁ f) := by apply_fun (tensorRightHomEquiv _ (ᘁY) Y _).symm simp #align category_theory.coevaluation_comp_left_adjoint_mate CategoryTheory.coevaluation_comp_leftAdjointMate theorem rightAdjointMate_comp_evaluation {X Y : C} [HasRightDual X] [HasRightDual Y] (f : X ⟢ Y) : - (fᘁ βŠ— πŸ™ X) ≫ Ξ΅_ X (Xᘁ) = (πŸ™ (Yᘁ) βŠ— f) ≫ Ξ΅_ Y (Yᘁ) := by + (fᘁ β–· X) ≫ Ξ΅_ X (Xᘁ) = ((Yᘁ) ◁ f) ≫ Ξ΅_ Y (Yᘁ) := by apply_fun tensorRightHomEquiv _ X (Xᘁ) _ simp #align category_theory.right_adjoint_mate_comp_evaluation CategoryTheory.rightAdjointMate_comp_evaluation /-- Transport an exact pairing across an isomorphism in the first argument. -/ def exactPairingCongrLeft {X X' Y : C} [ExactPairing X' Y] (i : X β‰… X') : ExactPairing X Y where - evaluation' := (πŸ™ Y βŠ— i.hom) ≫ Ξ΅_ _ _ - coevaluation' := Ξ·_ _ _ ≫ (i.inv βŠ— πŸ™ Y) + evaluation' := (Y ◁ i.hom) ≫ Ξ΅_ _ _ + coevaluation' := Ξ·_ _ _ ≫ (i.inv β–· Y) evaluation_coevaluation' := by + simp_rw [← id_tensorHom, ← tensorHom_id] rw [id_tensor_comp, comp_tensor_id] slice_lhs 2 3 => rw [associator_naturality] slice_lhs 3 4 => rw [tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] slice_lhs 4 5 => rw [tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] slice_lhs 2 3 => rw [← associator_naturality] slice_lhs 1 2 => rw [tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] - slice_lhs 2 4 => rw [evaluation_coevaluation] - slice_lhs 1 2 => rw [leftUnitor_naturality] - slice_lhs 3 4 => rw [← rightUnitor_inv_naturality] + slice_lhs 2 4 => rw [evaluation_coevaluation''] + slice_lhs 1 2 => rw [leftUnitor_naturality'] + slice_lhs 3 4 => rw [← rightUnitor_inv_naturality'] simp coevaluation_evaluation' := by + simp_rw [← id_tensorHom, ← tensorHom_id] rw [id_tensor_comp, comp_tensor_id] simp only [Iso.inv_hom_id_assoc, associator_conjugation, Category.assoc] slice_lhs 2 3 => @@ -589,9 +596,10 @@ def exactPairingCongrLeft {X X' Y : C} [ExactPairing X' Y] (i : X β‰… X') : Exac /-- Transport an exact pairing across an isomorphism in the second argument. -/ def exactPairingCongrRight {X Y Y' : C} [ExactPairing X Y'] (i : Y β‰… Y') : ExactPairing X Y where - evaluation' := (i.hom βŠ— πŸ™ X) ≫ Ξ΅_ _ _ - coevaluation' := Ξ·_ _ _ ≫ (πŸ™ X βŠ— i.inv) + evaluation' := (i.hom β–· X) ≫ Ξ΅_ _ _ + coevaluation' := Ξ·_ _ _ ≫ (X ◁ i.inv) evaluation_coevaluation' := by + simp_rw [← id_tensorHom, ← tensorHom_id] rw [id_tensor_comp, comp_tensor_id] simp only [Iso.inv_hom_id_assoc, associator_conjugation, Category.assoc] slice_lhs 3 4 => @@ -599,15 +607,16 @@ def exactPairingCongrRight {X Y Y' : C} [ExactPairing X Y'] (i : Y β‰… Y') : Exa simp simp coevaluation_evaluation' := by + simp_rw [← id_tensorHom, ← tensorHom_id] rw [id_tensor_comp, comp_tensor_id] slice_lhs 3 4 => rw [← associator_inv_naturality] slice_lhs 2 3 => rw [tensor_id, id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] slice_lhs 1 2 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] slice_lhs 3 4 => rw [associator_inv_naturality] slice_lhs 4 5 => rw [tensor_id, id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_lhs 2 4 => rw [coevaluation_evaluation] - slice_lhs 1 2 => rw [rightUnitor_naturality] - slice_lhs 3 4 => rw [← leftUnitor_inv_naturality] + slice_lhs 2 4 => rw [coevaluation_evaluation''] + slice_lhs 1 2 => rw [rightUnitor_naturality'] + slice_lhs 3 4 => rw [← leftUnitor_inv_naturality'] simp #align category_theory.exact_pairing_congr_right CategoryTheory.exactPairingCongrRight diff --git a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean index 9b30b98669b33..20d9b10462933 100644 --- a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean @@ -53,24 +53,52 @@ When `P` is a monoidal predicate, the full subcategory for `P` inherits the mono -/ instance fullMonoidalSubcategory : MonoidalCategory (FullSubcategory P) where tensorObj X Y := ⟨X.1 βŠ— Y.1, prop_tensor X.2 Y.2⟩ - tensorHom := @fun X₁ Y₁ Xβ‚‚ Yβ‚‚ f g => by - change X₁.1 βŠ— Xβ‚‚.1 ⟢ Y₁.1 βŠ— Yβ‚‚.1 - change X₁.1 ⟢ Y₁.1 at f; change Xβ‚‚.1 ⟢ Yβ‚‚.1 at g; exact f βŠ— g + whiskerLeft := fun X _ _ f ↦ X.1 ◁ f + whiskerRight := fun f Y ↦ (fun f ↦ f β–· Y.1) f tensorUnit' := βŸ¨πŸ™_ C, prop_id⟩ - associator X Y Z := - ⟨(Ξ±_ X.1 Y.1 Z.1).hom, (Ξ±_ X.1 Y.1 Z.1).inv, hom_inv_id (Ξ±_ X.1 Y.1 Z.1), - inv_hom_id (Ξ±_ X.1 Y.1 Z.1)⟩ leftUnitor X := ⟨(Ξ»_ X.1).hom, (Ξ»_ X.1).inv, hom_inv_id (Ξ»_ X.1), inv_hom_id (Ξ»_ X.1)⟩ rightUnitor X := ⟨(ρ_ X.1).hom, (ρ_ X.1).inv, hom_inv_id (ρ_ X.1), inv_hom_id (ρ_ X.1)⟩ - tensor_id X Y := tensor_id X.1 Y.1 - tensor_comp f₁ fβ‚‚ g₁ gβ‚‚ := @tensor_comp C _ _ _ _ _ _ _ _ f₁ fβ‚‚ g₁ gβ‚‚ - associator_naturality f₁ fβ‚‚ f₃ := @associator_naturality C _ _ _ _ _ _ _ _ f₁ fβ‚‚ f₃ - leftUnitor_naturality f := @leftUnitor_naturality C _ _ _ _ f - rightUnitor_naturality f := @rightUnitor_naturality C _ _ _ _ f + associator X Y Z := + ⟨(Ξ±_ X.1 Y.1 Z.1).hom, (Ξ±_ X.1 Y.1 Z.1).inv, hom_inv_id (Ξ±_ X.1 Y.1 Z.1), + inv_hom_id (Ξ±_ X.1 Y.1 Z.1)⟩ + whiskerLeft_id X Y := whiskerLeft_id X.1 Y.1 + whiskerLeft_comp _ _ _ _ f g := whiskerLeft_comp (C := C) _ f g + id_whiskerLeft f := id_whiskerLeft (C := C) f + tensor_whiskerLeft _ _ _ _ f := tensor_whiskerLeft (C := C) _ _ f + id_whiskerRight X Y := id_whiskerRight X.1 Y.1 + comp_whiskerRight f g _ := comp_whiskerRight (C := C) f g _ + whiskerRight_id f := whiskerRight_id (C := C) f + whiskerRight_tensor f _ _ := whiskerRight_tensor (C := C) f _ _ + whisker_assoc _ _ _ f _ := whisker_assoc (C := C) _ f _ + whisker_exchange f g := whisker_exchange (C := C) f g pentagon W X Y Z := pentagon W.1 X.1 Y.1 Z.1 triangle X Y := triangle X.1 Y.1 #align category_theory.monoidal_category.full_monoidal_subcategory CategoryTheory.MonoidalCategory.fullMonoidalSubcategory +-- /-- +-- When `P` is a monoidal predicate, the full subcategory for `P` inherits the monoidal structure of +-- `C`. +-- -/ +-- instance fullMonoidalSubcategory''' : MonoidalCategory (FullSubcategory P) where +-- tensorObj X Y := ⟨X.1 βŠ— Y.1, prop_tensor X.2 Y.2⟩ +-- whiskerLeft := @fun X Y₁ Yβ‚‚ f => by +-- change X₁.1 βŠ— Xβ‚‚.1 ⟢ Y₁.1 βŠ— Yβ‚‚.1 +-- change X₁.1 ⟢ Y₁.1 at f; change Xβ‚‚.1 ⟢ Yβ‚‚.1 at g; exact f βŠ— g +-- tensorUnit' := βŸ¨πŸ™_ C, prop_id⟩ +-- associator X Y Z := +-- ⟨(Ξ±_ X.1 Y.1 Z.1).hom, (Ξ±_ X.1 Y.1 Z.1).inv, hom_inv_id (Ξ±_ X.1 Y.1 Z.1), +-- inv_hom_id (Ξ±_ X.1 Y.1 Z.1)⟩ +-- leftUnitor X := ⟨(Ξ»_ X.1).hom, (Ξ»_ X.1).inv, hom_inv_id (Ξ»_ X.1), inv_hom_id (Ξ»_ X.1)⟩ +-- rightUnitor X := ⟨(ρ_ X.1).hom, (ρ_ X.1).inv, hom_inv_id (ρ_ X.1), inv_hom_id (ρ_ X.1)⟩ +-- tensor_id X Y := tensor_id X.1 Y.1 +-- tensor_comp f₁ fβ‚‚ g₁ gβ‚‚ := @tensor_comp C _ _ _ _ _ _ _ _ f₁ fβ‚‚ g₁ gβ‚‚ +-- associator_naturality f₁ fβ‚‚ f₃ := @associator_naturality C _ _ _ _ _ _ _ _ f₁ fβ‚‚ f₃ +-- leftUnitor_naturality f := @leftUnitor_naturality C _ _ _ _ f +-- rightUnitor_naturality f := @rightUnitor_naturality C _ _ _ _ f +-- pentagon W X Y Z := pentagon W.1 X.1 Y.1 Z.1 +-- triangle X Y := triangle X.1 Y.1 +-- #align category_theory.monoidal_category.full_monoidal_subcategory CategoryTheory.MonoidalCategory.fullMonoidalSubcategory + /-- The forgetful monoidal functor from a full monoidal subcategory into the original category ("forgetting" the condition). -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index 36c509c98a257..ec34ec0ff8e5c 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -35,25 +35,108 @@ variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] variable {D : Type uβ‚‚} [Category.{vβ‚‚} D] +-- -- porting note: it was @[simps {attrs := [`_refl_lemma]}] +-- /-- Transport a monoidal structure along an equivalence of (plain) categories. +-- -/ +-- @[simps] +-- def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D where +-- tensorObj X Y := e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y) +-- -- tensorHom f g := e.functor.map (e.inverse.map f βŠ— e.inverse.map g) +-- whiskerLeft X _ _ f := e.functor.map (e.inverse.obj X ◁ e.inverse.map f) +-- whiskerRight f Y := e.functor.map (e.inverse.map f β–· e.inverse.obj Y) +-- -- tensorHom_def := _ +-- tensorUnit' := e.functor.obj (πŸ™_ C) +-- associator X Y Z := +-- e.functor.mapIso +-- (((e.unitIso.app _).symm βŠ— Iso.refl _) β‰ͺ≫ +-- Ξ±_ (e.inverse.obj X) (e.inverse.obj Y) (e.inverse.obj Z) β‰ͺ≫ (Iso.refl _ βŠ— e.unitIso.app _)) +-- leftUnitor X := +-- e.functor.mapIso (((e.unitIso.app _).symm βŠ— Iso.refl _) β‰ͺ≫ Ξ»_ (e.inverse.obj X)) β‰ͺ≫ +-- e.counitIso.app _ +-- rightUnitor X := +-- e.functor.mapIso ((Iso.refl _ βŠ— (e.unitIso.app _).symm) β‰ͺ≫ ρ_ (e.inverse.obj X)) β‰ͺ≫ +-- e.counitIso.app _ +-- whiskerLeft_id := _ +-- whiskerLeft_comp := _ +-- id_whiskerLeft := _ +-- tensor_whiskerLeft := _ +-- id_whiskerRight := _ +-- comp_whiskerRight := _ +-- whiskerRight_id := _ +-- whiskerRight_tensor := _ +-- whisker_assoc := _ +-- whisker_exchange := _ +-- pentagon W X Y Z := by +-- dsimp +-- simp_rw [← id_tensorHom, ← tensorHom_id] +-- simp only [Iso.hom_inv_id_app_assoc, comp_tensor_id, assoc, Equivalence.inv_fun_map, +-- Functor.map_comp, id_tensor_comp, e.inverse.map_id] +-- simp only [← e.functor.map_comp] +-- congr 2 +-- slice_lhs 4 5 => +-- rw [← comp_tensor_id, Iso.hom_inv_id_app] +-- dsimp +-- rw [tensor_id] +-- simp only [Category.id_comp, Category.assoc] +-- slice_lhs 5 6 => +-- rw [← id_tensor_comp, Iso.hom_inv_id_app] +-- dsimp +-- rw [tensor_id] +-- simp only [Category.id_comp, Category.assoc] +-- slice_rhs 2 3 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] +-- slice_rhs 1 2 => rw [← tensor_id, ← associator_naturality] +-- slice_rhs 3 4 => rw [← tensor_id, associator_naturality] +-- slice_rhs 2 3 => rw [← pentagon] +-- simp only [Category.assoc] +-- congr 2 +-- simp_rw [← id_tensorHom, ← tensorHom_id] +-- slice_lhs 1 2 => rw [associator_naturality] +-- simp only [Category.assoc] +-- congr 1 +-- slice_lhs 1 2 => +-- rw [← id_tensor_comp, ← comp_tensor_id, Iso.hom_inv_id_app] +-- dsimp +-- rw [tensor_id, tensor_id] +-- simp only [Category.id_comp, Category.assoc] +-- simp [id_tensorHom, tensorHom_id] +-- triangle X Y := by +-- dsimp +-- simp_rw [← id_tensorHom, ← tensorHom_id] +-- simp only [Iso.hom_inv_id_app_assoc, comp_tensor_id, Equivalence.unit_inverse_comp, assoc, +-- Equivalence.inv_fun_map, comp_id, Functor.map_comp, id_tensor_comp, e.inverse.map_id] +-- simp only [← e.functor.map_comp] +-- congr 2 +-- slice_lhs 2 3 => +-- rw [← id_tensor_comp] +-- simp +-- simp [id_tensorHom, tensorHom_id] + +attribute [local simp] id_tensorHom tensorHom_id + -- porting note: it was @[simps {attrs := [`_refl_lemma]}] /-- Transport a monoidal structure along an equivalence of (plain) categories. -/ -@[simps] -def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D where - tensorObj X Y := e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y) - tensorHom f g := e.functor.map (e.inverse.map f βŠ— e.inverse.map g) - tensorUnit' := e.functor.obj (πŸ™_ C) - associator X Y Z := +def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D := .ofTensorHom + (tensorObj := fun X Y ↦ e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y)) + (tensorHom := fun f g ↦ e.functor.map (e.inverse.map f βŠ— e.inverse.map g)) + (tensor_comp := by + intro X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ f₁ fβ‚‚ g₁ gβ‚‚ + dsimp + simp only [← e.functor.map_comp] + congr 1 + simp [tensorHom_def, whisker_exchange_assoc]) + (tensorUnit' := e.functor.obj (πŸ™_ C)) + (associator := fun X Y Z ↦ e.functor.mapIso (((e.unitIso.app _).symm βŠ— Iso.refl _) β‰ͺ≫ - Ξ±_ (e.inverse.obj X) (e.inverse.obj Y) (e.inverse.obj Z) β‰ͺ≫ (Iso.refl _ βŠ— e.unitIso.app _)) - leftUnitor X := + Ξ±_ (e.inverse.obj X) (e.inverse.obj Y) (e.inverse.obj Z) β‰ͺ≫ (Iso.refl _ βŠ— e.unitIso.app _))) + (leftUnitor := fun X ↦ e.functor.mapIso (((e.unitIso.app _).symm βŠ— Iso.refl _) β‰ͺ≫ Ξ»_ (e.inverse.obj X)) β‰ͺ≫ - e.counitIso.app _ - rightUnitor X := + e.counitIso.app _) + (rightUnitor := fun X ↦ e.functor.mapIso ((Iso.refl _ βŠ— (e.unitIso.app _).symm) β‰ͺ≫ ρ_ (e.inverse.obj X)) β‰ͺ≫ - e.counitIso.app _ - triangle X Y := by + e.counitIso.app _) + (triangle := fun X Y ↦ by dsimp simp only [Iso.hom_inv_id_app_assoc, comp_tensor_id, Equivalence.unit_inverse_comp, assoc, Equivalence.inv_fun_map, comp_id, Functor.map_comp, id_tensor_comp, e.inverse.map_id] @@ -62,8 +145,8 @@ def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D where slice_lhs 2 3 => rw [← id_tensor_comp] simp - rw [Category.id_comp, ← associator_naturality_assoc, triangle] - pentagon W X Y Z := by + simp [id_tensorHom, tensorHom_id]) + (pentagon := fun W X Y Z ↦ by dsimp simp only [Iso.hom_inv_id_app_assoc, comp_tensor_id, assoc, Equivalence.inv_fun_map, Functor.map_comp, id_tensor_comp, e.inverse.map_id] @@ -82,7 +165,7 @@ def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D where slice_rhs 2 3 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] slice_rhs 1 2 => rw [← tensor_id, ← associator_naturality] slice_rhs 3 4 => rw [← tensor_id, associator_naturality] - slice_rhs 2 3 => rw [← pentagon] + slice_rhs 2 3 => rw [← pentagon'] simp only [Category.assoc] congr 2 slice_lhs 1 2 => rw [associator_naturality] @@ -92,24 +175,24 @@ def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D where rw [← id_tensor_comp, ← comp_tensor_id, Iso.hom_inv_id_app] dsimp rw [tensor_id, tensor_id] - simp only [Category.id_comp, Category.assoc] - leftUnitor_naturality f := by + simp only [Category.id_comp, Category.assoc]) + (leftUnitor_naturality := fun f ↦ by dsimp simp only [Functor.map_comp, Functor.map_id, Category.assoc] erw [← e.counitIso.hom.naturality] simp only [Functor.comp_map, ← e.functor.map_comp_assoc] congr 2 rw [id_tensor_comp_tensor_id_assoc, ← tensor_id_comp_id_tensor_assoc, - leftUnitor_naturality] - rightUnitor_naturality f := by + leftUnitor_naturality']) + (rightUnitor_naturality := fun f ↦ by dsimp simp only [Functor.map_comp, Functor.map_id, Category.assoc] erw [← e.counitIso.hom.naturality] simp only [Functor.comp_map, ← e.functor.map_comp_assoc] congr 2 erw [tensor_id_comp_id_tensor_assoc, ← id_tensor_comp_tensor_id_assoc, - rightUnitor_naturality] - associator_naturality f₁ fβ‚‚ f₃ := by + rightUnitor_naturality']) + (associator_naturality := fun f₁ fβ‚‚ f₃ ↦ by dsimp simp only [Equivalence.inv_fun_map, Functor.map_comp, Category.assoc] simp only [← e.functor.map_comp] @@ -142,7 +225,7 @@ def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D where simp only [Category.id_comp, Category.assoc] conv_rhs => rw [id_tensor_comp] slice_rhs 2 3 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_rhs 1 2 => rw [id_tensor_comp_tensor_id] + slice_rhs 1 2 => rw [id_tensor_comp_tensor_id]) #align category_theory.monoidal.transport CategoryTheory.Monoidal.transport /-- A type synonym for `D`, which will carry the transported monoidal structure. -/ @@ -152,12 +235,44 @@ def Transported (_ : C β‰Œ D) := D instance (e : C β‰Œ D) : Category (Transported e) := (inferInstance : Category D) -instance (e : C β‰Œ D) : MonoidalCategory (Transported e) := +-- @[simps!] +instance transportMonoidalCategory (e : C β‰Œ D) : MonoidalCategory (Transported e) := transport e instance (e : C β‰Œ D) : Inhabited (Transported e) := βŸ¨πŸ™_ _⟩ +theorem transport_tensorUnit' (e : C β‰Œ D) : πŸ™_ (Transported e) = e.functor.obj (πŸ™_ C) := rfl + +theorem transport_tensorObj (e : C β‰Œ D) (X Y : Transported e) : + X βŠ— Y = e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y) := + rfl + +theorem transport_tensorHom (e : C β‰Œ D) {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Transported e} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : + f βŠ— g = e.functor.map (e.inverse.map f βŠ— e.inverse.map g) := by + rfl + +theorem transport_associator (e : C β‰Œ D) (X Y Z : Transported e) : + Ξ±_ X Y Z = + e.functor.mapIso + (((e.unitIso.app (e.inverse.obj X βŠ— e.inverse.obj Y)).symm βŠ— + Iso.refl (e.inverse.obj Z)) β‰ͺ≫ + Ξ±_ (e.inverse.obj X) (e.inverse.obj Y) (e.inverse.obj Z) β‰ͺ≫ + (Iso.refl (e.inverse.obj X) βŠ— e.unitIso.app (e.inverse.obj Y βŠ— e.inverse.obj Z))) := + rfl + +theorem transport_leftUnitor (e : C β‰Œ D) (X : Transported e) : + Ξ»_ X = + e.functor.mapIso (((e.unitIso.app (πŸ™_ C)).symm βŠ— Iso.refl (e.inverse.obj X)) β‰ͺ≫ + Ξ»_ (e.inverse.obj X)) β‰ͺ≫ e.counitIso.app X := + rfl + +theorem transport_rightUnitor (e : C β‰Œ D) (X : Transported e) : + ρ_ X = + e.functor.mapIso ((Iso.refl (e.inverse.obj X) βŠ— (e.unitIso.app (πŸ™_ C)).symm) β‰ͺ≫ + ρ_ (e.inverse.obj X)) β‰ͺ≫ e.counitIso.app X := + rfl + section attribute [local simp] transport_tensorUnit' @@ -165,24 +280,25 @@ attribute [local simp] transport_tensorUnit' section attribute [local simp] - transport_tensorHom transport_associator transport_leftUnitor transport_rightUnitor + transport_tensorObj transport_tensorHom transport_associator + transport_leftUnitor transport_rightUnitor /-- We can upgrade `e.functor` to a lax monoidal functor from `C` to `D` with the transported structure. -/ -@[simps] -def laxToTransported (e : C β‰Œ D) : LaxMonoidalFunctor C (Transported e) where - toFunctor := e.functor - Ξ΅ := πŸ™ (e.functor.obj (πŸ™_ C)) - ΞΌ X Y := e.functor.map (e.unitInv.app X βŠ— e.unitInv.app Y) - ΞΌ_natural f g := by +@[simp] +def laxToTransported (e : C β‰Œ D) : LaxMonoidalFunctor C (Transported e) := .ofTensorHom + (F := e.functor) + (Ξ΅ := πŸ™ (e.functor.obj (πŸ™_ C))) + (ΞΌ := fun X Y ↦ e.functor.map (e.unitInv.app X βŠ— e.unitInv.app Y)) + (ΞΌ_natural := fun f g ↦ by dsimp rw [Equivalence.inv_fun_map, Equivalence.inv_fun_map, tensor_comp, Functor.map_comp, tensor_comp, ← e.functor.map_comp, ← e.functor.map_comp, ← e.functor.map_comp, assoc, assoc, ← tensor_comp, Iso.hom_inv_id_app, Iso.hom_inv_id_app, ← tensor_comp] dsimp - rw [comp_id, comp_id] - associativity X Y Z := by + rw [comp_id, comp_id]) + (associativity := fun X Y Z ↦ by dsimp rw [Equivalence.inv_fun_map, Equivalence.inv_fun_map, Functor.map_comp, Functor.map_comp, assoc, assoc, e.inverse.map_id, e.inverse.map_id, @@ -208,21 +324,21 @@ def laxToTransported (e : C β‰Œ D) : LaxMonoidalFunctor C (Transported e) where dsimp rw [tensor_id] simp only [associator_conjugation, ←tensor_id, ←tensor_comp, Iso.inv_hom_id, - Iso.inv_hom_id_assoc, assoc, id_comp, comp_id] - left_unitality X := by + Iso.inv_hom_id_assoc, assoc, id_comp, comp_id]) + (left_unitality := fun X ↦ by dsimp rw [e.inverse.map_id, e.inverse.map_id, tensor_id, Functor.map_comp, assoc, Equivalence.counit_app_functor, ← e.functor.map_comp, ← e.functor.map_comp, - ← e.functor.map_comp, ← e.functor.map_comp, ← leftUnitor_naturality, + ← e.functor.map_comp, ← e.functor.map_comp, ← leftUnitor_naturality', ← tensor_comp_assoc, comp_id, id_comp, id_comp] - rfl - right_unitality X := by + rfl) + (right_unitality := fun X ↦ by dsimp rw [Functor.map_comp, assoc, e.inverse.map_id, e.inverse.map_id, tensor_id, Functor.map_id, id_comp, Equivalence.counit_app_functor, ← e.functor.map_comp, - ← e.functor.map_comp, ← e.functor.map_comp, ← rightUnitor_naturality, + ← e.functor.map_comp, ← e.functor.map_comp, ← rightUnitor_naturality', ← tensor_comp_assoc, id_comp, comp_id] - rfl + rfl) #align category_theory.monoidal.lax_to_transported CategoryTheory.Monoidal.laxToTransported end @@ -242,9 +358,8 @@ def toTransported (e : C β‰Œ D) : MonoidalFunctor C (Transported e) where end -instance (e : C β‰Œ D) : IsEquivalence (toTransported e).toFunctor := by - dsimp - infer_instance +instance (e : C β‰Œ D) : IsEquivalence (toTransported e).toFunctor := + inferInstanceAs (IsEquivalence e.functor) /-- We can upgrade `e.inverse` to a monoidal functor from `D` with the transported structure to `C`. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean b/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean index 6638151e9fc14..a6ed99f917dac 100644 --- a/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean +++ b/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean @@ -36,23 +36,27 @@ open MonoidalCategory /-- `(πŸ™_ C ⟢ -)` is a lax monoidal functor to `Type`. -/ noncomputable def coyonedaTensorUnit (C : Type u) [Category.{v} C] [MonoidalCategory C] : - LaxMonoidalFunctor C (Type v) := - { coyoneda.obj (op (πŸ™_ C)) with - Ξ΅ := fun _p => πŸ™ _ - ΞΌ := fun X Y p => (Ξ»_ (πŸ™_ C)).inv ≫ (p.1 βŠ— p.2) - ΞΌ_natural := by aesop_cat - associativity := fun X Y Z => by + LaxMonoidalFunctor C (Type v) := .ofTensorHom + (F := coyoneda.obj (op (πŸ™_ C))) + (Ξ΅ := fun _p => πŸ™ _) + (ΞΌ := fun X Y p => (Ξ»_ (πŸ™_ C)).inv ≫ (p.1 βŠ— p.2)) + (ΞΌ_natural := by aesop_cat) + (associativity := fun X Y Z => by ext ⟨⟨f, g⟩, h⟩; dsimp at f g h dsimp; simp only [Iso.cancel_iso_inv_left, Category.assoc] conv_lhs => rw [← Category.id_comp h, tensor_comp, Category.assoc, associator_naturality, ← - Category.assoc, unitors_inv_equal, triangle_assoc_comp_right_inv] - conv_rhs => rw [← Category.id_comp f, tensor_comp] - left_unitality := by aesop_cat - right_unitality := fun X => by + Category.assoc, unitors_inv_equal, triangle_assoc_comp_right_inv'] + conv_rhs => rw [← Category.id_comp f, tensor_comp]) + (left_unitality := by + intros + ext ⟨⟨⟩, f⟩; dsimp at f + dsimp + coherence) + (right_unitality := fun X => by ext ⟨f, ⟨⟩⟩; dsimp at f - dsimp; simp only [Category.assoc] - rw [rightUnitor_naturality, unitors_inv_equal, Iso.inv_hom_id_assoc] } + dsimp + coherence) #align category_theory.coyoneda_tensor_unit CategoryTheory.coyonedaTensorUnit end CategoryTheory diff --git a/Mathlib/CategoryTheory/Shift/Basic.lean b/Mathlib/CategoryTheory/Shift/Basic.lean index 1c2ed8eb598aa..8dd2a01e380eb 100644 --- a/Mathlib/CategoryTheory/Shift/Basic.lean +++ b/Mathlib/CategoryTheory/Shift/Basic.lean @@ -128,8 +128,13 @@ def hasShiftMk (h : ShiftMkCore C A) : HasShift C A := ⟨{ Discrete.functor h.F with Ξ΅ := h.zero.inv ΞΌ := fun m n => (h.add m.as n.as).inv - ΞΌ_natural := by - rintro ⟨X⟩ ⟨Y⟩ ⟨X'⟩ ⟨Y'⟩ ⟨⟨⟨rfl⟩⟩⟩ ⟨⟨⟨rfl⟩⟩⟩ + ΞΌ_natural_left := by + rintro ⟨X⟩ ⟨Y⟩ ⟨⟨⟨rfl⟩⟩⟩ ⟨X'⟩ + ext + dsimp + simp + ΞΌ_natural_right := by + rintro ⟨X⟩ ⟨Y⟩ ⟨X'⟩ ⟨⟨⟨rfl⟩⟩⟩ ext dsimp simp diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index 1a9117095db97..46bfa19f418f3 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -85,9 +85,17 @@ instance LiftHom_comp {X Y Z : C} [LiftObj X] [LiftObj Y] [LiftObj Z] (f : X ⟢ [LiftHom f] [LiftHom g] : LiftHom (f ≫ g) where lift := LiftHom.lift f ≫ LiftHom.lift g -instance LiftHom_tensor {W X Y Z : C} [LiftObj W] [LiftObj X] [LiftObj Y] [LiftObj Z] - (f : W ⟢ X) (g : Y ⟢ Z) [LiftHom f] [LiftHom g] : LiftHom (f βŠ— g) where - lift := LiftHom.lift f βŠ— LiftHom.lift g +-- instance LiftHom_tensor {W X Y Z : C} [LiftObj W] [LiftObj X] [LiftObj Y] [LiftObj Z] +-- (f : W ⟢ X) (g : Y ⟢ Z) [LiftHom f] [LiftHom g] : LiftHom (f βŠ— g) where +-- lift := LiftHom.lift f βŠ— LiftHom.lift g + +instance liftHom_WhiskerLeft (X : C) [LiftObj X] {Y Z : C} [LiftObj Y] [LiftObj Z] + (f : Y ⟢ Z) [LiftHom f] : LiftHom (X ◁ f) where + lift := LiftObj.lift X ◁ LiftHom.lift f + +instance liftHom_WhiskerRight {X Y : C} (f : X ⟢ Y) [LiftObj X] [LiftObj Y] [LiftHom f] + {Z : C} [LiftObj Z] : LiftHom (f β–· Z) where + lift := LiftHom.lift f β–· LiftObj.lift Z /-- A typeclass carrying a choice of monoidal structural isomorphism between two objects. @@ -108,17 +116,17 @@ instance refl (X : C) [LiftObj X] : MonoidalCoherence X X := βŸ¨πŸ™ _⟩ @[simps] instance tensor (X Y Z : C) [LiftObj X] [LiftObj Y] [LiftObj Z] [MonoidalCoherence Y Z] : MonoidalCoherence (X βŠ— Y) (X βŠ— Z) := - βŸ¨πŸ™ X βŠ— MonoidalCoherence.hom⟩ + ⟨X ◁ MonoidalCoherence.hom⟩ @[simps] instance tensor_right (X Y : C) [LiftObj X] [LiftObj Y] [MonoidalCoherence (πŸ™_ C) Y] : MonoidalCoherence X (X βŠ— Y) := - ⟨(ρ_ X).inv ≫ (πŸ™ X βŠ— MonoidalCoherence.hom)⟩ + ⟨(ρ_ X).inv ≫ (X ◁ MonoidalCoherence.hom)⟩ @[simps] instance tensor_right' (X Y : C) [LiftObj X] [LiftObj Y] [MonoidalCoherence Y (πŸ™_ C)] : MonoidalCoherence (X βŠ— Y) X := - ⟨(πŸ™ X βŠ— MonoidalCoherence.hom) ≫ (ρ_ X).hom⟩ + ⟨(X ◁ MonoidalCoherence.hom) ≫ (ρ_ X).hom⟩ @[simps] instance left (X Y : C) [LiftObj X] [LiftObj Y] [MonoidalCoherence X Y] : @@ -194,7 +202,7 @@ example {W X Y Z : C} (f : W ⟢ (X βŠ— Y) βŠ— Z) : W ⟢ X βŠ— (Y βŠ— Z) := f example {U V W X Y : C} (f : U ⟢ V βŠ— (W βŠ— X)) (g : (V βŠ— W) βŠ— X ⟢ Y) : f βŠ—β‰« g = f ≫ (Ξ±_ _ _ _).inv ≫ g := by - simp [monoidalComp] + simp [MonoidalCategory.tensorHom_def, monoidalComp] end lifting @@ -222,7 +230,7 @@ def mkProjectMapExpr (e : Expr) : TermElabM Expr := do /-- Coherence tactic for monoidal categories. -/ def monoidal_coherence (g : MVarId) : TermElabM Unit := g.withContext do withOptions (fun opts => synthInstance.maxSize.set opts - (max 256 (synthInstance.maxSize.get opts))) do + (max 512 (synthInstance.maxSize.get opts))) do -- TODO: is this `dsimp only` step necessary? It doesn't appear to be in the tests below. let (ty, _) ← dsimp (← g.getType) (← Simp.Context.ofNames [] true) let some (_, lhs, rhs) := (← whnfR ty).eq? | exception g "Not an equation of morphisms." @@ -313,7 +321,7 @@ def insertTrailingIds (g : MVarId) : MetaM MVarId := do -- Porting note: this is an ugly port, using too many `evalTactic`s. -- We can refactor later into either a `macro` (but the flow control is awkward) -- or a `MetaM` tactic. -def coherence_loop (maxSteps := 37) : TacticM Unit := +def coherence_loop (maxSteps := 47) : TacticM Unit := match maxSteps with | 0 => exception' "`coherence` tactic reached iteration limit" | maxSteps' + 1 => do @@ -342,6 +350,22 @@ def coherence_loop (maxSteps := 37) : TacticM Unit := -- and whose second terms can be identified by recursively called `coherence`. coherence_loop maxSteps' +/-- +Simp lemmas for rewriting a hom in monoical categories into a normal form. +-/ +syntax (name := monoidal_simps) "monoidal_simps" : tactic + +@[inherit_doc monoidal_simps] +elab_rules : tactic +| `(tactic| monoidal_simps) => do + evalTactic (← `(tactic| + simp only [Category.assoc, MonoidalCategory.tensor_whiskerLeft, MonoidalCategory.id_whiskerLeft, + MonoidalCategory.whiskerRight_tensor, MonoidalCategory.whiskerRight_id, + MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_id, + MonoidalCategory.comp_whiskerRight, MonoidalCategory.id_whiskerRight, MonoidalCategory.whisker_assoc, + MonoidalCategory.id_tensorHom, MonoidalCategory.tensorHom_id]; simp only [MonoidalCategory.tensorHom_def] + )) + /-- Use the coherence theorem for monoidal categories to solve equations in a monoidal equation, where the two sides only differ by replacing strings of monoidal structural morphisms @@ -364,6 +388,6 @@ elab_rules : tactic evalTactic (← `(tactic| simp only [bicategoricalComp]; simp only [monoidalComp]; - try whisker_simps + try (first | whisker_simps | monoidal_simps); )) coherence_loop From f400ce77cdef319f06dfd3717e0f0cc182d03819 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 3 Aug 2023 06:01:22 +0900 Subject: [PATCH 02/62] braided wip --- Mathlib/CategoryTheory/Monoidal/Braided.lean | 432 ++++++++++--------- Mathlib/Tactic/CategoryTheory/Coherence.lean | 3 +- test/CategoryTheory/Coherence.lean | 4 +- 3 files changed, 225 insertions(+), 214 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 60e24c5a631a8..8650f59576671 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -43,34 +43,84 @@ and also satisfies the two hexagon identities. class BraidedCategory (C : Type u) [Category.{v} C] [MonoidalCategory.{v} C] where -- braiding natural iso: braiding : βˆ€ X Y : C, X βŠ— Y β‰… Y βŠ— X - braiding_naturality : - βˆ€ {X X' Y Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y'), - (f βŠ— g) ≫ (braiding Y Y').hom = (braiding X X').hom ≫ (g βŠ— f) := by + braiding_naturality_right : + βˆ€ (X : C) {Y Z : C} (f : Y ⟢ Z), + X ◁ f ≫ (braiding X Z).hom = (braiding X Y).hom ≫ f β–· X := by aesop_cat + braiding_naturality_left : + βˆ€ {X Y : C} (f : X ⟢ Y) (Z : C), + f β–· Z ≫ (braiding Y Z).hom = (braiding X Z).hom ≫ Z ◁ f := by + aesop_cat + -- braiding_naturality : + -- βˆ€ {X X' Y Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y'), + -- (f βŠ— g) ≫ (braiding Y Y').hom = (braiding X X').hom ≫ (g βŠ— f) := by + -- aesop_cat -- hexagon identities: hexagon_forward : βˆ€ X Y Z : C, (Ξ±_ X Y Z).hom ≫ (braiding X (Y βŠ— Z)).hom ≫ (Ξ±_ Y Z X).hom = - ((braiding X Y).hom βŠ— πŸ™ Z) ≫ (Ξ±_ Y X Z).hom ≫ (πŸ™ Y βŠ— (braiding X Z).hom) := by + ((braiding X Y).hom β–· Z) ≫ (Ξ±_ Y X Z).hom ≫ (Y ◁ (braiding X Z).hom) := by aesop_cat hexagon_reverse : βˆ€ X Y Z : C, (Ξ±_ X Y Z).inv ≫ (braiding (X βŠ— Y) Z).hom ≫ (Ξ±_ Z X Y).inv = - (πŸ™ X βŠ— (braiding Y Z).hom) ≫ (Ξ±_ X Z Y).inv ≫ ((braiding X Z).hom βŠ— πŸ™ Y) := by + (X ◁ (braiding Y Z).hom) ≫ (Ξ±_ X Z Y).inv ≫ ((braiding X Z).hom β–· Y) := by aesop_cat #align category_theory.braided_category CategoryTheory.BraidedCategory -attribute [reassoc (attr := simp)] BraidedCategory.braiding_naturality +attribute [reassoc (attr := simp)] + BraidedCategory.braiding_naturality_left + BraidedCategory.braiding_naturality_right attribute [reassoc] BraidedCategory.hexagon_forward BraidedCategory.hexagon_reverse +notation "Ξ²_" => BraidedCategory.braiding + +namespace BraidedCategory + +variable {C : Type u} [Category.{v} C] [MonoidalCategory.{v} C] [BraidedCategory.{v} C] + +@[simp] +theorem braiding_tensor_left (X Y Z : C) : + (Ξ²_ (X βŠ— Y) Z).hom = + (Ξ±_ X Y Z).hom ≫ X ◁ (Ξ²_ Y Z).hom ≫ (Ξ±_ X Z Y).inv ≫ + (Ξ²_ X Z).hom β–· Y ≫ (Ξ±_ Z X Y).hom := by + intros + apply (cancel_epi (Ξ±_ X Y Z).inv).1 + apply (cancel_mono (Ξ±_ Z X Y).inv).1 + simp [hexagon_reverse] + +@[simp] +theorem braiding_tensor_right (X Y Z : C) : + (Ξ²_ X (Y βŠ— Z)).hom = + (Ξ±_ X Y Z).inv ≫ (Ξ²_ X Y).hom β–· Z ≫ (Ξ±_ Y X Z).hom ≫ + Y ◁ (Ξ²_ X Z).hom ≫ (Ξ±_ Y Z X).inv := by + intros + apply (cancel_epi (Ξ±_ X Y Z).hom).1 + apply (cancel_mono (Ξ±_ Y Z X).hom).1 + simp [hexagon_forward] + +@[simp] +theorem braiding_inv_tensor_left (X Y Z : C) : + (Ξ²_ (X βŠ— Y) Z).inv = + (Ξ±_ Z X Y).inv ≫ (Ξ²_ X Z).inv β–· Y ≫ (Ξ±_ X Z Y).hom ≫ + X ◁ (Ξ²_ Y Z).inv ≫ (Ξ±_ X Y Z).inv := + eq_of_inv_eq_inv (by simp) + +@[simp] +theorem braiding_inv_tensor_right (X Y Z : C) : + (Ξ²_ X (Y βŠ— Z)).inv = + (Ξ±_ Y Z X).hom ≫ Y ◁ (Ξ²_ X Z).inv ≫ (Ξ±_ Y X Z).inv ≫ + (Ξ²_ X Y).inv β–· Z ≫ (Ξ±_ X Y Z).hom := + eq_of_inv_eq_inv (by simp) + +end BraidedCategory + open Category open MonoidalCategory open BraidedCategory -notation "Ξ²_" => braiding - /-- Verifying the axioms for a braiding by checking that the candidate braiding is sent to a braiding by a faithful monoidal functor. @@ -80,34 +130,41 @@ def braidedCategoryOfFaithful {C D : Type _} [Category C] [Category D] [Monoidal (Ξ² : βˆ€ X Y : C, X βŠ— Y β‰… Y βŠ— X) (w : βˆ€ X Y, F.ΞΌ _ _ ≫ F.map (Ξ² X Y).hom = (Ξ²_ _ _).hom ≫ F.ΞΌ _ _) : BraidedCategory C where braiding := Ξ² - braiding_naturality := by + braiding_naturality_left := by intros apply F.map_injective refine (cancel_epi (F.ΞΌ ?_ ?_)).1 ?_ - rw [Functor.map_comp, ← LaxMonoidalFunctor.ΞΌ_natural_assoc, w, Functor.map_comp, reassoc_of% w, - braiding_naturality_assoc, LaxMonoidalFunctor.ΞΌ_natural] + rw [Functor.map_comp, ← LaxMonoidalFunctor.ΞΌ_natural_left_assoc, w, Functor.map_comp, + reassoc_of% w, braiding_naturality_left_assoc, LaxMonoidalFunctor.ΞΌ_natural_right] + braiding_naturality_right := by + intros + apply F.map_injective + refine (cancel_epi (F.ΞΌ ?_ ?_)).1 ?_ + rw [Functor.map_comp, ← LaxMonoidalFunctor.ΞΌ_natural_right_assoc, w, Functor.map_comp, + reassoc_of% w, braiding_naturality_right_assoc, LaxMonoidalFunctor.ΞΌ_natural_left] hexagon_forward := by intros apply F.map_injective refine (cancel_epi (F.ΞΌ _ _)).1 ?_ - refine (cancel_epi (F.ΞΌ _ _ βŠ— πŸ™ _)).1 ?_ + refine (cancel_epi (F.ΞΌ _ _ β–· _)).1 ?_ rw [Functor.map_comp, Functor.map_comp, Functor.map_comp, Functor.map_comp, ← - LaxMonoidalFunctor.ΞΌ_natural_assoc, Functor.map_id, ← comp_tensor_id_assoc, w, - comp_tensor_id, assoc, LaxMonoidalFunctor.associativity'_assoc, - LaxMonoidalFunctor.associativity'_assoc, ← LaxMonoidalFunctor.ΞΌ_natural, Functor.map_id, ← - id_tensor_comp_assoc, w, id_tensor_comp_assoc, reassoc_of% w, braiding_naturality_assoc, - LaxMonoidalFunctor.associativity', hexagon_forward_assoc] + LaxMonoidalFunctor.ΞΌ_natural_left_assoc, ← comp_whiskerRight_assoc, w, + comp_whiskerRight_assoc, LaxMonoidalFunctor.associativity_assoc, + LaxMonoidalFunctor.associativity_assoc, ← LaxMonoidalFunctor.ΞΌ_natural_right, ← + MonoidalCategory.whiskerLeft_comp_assoc, w, MonoidalCategory.whiskerLeft_comp_assoc, + reassoc_of% w, braiding_naturality_right_assoc, + LaxMonoidalFunctor.associativity, hexagon_forward_assoc] hexagon_reverse := by intros apply F.toFunctor.map_injective refine (cancel_epi (F.ΞΌ _ _)).1 ?_ - refine (cancel_epi (πŸ™ _ βŠ— F.ΞΌ _ _)).1 ?_ + refine (cancel_epi (_ ◁ F.ΞΌ _ _)).1 ?_ rw [Functor.map_comp, Functor.map_comp, Functor.map_comp, Functor.map_comp, ← - LaxMonoidalFunctor.ΞΌ_natural_assoc, Functor.map_id, ← id_tensor_comp_assoc, w, - id_tensor_comp_assoc, LaxMonoidalFunctor.associativity_inv'_assoc, - LaxMonoidalFunctor.associativity_inv'_assoc, ← LaxMonoidalFunctor.ΞΌ_natural, - Functor.map_id, ← comp_tensor_id_assoc, w, comp_tensor_id_assoc, reassoc_of% w, - braiding_naturality_assoc, LaxMonoidalFunctor.associativity_inv', hexagon_reverse_assoc] + LaxMonoidalFunctor.ΞΌ_natural_right_assoc, ← MonoidalCategory.whiskerLeft_comp_assoc, w, + MonoidalCategory.whiskerLeft_comp_assoc, LaxMonoidalFunctor.associativity_inv_assoc, + LaxMonoidalFunctor.associativity_inv_assoc, ← LaxMonoidalFunctor.ΞΌ_natural_left, + ← comp_whiskerRight_assoc, w, comp_whiskerRight_assoc, reassoc_of% w, + braiding_naturality_left_assoc, LaxMonoidalFunctor.associativity_inv, hexagon_reverse_assoc] #align category_theory.braided_category_of_faithful CategoryTheory.braidedCategoryOfFaithful /-- Pull back a braiding along a fully faithful monoidal functor. -/ @@ -139,68 +196,68 @@ variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory C] [BraidedCatego theorem braiding_leftUnitor_aux₁ (X : C) : (Ξ±_ (πŸ™_ C) (πŸ™_ C) X).hom ≫ - (πŸ™ (πŸ™_ C) βŠ— (Ξ²_ X (πŸ™_ C)).inv) ≫ (Ξ±_ _ X _).inv ≫ ((Ξ»_ X).hom βŠ— πŸ™ _) = - ((Ξ»_ _).hom βŠ— πŸ™ X) ≫ (Ξ²_ X (πŸ™_ C)).inv := by + (πŸ™_ C ◁ (Ξ²_ X (πŸ™_ C)).inv) ≫ (Ξ±_ _ X _).inv ≫ ((Ξ»_ X).hom β–· _) = + ((Ξ»_ _).hom β–· X) ≫ (Ξ²_ X (πŸ™_ C)).inv := by coherence #align category_theory.braiding_left_unitor_aux₁ CategoryTheory.braiding_leftUnitor_aux₁ theorem braiding_leftUnitor_auxβ‚‚ (X : C) : - ((Ξ²_ X (πŸ™_ C)).hom βŠ— πŸ™ (πŸ™_ C)) ≫ ((Ξ»_ X).hom βŠ— πŸ™ (πŸ™_ C)) = (ρ_ X).hom βŠ— πŸ™ (πŸ™_ C) := + ((Ξ²_ X (πŸ™_ C)).hom β–· πŸ™_ C) ≫ ((Ξ»_ X).hom β–· πŸ™_ C) = (ρ_ X).hom β–· πŸ™_ C := calc - ((Ξ²_ X (πŸ™_ C)).hom βŠ— πŸ™ (πŸ™_ C)) ≫ ((Ξ»_ X).hom βŠ— πŸ™ (πŸ™_ C)) = - ((Ξ²_ X (πŸ™_ C)).hom βŠ— πŸ™ (πŸ™_ C)) ≫ (Ξ±_ _ _ _).hom ≫ (Ξ±_ _ _ _).inv ≫ ((Ξ»_ X).hom βŠ— πŸ™ (πŸ™_ C)) := + ((Ξ²_ X (πŸ™_ C)).hom β–· πŸ™_ C) ≫ ((Ξ»_ X).hom β–· πŸ™_ C) = + ((Ξ²_ X (πŸ™_ C)).hom β–· πŸ™_ C) ≫ (Ξ±_ _ _ _).hom ≫ (Ξ±_ _ _ _).inv ≫ ((Ξ»_ X).hom β–· πŸ™_ C) := by coherence - _ = ((Ξ²_ X (πŸ™_ C)).hom βŠ— πŸ™ (πŸ™_ C)) ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ _ βŠ— (Ξ²_ X _).hom) ≫ - (πŸ™ _ βŠ— (Ξ²_ X _).inv) ≫ (Ξ±_ _ _ _).inv ≫ ((Ξ»_ X).hom βŠ— πŸ™ (πŸ™_ C)) := - by (slice_rhs 3 4 => rw [← id_tensor_comp, Iso.hom_inv_id, tensor_id]); rw [id_comp] - _ = (Ξ±_ _ _ _).hom ≫ (Ξ²_ _ _).hom ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ _ βŠ— (Ξ²_ X _).inv) ≫ (Ξ±_ _ _ _).inv ≫ - ((Ξ»_ X).hom βŠ— πŸ™ (πŸ™_ C)) := + _ = ((Ξ²_ X (πŸ™_ C)).hom β–· πŸ™_ C) ≫ (Ξ±_ _ _ _).hom ≫ (_ ◁ (Ξ²_ X _).hom) ≫ + (_ ◁ (Ξ²_ X _).inv) ≫ (Ξ±_ _ _ _).inv ≫ ((Ξ»_ X).hom β–· πŸ™_ C) := + by simp + _ = (Ξ±_ _ _ _).hom ≫ (Ξ²_ _ _).hom ≫ (Ξ±_ _ _ _).hom ≫ (_ ◁ (Ξ²_ X _).inv) ≫ (Ξ±_ _ _ _).inv ≫ + ((Ξ»_ X).hom β–· πŸ™_ C) := by (slice_lhs 1 3 => rw [← hexagon_forward]); simp only [assoc] - _ = (Ξ±_ _ _ _).hom ≫ (Ξ²_ _ _).hom ≫ ((Ξ»_ _).hom βŠ— πŸ™ X) ≫ (Ξ²_ X _).inv := + _ = (Ξ±_ _ _ _).hom ≫ (Ξ²_ _ _).hom ≫ ((Ξ»_ _).hom β–· X) ≫ (Ξ²_ X _).inv := by rw [braiding_leftUnitor_aux₁] - _ = (Ξ±_ _ _ _).hom ≫ (πŸ™ _ βŠ— (Ξ»_ _).hom) ≫ (Ξ²_ _ _).hom ≫ (Ξ²_ X _).inv := - by (slice_lhs 2 3 => rw [← braiding_naturality]); simp only [assoc] - _ = (Ξ±_ _ _ _).hom ≫ (πŸ™ _ βŠ— (Ξ»_ _).hom) := by rw [Iso.hom_inv_id, comp_id] - _ = (ρ_ X).hom βŠ— πŸ™ (πŸ™_ C) := by rw [triangle'] + _ = (Ξ±_ _ _ _).hom ≫ (_ ◁ (Ξ»_ _).hom) ≫ (Ξ²_ _ _).hom ≫ (Ξ²_ X _).inv := + by (slice_lhs 2 3 => rw [← braiding_naturality_right]); simp only [assoc] + _ = (Ξ±_ _ _ _).hom ≫ (_ ◁ (Ξ»_ _).hom) := by rw [Iso.hom_inv_id, comp_id] + _ = (ρ_ X).hom β–· πŸ™_ C := by rw [triangle] #align category_theory.braiding_left_unitor_auxβ‚‚ CategoryTheory.braiding_leftUnitor_auxβ‚‚ @[simp] theorem braiding_leftUnitor (X : C) : (Ξ²_ X (πŸ™_ C)).hom ≫ (Ξ»_ X).hom = (ρ_ X).hom := by - rw [← tensor_right_iff, comp_tensor_id, braiding_leftUnitor_auxβ‚‚] + rw [← whiskerRight_iff, comp_whiskerRight, braiding_leftUnitor_auxβ‚‚] #align category_theory.braiding_left_unitor CategoryTheory.braiding_leftUnitor theorem braiding_rightUnitor_aux₁ (X : C) : (Ξ±_ X (πŸ™_ C) (πŸ™_ C)).inv ≫ - ((Ξ²_ (πŸ™_ C) X).inv βŠ— πŸ™ (πŸ™_ C)) ≫ (Ξ±_ _ X _).hom ≫ (πŸ™ _ βŠ— (ρ_ X).hom) = - (πŸ™ X βŠ— (ρ_ _).hom) ≫ (Ξ²_ (πŸ™_ C) X).inv := by + ((Ξ²_ (πŸ™_ C) X).inv β–· πŸ™_ C) ≫ (Ξ±_ _ X _).hom ≫ (_ ◁ (ρ_ X).hom) = + (X ◁ (ρ_ _).hom) ≫ (Ξ²_ (πŸ™_ C) X).inv := by coherence #align category_theory.braiding_right_unitor_aux₁ CategoryTheory.braiding_rightUnitor_aux₁ theorem braiding_rightUnitor_auxβ‚‚ (X : C) : - (πŸ™ (πŸ™_ C) βŠ— (Ξ²_ (πŸ™_ C) X).hom) ≫ (πŸ™ (πŸ™_ C) βŠ— (ρ_ X).hom) = πŸ™ (πŸ™_ C) βŠ— (Ξ»_ X).hom := + (πŸ™_ C ◁ (Ξ²_ (πŸ™_ C) X).hom) ≫ (πŸ™_ C ◁ (ρ_ X).hom) = πŸ™_ C ◁ (Ξ»_ X).hom := calc - (πŸ™ (πŸ™_ C) βŠ— (Ξ²_ (πŸ™_ C) X).hom) ≫ (πŸ™ (πŸ™_ C) βŠ— (ρ_ X).hom) = - (πŸ™ (πŸ™_ C) βŠ— (Ξ²_ (πŸ™_ C) X).hom) ≫ (Ξ±_ _ _ _).inv ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ (πŸ™_ C) βŠ— (ρ_ X).hom) := + (πŸ™_ C ◁ (Ξ²_ (πŸ™_ C) X).hom) ≫ (πŸ™_ C ◁ (ρ_ X).hom) = + (πŸ™_ C ◁ (Ξ²_ (πŸ™_ C) X).hom) ≫ (Ξ±_ _ _ _).inv ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™_ C ◁ (ρ_ X).hom) := by coherence - _ = (πŸ™ (πŸ™_ C) βŠ— (Ξ²_ (πŸ™_ C) X).hom) ≫ (Ξ±_ _ _ _).inv ≫ ((Ξ²_ _ X).hom βŠ— πŸ™ _) ≫ - ((Ξ²_ _ X).inv βŠ— πŸ™ _) ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ (πŸ™_ C) βŠ— (ρ_ X).hom) := - by (slice_rhs 3 4 => rw [← comp_tensor_id, Iso.hom_inv_id, tensor_id]); rw [id_comp] - _ = (Ξ±_ _ _ _).inv ≫ (Ξ²_ _ _).hom ≫ (Ξ±_ _ _ _).inv ≫ ((Ξ²_ _ X).inv βŠ— πŸ™ _) ≫ (Ξ±_ _ _ _).hom ≫ - (πŸ™ (πŸ™_ C) βŠ— (ρ_ X).hom) := + _ = (πŸ™_ C ◁ (Ξ²_ (πŸ™_ C) X).hom) ≫ (Ξ±_ _ _ _).inv ≫ ((Ξ²_ _ X).hom β–· _) ≫ + ((Ξ²_ _ X).inv β–· _) ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™_ C ◁ (ρ_ X).hom) := + by simp + _ = (Ξ±_ _ _ _).inv ≫ (Ξ²_ _ _).hom ≫ (Ξ±_ _ _ _).inv ≫ ((Ξ²_ _ X).inv β–· _) ≫ (Ξ±_ _ _ _).hom ≫ + (πŸ™_ C ◁ (ρ_ X).hom) := by (slice_lhs 1 3 => rw [← hexagon_reverse]); simp only [assoc] - _ = (Ξ±_ _ _ _).inv ≫ (Ξ²_ _ _).hom ≫ (πŸ™ X βŠ— (ρ_ _).hom) ≫ (Ξ²_ _ X).inv := + _ = (Ξ±_ _ _ _).inv ≫ (Ξ²_ _ _).hom ≫ (X ◁ (ρ_ _).hom) ≫ (Ξ²_ _ X).inv := by rw [braiding_rightUnitor_aux₁] - _ = (Ξ±_ _ _ _).inv ≫ ((ρ_ _).hom βŠ— πŸ™ _) ≫ (Ξ²_ _ X).hom ≫ (Ξ²_ _ _).inv := - by (slice_lhs 2 3 => rw [← braiding_naturality]); simp only [assoc] - _ = (Ξ±_ _ _ _).inv ≫ ((ρ_ _).hom βŠ— πŸ™ _) := by rw [Iso.hom_inv_id, comp_id] - _ = πŸ™ (πŸ™_ C) βŠ— (Ξ»_ X).hom := by rw [triangle_assoc_comp_right'] + _ = (Ξ±_ _ _ _).inv ≫ ((ρ_ _).hom β–· _) ≫ (Ξ²_ _ X).hom ≫ (Ξ²_ _ _).inv := + by (slice_lhs 2 3 => rw [← braiding_naturality_left]); simp only [assoc] + _ = (Ξ±_ _ _ _).inv ≫ ((ρ_ _).hom β–· _) := by rw [Iso.hom_inv_id, comp_id] + _ = πŸ™_ C ◁ (Ξ»_ X).hom := by rw [triangle_assoc_comp_right] #align category_theory.braiding_right_unitor_auxβ‚‚ CategoryTheory.braiding_rightUnitor_auxβ‚‚ @[simp] theorem braiding_rightUnitor (X : C) : (Ξ²_ (πŸ™_ C) X).hom ≫ (ρ_ X).hom = (Ξ»_ X).hom := by - rw [← tensor_left_iff, id_tensor_comp, braiding_rightUnitor_auxβ‚‚] + rw [← whiskerLeft_iff, MonoidalCategory.whiskerLeft_comp, braiding_rightUnitor_auxβ‚‚] #align category_theory.braiding_right_unitor CategoryTheory.braiding_rightUnitor @[simp] @@ -389,23 +446,44 @@ section Tensor /-- The strength of the tensor product functor from `C Γ— C` to `C`. -/ def tensor_ΞΌ (X Y : C Γ— C) : (tensor C).obj X βŠ— (tensor C).obj Y ⟢ (tensor C).obj (X βŠ— Y) := (Ξ±_ X.1 X.2 (Y.1 βŠ— Y.2)).hom ≫ - (πŸ™ X.1 βŠ— (Ξ±_ X.2 Y.1 Y.2).inv) ≫ - (πŸ™ X.1 βŠ— (Ξ²_ X.2 Y.1).hom βŠ— πŸ™ Y.2) ≫ - (πŸ™ X.1 βŠ— (Ξ±_ Y.1 X.2 Y.2).hom) ≫ (Ξ±_ X.1 Y.1 (X.2 βŠ— Y.2)).inv + (X.1 ◁ (Ξ±_ X.2 Y.1 Y.2).inv) ≫ + (X.1 ◁ (Ξ²_ X.2 Y.1).hom β–· Y.2) ≫ + (X.1 ◁ (Ξ±_ Y.1 X.2 Y.2).hom) ≫ (Ξ±_ X.1 Y.1 (X.2 βŠ— Y.2)).inv #align category_theory.tensor_ΞΌ CategoryTheory.tensor_ΞΌ theorem tensor_ΞΌ_def₁ (X₁ Xβ‚‚ Y₁ Yβ‚‚ : C) : - tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) ≫ (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— Yβ‚‚)).hom ≫ (πŸ™ X₁ βŠ— (Ξ±_ Y₁ Xβ‚‚ Yβ‚‚).inv) = - (Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— Yβ‚‚)).hom ≫ (πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ Y₁ Yβ‚‚).inv) ≫ (πŸ™ X₁ βŠ— (Ξ²_ Xβ‚‚ Y₁).hom βŠ— πŸ™ Yβ‚‚) := + tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) ≫ (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— Yβ‚‚)).hom ≫ (X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ Yβ‚‚).inv) = + (Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— Yβ‚‚)).hom ≫ (X₁ ◁ (Ξ±_ Xβ‚‚ Y₁ Yβ‚‚).inv) ≫ (X₁ ◁ (Ξ²_ Xβ‚‚ Y₁).hom β–· Yβ‚‚) := by dsimp [tensor_ΞΌ]; simp #align category_theory.tensor_ΞΌ_def₁ CategoryTheory.tensor_ΞΌ_def₁ theorem tensor_ΞΌ_defβ‚‚ (X₁ Xβ‚‚ Y₁ Yβ‚‚ : C) : - (πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ Y₁ Yβ‚‚).hom) ≫ (Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— Yβ‚‚)).inv ≫ tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) = - (πŸ™ X₁ βŠ— (Ξ²_ Xβ‚‚ Y₁).hom βŠ— πŸ™ Yβ‚‚) ≫ (πŸ™ X₁ βŠ— (Ξ±_ Y₁ Xβ‚‚ Yβ‚‚).hom) ≫ (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— Yβ‚‚)).inv := + (X₁ ◁ (Ξ±_ Xβ‚‚ Y₁ Yβ‚‚).hom) ≫ (Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— Yβ‚‚)).inv ≫ tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) = + (X₁ ◁ (Ξ²_ Xβ‚‚ Y₁).hom β–· Yβ‚‚) ≫ (X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ Yβ‚‚).hom) ≫ (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— Yβ‚‚)).inv := by dsimp [tensor_ΞΌ]; simp #align category_theory.tensor_ΞΌ_defβ‚‚ CategoryTheory.tensor_ΞΌ_defβ‚‚ +theorem tensor_ΞΌ_natural_left {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁: X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (Z₁ Zβ‚‚ : C) : + (f₁ βŠ— fβ‚‚) β–· (Z₁ βŠ— Zβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚) = + tensor_ΞΌ C (X₁, Xβ‚‚) (Z₁, Zβ‚‚) ≫ (f₁ β–· Z₁ βŠ— fβ‚‚ β–· Zβ‚‚) := by + dsimp only [tensor_ΞΌ, prodMonoidal_tensorObj, tensor_obj] + calc + _ = πŸ™ _ βŠ—β‰« + f₁ β–· Xβ‚‚ β–· Z₁ β–· Zβ‚‚ βŠ—β‰« Y₁ ◁ (fβ‚‚ β–· Z₁ ≫ (Ξ²_ Yβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := ?eq1 + _ = πŸ™ _ βŠ—β‰« + (f₁ β–· (Xβ‚‚ βŠ— Z₁) ≫ Y₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« Y₁ ◁ Z₁ ◁ fβ‚‚ β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := ?eq2 + _ = πŸ™ _ βŠ—β‰« + X₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom β–· Zβ‚‚ βŠ—β‰« (f₁ β–· Z₁ β–· (Xβ‚‚ βŠ— Zβ‚‚) ≫ (Y₁ βŠ— Z₁) ◁ fβ‚‚ β–· Zβ‚‚) βŠ—β‰« πŸ™ _ := ?eq3 + _ = _ := ?eq4 + case eq1 => + rw [tensorHom_def']; coherence + case eq2 => + rw [braiding_naturality_left]; coherence + case eq3 => + rw [← whisker_exchange]; coherence + case eq4 => + rw [tensorHom_def']; coherence + theorem tensor_ΞΌ_natural {X₁ Xβ‚‚ Y₁ Yβ‚‚ U₁ Uβ‚‚ V₁ Vβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : U₁ ⟢ V₁) (gβ‚‚ : Uβ‚‚ ⟢ Vβ‚‚) : ((f₁ βŠ— fβ‚‚) βŠ— g₁ βŠ— gβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (V₁, Vβ‚‚) = @@ -425,154 +503,84 @@ theorem tensor_ΞΌ_natural {X₁ Xβ‚‚ Y₁ Yβ‚‚ U₁ Uβ‚‚ V₁ Vβ‚‚ : C} (f₁ : theorem tensor_left_unitality (X₁ Xβ‚‚ : C) : (Ξ»_ (X₁ βŠ— Xβ‚‚)).hom = - ((Ξ»_ (πŸ™_ C)).inv βŠ— πŸ™ (X₁ βŠ— Xβ‚‚)) ≫ + ((Ξ»_ (πŸ™_ C)).inv β–· (X₁ βŠ— Xβ‚‚)) ≫ tensor_ΞΌ C (πŸ™_ C, πŸ™_ C) (X₁, Xβ‚‚) ≫ ((Ξ»_ X₁).hom βŠ— (Ξ»_ Xβ‚‚).hom) := by dsimp only [tensor_ΞΌ] have : - ((Ξ»_ (πŸ™_ C)).inv βŠ— πŸ™ (X₁ βŠ— Xβ‚‚)) ≫ - (Ξ±_ (πŸ™_ C) (πŸ™_ C) (X₁ βŠ— Xβ‚‚)).hom ≫ (πŸ™ (πŸ™_ C) βŠ— (Ξ±_ (πŸ™_ C) X₁ Xβ‚‚).inv) = - πŸ™ (πŸ™_ C) βŠ— (Ξ»_ X₁).inv βŠ— πŸ™ Xβ‚‚ := + ((Ξ»_ (πŸ™_ C)).inv β–· (X₁ βŠ— Xβ‚‚)) ≫ + (Ξ±_ (πŸ™_ C) (πŸ™_ C) (X₁ βŠ— Xβ‚‚)).hom ≫ (πŸ™_ C ◁ (Ξ±_ (πŸ™_ C) X₁ Xβ‚‚).inv) = + πŸ™_ C ◁ (Ξ»_ X₁).inv β–· Xβ‚‚ := by coherence slice_rhs 1 3 => rw [this] clear this - slice_rhs 1 2 => rw [← tensor_comp, ← tensor_comp, comp_id, comp_id, leftUnitor_inv_braiding] + slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, ← comp_whiskerRight, + leftUnitor_inv_braiding] simp [tensorHom_id, id_tensorHom, tensorHom_def] #align category_theory.tensor_left_unitality CategoryTheory.tensor_left_unitality theorem tensor_right_unitality (X₁ Xβ‚‚ : C) : (ρ_ (X₁ βŠ— Xβ‚‚)).hom = - (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— (Ξ»_ (πŸ™_ C)).inv) ≫ + ((X₁ βŠ— Xβ‚‚) ◁ (Ξ»_ (πŸ™_ C)).inv) ≫ tensor_ΞΌ C (X₁, Xβ‚‚) (πŸ™_ C, πŸ™_ C) ≫ ((ρ_ X₁).hom βŠ— (ρ_ Xβ‚‚).hom) := by dsimp only [tensor_ΞΌ] have : - (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— (Ξ»_ (πŸ™_ C)).inv) ≫ - (Ξ±_ X₁ Xβ‚‚ (πŸ™_ C βŠ— πŸ™_ C)).hom ≫ (πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ (πŸ™_ C) (πŸ™_ C)).inv) = - (Ξ±_ X₁ Xβ‚‚ (πŸ™_ C)).hom ≫ (πŸ™ X₁ βŠ— (ρ_ Xβ‚‚).inv βŠ— πŸ™ (πŸ™_ C)) := + ((X₁ βŠ— Xβ‚‚) ◁ (Ξ»_ (πŸ™_ C)).inv) ≫ + (Ξ±_ X₁ Xβ‚‚ (πŸ™_ C βŠ— πŸ™_ C)).hom ≫ (X₁ ◁ (Ξ±_ Xβ‚‚ (πŸ™_ C) (πŸ™_ C)).inv) = + (Ξ±_ X₁ Xβ‚‚ (πŸ™_ C)).hom ≫ (X₁ ◁ (ρ_ Xβ‚‚).inv β–· πŸ™_ C) := by coherence slice_rhs 1 3 => rw [this] clear this - slice_rhs 2 3 => rw [← tensor_comp, ← tensor_comp, comp_id, comp_id, rightUnitor_inv_braiding] + slice_rhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, ← comp_whiskerRight, + rightUnitor_inv_braiding] simp [tensorHom_id, id_tensorHom, tensorHom_def] #align category_theory.tensor_right_unitality CategoryTheory.tensor_right_unitality -/- -Diagram B6 from Proposition 1 of [Joyal and Street, *Braided monoidal categories*][Joyal_Street]. --/ -theorem tensor_associativity_aux (W X Y Z : C) : - ((Ξ²_ W X).hom βŠ— πŸ™ (Y βŠ— Z)) ≫ - (Ξ±_ X W (Y βŠ— Z)).hom ≫ - (πŸ™ X βŠ— (Ξ±_ W Y Z).inv) ≫ (πŸ™ X βŠ— (Ξ²_ (W βŠ— Y) Z).hom) ≫ (πŸ™ X βŠ— (Ξ±_ Z W Y).inv) = - (πŸ™ (W βŠ— X) βŠ— (Ξ²_ Y Z).hom) ≫ - (Ξ±_ (W βŠ— X) Z Y).inv ≫ - ((Ξ±_ W X Z).hom βŠ— πŸ™ Y) ≫ - ((Ξ²_ W (X βŠ— Z)).hom βŠ— πŸ™ Y) ≫ ((Ξ±_ X Z W).hom βŠ— πŸ™ Y) ≫ (Ξ±_ X (Z βŠ— W) Y).hom := by - slice_rhs 3 5 => rw [← tensor_comp, ← tensor_comp, hexagon_forward, tensor_comp, tensor_comp] - slice_rhs 5 6 => rw [associator_naturality] - slice_rhs 2 3 => rw [← associator_inv_naturality] - slice_rhs 3 5 => rw [← pentagon_hom_inv] - slice_rhs 1 2 => rw [tensor_id, id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_rhs 2 3 => rw [← tensor_id, associator_naturality] - slice_rhs 3 5 => rw [← tensor_comp, ← tensor_comp, ← hexagon_reverse, tensor_comp, tensor_comp] -#align category_theory.tensor_associativity_aux CategoryTheory.tensor_associativity_aux - -set_option maxHeartbeats 400000 in theorem tensor_associativity (X₁ Xβ‚‚ Y₁ Yβ‚‚ Z₁ Zβ‚‚ : C) : - (tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) βŠ— πŸ™ (Z₁ βŠ— Zβ‚‚)) ≫ + (tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) β–· (Z₁ βŠ— Zβ‚‚)) ≫ tensor_ΞΌ C (X₁ βŠ— Y₁, Xβ‚‚ βŠ— Yβ‚‚) (Z₁, Zβ‚‚) ≫ ((Ξ±_ X₁ Y₁ Z₁).hom βŠ— (Ξ±_ Xβ‚‚ Yβ‚‚ Zβ‚‚).hom) = (Ξ±_ (X₁ βŠ— Xβ‚‚) (Y₁ βŠ— Yβ‚‚) (Z₁ βŠ— Zβ‚‚)).hom ≫ - (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚)) ≫ tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁ βŠ— Z₁, Yβ‚‚ βŠ— Zβ‚‚) := by - have : - (Ξ±_ X₁ Y₁ Z₁).hom βŠ— (Ξ±_ Xβ‚‚ Yβ‚‚ Zβ‚‚).hom = - (Ξ±_ (X₁ βŠ— Y₁) Z₁ ((Xβ‚‚ βŠ— Yβ‚‚) βŠ— Zβ‚‚)).hom ≫ - (πŸ™ (X₁ βŠ— Y₁) βŠ— (Ξ±_ Z₁ (Xβ‚‚ βŠ— Yβ‚‚) Zβ‚‚).inv) ≫ - (Ξ±_ X₁ Y₁ ((Z₁ βŠ— Xβ‚‚ βŠ— Yβ‚‚) βŠ— Zβ‚‚)).hom ≫ - (πŸ™ X₁ βŠ— (Ξ±_ Y₁ (Z₁ βŠ— Xβ‚‚ βŠ— Yβ‚‚) Zβ‚‚).inv) ≫ - (Ξ±_ X₁ (Y₁ βŠ— Z₁ βŠ— Xβ‚‚ βŠ— Yβ‚‚) Zβ‚‚).inv ≫ - ((πŸ™ X₁ βŠ— πŸ™ Y₁ βŠ— (Ξ±_ Z₁ Xβ‚‚ Yβ‚‚).inv) βŠ— πŸ™ Zβ‚‚) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ Y₁ (Z₁ βŠ— Xβ‚‚) Yβ‚‚).inv) βŠ— πŸ™ Zβ‚‚) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ Y₁ Z₁ Xβ‚‚).inv βŠ— πŸ™ Yβ‚‚) βŠ— πŸ™ Zβ‚‚) ≫ - (Ξ±_ X₁ (((Y₁ βŠ— Z₁) βŠ— Xβ‚‚) βŠ— Yβ‚‚) Zβ‚‚).hom ≫ - (πŸ™ X₁ βŠ— (Ξ±_ ((Y₁ βŠ— Z₁) βŠ— Xβ‚‚) Yβ‚‚ Zβ‚‚).hom) ≫ - (πŸ™ X₁ βŠ— (Ξ±_ (Y₁ βŠ— Z₁) Xβ‚‚ (Yβ‚‚ βŠ— Zβ‚‚)).hom) ≫ - (Ξ±_ X₁ (Y₁ βŠ— Z₁) (Xβ‚‚ βŠ— Yβ‚‚ βŠ— Zβ‚‚)).inv := - by monoidal_simps; pure_coherence - rw [this]; clear this - slice_lhs 2 4 => rw [tensor_ΞΌ_def₁] - slice_lhs 4 5 => rw [← tensor_id, associator_naturality] - slice_lhs 5 6 => rw [← tensor_comp, associator_inv_naturality, tensor_comp] - slice_lhs 6 7 => rw [associator_inv_naturality] - have : - (Ξ±_ (X₁ βŠ— Y₁) (Xβ‚‚ βŠ— Yβ‚‚) (Z₁ βŠ— Zβ‚‚)).hom ≫ - (πŸ™ (X₁ βŠ— Y₁) βŠ— (Ξ±_ (Xβ‚‚ βŠ— Yβ‚‚) Z₁ Zβ‚‚).inv) ≫ - (Ξ±_ X₁ Y₁ (((Xβ‚‚ βŠ— Yβ‚‚) βŠ— Z₁) βŠ— Zβ‚‚)).hom ≫ - (πŸ™ X₁ βŠ— (Ξ±_ Y₁ ((Xβ‚‚ βŠ— Yβ‚‚) βŠ— Z₁) Zβ‚‚).inv) ≫ (Ξ±_ X₁ (Y₁ βŠ— (Xβ‚‚ βŠ— Yβ‚‚) βŠ— Z₁) Zβ‚‚).inv = - ((Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— Yβ‚‚)).hom βŠ— πŸ™ (Z₁ βŠ— Zβ‚‚)) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ Y₁ Xβ‚‚ Yβ‚‚).inv) βŠ— πŸ™ (Z₁ βŠ— Zβ‚‚)) ≫ - (Ξ±_ (X₁ βŠ— (Y₁ βŠ— Xβ‚‚) βŠ— Yβ‚‚) Z₁ Zβ‚‚).inv ≫ - ((Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚) βŠ— Yβ‚‚) Z₁).hom βŠ— πŸ™ Zβ‚‚) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ (Y₁ βŠ— Xβ‚‚) Yβ‚‚ Z₁).hom) βŠ— πŸ™ Zβ‚‚) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ Y₁ Xβ‚‚ (Yβ‚‚ βŠ— Z₁)).hom) βŠ— πŸ™ Zβ‚‚) ≫ - ((πŸ™ X₁ βŠ— πŸ™ Y₁ βŠ— (Ξ±_ Xβ‚‚ Yβ‚‚ Z₁).inv) βŠ— πŸ™ Zβ‚‚) := - by monoidal_simps; pure_coherence - slice_lhs 2 6 => rw [this] - clear this - slice_lhs 1 3 => rw [← tensor_comp, ← tensor_comp, tensor_ΞΌ_def₁, tensor_comp, tensor_comp] - slice_lhs 3 4 => rw [← tensor_id, associator_inv_naturality] - slice_lhs 4 5 => rw [← tensor_comp, associator_naturality, tensor_comp] - slice_lhs 5 6 => - rw [← tensor_comp, ← tensor_comp, associator_naturality, tensor_comp, tensor_comp] - slice_lhs 6 10 => - rw [← tensor_comp, ← tensor_comp, ← tensor_comp, ← tensor_comp, ← tensor_comp, ← tensor_comp, ← - tensor_comp, ← tensor_comp, tensor_id, tensor_associativity_aux, ← tensor_id, ← - id_comp (πŸ™ X₁ ≫ πŸ™ X₁ ≫ πŸ™ X₁ ≫ πŸ™ X₁ ≫ πŸ™ X₁), ← id_comp (πŸ™ Zβ‚‚ ≫ πŸ™ Zβ‚‚ ≫ πŸ™ Zβ‚‚ ≫ πŸ™ Zβ‚‚ ≫ πŸ™ Zβ‚‚), - tensor_comp, tensor_comp, tensor_comp, tensor_comp, tensor_comp, tensor_comp, tensor_comp, - tensor_comp, tensor_comp, tensor_comp] - slice_lhs 11 12 => - rw [← tensor_comp, ← tensor_comp, Iso.hom_inv_id] - rw [comp_id, id_comp, id_tensorHom, tensorHom_id, MonoidalCategory.whiskerLeft_id, id_whiskerRight] - simp only [assoc, id_comp] - slice_lhs 10 11 => - rw [← tensor_comp, ← tensor_comp, ← tensor_comp, Iso.hom_inv_id] - simp only [id_tensorHom, tensorHom_id, comp_id, id_comp] - simp - simp only [assoc, id_comp] - slice_lhs 9 10 => rw [associator_naturality] - slice_lhs 10 11 => rw [← tensor_comp, associator_naturality, tensor_comp] - slice_lhs 11 13 => rw [tensor_id, ← tensor_ΞΌ_defβ‚‚] - have : - ((πŸ™ X₁ βŠ— (Ξ±_ (Xβ‚‚ βŠ— Y₁) Z₁ Yβ‚‚).inv) βŠ— πŸ™ Zβ‚‚) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ Y₁ Z₁).hom βŠ— πŸ™ Yβ‚‚) βŠ— πŸ™ Zβ‚‚) ≫ - (Ξ±_ X₁ ((Xβ‚‚ βŠ— Y₁ βŠ— Z₁) βŠ— Yβ‚‚) Zβ‚‚).hom ≫ - (πŸ™ X₁ βŠ— (Ξ±_ (Xβ‚‚ βŠ— Y₁ βŠ— Z₁) Yβ‚‚ Zβ‚‚).hom) ≫ - (πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ (Y₁ βŠ— Z₁) (Yβ‚‚ βŠ— Zβ‚‚)).hom) ≫ (Ξ±_ X₁ Xβ‚‚ ((Y₁ βŠ— Z₁) βŠ— Yβ‚‚ βŠ— Zβ‚‚)).inv = - (Ξ±_ X₁ ((Xβ‚‚ βŠ— Y₁) βŠ— Z₁ βŠ— Yβ‚‚) Zβ‚‚).hom ≫ - (πŸ™ X₁ βŠ— (Ξ±_ (Xβ‚‚ βŠ— Y₁) (Z₁ βŠ— Yβ‚‚) Zβ‚‚).hom) ≫ - (πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ Y₁ ((Z₁ βŠ— Yβ‚‚) βŠ— Zβ‚‚)).hom) ≫ - (Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— (Z₁ βŠ— Yβ‚‚) βŠ— Zβ‚‚)).inv ≫ - (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— πŸ™ Y₁ βŠ— (Ξ±_ Z₁ Yβ‚‚ Zβ‚‚).hom) ≫ (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— (Ξ±_ Y₁ Z₁ (Yβ‚‚ βŠ— Zβ‚‚)).inv) := - by monoidal_simps; pure_coherence - slice_lhs 7 12 => rw [this] - clear this - slice_lhs 6 7 => rw [associator_naturality] - slice_lhs 7 8 => rw [← tensor_comp, associator_naturality, tensor_comp] - slice_lhs 8 9 => rw [← tensor_comp, associator_naturality, tensor_comp] - slice_lhs 9 10 => rw [associator_inv_naturality] - slice_lhs 10 12 => rw [← tensor_comp, ← tensor_comp, ← tensor_ΞΌ_defβ‚‚, tensor_comp, tensor_comp] - dsimp only [tensor_obj, prodMonoidal_tensorObj] - simp_rw [← Category.assoc] - congr 2 - coherence + ((X₁ βŠ— Xβ‚‚) ◁ tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚)) ≫ tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁ βŠ— Z₁, Yβ‚‚ βŠ— Zβ‚‚) := by + dsimp only [tensor_obj, prodMonoidal_tensorObj, tensor_ΞΌ] + simp only [whiskerRight_tensor, comp_whiskerRight, whisker_assoc, assoc, Iso.inv_hom_id_assoc, tensor_whiskerLeft, + braiding_tensor_left, MonoidalCategory.whiskerLeft_comp, braiding_tensor_right] + calc + _ = πŸ™ _ βŠ—β‰« + X₁ ◁ ((Ξ²_ Xβ‚‚ Y₁).hom β–· (Yβ‚‚ βŠ— Z₁) ≫ (Y₁ βŠ— Xβ‚‚) ◁ (Ξ²_ Yβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« + X₁ ◁ Y₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom β–· Yβ‚‚ β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := ?eq1 + _ = _ := ?eq2 + case eq1 => coherence + case eq2 => rw [← whisker_exchange]; coherence #align category_theory.tensor_associativity CategoryTheory.tensor_associativity @[simp] -def tensorLaxMonoidal : LaxMonoidalFunctor (C Γ— C) C := LaxMonoidalFunctor.ofTensorHom - (F := tensor C) - (Ξ΅ := (Ξ»_ (πŸ™_ C)).inv) - (ΞΌ := fun X Y => tensor_ΞΌ C X Y) - (ΞΌ_natural := fun f g => tensor_ΞΌ_natural C f.1 f.2 g.1 g.2) - (associativity := fun X Y Z => tensor_associativity C X.1 X.2 Y.1 Y.2 Z.1 Z.2) - (left_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_left_unitality C X₁ Xβ‚‚) - (right_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_right_unitality C X₁ Xβ‚‚) +def tensorLaxMonoidal : MonoidalFunctor (C Γ— C) C := + { tensor C with + Ξ΅ := (Ξ»_ (πŸ™_ C)).inv + ΞΌ := tensor_ΞΌ C + ΞΌ_natural_left := fun f Z ↦ tensor_ΞΌ_natural_left C f.1 f.2 Z.1 Z.2 + ΞΌ_natural_right := _ + associativity := fun X Y Z => tensor_associativity C X.1 X.2 Y.1 Y.2 Z.1 Z.2 + left_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_left_unitality C X₁ Xβ‚‚ + right_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_right_unitality C X₁ Xβ‚‚ + Ξ΅_isIso := inferInstanceAs (IsIso (Ξ»_ (πŸ™_ C)).inv) + ΞΌ_isIso := by dsimp [tensor_ΞΌ]; infer_instance } + -- (F := tensor C) + -- (Ξ΅ := (Ξ»_ (πŸ™_ C)).inv) + -- (ΞΌ := fun X Y => tensor_ΞΌ C X Y) + -- (ΞΌ_natural := fun f g => tensor_ΞΌ_natural C f.1 f.2 g.1 g.2) + -- (associativity := fun X Y Z => tensor_associativity C X.1 X.2 Y.1 Y.2 Z.1 Z.2) + -- (left_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_left_unitality C X₁ Xβ‚‚) + -- (right_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_right_unitality C X₁ Xβ‚‚) + +-- @[simp] +-- def tensorLaxMonoidal : LaxMonoidalFunctor (C Γ— C) C := LaxMonoidalFunctor.ofTensorHom +-- (F := tensor C) +-- (Ξ΅ := (Ξ»_ (πŸ™_ C)).inv) +-- (ΞΌ := fun X Y => tensor_ΞΌ C X Y) +-- (ΞΌ_natural := fun f g => tensor_ΞΌ_natural C f.1 f.2 g.1 g.2) +-- (associativity := fun X Y Z => tensor_associativity C X.1 X.2 Y.1 Y.2 Z.1 Z.2) +-- (left_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_left_unitality C X₁ Xβ‚‚) +-- (right_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_right_unitality C X₁ Xβ‚‚) /-- The tensor product functor from `C Γ— C` to `C` as a monoidal functor. -/ @[simps!] @@ -584,12 +592,12 @@ def tensorMonoidal : MonoidalFunctor (C Γ— C) C := theorem leftUnitor_monoidal (X₁ Xβ‚‚ : C) : (Ξ»_ X₁).hom βŠ— (Ξ»_ Xβ‚‚).hom = - tensor_ΞΌ C (πŸ™_ C, X₁) (πŸ™_ C, Xβ‚‚) ≫ ((Ξ»_ (πŸ™_ C)).hom βŠ— πŸ™ (X₁ βŠ— Xβ‚‚)) ≫ (Ξ»_ (X₁ βŠ— Xβ‚‚)).hom := by + tensor_ΞΌ C (πŸ™_ C, X₁) (πŸ™_ C, Xβ‚‚) ≫ ((Ξ»_ (πŸ™_ C)).hom β–· (X₁ βŠ— Xβ‚‚)) ≫ (Ξ»_ (X₁ βŠ— Xβ‚‚)).hom := by dsimp only [tensor_ΞΌ] have : (Ξ»_ X₁).hom βŠ— (Ξ»_ Xβ‚‚).hom = (Ξ±_ (πŸ™_ C) X₁ (πŸ™_ C βŠ— Xβ‚‚)).hom ≫ - (πŸ™ (πŸ™_ C) βŠ— (Ξ±_ X₁ (πŸ™_ C) Xβ‚‚).inv) ≫ (Ξ»_ ((X₁ βŠ— πŸ™_ C) βŠ— Xβ‚‚)).hom ≫ ((ρ_ X₁).hom βŠ— πŸ™ Xβ‚‚) := + (πŸ™_ C ◁ (Ξ±_ X₁ (πŸ™_ C) Xβ‚‚).inv) ≫ (Ξ»_ ((X₁ βŠ— πŸ™_ C) βŠ— Xβ‚‚)).hom ≫ ((ρ_ X₁).hom β–· Xβ‚‚) := by coherence rw [this]; clear this rw [← braiding_leftUnitor] @@ -599,12 +607,12 @@ theorem leftUnitor_monoidal (X₁ Xβ‚‚ : C) : theorem rightUnitor_monoidal (X₁ Xβ‚‚ : C) : (ρ_ X₁).hom βŠ— (ρ_ Xβ‚‚).hom = - tensor_ΞΌ C (X₁, πŸ™_ C) (Xβ‚‚, πŸ™_ C) ≫ (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— (Ξ»_ (πŸ™_ C)).hom) ≫ (ρ_ (X₁ βŠ— Xβ‚‚)).hom := by + tensor_ΞΌ C (X₁, πŸ™_ C) (Xβ‚‚, πŸ™_ C) ≫ ((X₁ βŠ— Xβ‚‚) ◁ (Ξ»_ (πŸ™_ C)).hom) ≫ (ρ_ (X₁ βŠ— Xβ‚‚)).hom := by dsimp only [tensor_ΞΌ] have : (ρ_ X₁).hom βŠ— (ρ_ Xβ‚‚).hom = (Ξ±_ X₁ (πŸ™_ C) (Xβ‚‚ βŠ— πŸ™_ C)).hom ≫ - (πŸ™ X₁ βŠ— (Ξ±_ (πŸ™_ C) Xβ‚‚ (πŸ™_ C)).inv) ≫ (πŸ™ X₁ βŠ— (ρ_ (πŸ™_ C βŠ— Xβ‚‚)).hom) ≫ (πŸ™ X₁ βŠ— (Ξ»_ Xβ‚‚).hom) := + (X₁ ◁ (Ξ±_ (πŸ™_ C) Xβ‚‚ (πŸ™_ C)).inv) ≫ (X₁ ◁ (ρ_ (πŸ™_ C βŠ— Xβ‚‚)).hom) ≫ (X₁ ◁ (Ξ»_ Xβ‚‚).hom) := by coherence rw [this]; clear this rw [← braiding_rightUnitor] @@ -633,22 +641,22 @@ theorem associator_monoidal_aux (W X Y Z : C) : set_option maxHeartbeats 400000 in theorem associator_monoidal (X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C) : tensor_ΞΌ C (X₁ βŠ— Xβ‚‚, X₃) (Y₁ βŠ— Yβ‚‚, Y₃) ≫ - (tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) βŠ— πŸ™ (X₃ βŠ— Y₃)) ≫ (Ξ±_ (X₁ βŠ— Y₁) (Xβ‚‚ βŠ— Yβ‚‚) (X₃ βŠ— Y₃)).hom = + (tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) β–· (X₃ βŠ— Y₃)) ≫ (Ξ±_ (X₁ βŠ— Y₁) (Xβ‚‚ βŠ— Yβ‚‚) (X₃ βŠ— Y₃)).hom = ((Ξ±_ X₁ Xβ‚‚ X₃).hom βŠ— (Ξ±_ Y₁ Yβ‚‚ Y₃).hom) ≫ - tensor_ΞΌ C (X₁, Xβ‚‚ βŠ— X₃) (Y₁, Yβ‚‚ βŠ— Y₃) ≫ (πŸ™ (X₁ βŠ— Y₁) βŠ— tensor_ΞΌ C (Xβ‚‚, X₃) (Yβ‚‚, Y₃)) := by + tensor_ΞΌ C (X₁, Xβ‚‚ βŠ— X₃) (Y₁, Yβ‚‚ βŠ— Y₃) ≫ ((X₁ βŠ— Y₁) ◁ tensor_ΞΌ C (Xβ‚‚, X₃) (Yβ‚‚, Y₃)) := by have : (Ξ±_ (X₁ βŠ— Y₁) (Xβ‚‚ βŠ— Yβ‚‚) (X₃ βŠ— Y₃)).hom = - ((Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— Yβ‚‚)).hom βŠ— πŸ™ (X₃ βŠ— Y₃)) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ Y₁ Xβ‚‚ Yβ‚‚).inv) βŠ— πŸ™ (X₃ βŠ— Y₃)) ≫ + ((Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— Yβ‚‚)).hom β–· (X₃ βŠ— Y₃)) ≫ + ((X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ Yβ‚‚).inv) β–· (X₃ βŠ— Y₃)) ≫ (Ξ±_ (X₁ βŠ— (Y₁ βŠ— Xβ‚‚) βŠ— Yβ‚‚) X₃ Y₃).inv ≫ - ((Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚) βŠ— Yβ‚‚) X₃).hom βŠ— πŸ™ Y₃) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ (Y₁ βŠ— Xβ‚‚) Yβ‚‚ X₃).hom) βŠ— πŸ™ Y₃) ≫ + ((Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚) βŠ— Yβ‚‚) X₃).hom β–· Y₃) ≫ + ((X₁ ◁ (Ξ±_ (Y₁ βŠ— Xβ‚‚) Yβ‚‚ X₃).hom) β–· Y₃) ≫ (Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚) βŠ— Yβ‚‚ βŠ— X₃) Y₃).hom ≫ - (πŸ™ X₁ βŠ— (Ξ±_ (Y₁ βŠ— Xβ‚‚) (Yβ‚‚ βŠ— X₃) Y₃).hom) ≫ - (πŸ™ X₁ βŠ— (Ξ±_ Y₁ Xβ‚‚ ((Yβ‚‚ βŠ— X₃) βŠ— Y₃)).hom) ≫ + (X₁ ◁ (Ξ±_ (Y₁ βŠ— Xβ‚‚) (Yβ‚‚ βŠ— X₃) Y₃).hom) ≫ + (X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ ((Yβ‚‚ βŠ— X₃) βŠ— Y₃)).hom) ≫ (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— (Yβ‚‚ βŠ— X₃) βŠ— Y₃)).inv ≫ - (πŸ™ (X₁ βŠ— Y₁) βŠ— πŸ™ Xβ‚‚ βŠ— (Ξ±_ Yβ‚‚ X₃ Y₃).hom) ≫ - (πŸ™ (X₁ βŠ— Y₁) βŠ— (Ξ±_ Xβ‚‚ Yβ‚‚ (X₃ βŠ— Y₃)).inv) := + ((X₁ βŠ— Y₁) ◁ Xβ‚‚ ◁ (Ξ±_ Yβ‚‚ X₃ Y₃).hom) ≫ + ((X₁ βŠ— Y₁) ◁ (Ξ±_ Xβ‚‚ Yβ‚‚ (X₃ βŠ— Y₃)).inv) := by coherence rw [this]; clear this slice_lhs 2 4 => rw [← tensor_comp, ← tensor_comp, tensor_ΞΌ_def₁, tensor_comp, tensor_comp] @@ -657,17 +665,17 @@ theorem associator_monoidal (X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C) : slice_lhs 6 7 => rw [← tensor_comp, ← tensor_comp, associator_naturality, tensor_comp, tensor_comp] have : - ((Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— Yβ‚‚)).hom βŠ— πŸ™ (X₃ βŠ— Y₃)) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ Y₁ Yβ‚‚).inv) βŠ— πŸ™ (X₃ βŠ— Y₃)) ≫ + ((Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— Yβ‚‚)).hom β–· (X₃ βŠ— Y₃)) ≫ + ((X₁ ◁ (Ξ±_ Xβ‚‚ Y₁ Yβ‚‚).inv) β–· (X₃ βŠ— Y₃)) ≫ (Ξ±_ (X₁ βŠ— (Xβ‚‚ βŠ— Y₁) βŠ— Yβ‚‚) X₃ Y₃).inv ≫ - ((Ξ±_ X₁ ((Xβ‚‚ βŠ— Y₁) βŠ— Yβ‚‚) X₃).hom βŠ— πŸ™ Y₃) ≫ ((πŸ™ X₁ βŠ— (Ξ±_ (Xβ‚‚ βŠ— Y₁) Yβ‚‚ X₃).hom) βŠ— πŸ™ Y₃) = + ((Ξ±_ X₁ ((Xβ‚‚ βŠ— Y₁) βŠ— Yβ‚‚) X₃).hom β–· Y₃) ≫ ((X₁ ◁ (Ξ±_ (Xβ‚‚ βŠ— Y₁) Yβ‚‚ X₃).hom) β–· Y₃) = (Ξ±_ (X₁ βŠ— Xβ‚‚) (Y₁ βŠ— Yβ‚‚) (X₃ βŠ— Y₃)).hom ≫ - (πŸ™ (X₁ βŠ— Xβ‚‚) βŠ— (Ξ±_ (Y₁ βŠ— Yβ‚‚) X₃ Y₃).inv) ≫ + ((X₁ βŠ— Xβ‚‚) ◁ (Ξ±_ (Y₁ βŠ— Yβ‚‚) X₃ Y₃).inv) ≫ (Ξ±_ X₁ Xβ‚‚ (((Y₁ βŠ— Yβ‚‚) βŠ— X₃) βŠ— Y₃)).hom ≫ - (πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ ((Y₁ βŠ— Yβ‚‚) βŠ— X₃) Y₃).inv) ≫ + (X₁ ◁ (Ξ±_ Xβ‚‚ ((Y₁ βŠ— Yβ‚‚) βŠ— X₃) Y₃).inv) ≫ (Ξ±_ X₁ (Xβ‚‚ βŠ— (Y₁ βŠ— Yβ‚‚) βŠ— X₃) Y₃).inv ≫ - ((πŸ™ X₁ βŠ— πŸ™ Xβ‚‚ βŠ— (Ξ±_ Y₁ Yβ‚‚ X₃).hom) βŠ— πŸ™ Y₃) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ Xβ‚‚ Y₁ (Yβ‚‚ βŠ— X₃)).inv) βŠ— πŸ™ Y₃) := + ((πŸ™ X₁ β–· Xβ‚‚ βŠ— (Ξ±_ Y₁ Yβ‚‚ X₃).hom) β–· Y₃) ≫ + ((X₁ ◁ (Ξ±_ Xβ‚‚ Y₁ (Yβ‚‚ βŠ— X₃)).inv) β–· Y₃) := by coherence slice_lhs 2 6 => rw [this] clear this @@ -688,17 +696,17 @@ theorem associator_monoidal (X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C) : slice_lhs 15 17 => rw [tensor_id, ← tensor_comp, ← tensor_comp, ← tensor_ΞΌ_defβ‚‚, tensor_comp, tensor_comp] have : - ((πŸ™ X₁ βŠ— (Ξ±_ Y₁ Xβ‚‚ X₃).inv βŠ— πŸ™ Yβ‚‚) βŠ— πŸ™ Y₃) ≫ - ((πŸ™ X₁ βŠ— (Ξ±_ (Y₁ βŠ— Xβ‚‚) X₃ Yβ‚‚).hom) βŠ— πŸ™ Y₃) ≫ + ((X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ X₃).inv β–· Yβ‚‚) β–· Y₃) ≫ + ((X₁ ◁ (Ξ±_ (Y₁ βŠ— Xβ‚‚) X₃ Yβ‚‚).hom) β–· Y₃) ≫ (Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚) βŠ— X₃ βŠ— Yβ‚‚) Y₃).hom ≫ - (πŸ™ X₁ βŠ— (Ξ±_ (Y₁ βŠ— Xβ‚‚) (X₃ βŠ— Yβ‚‚) Y₃).hom) ≫ - (πŸ™ X₁ βŠ— (Ξ±_ Y₁ Xβ‚‚ ((X₃ βŠ— Yβ‚‚) βŠ— Y₃)).hom) ≫ + (X₁ ◁ (Ξ±_ (Y₁ βŠ— Xβ‚‚) (X₃ βŠ— Yβ‚‚) Y₃).hom) ≫ + (X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ ((X₃ βŠ— Yβ‚‚) βŠ— Y₃)).hom) ≫ (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— (X₃ βŠ— Yβ‚‚) βŠ— Y₃)).inv ≫ - (πŸ™ (X₁ βŠ— Y₁) βŠ— πŸ™ Xβ‚‚ βŠ— (Ξ±_ X₃ Yβ‚‚ Y₃).hom) ≫ + (πŸ™ (X₁ βŠ— Y₁) β–· Xβ‚‚ βŠ— (Ξ±_ X₃ Yβ‚‚ Y₃).hom) ≫ (πŸ™ (X₁ βŠ— Y₁) βŠ— (Ξ±_ Xβ‚‚ X₃ (Yβ‚‚ βŠ— Y₃)).inv) = (Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚ βŠ— X₃) βŠ— Yβ‚‚) Y₃).hom ≫ - (πŸ™ X₁ βŠ— (Ξ±_ (Y₁ βŠ— Xβ‚‚ βŠ— X₃) Yβ‚‚ Y₃).hom) ≫ - (πŸ™ X₁ βŠ— (Ξ±_ Y₁ (Xβ‚‚ βŠ— X₃) (Yβ‚‚ βŠ— Y₃)).hom) ≫ (Ξ±_ X₁ Y₁ ((Xβ‚‚ βŠ— X₃) βŠ— Yβ‚‚ βŠ— Y₃)).inv := + (X₁ ◁ (Ξ±_ (Y₁ βŠ— Xβ‚‚ βŠ— X₃) Yβ‚‚ Y₃).hom) ≫ + (X₁ ◁ (Ξ±_ Y₁ (Xβ‚‚ βŠ— X₃) (Yβ‚‚ βŠ— Y₃)).hom) ≫ (Ξ±_ X₁ Y₁ ((Xβ‚‚ βŠ— X₃) βŠ— Yβ‚‚ βŠ— Y₃)).inv := by coherence slice_lhs 9 16 => rw [this] clear this diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index 46bfa19f418f3..14fd9fede21ef 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -363,7 +363,8 @@ elab_rules : tactic MonoidalCategory.whiskerRight_tensor, MonoidalCategory.whiskerRight_id, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_id, MonoidalCategory.comp_whiskerRight, MonoidalCategory.id_whiskerRight, MonoidalCategory.whisker_assoc, - MonoidalCategory.id_tensorHom, MonoidalCategory.tensorHom_id]; simp only [MonoidalCategory.tensorHom_def] + MonoidalCategory.id_tensorHom, MonoidalCategory.tensorHom_id]; + try simp only [MonoidalCategory.tensorHom_def] )) /-- diff --git a/test/CategoryTheory/Coherence.lean b/test/CategoryTheory/Coherence.lean index 315dc57a504f1..89d7330089fef 100644 --- a/test/CategoryTheory/Coherence.lean +++ b/test/CategoryTheory/Coherence.lean @@ -10,11 +10,13 @@ open scoped MonoidalCategory -- Internal tactics +example (X₁ Xβ‚‚ : C) : (πŸ™ X₁) β–· Xβ‚‚ = πŸ™ (X₁ βŠ— Xβ‚‚) := by monoidal_simps + example (X₁ Xβ‚‚ : C) : ((Ξ»_ (πŸ™_ C)).inv βŠ— πŸ™ (X₁ βŠ— Xβ‚‚)) ≫ (Ξ±_ (πŸ™_ C) (πŸ™_ C) (X₁ βŠ— Xβ‚‚)).hom ≫ (πŸ™ (πŸ™_ C) βŠ— (Ξ±_ (πŸ™_ C) X₁ Xβ‚‚).inv) = πŸ™ (πŸ™_ C) βŠ— ((Ξ»_ X₁).inv βŠ— πŸ™ Xβ‚‚) := by - pure_coherence + coherence -- This is just running: -- change projectMap id _ _ (LiftHom.lift (((Ξ»_ (πŸ™_ C)).inv βŠ— πŸ™ (X₁ βŠ— Xβ‚‚)) ≫ -- (Ξ±_ (πŸ™_ C) (πŸ™_ C) (X₁ βŠ— Xβ‚‚)).hom ≫ (πŸ™ (πŸ™_ C) βŠ— (Ξ±_ (πŸ™_ C) X₁ Xβ‚‚).inv))) = From a9529a25a58a5b8be02ec1fefc805a16042b3ea6 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 3 Aug 2023 06:38:04 +0900 Subject: [PATCH 03/62] braided done --- Mathlib/CategoryTheory/Monoidal/Braided.lean | 200 +++++-------------- 1 file changed, 46 insertions(+), 154 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 8650f59576671..b8e9ceeadd7bc 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -113,6 +113,11 @@ theorem braiding_inv_tensor_right (X Y Z : C) : (Ξ²_ X Y).inv β–· Z ≫ (Ξ±_ X Y Z).hom := eq_of_inv_eq_inv (by simp) +theorem braiding_naturality {X X' Y Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y') : + (f βŠ— g) ≫ (braiding Y Y').hom = (braiding X X').hom ≫ (g βŠ— f) := by + rw [tensorHom_def f g, tensorHom_def' g f] + simp_rw [Category.assoc, braiding_naturality_left, braiding_naturality_right_assoc] + end BraidedCategory open Category @@ -451,38 +456,22 @@ def tensor_ΞΌ (X Y : C Γ— C) : (tensor C).obj X βŠ— (tensor C).obj Y ⟢ (tensor (X.1 ◁ (Ξ±_ Y.1 X.2 Y.2).hom) ≫ (Ξ±_ X.1 Y.1 (X.2 βŠ— Y.2)).inv #align category_theory.tensor_ΞΌ CategoryTheory.tensor_ΞΌ -theorem tensor_ΞΌ_def₁ (X₁ Xβ‚‚ Y₁ Yβ‚‚ : C) : - tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) ≫ (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— Yβ‚‚)).hom ≫ (X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ Yβ‚‚).inv) = - (Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— Yβ‚‚)).hom ≫ (X₁ ◁ (Ξ±_ Xβ‚‚ Y₁ Yβ‚‚).inv) ≫ (X₁ ◁ (Ξ²_ Xβ‚‚ Y₁).hom β–· Yβ‚‚) := - by dsimp [tensor_ΞΌ]; simp -#align category_theory.tensor_ΞΌ_def₁ CategoryTheory.tensor_ΞΌ_def₁ - -theorem tensor_ΞΌ_defβ‚‚ (X₁ Xβ‚‚ Y₁ Yβ‚‚ : C) : - (X₁ ◁ (Ξ±_ Xβ‚‚ Y₁ Yβ‚‚).hom) ≫ (Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— Yβ‚‚)).inv ≫ tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) = - (X₁ ◁ (Ξ²_ Xβ‚‚ Y₁).hom β–· Yβ‚‚) ≫ (X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ Yβ‚‚).hom) ≫ (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— Yβ‚‚)).inv := - by dsimp [tensor_ΞΌ]; simp -#align category_theory.tensor_ΞΌ_defβ‚‚ CategoryTheory.tensor_ΞΌ_defβ‚‚ - -theorem tensor_ΞΌ_natural_left {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁: X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (Z₁ Zβ‚‚ : C) : - (f₁ βŠ— fβ‚‚) β–· (Z₁ βŠ— Zβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚) = - tensor_ΞΌ C (X₁, Xβ‚‚) (Z₁, Zβ‚‚) ≫ (f₁ β–· Z₁ βŠ— fβ‚‚ β–· Zβ‚‚) := by - dsimp only [tensor_ΞΌ, prodMonoidal_tensorObj, tensor_obj] - calc - _ = πŸ™ _ βŠ—β‰« - f₁ β–· Xβ‚‚ β–· Z₁ β–· Zβ‚‚ βŠ—β‰« Y₁ ◁ (fβ‚‚ β–· Z₁ ≫ (Ξ²_ Yβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := ?eq1 - _ = πŸ™ _ βŠ—β‰« - (f₁ β–· (Xβ‚‚ βŠ— Z₁) ≫ Y₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« Y₁ ◁ Z₁ ◁ fβ‚‚ β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := ?eq2 - _ = πŸ™ _ βŠ—β‰« - X₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom β–· Zβ‚‚ βŠ—β‰« (f₁ β–· Z₁ β–· (Xβ‚‚ βŠ— Zβ‚‚) ≫ (Y₁ βŠ— Z₁) ◁ fβ‚‚ β–· Zβ‚‚) βŠ—β‰« πŸ™ _ := ?eq3 - _ = _ := ?eq4 - case eq1 => - rw [tensorHom_def']; coherence - case eq2 => - rw [braiding_naturality_left]; coherence - case eq3 => - rw [← whisker_exchange]; coherence - case eq4 => - rw [tensorHom_def']; coherence +-- theorem tensor_ΞΌ_natural_left {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁: X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (Z₁ Zβ‚‚ : C) : +-- (f₁ βŠ— fβ‚‚) β–· (Z₁ βŠ— Zβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚) = +-- tensor_ΞΌ C (X₁, Xβ‚‚) (Z₁, Zβ‚‚) ≫ (f₁ β–· Z₁ βŠ— fβ‚‚ β–· Zβ‚‚) := by +-- dsimp only [tensor_ΞΌ, prodMonoidal_tensorObj, tensor_obj] +-- calc +-- _ = πŸ™ _ βŠ—β‰« +-- f₁ β–· Xβ‚‚ β–· Z₁ β–· Zβ‚‚ βŠ—β‰« Y₁ ◁ (fβ‚‚ β–· Z₁ ≫ (Ξ²_ Yβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := ?eq1 +-- _ = πŸ™ _ βŠ—β‰« +-- (f₁ β–· (Xβ‚‚ βŠ— Z₁) ≫ Y₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« Y₁ ◁ Z₁ ◁ fβ‚‚ β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := ?eq2 +-- _ = πŸ™ _ βŠ—β‰« +-- X₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom β–· Zβ‚‚ βŠ—β‰« (f₁ β–· Z₁ β–· (Xβ‚‚ βŠ— Zβ‚‚) ≫ (Y₁ βŠ— Z₁) ◁ fβ‚‚ β–· Zβ‚‚) βŠ—β‰« πŸ™ _ := ?eq3 +-- _ = _ := ?eq4 +-- case eq1 => rw [tensorHom_def']; coherence +-- case eq2 => rw [braiding_naturality_left]; coherence +-- case eq3 => rw [← whisker_exchange]; coherence +-- case eq4 => rw [tensorHom_def']; coherence theorem tensor_ΞΌ_natural {X₁ Xβ‚‚ Y₁ Yβ‚‚ U₁ Uβ‚‚ V₁ Vβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : U₁ ⟢ V₁) (gβ‚‚ : Uβ‚‚ ⟢ Vβ‚‚) : @@ -501,6 +490,16 @@ theorem tensor_ΞΌ_natural {X₁ Xβ‚‚ Y₁ Yβ‚‚ U₁ Uβ‚‚ V₁ Vβ‚‚ : C} (f₁ : simp only [assoc] #align category_theory.tensor_ΞΌ_natural CategoryTheory.tensor_ΞΌ_natural +theorem tensor_ΞΌ_natural_left {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁: X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (Z₁ Zβ‚‚ : C) : + (f₁ βŠ— fβ‚‚) β–· (Z₁ βŠ— Zβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚) = + tensor_ΞΌ C (X₁, Xβ‚‚) (Z₁, Zβ‚‚) ≫ (f₁ β–· Z₁ βŠ— fβ‚‚ β–· Zβ‚‚) := by + convert tensor_ΞΌ_natural C f₁ fβ‚‚ (πŸ™ Z₁) (πŸ™ Zβ‚‚) using 1 <;> simp [id_tensorHom, tensorHom_id] + +theorem tensor_ΞΌ_natural_right (Z₁ Zβ‚‚ : C) {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) : + (Z₁ βŠ— Zβ‚‚) ◁ (f₁ βŠ— fβ‚‚) ≫ tensor_ΞΌ C (Z₁, Zβ‚‚) (Y₁, Yβ‚‚) = + tensor_ΞΌ C (Z₁, Zβ‚‚) (X₁, Xβ‚‚) ≫ (Z₁ ◁ f₁ βŠ— Zβ‚‚ ◁ fβ‚‚) := by + convert tensor_ΞΌ_natural C (πŸ™ Z₁) (πŸ™ Zβ‚‚) f₁ fβ‚‚ using 1 <;> simp [id_tensorHom, tensorHom_id] + theorem tensor_left_unitality (X₁ Xβ‚‚ : C) : (Ξ»_ (X₁ βŠ— Xβ‚‚)).hom = ((Ξ»_ (πŸ™_ C)).inv β–· (X₁ βŠ— Xβ‚‚)) ≫ @@ -552,43 +551,18 @@ theorem tensor_associativity (X₁ Xβ‚‚ Y₁ Yβ‚‚ Z₁ Zβ‚‚ : C) : case eq2 => rw [← whisker_exchange]; coherence #align category_theory.tensor_associativity CategoryTheory.tensor_associativity -@[simp] -def tensorLaxMonoidal : MonoidalFunctor (C Γ— C) C := +/-- The tensor product functor from `C Γ— C` to `C` as a monoidal functor. -/ +@[simps!] +def tensorMonoidal : MonoidalFunctor (C Γ— C) C := { tensor C with Ξ΅ := (Ξ»_ (πŸ™_ C)).inv ΞΌ := tensor_ΞΌ C - ΞΌ_natural_left := fun f Z ↦ tensor_ΞΌ_natural_left C f.1 f.2 Z.1 Z.2 - ΞΌ_natural_right := _ + ΞΌ_natural_left := fun f Z => tensor_ΞΌ_natural_left C f.1 f.2 Z.1 Z.2 + ΞΌ_natural_right := fun Z f => tensor_ΞΌ_natural_right C Z.1 Z.2 f.1 f.2 associativity := fun X Y Z => tensor_associativity C X.1 X.2 Y.1 Y.2 Z.1 Z.2 left_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_left_unitality C X₁ Xβ‚‚ right_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_right_unitality C X₁ Xβ‚‚ - Ξ΅_isIso := inferInstanceAs (IsIso (Ξ»_ (πŸ™_ C)).inv) ΞΌ_isIso := by dsimp [tensor_ΞΌ]; infer_instance } - -- (F := tensor C) - -- (Ξ΅ := (Ξ»_ (πŸ™_ C)).inv) - -- (ΞΌ := fun X Y => tensor_ΞΌ C X Y) - -- (ΞΌ_natural := fun f g => tensor_ΞΌ_natural C f.1 f.2 g.1 g.2) - -- (associativity := fun X Y Z => tensor_associativity C X.1 X.2 Y.1 Y.2 Z.1 Z.2) - -- (left_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_left_unitality C X₁ Xβ‚‚) - -- (right_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_right_unitality C X₁ Xβ‚‚) - --- @[simp] --- def tensorLaxMonoidal : LaxMonoidalFunctor (C Γ— C) C := LaxMonoidalFunctor.ofTensorHom --- (F := tensor C) --- (Ξ΅ := (Ξ»_ (πŸ™_ C)).inv) --- (ΞΌ := fun X Y => tensor_ΞΌ C X Y) --- (ΞΌ_natural := fun f g => tensor_ΞΌ_natural C f.1 f.2 g.1 g.2) --- (associativity := fun X Y Z => tensor_associativity C X.1 X.2 Y.1 Y.2 Z.1 Z.2) --- (left_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_left_unitality C X₁ Xβ‚‚) --- (right_unitality := fun ⟨X₁, Xβ‚‚βŸ© => tensor_right_unitality C X₁ Xβ‚‚) - -/-- The tensor product functor from `C Γ— C` to `C` as a monoidal functor. -/ -@[simps!] -def tensorMonoidal : MonoidalFunctor (C Γ— C) C := - { tensorLaxMonoidal C with - ΞΌ_isIso := by dsimp [tensor_ΞΌ]; infer_instance - Ξ΅_isIso := by dsimp; infer_instance } -#align category_theory.tensor_monoidal CategoryTheory.tensorMonoidal theorem leftUnitor_monoidal (X₁ Xβ‚‚ : C) : (Ξ»_ X₁).hom βŠ— (Ξ»_ Xβ‚‚).hom = @@ -620,103 +594,21 @@ theorem rightUnitor_monoidal (X₁ Xβ‚‚ : C) : coherence #align category_theory.right_unitor_monoidal CategoryTheory.rightUnitor_monoidal -theorem associator_monoidal_aux (W X Y Z : C) : - (πŸ™ W βŠ— (Ξ²_ X (Y βŠ— Z)).hom) ≫ - (πŸ™ W βŠ— (Ξ±_ Y Z X).hom) ≫ (Ξ±_ W Y (Z βŠ— X)).inv ≫ ((Ξ²_ W Y).hom βŠ— πŸ™ (Z βŠ— X)) = - (Ξ±_ W X (Y βŠ— Z)).inv ≫ - (Ξ±_ (W βŠ— X) Y Z).inv ≫ - ((Ξ²_ (W βŠ— X) Y).hom βŠ— πŸ™ Z) ≫ - ((Ξ±_ Y W X).inv βŠ— πŸ™ Z) ≫ (Ξ±_ (Y βŠ— W) X Z).hom ≫ (πŸ™ (Y βŠ— W) βŠ— (Ξ²_ X Z).hom) := by - slice_rhs 1 2 => rw [← pentagon_inv'] - slice_rhs 3 5 => rw [← tensor_comp, ← tensor_comp, hexagon_reverse, tensor_comp, tensor_comp] - slice_rhs 5 6 => rw [associator_naturality] - slice_rhs 6 7 => rw [tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] - slice_rhs 2 3 => rw [← associator_inv_naturality] - slice_rhs 3 5 => rw [pentagon_inv_inv_hom] - slice_rhs 4 5 => rw [← tensor_id, ← associator_inv_naturality] - slice_rhs 2 4 => rw [← tensor_comp, ← tensor_comp, ← hexagon_forward, tensor_comp, tensor_comp] - simp -#align category_theory.associator_monoidal_aux CategoryTheory.associator_monoidal_aux - -set_option maxHeartbeats 400000 in theorem associator_monoidal (X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C) : tensor_ΞΌ C (X₁ βŠ— Xβ‚‚, X₃) (Y₁ βŠ— Yβ‚‚, Y₃) ≫ (tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁, Yβ‚‚) β–· (X₃ βŠ— Y₃)) ≫ (Ξ±_ (X₁ βŠ— Y₁) (Xβ‚‚ βŠ— Yβ‚‚) (X₃ βŠ— Y₃)).hom = ((Ξ±_ X₁ Xβ‚‚ X₃).hom βŠ— (Ξ±_ Y₁ Yβ‚‚ Y₃).hom) ≫ tensor_ΞΌ C (X₁, Xβ‚‚ βŠ— X₃) (Y₁, Yβ‚‚ βŠ— Y₃) ≫ ((X₁ βŠ— Y₁) ◁ tensor_ΞΌ C (Xβ‚‚, X₃) (Yβ‚‚, Y₃)) := by - have : - (Ξ±_ (X₁ βŠ— Y₁) (Xβ‚‚ βŠ— Yβ‚‚) (X₃ βŠ— Y₃)).hom = - ((Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— Yβ‚‚)).hom β–· (X₃ βŠ— Y₃)) ≫ - ((X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ Yβ‚‚).inv) β–· (X₃ βŠ— Y₃)) ≫ - (Ξ±_ (X₁ βŠ— (Y₁ βŠ— Xβ‚‚) βŠ— Yβ‚‚) X₃ Y₃).inv ≫ - ((Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚) βŠ— Yβ‚‚) X₃).hom β–· Y₃) ≫ - ((X₁ ◁ (Ξ±_ (Y₁ βŠ— Xβ‚‚) Yβ‚‚ X₃).hom) β–· Y₃) ≫ - (Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚) βŠ— Yβ‚‚ βŠ— X₃) Y₃).hom ≫ - (X₁ ◁ (Ξ±_ (Y₁ βŠ— Xβ‚‚) (Yβ‚‚ βŠ— X₃) Y₃).hom) ≫ - (X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ ((Yβ‚‚ βŠ— X₃) βŠ— Y₃)).hom) ≫ - (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— (Yβ‚‚ βŠ— X₃) βŠ— Y₃)).inv ≫ - ((X₁ βŠ— Y₁) ◁ Xβ‚‚ ◁ (Ξ±_ Yβ‚‚ X₃ Y₃).hom) ≫ - ((X₁ βŠ— Y₁) ◁ (Ξ±_ Xβ‚‚ Yβ‚‚ (X₃ βŠ— Y₃)).inv) := - by coherence - rw [this]; clear this - slice_lhs 2 4 => rw [← tensor_comp, ← tensor_comp, tensor_ΞΌ_def₁, tensor_comp, tensor_comp] - slice_lhs 4 5 => rw [← tensor_id, associator_inv_naturality] - slice_lhs 5 6 => rw [← tensor_comp, associator_naturality, tensor_comp] - slice_lhs 6 7 => - rw [← tensor_comp, ← tensor_comp, associator_naturality, tensor_comp, tensor_comp] - have : - ((Ξ±_ X₁ Xβ‚‚ (Y₁ βŠ— Yβ‚‚)).hom β–· (X₃ βŠ— Y₃)) ≫ - ((X₁ ◁ (Ξ±_ Xβ‚‚ Y₁ Yβ‚‚).inv) β–· (X₃ βŠ— Y₃)) ≫ - (Ξ±_ (X₁ βŠ— (Xβ‚‚ βŠ— Y₁) βŠ— Yβ‚‚) X₃ Y₃).inv ≫ - ((Ξ±_ X₁ ((Xβ‚‚ βŠ— Y₁) βŠ— Yβ‚‚) X₃).hom β–· Y₃) ≫ ((X₁ ◁ (Ξ±_ (Xβ‚‚ βŠ— Y₁) Yβ‚‚ X₃).hom) β–· Y₃) = - (Ξ±_ (X₁ βŠ— Xβ‚‚) (Y₁ βŠ— Yβ‚‚) (X₃ βŠ— Y₃)).hom ≫ - ((X₁ βŠ— Xβ‚‚) ◁ (Ξ±_ (Y₁ βŠ— Yβ‚‚) X₃ Y₃).inv) ≫ - (Ξ±_ X₁ Xβ‚‚ (((Y₁ βŠ— Yβ‚‚) βŠ— X₃) βŠ— Y₃)).hom ≫ - (X₁ ◁ (Ξ±_ Xβ‚‚ ((Y₁ βŠ— Yβ‚‚) βŠ— X₃) Y₃).inv) ≫ - (Ξ±_ X₁ (Xβ‚‚ βŠ— (Y₁ βŠ— Yβ‚‚) βŠ— X₃) Y₃).inv ≫ - ((πŸ™ X₁ β–· Xβ‚‚ βŠ— (Ξ±_ Y₁ Yβ‚‚ X₃).hom) β–· Y₃) ≫ - ((X₁ ◁ (Ξ±_ Xβ‚‚ Y₁ (Yβ‚‚ βŠ— X₃)).inv) β–· Y₃) := - by coherence - slice_lhs 2 6 => rw [this] - clear this - slice_lhs 1 3 => rw [tensor_ΞΌ_def₁] - slice_lhs 3 4 => rw [← tensor_id, associator_naturality] - slice_lhs 4 5 => rw [← tensor_comp, associator_inv_naturality, tensor_comp] - slice_lhs 5 6 => rw [associator_inv_naturality] - slice_lhs 6 9 => - rw [← tensor_comp, ← tensor_comp, ← tensor_comp, ← tensor_comp, ← tensor_comp, ← tensor_comp, - tensor_id, associator_monoidal_aux, ← id_comp (πŸ™ X₁ ≫ πŸ™ X₁ ≫ πŸ™ X₁ ≫ πŸ™ X₁), ← - id_comp (πŸ™ X₁ ≫ πŸ™ X₁ ≫ πŸ™ X₁ ≫ πŸ™ X₁ ≫ πŸ™ X₁), ← id_comp (πŸ™ Y₃ ≫ πŸ™ Y₃ ≫ πŸ™ Y₃ ≫ πŸ™ Y₃), ← - id_comp (πŸ™ Y₃ ≫ πŸ™ Y₃ ≫ πŸ™ Y₃ ≫ πŸ™ Y₃ ≫ πŸ™ Y₃), tensor_comp, tensor_comp, tensor_comp, - tensor_comp, tensor_comp, tensor_comp, tensor_comp, tensor_comp, tensor_comp, tensor_comp] - slice_lhs 11 12 => rw [associator_naturality] - slice_lhs 12 13 => rw [← tensor_comp, associator_naturality, tensor_comp] - slice_lhs 13 14 => rw [← tensor_comp, ← tensor_id, associator_naturality, tensor_comp] - slice_lhs 14 15 => rw [associator_inv_naturality] - slice_lhs 15 17 => - rw [tensor_id, ← tensor_comp, ← tensor_comp, ← tensor_ΞΌ_defβ‚‚, tensor_comp, tensor_comp] - have : - ((X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ X₃).inv β–· Yβ‚‚) β–· Y₃) ≫ - ((X₁ ◁ (Ξ±_ (Y₁ βŠ— Xβ‚‚) X₃ Yβ‚‚).hom) β–· Y₃) ≫ - (Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚) βŠ— X₃ βŠ— Yβ‚‚) Y₃).hom ≫ - (X₁ ◁ (Ξ±_ (Y₁ βŠ— Xβ‚‚) (X₃ βŠ— Yβ‚‚) Y₃).hom) ≫ - (X₁ ◁ (Ξ±_ Y₁ Xβ‚‚ ((X₃ βŠ— Yβ‚‚) βŠ— Y₃)).hom) ≫ - (Ξ±_ X₁ Y₁ (Xβ‚‚ βŠ— (X₃ βŠ— Yβ‚‚) βŠ— Y₃)).inv ≫ - (πŸ™ (X₁ βŠ— Y₁) β–· Xβ‚‚ βŠ— (Ξ±_ X₃ Yβ‚‚ Y₃).hom) ≫ - (πŸ™ (X₁ βŠ— Y₁) βŠ— (Ξ±_ Xβ‚‚ X₃ (Yβ‚‚ βŠ— Y₃)).inv) = - (Ξ±_ X₁ ((Y₁ βŠ— Xβ‚‚ βŠ— X₃) βŠ— Yβ‚‚) Y₃).hom ≫ - (X₁ ◁ (Ξ±_ (Y₁ βŠ— Xβ‚‚ βŠ— X₃) Yβ‚‚ Y₃).hom) ≫ - (X₁ ◁ (Ξ±_ Y₁ (Xβ‚‚ βŠ— X₃) (Yβ‚‚ βŠ— Y₃)).hom) ≫ (Ξ±_ X₁ Y₁ ((Xβ‚‚ βŠ— X₃) βŠ— Yβ‚‚ βŠ— Y₃)).inv := - by coherence - slice_lhs 9 16 => rw [this] - clear this - slice_lhs 8 9 => rw [associator_naturality] - slice_lhs 9 10 => rw [← tensor_comp, associator_naturality, tensor_comp] - slice_lhs 10 12 => rw [tensor_id, ← tensor_ΞΌ_defβ‚‚] - dsimp only [tensor_obj, prodMonoidal_tensorObj] - simp_rw [← Category.assoc] - congr 2 - coherence + dsimp [tensor_ΞΌ] + simp only [tensor_whiskerLeft, braiding_tensor_right, comp_whiskerRight, whisker_assoc, assoc, + MonoidalCategory.whiskerLeft_comp, Iso.inv_hom_id_assoc, whiskerRight_tensor, braiding_tensor_left] + calc + _ = πŸ™ _ βŠ—β‰« + X₁ ◁ Xβ‚‚ ◁ (Ξ²_ X₃ Y₁).hom β–· Yβ‚‚ β–· Y₃ βŠ—β‰« + X₁ ◁ ((Xβ‚‚ βŠ— Y₁) ◁ (Ξ²_ X₃ Yβ‚‚).hom ≫ (Ξ²_ Xβ‚‚ Y₁).hom β–· (Yβ‚‚ βŠ— X₃)) β–· Y₃ βŠ—β‰« πŸ™ _ := ?eq1 + _ = _ := ?eq2 + case eq1 => coherence + case eq2 => rw [whisker_exchange]; coherence #align category_theory.associator_monoidal CategoryTheory.associator_monoidal end Tensor From 1b888e4bb6477e61be04543497fbd5bd900ab3d5 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 3 Aug 2023 17:07:12 +0900 Subject: [PATCH 04/62] wip --- .../ModuleCat/Monoidal/Symmetric.lean | 15 +- .../Closed/FunctorCategory.lean | 8 +- Mathlib/CategoryTheory/Closed/Ideal.lean | 2 +- Mathlib/CategoryTheory/Enriched/Basic.lean | 42 +++-- Mathlib/CategoryTheory/Monoidal/Braided.lean | 25 ++- Mathlib/CategoryTheory/Monoidal/Center.lean | 53 ++---- Mathlib/CategoryTheory/Monoidal/CommMon_.lean | 2 + .../Monoidal/FunctorCategory.lean | 37 ++++- Mathlib/CategoryTheory/Monoidal/Limits.lean | 23 +-- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 152 +++++++++++------- .../OfChosenFiniteProducts/Symmetric.lean | 3 +- .../Monoidal/OfHasFiniteProducts.lean | 22 ++- Mathlib/CategoryTheory/Monoidal/Opposite.lean | 40 +++-- .../CategoryTheory/Monoidal/Rigid/Basic.lean | 4 + .../Monoidal/Rigid/FunctorCategory.lean | 24 +-- .../Monoidal/Rigid/OfEquivalence.lean | 21 ++- .../CategoryTheory/Monoidal/Transport.lean | 17 +- .../CategoryTheory/Monoidal/Types/Basic.lean | 10 ++ Mathlib/RepresentationTheory/Action.lean | 51 +++--- Mathlib/Tactic/CategoryTheory/Coherence.lean | 11 +- 20 files changed, 361 insertions(+), 201 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean index 63f531b5244da..ca76373143f9f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean @@ -38,6 +38,18 @@ theorem braiding_naturality {X₁ Xβ‚‚ Y₁ Yβ‚‚ : ModuleCat.{u} R} (f : X₁ set_option linter.uppercaseLean3 false in #align Module.monoidal_category.braiding_naturality ModuleCat.MonoidalCategory.braiding_naturality +@[simp] +theorem braiding_naturality_left {X Y : ModuleCat R} (f : X ⟢ Y) (Z : ModuleCat R) : + f β–· Z ≫ (braiding Y Z).hom = (braiding X Z).hom ≫ Z ◁ f := by + simp_rw [← id_tensorHom, tensorHom_id] + apply braiding_naturality + +@[simp] +theorem braiding_naturality_right (X : ModuleCat R) {Y Z : ModuleCat R} (f : Y ⟢ Z) : + X ◁ f ≫ (braiding X Z).hom = (braiding X Y).hom ≫ f β–· X := by + simp_rw [← id_tensorHom, tensorHom_id] + apply braiding_naturality + @[simp] theorem hexagon_forward (X Y Z : ModuleCat.{u} R) : (Ξ±_ X Y Z).hom ≫ (braiding X _).hom ≫ (Ξ±_ Y Z X).hom = @@ -64,7 +76,8 @@ attribute [local ext] TensorProduct.ext /-- The symmetric monoidal structure on `Module R`. -/ instance symmetricCategory : SymmetricCategory (ModuleCat.{u} R) where braiding := braiding - braiding_naturality f g := braiding_naturality f g + braiding_naturality_left := braiding_naturality_left + braiding_naturality_right := braiding_naturality_right hexagon_forward := hexagon_forward hexagon_reverse := hexagon_reverse -- porting note: this proof was automatic in Lean3 diff --git a/Mathlib/CategoryTheory/Closed/FunctorCategory.lean b/Mathlib/CategoryTheory/Closed/FunctorCategory.lean index daa98d171a481..10114ac1ede7c 100644 --- a/Mathlib/CategoryTheory/Closed/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Closed/FunctorCategory.lean @@ -42,7 +42,9 @@ def closedUnit (F : D β₯€ C) : 𝟭 (D β₯€ C) ⟢ tensorLeft F β‹™ closedIhom F simp only [ihom.coev_naturality, closedIhom_obj_map, Monoidal.tensorObj_map] dsimp rw [coev_app_comp_pre_app_assoc, ← Functor.map_comp] - simp } + simp + sorry + } #align category_theory.functor.closed_unit CategoryTheory.Functor.closedUnit /-- Auxiliary definition for `CategoryTheory.Functor.closed`. @@ -56,7 +58,9 @@ def closedCounit (F : D β₯€ C) : closedIhom F β‹™ tensorLeft F ⟢ 𝟭 (D β₯€ C dsimp simp only [closedIhom_obj_map, pre_comm_ihom_map] rw [← tensor_id_comp_id_tensor, id_tensor_comp] - simp } + simp + sorry + } #align category_theory.functor.closed_counit CategoryTheory.Functor.closedCounit /-- If `C` is a monoidal closed category and `D` is a groupoid, then every functor `F : D β₯€ C` is diff --git a/Mathlib/CategoryTheory/Closed/Ideal.lean b/Mathlib/CategoryTheory/Closed/Ideal.lean index 7a1a26c83e924..06a06041c74ff 100644 --- a/Mathlib/CategoryTheory/Closed/Ideal.lean +++ b/Mathlib/CategoryTheory/Closed/Ideal.lean @@ -161,7 +161,7 @@ def cartesianClosedOfReflective : CartesianClosed D := Adjunction.rightAdjointPreservesLimits.{0, 0} (Adjunction.ofRightAdjoint i) apply asIso (prodComparison i B X) Β· dsimp [asIso] - rw [prodComparison_natural, Functor.map_id] + erw [prodComparison_natural, Functor.map_id] Β· apply (exponentialIdealReflective i _).symm } } } #align category_theory.cartesian_closed_of_reflective CategoryTheory.cartesianClosedOfReflective diff --git a/Mathlib/CategoryTheory/Enriched/Basic.lean b/Mathlib/CategoryTheory/Enriched/Basic.lean index 5499018e5c206..9dd6a928a0456 100644 --- a/Mathlib/CategoryTheory/Enriched/Basic.lean +++ b/Mathlib/CategoryTheory/Enriched/Basic.lean @@ -53,10 +53,10 @@ class EnrichedCategory (C : Type u₁) where Hom : C β†’ C β†’ V id (X : C) : πŸ™_ V ⟢ Hom X X comp (X Y Z : C) : Hom X Y βŠ— Hom Y Z ⟢ Hom X Z - id_comp (X Y : C) : (Ξ»_ (Hom X Y)).inv ≫ (id X βŠ— πŸ™ _) ≫ comp X X Y = πŸ™ _ := by aesop_cat - comp_id (X Y : C) : (ρ_ (Hom X Y)).inv ≫ (πŸ™ _ βŠ— id Y) ≫ comp X Y Y = πŸ™ _ := by aesop_cat - assoc (W X Y Z : C) : (Ξ±_ _ _ _).inv ≫ (comp W X Y βŠ— πŸ™ _) ≫ comp W Y Z = - (πŸ™ _ βŠ— comp X Y Z) ≫ comp W X Z := by aesop_cat + id_comp (X Y : C) : (Ξ»_ (Hom X Y)).inv ≫ (id X β–· _) ≫ comp X X Y = πŸ™ _ := by aesop_cat + comp_id (X Y : C) : (ρ_ (Hom X Y)).inv ≫ (_ ◁ id Y) ≫ comp X Y Y = πŸ™ _ := by aesop_cat + assoc (W X Y Z : C) : (Ξ±_ _ _ _).inv ≫ (comp W X Y β–· _) ≫ comp W Y Z = + (_ ◁ comp X Y Z) ≫ comp W X Z := by aesop_cat #align category_theory.enriched_category CategoryTheory.EnrichedCategory notation X " ⟢[" V "] " Y:10 => (EnrichedCategory.Hom X Y : V) @@ -78,20 +78,20 @@ def eComp (X Y Z : C) : ((X ⟢[V] Y) βŠ— Y ⟢[V] Z) ⟢ X ⟢[V] Z := -- We don't just use `restate_axiom` here; that would leave `V` as an implicit argument. @[reassoc (attr := simp)] theorem e_id_comp (X Y : C) : - (Ξ»_ (X ⟢[V] Y)).inv ≫ (eId V X βŠ— πŸ™ _) ≫ eComp V X X Y = πŸ™ (X ⟢[V] Y) := + (Ξ»_ (X ⟢[V] Y)).inv ≫ (eId V X β–· _) ≫ eComp V X X Y = πŸ™ (X ⟢[V] Y) := EnrichedCategory.id_comp X Y #align category_theory.e_id_comp CategoryTheory.e_id_comp @[reassoc (attr := simp)] theorem e_comp_id (X Y : C) : - (ρ_ (X ⟢[V] Y)).inv ≫ (πŸ™ _ βŠ— eId V Y) ≫ eComp V X Y Y = πŸ™ (X ⟢[V] Y) := + (ρ_ (X ⟢[V] Y)).inv ≫ (_ ◁ eId V Y) ≫ eComp V X Y Y = πŸ™ (X ⟢[V] Y) := EnrichedCategory.comp_id X Y #align category_theory.e_comp_id CategoryTheory.e_comp_id @[reassoc (attr := simp)] theorem e_assoc (W X Y Z : C) : - (Ξ±_ _ _ _).inv ≫ (eComp V W X Y βŠ— πŸ™ _) ≫ eComp V W Y Z = - (πŸ™ _ βŠ— eComp V X Y Z) ≫ eComp V W X Z := + (Ξ±_ _ _ _).inv ≫ (eComp V W X Y β–· _) ≫ eComp V W Y Z = + (_ ◁ eComp V X Y Z) ≫ eComp V W X Z := EnrichedCategory.assoc W X Y Z #align category_theory.e_assoc CategoryTheory.e_assoc @@ -114,18 +114,26 @@ instance (F : LaxMonoidalFunctor V W) : EnrichedCategory W (TransportEnrichment id := fun X : C => F.Ξ΅ ≫ F.map (eId V X) comp := fun X Y Z : C => F.ΞΌ _ _ ≫ F.map (eComp V X Y Z) id_comp X Y := by - rw [comp_tensor_id, Category.assoc, ← F.toFunctor.map_id, F.ΞΌ_natural_assoc, - F.toFunctor.map_id, F.left_unitality_inv_assoc, ← F.toFunctor.map_comp, ← - F.toFunctor.map_comp, e_id_comp, F.toFunctor.map_id] + simp only [comp_whiskerRight, Category.assoc, LaxMonoidalFunctor.ΞΌ_natural_left_assoc, + LaxMonoidalFunctor.left_unitality_inv_assoc] + simp_rw [← F.map_comp] + convert F.map_id _ + simp comp_id X Y := by - rw [id_tensor_comp, Category.assoc, ← F.toFunctor.map_id, F.ΞΌ_natural_assoc, - F.toFunctor.map_id, F.right_unitality_inv_assoc, ← F.toFunctor.map_comp, ← - F.toFunctor.map_comp, e_comp_id, F.toFunctor.map_id] + simp only [MonoidalCategory.whiskerLeft_comp, Category.assoc, + LaxMonoidalFunctor.ΞΌ_natural_right_assoc, + LaxMonoidalFunctor.right_unitality_inv_assoc] + simp_rw [← F.map_comp] + convert F.map_id _ + simp assoc P Q R S := by + simp only [← id_tensorHom, ←tensorHom_id] rw [comp_tensor_id, Category.assoc, ← F.toFunctor.map_id, F.ΞΌ_natural_assoc, - F.toFunctor.map_id, ← F.associativity_inv_assoc, ← F.toFunctor.map_comp, ← - F.toFunctor.map_comp, e_assoc, id_tensor_comp, Category.assoc, ← F.toFunctor.map_id, + F.toFunctor.map_id, ← F.associativity_inv'_assoc, ← F.toFunctor.map_comp, ← + F.toFunctor.map_comp, id_tensorHom, tensorHom_id, e_assoc, id_tensor_comp, + Category.assoc, ← F.toFunctor.map_id, F.ΞΌ_natural_assoc, F.toFunctor.map_comp] + simp [id_tensorHom, tensorHom_id] end @@ -141,6 +149,8 @@ def categoryOfEnrichedCategoryType (C : Type u₁) [π’ž : EnrichedCategory (Typ assoc f g h := (congr_fun (e_assoc (Type v) _ _ _ _) ⟨f, g, h⟩ : _) #align category_theory.category_of_enriched_category_Type CategoryTheory.categoryOfEnrichedCategoryType +attribute [local simp] id_tensorHom tensorHom_id + /-- Construct a `Type v`-enriched category from an honest category. -/ def enrichedCategoryTypeOfCategory (C : Type u₁) [π’ž : Category.{v} C] : EnrichedCategory (Type v) C diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index b8e9ceeadd7bc..1fb0f659eabcf 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -113,6 +113,7 @@ theorem braiding_inv_tensor_right (X Y Z : C) : (Ξ²_ X Y).inv β–· Z ≫ (Ξ±_ X Y Z).hom := eq_of_inv_eq_inv (by simp) +@[reassoc (attr := simp)] theorem braiding_naturality {X X' Y Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y') : (f βŠ— g) ≫ (braiding Y Y').hom = (braiding X X').hom ≫ (g βŠ— f) := by rw [tensorHom_def f g, tensorHom_def' g f] @@ -260,23 +261,27 @@ theorem braiding_rightUnitor_auxβ‚‚ (X : C) : #align category_theory.braiding_right_unitor_auxβ‚‚ CategoryTheory.braiding_rightUnitor_auxβ‚‚ -@[simp] theorem braiding_rightUnitor (X : C) : (Ξ²_ (πŸ™_ C) X).hom ≫ (ρ_ X).hom = (Ξ»_ X).hom := by rw [← whiskerLeft_iff, MonoidalCategory.whiskerLeft_comp, braiding_rightUnitor_auxβ‚‚] #align category_theory.braiding_right_unitor CategoryTheory.braiding_rightUnitor @[simp] +theorem braiding_tensorUnit_left (X : C) : (Ξ²_ (πŸ™_ C) X).hom = (Ξ»_ X).hom ≫ (ρ_ X).inv := by + simp [← braiding_rightUnitor] + theorem leftUnitor_inv_braiding (X : C) : (Ξ»_ X).inv ≫ (Ξ²_ (πŸ™_ C) X).hom = (ρ_ X).inv := by - apply (cancel_mono (ρ_ X).hom).1 - simp only [assoc, braiding_rightUnitor, Iso.inv_hom_id] + simp #align category_theory.left_unitor_inv_braiding CategoryTheory.leftUnitor_inv_braiding -@[simp] theorem rightUnitor_inv_braiding (X : C) : (ρ_ X).inv ≫ (Ξ²_ X (πŸ™_ C)).hom = (Ξ»_ X).inv := by apply (cancel_mono (Ξ»_ X).hom).1 simp only [assoc, braiding_leftUnitor, Iso.inv_hom_id] #align category_theory.right_unitor_inv_braiding CategoryTheory.rightUnitor_inv_braiding +@[simp] +theorem braiding_tensorUnit_right (X : C) : (Ξ²_ X (πŸ™_ C)).hom = (ρ_ X).hom ≫ (Ξ»_ X).inv := by + simp [← rightUnitor_inv_braiding] + end /-- @@ -448,8 +453,16 @@ end CommMonoid section Tensor +-- /-- The strength of the tensor product functor from `C Γ— C` to `C`. -/ +-- def tensor_ΞΌ (X Y : C Γ— C) : (tensor C).obj X βŠ— (tensor C).obj Y ⟢ (tensor C).obj (X βŠ— Y) := +-- (Ξ±_ X.1 X.2 (Y.1 βŠ— Y.2)).hom ≫ +-- (X.1 ◁ (Ξ±_ X.2 Y.1 Y.2).inv) ≫ +-- (X.1 ◁ (Ξ²_ X.2 Y.1).hom β–· Y.2) ≫ +-- (X.1 ◁ (Ξ±_ Y.1 X.2 Y.2).hom) ≫ (Ξ±_ X.1 Y.1 (X.2 βŠ— Y.2)).inv +-- #align category_theory.tensor_ΞΌ CategoryTheory.tensor_ΞΌ + /-- The strength of the tensor product functor from `C Γ— C` to `C`. -/ -def tensor_ΞΌ (X Y : C Γ— C) : (tensor C).obj X βŠ— (tensor C).obj Y ⟢ (tensor C).obj (X βŠ— Y) := +def tensor_ΞΌ (X Y : C Γ— C) : (X.1 βŠ— X.2) βŠ— Y.1 βŠ— Y.2 ⟢ (X.1 βŠ— Y.1) βŠ— X.2 βŠ— Y.2 := (Ξ±_ X.1 X.2 (Y.1 βŠ— Y.2)).hom ≫ (X.1 ◁ (Ξ±_ X.2 Y.1 Y.2).inv) ≫ (X.1 ◁ (Ξ²_ X.2 Y.1).hom β–· Y.2) ≫ @@ -490,11 +503,13 @@ theorem tensor_ΞΌ_natural {X₁ Xβ‚‚ Y₁ Yβ‚‚ U₁ Uβ‚‚ V₁ Vβ‚‚ : C} (f₁ : simp only [assoc] #align category_theory.tensor_ΞΌ_natural CategoryTheory.tensor_ΞΌ_natural +@[reassoc] theorem tensor_ΞΌ_natural_left {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁: X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (Z₁ Zβ‚‚ : C) : (f₁ βŠ— fβ‚‚) β–· (Z₁ βŠ— Zβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚) = tensor_ΞΌ C (X₁, Xβ‚‚) (Z₁, Zβ‚‚) ≫ (f₁ β–· Z₁ βŠ— fβ‚‚ β–· Zβ‚‚) := by convert tensor_ΞΌ_natural C f₁ fβ‚‚ (πŸ™ Z₁) (πŸ™ Zβ‚‚) using 1 <;> simp [id_tensorHom, tensorHom_id] +@[reassoc] theorem tensor_ΞΌ_natural_right (Z₁ Zβ‚‚ : C) {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) : (Z₁ βŠ— Zβ‚‚) ◁ (f₁ βŠ— fβ‚‚) ≫ tensor_ΞΌ C (Z₁, Zβ‚‚) (Y₁, Yβ‚‚) = tensor_ΞΌ C (Z₁, Zβ‚‚) (X₁, Xβ‚‚) ≫ (Z₁ ◁ f₁ βŠ— Zβ‚‚ ◁ fβ‚‚) := by diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index f5ac70d24356e..84eb1981adaba 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -185,14 +185,14 @@ def whiskerLeft (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) : calc _ = πŸ™ _ βŠ—β‰« X.fst ◁ (f.f β–· U ≫ (HalfBraiding.Ξ² Yβ‚‚.snd U).hom) βŠ—β‰« - (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst βŠ—β‰« πŸ™ _ := ?_ + (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst βŠ—β‰« πŸ™ _ := ?eq1 _ = πŸ™ _ βŠ—β‰« X.fst ◁ (HalfBraiding.Ξ² Y₁.snd U).hom βŠ—β‰« - ((X.fst βŠ— U) ◁ f.f ≫ (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst) βŠ—β‰« πŸ™ _ := ?eq1 - _ = _ := ?eq2 - case eq1 => rw [f.comm]; coherence - case eq2 => rw [whisker_exchange]; coherence - any_goals coherence + ((X.fst βŠ— U) ◁ f.f ≫ (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst) βŠ—β‰« πŸ™ _ := ?eq2 + _ = _ := ?eq3 + case eq1 => coherence + case eq2 => rw [f.comm]; coherence + case eq3 => rw [whisker_exchange]; coherence /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ def whiskerRight {X₁ Xβ‚‚ : Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) : @@ -204,14 +204,14 @@ def whiskerRight {X₁ Xβ‚‚ : Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) : calc _ = πŸ™ _ βŠ—β‰« (f.f β–· (Y.fst βŠ— U) ≫ Xβ‚‚.fst ◁ (HalfBraiding.Ξ² Y.snd U).hom) βŠ—β‰« - (HalfBraiding.Ξ² Xβ‚‚.snd U).hom β–· Y.fst βŠ—β‰« πŸ™ _ := ?_ + (HalfBraiding.Ξ² Xβ‚‚.snd U).hom β–· Y.fst βŠ—β‰« πŸ™ _ := ?eq1 _ = πŸ™ _ βŠ—β‰« X₁.fst ◁ (HalfBraiding.Ξ² Y.snd U).hom βŠ—β‰« - (f.f β–· U ≫ (HalfBraiding.Ξ² Xβ‚‚.snd U).hom) β–· Y.fst βŠ—β‰« πŸ™ _ := ?eq1 - _ = _ := ?eq2 - case eq1 => rw [← whisker_exchange]; coherence - case eq2 => rw [f.comm]; coherence - any_goals coherence + (f.f β–· U ≫ (HalfBraiding.Ξ² Xβ‚‚.snd U).hom) β–· Y.fst βŠ—β‰« πŸ™ _ := ?eq2 + _ = _ := ?eq3 + case eq1 => coherence + case eq2 => rw [← whisker_exchange]; coherence + case eq3 => rw [f.comm]; coherence /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ @[simps] @@ -353,11 +353,6 @@ def braiding (X Y : Center C) : X βŠ— Y β‰… Y βŠ— X := instance braidedCategoryCenter : BraidedCategory (Center C) where braiding := braiding - braiding_naturality f g := by - ext - dsimp only [comp_f, braiding_hom_f] - erw [tensorHom_def', Category.assoc, HalfBraiding.naturality, f.comm_assoc, ← tensorHom_def, - tensor_f] #align category_theory.center.braided_category_center CategoryTheory.Center.braidedCategoryCenter -- `aesop_cat` handles the hexagon axioms @@ -370,12 +365,7 @@ open BraidedCategory /-- Auxiliary construction for `ofBraided`. -/ @[simps] def ofBraidedObj (X : C) : Center C := - ⟨X, { - Ξ² := fun Y => Ξ²_ X Y - monoidal := fun U U' => by - rw [Iso.eq_inv_comp, ← Category.assoc, ← Category.assoc, Iso.eq_comp_inv, Category.assoc, - Category.assoc] - exact hexagon_forward X U U' }⟩ + ⟨X, { Ξ² := fun Y => Ξ²_ X Y}⟩ #align category_theory.center.of_braided_obj CategoryTheory.Center.ofBraidedObj variable (C) @@ -387,20 +377,9 @@ def ofBraided : MonoidalFunctor C (Center C) where obj := ofBraidedObj map f := { f - comm := fun U => braiding_naturality _ _ } - Ξ΅ := - { f := πŸ™ _ - comm := fun U => by - dsimp - rw [tensor_id, Category.id_comp, tensor_id, Category.comp_id, ← braiding_rightUnitor, - Category.assoc, Iso.hom_inv_id, Category.comp_id] } - ΞΌ X Y := - { f := πŸ™ _ - comm := fun U => by - dsimp - rw [tensor_id, tensor_id, Category.id_comp, Category.comp_id, ← Iso.inv_comp_eq, - ← Category.assoc, ← Category.assoc, ← Iso.comp_inv_eq, Category.assoc, hexagon_reverse, - Category.assoc] } + comm := fun U => braiding_naturality_left f U } + Ξ΅ := { f := πŸ™ _ } + ΞΌ X Y := { f := πŸ™ _ } #align category_theory.center.of_braided CategoryTheory.Center.ofBraided end diff --git a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean index a248a36d13e5a..6a2957ca5be9e 100644 --- a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean @@ -189,6 +189,8 @@ def commMonToLaxBraided : CommMon_ C β₯€ LaxBraidedFunctor (Discrete PUnit.{u + set_option linter.uppercaseLean3 false in #align CommMon_.equiv_lax_braided_functor_punit.CommMon_to_lax_braided CommMon_.EquivLaxBraidedFunctorPunit.commMonToLaxBraided +attribute [local simp] id_tensorHom tensorHom_id + /-- Implementation of `CommMon_.equivLaxBraidedFunctorPunit`. -/ @[simps!] def unitIso : diff --git a/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean b/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean index dfbe63d4ad3e6..6de4a575549a6 100644 --- a/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean @@ -33,6 +33,8 @@ variable {D : Type uβ‚‚} [Category.{vβ‚‚} D] [MonoidalCategory.{vβ‚‚} D] namespace FunctorCategory +attribute [local simp] id_tensorHom tensorHom_id + variable (F G F' G' : C β₯€ D) /-- (An auxiliary definition for `functorCategoryMonoidal`.) @@ -57,10 +59,28 @@ def tensorHom : tensorObj F F' ⟢ tensorObj G G' where naturality X Y f := by dsimp; rw [← tensor_comp, Ξ±.naturality, Ξ².naturality, tensor_comp] #align category_theory.monoidal.functor_category.tensor_hom CategoryTheory.Monoidal.FunctorCategory.tensorHom +/-- (An auxiliary definition for `functorCategoryMonoidal`.) -/ +@[simps] +def whiskerLeft (F) (Ξ² : F' ⟢ G') : tensorObj F F' ⟢ tensorObj F G' where + app X := F.obj X ◁ Ξ².app X + naturality X Y f := by + simp only [← id_tensorHom] + apply (tensorHom (πŸ™ F) Ξ²).naturality + +/-- (An auxiliary definition for `functorCategoryMonoidal`.) -/ +@[simps] +def whiskerRight (F') : tensorObj F F' ⟢ tensorObj G F' where + app X := Ξ±.app X β–· F'.obj X + naturality X Y f := by + simp only [← tensorHom_id] + apply (tensorHom Ξ± (πŸ™ F')).naturality + end FunctorCategory open CategoryTheory.Monoidal.FunctorCategory +attribute [local simp] id_tensorHom tensorHom_id tensorHom_def + /-- When `C` is any category, and `D` is a monoidal category, the functor category `C β₯€ D` has a natural pointwise monoidal structure, where `(F βŠ— G).obj X = F.obj X βŠ— G.obj X`. @@ -68,11 +88,14 @@ where `(F βŠ— G).obj X = F.obj X βŠ— G.obj X`. instance functorCategoryMonoidal : MonoidalCategory (C β₯€ D) where tensorObj F G := tensorObj F G tensorHom Ξ± Ξ² := tensorHom Ξ± Ξ² + whiskerLeft F _ _ Ξ± := FunctorCategory.whiskerLeft F Ξ± + whiskerRight Ξ± F := FunctorCategory.whiskerRight Ξ± F tensorUnit' := (CategoryTheory.Functor.const C).obj (πŸ™_ D) leftUnitor F := NatIso.ofComponents fun X => Ξ»_ (F.obj X) rightUnitor F := NatIso.ofComponents fun X => ρ_ (F.obj X) associator F G H := NatIso.ofComponents fun X => Ξ±_ (F.obj X) (G.obj X) (H.obj X) - pentagon F G H K := by ext X; dsimp; rw [pentagon] + whisker_exchange := by intros; ext; simp [whisker_exchange] + -- pentagon F G H K := by ext X; dsimp; rw [pentagon] #align category_theory.monoidal.functor_category_monoidal CategoryTheory.Monoidal.functorCategoryMonoidal @[simp] @@ -101,6 +124,16 @@ theorem tensorHom_app {F G F' G' : C β₯€ D} {Ξ± : F ⟢ G} {Ξ² : F' ⟢ G'} {X} rfl #align category_theory.monoidal.tensor_hom_app CategoryTheory.Monoidal.tensorHom_app +@[simp] +theorem whiskerLeft_app (F : C β₯€ D) {G H : C β₯€ D} (Ξ± : G ⟢ H) (X : C) : + (F ◁ Ξ±).app X = F.obj X ◁ Ξ±.app X := + rfl + +@[simp] +theorem whiskerRight_app {G H : C β₯€ D} (Ξ± : G ⟢ H) (F : C β₯€ D) (X : C) : + (Ξ± β–· F).app X = Ξ±.app X β–· F.obj X := + rfl + @[simp] theorem leftUnitor_hom_app {F : C β₯€ D} {X} : ((Ξ»_ F).hom : πŸ™_ _ βŠ— F ⟢ F).app X = (Ξ»_ (F.obj X)).hom := @@ -148,7 +181,7 @@ the natural pointwise monoidal structure on the functor category `C β₯€ D` is also braided. -/ instance functorCategoryBraided : BraidedCategory (C β₯€ D) where - braiding F G := NatIso.ofComponents fun X => Ξ²_ _ _ + braiding F G := NatIso.ofComponents (fun X ↦ Ξ²_ _ _) (by simp [whisker_exchange]) hexagon_forward F G H := by ext X; apply hexagon_forward hexagon_reverse F G H := by ext X; apply hexagon_reverse #align category_theory.monoidal.functor_category_braided CategoryTheory.Monoidal.functorCategoryBraided diff --git a/Mathlib/CategoryTheory/Monoidal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Limits.lean index 1aac3fbfaf3c9..6cbe45e0c14aa 100644 --- a/Mathlib/CategoryTheory/Monoidal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Limits.lean @@ -47,7 +47,7 @@ theorem limitFunctorial_map {F G : J β₯€ C} (Ξ± : F ⟢ G) : variable [MonoidalCategory.{v} C] @[simps] -instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F where +instance limitLaxMonoidalStruct : LaxMonoidalStruct fun F : J β₯€ C => limit F where Ξ΅ := limit.lift _ { pt := _ @@ -60,11 +60,12 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F where naturality := fun j j' f => by dsimp simp only [Category.id_comp, ← tensor_comp, limit.w] } } - ΞΌ_natural f g := by +instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorHom + (ΞΌ_natural:= fun f g => by ext; dsimp simp only [limit.lift_Ο€, Cones.postcompose_obj_Ο€, Monoidal.tensorHom_app, limit.lift_map, - NatTrans.comp_app, Category.assoc, ← tensor_comp, limMap_Ο€] - associativity X Y Z := by + NatTrans.comp_app, Category.assoc, ← tensor_comp, limMap_Ο€]) + (associativity := fun X Y Z => by ext j; dsimp simp only [limit.lift_Ο€, Cones.postcompose_obj_Ο€, Monoidal.associator_hom_app, limit.lift_map, NatTrans.comp_app, Category.assoc] @@ -78,8 +79,8 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F where slice_rhs 2 3 => rw [← id_tensor_comp, limit.lift_Ο€] dsimp - dsimp; simp - left_unitality X := by + dsimp; simp) + (left_unitality := fun X => by ext j; dsimp simp conv_rhs => rw [← tensor_id_comp_id_tensor (limit.Ο€ X j)] @@ -87,9 +88,9 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F where rw [← comp_tensor_id] erw [limit.lift_Ο€] dsimp - slice_rhs 2 3 => rw [leftUnitor_naturality] - simp - right_unitality X := by + slice_rhs 2 3 => rw [leftUnitor_naturality'] + simp [id_tensorHom, tensorHom_id]) + (right_unitality := fun X => by ext j; dsimp simp conv_rhs => rw [← id_tensor_comp_tensor_id _ (limit.Ο€ X j)] @@ -97,8 +98,8 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F where rw [← id_tensor_comp] erw [limit.lift_Ο€] dsimp - slice_rhs 2 3 => rw [rightUnitor_naturality] - simp + slice_rhs 2 3 => rw [rightUnitor_naturality'] + simp [id_tensorHom, tensorHom_id]) #align category_theory.limits.limit_lax_monoidal CategoryTheory.Limits.limitLaxMonoidal /-- The limit functor `F ↦ limit F` bundled as a lax monoidal functor. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index f5263d9c40d90..cac6a84e2ec5f 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -36,17 +36,20 @@ structure Mon_ where X : C one : πŸ™_ C ⟢ X mul : X βŠ— X ⟢ X - one_mul : (one βŠ— πŸ™ X) ≫ mul = (Ξ»_ X).hom := by aesop_cat - mul_one : (πŸ™ X βŠ— one) ≫ mul = (ρ_ X).hom := by aesop_cat + one_mul : (one β–· X) ≫ mul = (Ξ»_ X).hom := by aesop_cat + mul_one : (X ◁ one) ≫ mul = (ρ_ X).hom := by aesop_cat -- Obviously there is some flexibility stating this axiom. -- This one has left- and right-hand sides matching the statement of `Monoid.mul_assoc`, -- and chooses to place the associator on the right-hand side. -- The heuristic is that unitors and associators "don't have much weight". - mul_assoc : (mul βŠ— πŸ™ X) ≫ mul = (Ξ±_ X X X).hom ≫ (πŸ™ X βŠ— mul) ≫ mul := by aesop_cat + mul_assoc : (mul β–· X) ≫ mul = (Ξ±_ X X X).hom ≫ (X ◁ mul) ≫ mul := by aesop_cat #align Mon_ Mon_ attribute [reassoc] Mon_.one_mul Mon_.mul_one +attribute [simp] Mon_.one_mul Mon_.mul_one + + -- We prove a more general `@[simp]` lemma below. attribute [reassoc (attr := simp)] Mon_.mul_assoc @@ -71,16 +74,16 @@ variable {M : Mon_ C} @[simp] theorem one_mul_hom {Z : C} (f : Z ⟢ M.X) : (M.one βŠ— f) ≫ M.mul = (Ξ»_ Z).hom ≫ f := by - rw [← id_tensor_comp_tensor_id, Category.assoc, M.one_mul, leftUnitor_naturality] + rw [tensorHom_def_assoc, M.one_mul, leftUnitor_naturality] #align Mon_.one_mul_hom Mon_.one_mul_hom @[simp] theorem mul_one_hom {Z : C} (f : Z ⟢ M.X) : (f βŠ— M.one) ≫ M.mul = (ρ_ Z).hom ≫ f := by - rw [← tensor_id_comp_id_tensor, Category.assoc, M.mul_one, rightUnitor_naturality] + rw [tensorHom_def'_assoc, M.mul_one, rightUnitor_naturality] #align Mon_.mul_one_hom Mon_.mul_one_hom theorem assoc_flip : - (πŸ™ M.X βŠ— M.mul) ≫ M.mul = (Ξ±_ M.X M.X M.X).inv ≫ (M.mul βŠ— πŸ™ M.X) ≫ M.mul := by simp + (M.X ◁ M.mul) ≫ M.mul = (Ξ±_ M.X M.X M.X).inv ≫ (M.mul β–· M.X) ≫ M.mul := by simp #align Mon_.assoc_flip Mon_.assoc_flip /-- A morphism of monoid objects. -/ @@ -93,6 +96,8 @@ structure Hom (M N : Mon_ C) where attribute [reassoc (attr := simp)] Hom.one_hom Hom.mul_hom +attribute [local simp] id_tensorHom tensorHom_id + /-- The identity morphism on a monoid object. -/ @[simps] def id (M : Mon_ C) : Hom M M where @@ -210,28 +215,17 @@ def mapMon (F : LaxMonoidalFunctor C D) : Mon_ C β₯€ Mon_ D where one := F.Ξ΅ ≫ F.map A.one mul := F.ΞΌ _ _ ≫ F.map A.mul one_mul := by - conv_lhs => rw [comp_tensor_id, ← F.toFunctor.map_id] - slice_lhs 2 3 => rw [F.ΞΌ_natural] + simp only [comp_whiskerRight, Category.assoc, ΞΌ_natural_left_assoc, left_unitality] slice_lhs 3 4 => rw [← F.toFunctor.map_comp, A.one_mul] - rw [F.toFunctor.map_id] - rw [F.left_unitality] mul_one := by - conv_lhs => rw [id_tensor_comp, ← F.toFunctor.map_id] - slice_lhs 2 3 => rw [F.ΞΌ_natural] + simp only [MonoidalCategory.whiskerLeft_comp, Category.assoc, ΞΌ_natural_right_assoc, + right_unitality] slice_lhs 3 4 => rw [← F.toFunctor.map_comp, A.mul_one] - rw [F.toFunctor.map_id] - rw [F.right_unitality] mul_assoc := by - conv_lhs => rw [comp_tensor_id, ← F.toFunctor.map_id] - slice_lhs 2 3 => rw [F.ΞΌ_natural] + simp only [comp_whiskerRight, Category.assoc, ΞΌ_natural_left_assoc, + MonoidalCategory.whiskerLeft_comp, ΞΌ_natural_right_assoc] slice_lhs 3 4 => rw [← F.toFunctor.map_comp, A.mul_assoc] - conv_lhs => rw [F.toFunctor.map_id] - conv_lhs => rw [F.toFunctor.map_comp, F.toFunctor.map_comp] - conv_rhs => rw [id_tensor_comp, ← F.toFunctor.map_id] - slice_rhs 3 4 => rw [F.ΞΌ_natural] - conv_rhs => rw [F.toFunctor.map_id] - slice_rhs 1 3 => rw [← F.associativity] - simp only [Category.assoc] } + simp } map f := { hom := F.map f.hom one_hom := by dsimp; rw [Category.assoc, ← F.toFunctor.map_comp, f.one_hom] @@ -289,6 +283,8 @@ attribute [local aesop safe tactic (rule_sets [CategoryTheory])] attribute [local simp] eqToIso_map +attribute [local simp] id_tensorHom tensorHom_id + /-- Implementation of `Mon_.equivLaxMonoidalFunctorPUnit`. -/ @[simps!] def unitIso : @@ -376,70 +372,101 @@ which have also been proved in `Mathlib.CategoryTheory.Monoidal.Braided`. variable {C} +attribute [local simp] id_tensorHom tensorHom_id + -- The proofs that associators and unitors preserve monoid units don't require braiding. theorem one_associator {M N P : Mon_ C} : ((Ξ»_ (πŸ™_ C)).inv ≫ ((Ξ»_ (πŸ™_ C)).inv ≫ (M.one βŠ— N.one) βŠ— P.one)) ≫ (Ξ±_ M.X N.X P.X).hom = (Ξ»_ (πŸ™_ C)).inv ≫ (M.one βŠ— (Ξ»_ (πŸ™_ C)).inv ≫ (N.one βŠ— P.one)) := by - simp + simp only [Category.assoc, Iso.cancel_iso_inv_left] slice_lhs 1 3 => rw [← Category.id_comp P.one, tensor_comp] slice_lhs 2 3 => rw [associator_naturality] slice_rhs 1 2 => rw [← Category.id_comp M.one, tensor_comp] - slice_lhs 1 2 => rw [← leftUnitor_tensor_inv] + slice_lhs 1 2 => rw [← leftUnitor_tensor_inv'] rw [← cancel_epi (Ξ»_ (πŸ™_ C)).inv] slice_lhs 1 2 => rw [leftUnitor_inv_naturality] - simp only [Category.assoc] + simp #align Mon_.one_associator Mon_.one_associator theorem one_leftUnitor {M : Mon_ C} : ((Ξ»_ (πŸ™_ C)).inv ≫ (πŸ™ (πŸ™_ C) βŠ— M.one)) ≫ (Ξ»_ M.X).hom = M.one := by - slice_lhs 2 3 => rw [leftUnitor_naturality] simp #align Mon_.one_left_unitor Mon_.one_leftUnitor theorem one_rightUnitor {M : Mon_ C} : ((Ξ»_ (πŸ™_ C)).inv ≫ (M.one βŠ— πŸ™ (πŸ™_ C))) ≫ (ρ_ M.X).hom = M.one := by - slice_lhs 2 3 => rw [rightUnitor_naturality, ← unitors_equal] - simp + simp [← unitors_equal] #align Mon_.one_right_unitor Mon_.one_rightUnitor variable [BraidedCategory C] +-- open Mathlib.Tactic.Coherence in +-- example {C : Type u} [Category.{v} C] [MonoidalCategory C] (M N : C) [LiftObj M] [LiftObj N] : +-- LiftObj ((tensor C).obj (M, N)) = LiftObj (M βŠ— N) := by +-- rfl + +-- #check _root_.id + +-- open Mathlib.Tactic.Coherence in +-- example {C : Type u} [Category.{v} C] [MonoidalCategory C] (M N : C) [LiftObj M] [LiftObj N] : +-- FreeMonoidalCategory.projectObj _root_.id (LiftObj.lift ((tensor C).obj (M, N))) = ((tensor C).obj (M, N)) := by +-- rfl + +-- open Mathlib.Tactic.Coherence in +-- instance {C : Type u} [Category.{v} C] [MonoidalCategory C] (M N : C) [LiftObj M] [LiftObj N] : +-- MonoidalCoherence ((M βŠ— N)) ((tensor C).obj (M, N)) := by + +-- infer_instance -- fails + +-- /- +-- To use `monoidalComp` in expressions containing `tensor_ΞΌ`, we need, for example, the following +-- instance. However, it is not automatically inferred. + +-- open Mathlib.Tactic.Coherence in +-- instance {C : Type u} [Category.{v} C] [MonoidalCategory C] (M N : C) : +-- LiftObj ((tensor C).obj (M, N)) := by +-- infer_instance -- fails + +-- open Mathlib.Tactic.Coherence in +-- instance {C : Type u} [Category.{v} C] [MonoidalCategory C] (M N : C) : +-- MonoidalCoherence ((M βŠ— N)) ((tensor C).obj (M, N)) := by +-- infer_instance -- fails +-- -/ + theorem Mon_tensor_one_mul (M N : Mon_ C) : - ((Ξ»_ (πŸ™_ C)).inv ≫ (M.one βŠ— N.one) βŠ— πŸ™ (M.X βŠ— N.X)) ≫ + (((Ξ»_ (πŸ™_ C)).inv ≫ (M.one βŠ— N.one)) β–· (M.X βŠ— N.X)) ≫ tensor_ΞΌ C (M.X, N.X) (M.X, N.X) ≫ (M.mul βŠ— N.mul) = (Ξ»_ (M.X βŠ— N.X)).hom := by - rw [← Category.id_comp (πŸ™ (M.X βŠ— N.X)), tensor_comp] - slice_lhs 2 3 => rw [← tensor_id, tensor_ΞΌ_natural] + simp only [comp_whiskerRight_assoc] + slice_lhs 2 3 => rw [tensor_ΞΌ_natural_left] slice_lhs 3 4 => rw [← tensor_comp, one_mul M, one_mul N] symm exact tensor_left_unitality C M.X N.X #align Mon_.Mon_tensor_one_mul Mon_.Mon_tensor_one_mul theorem Mon_tensor_mul_one (M N : Mon_ C) : - (πŸ™ (M.X βŠ— N.X) βŠ— (Ξ»_ (πŸ™_ C)).inv ≫ (M.one βŠ— N.one)) ≫ + (M.X βŠ— N.X) ◁ ((Ξ»_ (πŸ™_ C)).inv ≫ (M.one βŠ— N.one)) ≫ tensor_ΞΌ C (M.X, N.X) (M.X, N.X) ≫ (M.mul βŠ— N.mul) = (ρ_ (M.X βŠ— N.X)).hom := by - rw [← Category.id_comp (πŸ™ (M.X βŠ— N.X)), tensor_comp] - slice_lhs 2 3 => rw [← tensor_id, tensor_ΞΌ_natural] + simp only [whiskerLeft_comp_assoc] + slice_lhs 2 3 => rw [tensor_ΞΌ_natural_right] slice_lhs 3 4 => rw [← tensor_comp, mul_one M, mul_one N] symm exact tensor_right_unitality C M.X N.X #align Mon_.Mon_tensor_mul_one Mon_.Mon_tensor_mul_one theorem Mon_tensor_mul_assoc (M N : Mon_ C) : - (tensor_ΞΌ C (M.X, N.X) (M.X, N.X) ≫ (M.mul βŠ— N.mul) βŠ— πŸ™ (M.X βŠ— N.X)) ≫ + ((tensor_ΞΌ C (M.X, N.X) (M.X, N.X) ≫ (M.mul βŠ— N.mul)) β–· (M.X βŠ— N.X)) ≫ tensor_ΞΌ C (M.X, N.X) (M.X, N.X) ≫ (M.mul βŠ— N.mul) = (Ξ±_ (M.X βŠ— N.X) (M.X βŠ— N.X) (M.X βŠ— N.X)).hom ≫ - (πŸ™ (M.X βŠ— N.X) βŠ— tensor_ΞΌ C (M.X, N.X) (M.X, N.X) ≫ (M.mul βŠ— N.mul)) ≫ + ((M.X βŠ— N.X) ◁ (tensor_ΞΌ C (M.X, N.X) (M.X, N.X) ≫ (M.mul βŠ— N.mul))) ≫ tensor_ΞΌ C (M.X, N.X) (M.X, N.X) ≫ (M.mul βŠ— N.mul) := by - rw [← Category.id_comp (πŸ™ (M.X βŠ— N.X)), tensor_comp] - slice_lhs 2 3 => rw [← tensor_id, tensor_ΞΌ_natural] + simp only [comp_whiskerRight_assoc, whiskerLeft_comp_assoc] + slice_lhs 2 3 => rw [tensor_ΞΌ_natural_left] slice_lhs 3 4 => rw [← tensor_comp, mul_assoc M, mul_assoc N, tensor_comp, tensor_comp] - -- Porting note: needed to add `dsimp` here. - slice_lhs 1 3 => dsimp; rw [tensor_associativity] - slice_lhs 3 4 => rw [← tensor_ΞΌ_natural] - slice_lhs 2 3 => rw [← tensor_comp, tensor_id] - simp only [Category.assoc] + slice_lhs 1 3 => rw [tensor_associativity] + slice_lhs 3 4 => rw [← tensor_ΞΌ_natural_right] + simp #align Mon_.Mon_tensor_mul_assoc Mon_.Mon_tensor_mul_assoc theorem mul_associator {M N P : Mon_ C} : @@ -449,10 +476,11 @@ theorem mul_associator {M N P : Mon_ C} : ((Ξ±_ M.X N.X P.X).hom βŠ— (Ξ±_ M.X N.X P.X).hom) ≫ tensor_ΞΌ C (M.X, N.X βŠ— P.X) (M.X, N.X βŠ— P.X) ≫ (M.mul βŠ— tensor_ΞΌ C (N.X, P.X) (N.X, P.X) ≫ (N.mul βŠ— P.mul)) := by - simp + simp only [Category.assoc] slice_lhs 2 3 => rw [← Category.id_comp P.mul, tensor_comp] slice_lhs 3 4 => rw [associator_naturality] slice_rhs 3 4 => rw [← Category.id_comp M.mul, tensor_comp] + simp only [tensorHom_id, id_tensorHom] slice_lhs 1 3 => rw [associator_monoidal] simp only [Category.assoc] #align Mon_.mul_associator Mon_.mul_associator @@ -461,6 +489,7 @@ theorem mul_leftUnitor {M : Mon_ C} : (tensor_ΞΌ C (πŸ™_ C, M.X) (πŸ™_ C, M.X) ≫ ((Ξ»_ (πŸ™_ C)).hom βŠ— M.mul)) ≫ (Ξ»_ M.X).hom = ((Ξ»_ M.X).hom βŠ— (Ξ»_ M.X).hom) ≫ M.mul := by rw [← Category.comp_id (Ξ»_ (πŸ™_ C)).hom, ← Category.id_comp M.mul, tensor_comp] + simp only [tensorHom_id, id_tensorHom] slice_lhs 3 4 => rw [leftUnitor_naturality] slice_lhs 1 3 => rw [← leftUnitor_monoidal] simp only [Category.assoc, Category.id_comp] @@ -470,20 +499,21 @@ theorem mul_rightUnitor {M : Mon_ C} : (tensor_ΞΌ C (M.X, πŸ™_ C) (M.X, πŸ™_ C) ≫ (M.mul βŠ— (Ξ»_ (πŸ™_ C)).hom)) ≫ (ρ_ M.X).hom = ((ρ_ M.X).hom βŠ— (ρ_ M.X).hom) ≫ M.mul := by rw [← Category.id_comp M.mul, ← Category.comp_id (Ξ»_ (πŸ™_ C)).hom, tensor_comp] + simp only [tensorHom_id, id_tensorHom] slice_lhs 3 4 => rw [rightUnitor_naturality] slice_lhs 1 3 => rw [← rightUnitor_monoidal] simp only [Category.assoc, Category.id_comp] #align Mon_.mul_right_unitor Mon_.mul_rightUnitor -instance monMonoidal : MonoidalCategory (Mon_ C) where - tensorObj M N := +instance monMonoidal : MonoidalCategory (Mon_ C) := .ofTensorHom + (tensorObj := fun M N ↦ { X := M.X βŠ— N.X one := (Ξ»_ (πŸ™_ C)).inv ≫ (M.one βŠ— N.one) mul := tensor_ΞΌ C (M.X, N.X) (M.X, N.X) ≫ (M.mul βŠ— N.mul) one_mul := Mon_tensor_one_mul M N mul_one := Mon_tensor_mul_one M N - mul_assoc := Mon_tensor_mul_assoc M N } - tensorHom f g := + mul_assoc := Mon_tensor_mul_assoc M N }) + (tensorHom := fun f g ↦ { hom := f.hom βŠ— g.hom one_hom := by dsimp @@ -492,18 +522,18 @@ instance monMonoidal : MonoidalCategory (Mon_ C) where dsimp slice_rhs 1 2 => rw [tensor_ΞΌ_natural] slice_lhs 2 3 => rw [← tensor_comp, Hom.mul_hom f, Hom.mul_hom g, tensor_comp] - simp only [Category.assoc] } - tensor_id := by intros; ext; apply tensor_id - tensor_comp := by intros; ext; apply tensor_comp - tensorUnit' := trivial C - associator M N P := isoOfIso (Ξ±_ M.X N.X P.X) one_associator mul_associator - associator_naturality := by intros; ext; dsimp; apply associator_naturality - leftUnitor M := isoOfIso (Ξ»_ M.X) one_leftUnitor mul_leftUnitor - leftUnitor_naturality := by intros; ext; dsimp; apply leftUnitor_naturality - rightUnitor M := isoOfIso (ρ_ M.X) one_rightUnitor mul_rightUnitor - rightUnitor_naturality := by intros; ext; dsimp; apply rightUnitor_naturality - pentagon := by intros; ext; dsimp; apply pentagon - triangle := by intros; ext; dsimp; apply triangle + simp only [Category.assoc] }) + (tensor_id := by intros; ext; apply tensor_id) + (tensor_comp := by intros; ext; apply tensor_comp) + (tensorUnit' := trivial C) + (associator := fun M N P ↦ isoOfIso (Ξ±_ M.X N.X P.X) one_associator mul_associator) + (associator_naturality := by intros; ext; dsimp; apply associator_naturality) + (leftUnitor := fun M ↦ isoOfIso (Ξ»_ M.X) one_leftUnitor mul_leftUnitor) + (leftUnitor_naturality := by intros; ext; dsimp; apply leftUnitor_naturality') + (rightUnitor := fun M ↦ isoOfIso (ρ_ M.X) one_rightUnitor mul_rightUnitor) + (rightUnitor_naturality := by intros; ext; dsimp; apply rightUnitor_naturality') + (pentagon := by intros; ext; dsimp; apply pentagon') + (triangle := by intros; ext; dsimp; apply triangle') #align Mon_.Mon_monoidal Mon_.monMonoidal end Mon_ diff --git a/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Symmetric.lean b/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Symmetric.lean index c8a1421d75c15..17d61d3a2ad54 100644 --- a/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Symmetric.lean +++ b/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Symmetric.lean @@ -93,7 +93,8 @@ open MonoidalOfChosenFiniteProducts def symmetricOfChosenFiniteProducts : SymmetricCategory (MonoidalOfChosenFiniteProductsSynonym 𝒯 ℬ) where braiding _ _ := Limits.BinaryFan.braiding (ℬ _ _).isLimit (ℬ _ _).isLimit - braiding_naturality f g := braiding_naturality ℬ f g + braiding_naturality_left f X := braiding_naturality ℬ f (πŸ™ X) + braiding_naturality_right X _ _ f := braiding_naturality ℬ (πŸ™ X) f hexagon_forward X Y Z := hexagon_forward ℬ X Y Z hexagon_reverse X Y Z := hexagon_reverse ℬ X Y Z symmetry X Y := symmetry ℬ X Y diff --git a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean index bfc462eeee96f..d43e1d90eae3d 100644 --- a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean +++ b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean @@ -126,9 +126,14 @@ open MonoidalCategory @[simps] def symmetricOfHasFiniteProducts [HasTerminal C] [HasBinaryProducts C] : SymmetricCategory C where braiding X Y := Limits.prod.braiding X Y - braiding_naturality f g := by dsimp; simp - hexagon_forward X Y Z := by dsimp [monoidalOfHasFiniteProducts.associator_hom]; simp - hexagon_reverse X Y Z := by dsimp [monoidalOfHasFiniteProducts.associator_inv]; simp + braiding_naturality_left f X := by dsimp; simp [← id_tensorHom, ← tensorHom_id] + braiding_naturality_right X _ _ f := by dsimp; simp [← id_tensorHom, ← tensorHom_id] + hexagon_forward X Y Z := by + dsimp [monoidalOfHasFiniteProducts.associator_hom] + simp [← id_tensorHom, ← tensorHom_id] + hexagon_reverse X Y Z := by + dsimp [monoidalOfHasFiniteProducts.associator_inv] + simp [← id_tensorHom, ← tensorHom_id] symmetry X Y := by dsimp; simp #align category_theory.symmetric_of_has_finite_products CategoryTheory.symmetricOfHasFiniteProducts @@ -218,9 +223,14 @@ open MonoidalCategory def symmetricOfHasFiniteCoproducts [HasInitial C] [HasBinaryCoproducts C] : SymmetricCategory C where braiding := Limits.coprod.braiding - braiding_naturality f g := by dsimp [tensorHom]; simp - hexagon_forward X Y Z := by dsimp [monoidalOfHasFiniteCoproducts.associator_hom]; simp - hexagon_reverse X Y Z := by dsimp [monoidalOfHasFiniteCoproducts.associator_inv]; simp + braiding_naturality_left f g := by dsimp [tensorHom]; simp [← id_tensorHom, ← tensorHom_id] + braiding_naturality_right f g := by dsimp [tensorHom]; simp [← id_tensorHom, ← tensorHom_id] + hexagon_forward X Y Z := by + dsimp [monoidalOfHasFiniteCoproducts.associator_hom] + simp [← id_tensorHom, ← tensorHom_id] + hexagon_reverse X Y Z := by + dsimp [monoidalOfHasFiniteCoproducts.associator_inv] + simp [← id_tensorHom, ← tensorHom_id] symmetry X Y := by dsimp; simp #align category_theory.symmetric_of_has_finite_coproducts CategoryTheory.symmetricOfHasFiniteCoproducts diff --git a/Mathlib/CategoryTheory/Monoidal/Opposite.lean b/Mathlib/CategoryTheory/Monoidal/Opposite.lean index e6ddb78353085..0e7ccf741a8a5 100644 --- a/Mathlib/CategoryTheory/Monoidal/Opposite.lean +++ b/Mathlib/CategoryTheory/Monoidal/Opposite.lean @@ -170,16 +170,26 @@ open Opposite MonoidalCategory instance monoidalCategoryOp : MonoidalCategory Cα΅’α΅– where tensorObj X Y := op (unop X βŠ— unop Y) + whiskerLeft X _ _ f := (X.unop ◁ f.unop).op + whiskerRight f X := (f.unop β–· X.unop).op tensorHom f g := (f.unop βŠ— g.unop).op + tensorHom_def f g := Quiver.Hom.unop_inj (tensorHom_def' _ _) tensorUnit' := op (πŸ™_ C) - associator X Y Z := (Ξ±_ (unop X) (unop Y) (unop Z)).symm.op leftUnitor X := (Ξ»_ (unop X)).symm.op rightUnitor X := (ρ_ (unop X)).symm.op - associator_naturality f g h := Quiver.Hom.unop_inj (by simp) - leftUnitor_naturality f := Quiver.Hom.unop_inj (by simp) - rightUnitor_naturality f := Quiver.Hom.unop_inj (by simp) - triangle X Y := Quiver.Hom.unop_inj (by dsimp; coherence) + associator X Y Z := (Ξ±_ (unop X) (unop Y) (unop Z)).symm.op + whiskerLeft_id X Y := Quiver.Hom.unop_inj (by simp) + whiskerLeft_comp W X Y Z f g := Quiver.Hom.unop_inj (by simp) + id_whiskerLeft f := Quiver.Hom.unop_inj (by simp) + tensor_whiskerLeft W X Y Z f := Quiver.Hom.unop_inj (by simp) + id_whiskerRight X Y := Quiver.Hom.unop_inj (by simp) + comp_whiskerRight f g X := Quiver.Hom.unop_inj (by simp) + whiskerRight_id f := Quiver.Hom.unop_inj (by simp) + whiskerRight_tensor f X Y := Quiver.Hom.unop_inj (by simp) + whisker_assoc W X Y f Z := Quiver.Hom.unop_inj (by simp) + whisker_exchange _ _ := Quiver.Hom.unop_inj (whisker_exchange _ _).symm pentagon W X Y Z := Quiver.Hom.unop_inj (by dsimp; coherence) + triangle X Y := Quiver.Hom.unop_inj (by dsimp; coherence) #align category_theory.monoidal_category_op CategoryTheory.monoidalCategoryOp theorem op_tensorObj (X Y : Cα΅’α΅–) : X βŠ— Y = op (unop X βŠ— unop Y) := @@ -192,16 +202,26 @@ theorem op_tensorUnit : πŸ™_ Cα΅’α΅– = op (πŸ™_ C) := instance monoidalCategoryMop : MonoidalCategory Cα΄Ήα΅’α΅– where tensorObj X Y := mop (unmop Y βŠ— unmop X) + whiskerLeft X _ _ f := (f.unmop β–· X.unmop).mop + whiskerRight f X := (X.unmop ◁ f.unmop).mop tensorHom f g := (g.unmop βŠ— f.unmop).mop + tensorHom_def f g := unmop_inj (tensorHom_def' _ _) tensorUnit' := mop (πŸ™_ C) - associator X Y Z := (Ξ±_ (unmop Z) (unmop Y) (unmop X)).symm.mop leftUnitor X := (ρ_ (unmop X)).mop rightUnitor X := (Ξ»_ (unmop X)).mop - associator_naturality f g h := unmop_inj (by simp) - leftUnitor_naturality f := unmop_inj (by simp) - rightUnitor_naturality f := unmop_inj (by simp) - triangle X Y := unmop_inj (by simp) -- Porting note: Changed `by coherence` to `by simp` + associator X Y Z := (Ξ±_ (unmop Z) (unmop Y) (unmop X)).symm.mop + whiskerLeft_id X Y := unmop_inj (by simp) + whiskerLeft_comp W X Y Z f g := unmop_inj (by simp) + id_whiskerLeft f := unmop_inj (by simp) + tensor_whiskerLeft W X Y Z f := unmop_inj (by simp) + id_whiskerRight X Y := unmop_inj (by simp) + comp_whiskerRight f g X := unmop_inj (by simp) + whiskerRight_id f := unmop_inj (by simp) + whiskerRight_tensor f X Y := unmop_inj (by simp) + whisker_assoc W X Y f Z := unmop_inj (by simp) + whisker_exchange f g := unmop_inj (whisker_exchange g.unmop f.unmop).symm pentagon W X Y Z := unmop_inj (by dsimp; coherence) + triangle X Y := unmop_inj (by dsimp; coherence) #align category_theory.monoidal_category_mop CategoryTheory.monoidalCategoryMop theorem mop_tensorObj (X Y : Cα΄Ήα΅’α΅–) : X βŠ— Y = mop (unmop Y βŠ— unmop X) := diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean index 8e3973e93919a..854c7ad7e758c 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean @@ -544,24 +544,28 @@ theorem tensorRightHomEquiv_tensor_id_comp_evaluation {X Y : C} [HasRightDual X] #align category_theory.tensor_right_hom_equiv_tensor_id_comp_evaluation CategoryTheory.tensorRightHomEquiv_tensor_id_comp_evaluation -- Next four lemmas passing `fᘁ` or `ᘁf` through (co)evaluations. +@[reassoc] theorem coevaluation_comp_rightAdjointMate {X Y : C} [HasRightDual X] [HasRightDual Y] (f : X ⟢ Y) : Ξ·_ Y (Yᘁ) ≫ (_ ◁ fᘁ) = Ξ·_ _ _ ≫ (f β–· _) := by apply_fun (tensorLeftHomEquiv _ Y (Yᘁ) _).symm simp #align category_theory.coevaluation_comp_right_adjoint_mate CategoryTheory.coevaluation_comp_rightAdjointMate +@[reassoc] theorem leftAdjointMate_comp_evaluation {X Y : C} [HasLeftDual X] [HasLeftDual Y] (f : X ⟢ Y) : (X ◁ ᘁf) ≫ Ξ΅_ _ _ = (f β–· _) ≫ Ξ΅_ _ _ := by apply_fun tensorLeftHomEquiv _ (ᘁX) X _ simp #align category_theory.left_adjoint_mate_comp_evaluation CategoryTheory.leftAdjointMate_comp_evaluation +@[reassoc] theorem coevaluation_comp_leftAdjointMate {X Y : C} [HasLeftDual X] [HasLeftDual Y] (f : X ⟢ Y) : Ξ·_ (ᘁY) Y ≫ ((ᘁf) β–· Y) = Ξ·_ (ᘁX) X ≫ ((ᘁX) ◁ f) := by apply_fun (tensorRightHomEquiv _ (ᘁY) Y _).symm simp #align category_theory.coevaluation_comp_left_adjoint_mate CategoryTheory.coevaluation_comp_leftAdjointMate +@[reassoc] theorem rightAdjointMate_comp_evaluation {X Y : C} [HasRightDual X] [HasRightDual Y] (f : X ⟢ Y) : (fᘁ β–· X) ≫ Ξ΅_ X (Xᘁ) = ((Yᘁ) ◁ f) ≫ Ξ΅_ Y (Yᘁ) := by apply_fun tensorRightHomEquiv _ X (Xᘁ) _ diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/FunctorCategory.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/FunctorCategory.lean index 197903c677df5..bbeb2d8fc4cce 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/FunctorCategory.lean @@ -36,15 +36,17 @@ instance functorHasRightDual [RightRigidCategory D] (F : C β₯€ D) : HasRightDual naturality := fun X Y f => by dsimp rw [Category.comp_id, Functor.map_inv, ← id_tensor_comp_tensor_id, Category.assoc, - rightAdjointMate_comp_evaluation, ← Category.assoc, ← id_tensor_comp, - IsIso.hom_inv_id, tensor_id, Category.id_comp] } + id_tensorHom, tensorHom_id, + rightAdjointMate_comp_evaluation, ← whiskerLeft_comp_assoc, + IsIso.hom_inv_id, MonoidalCategory.whiskerLeft_id, Category.id_comp] } coevaluation' := { app := fun X => Ξ·_ _ _ naturality := fun X Y f => by dsimp - rw [Functor.map_inv, Category.id_comp, ← id_tensor_comp_tensor_id, ← Category.assoc, - coevaluation_comp_rightAdjointMate, Category.assoc, ← comp_tensor_id, - IsIso.inv_hom_id, tensor_id, Category.comp_id] } } + rw [Functor.map_inv, Category.id_comp, ← id_tensor_comp_tensor_id, + id_tensorHom, tensorHom_id, ← Category.assoc, + coevaluation_comp_rightAdjointMate, Category.assoc, ← comp_whiskerRight, + IsIso.inv_hom_id, id_whiskerRight, Category.comp_id] } } #align category_theory.monoidal.functor_has_right_dual CategoryTheory.Monoidal.functorHasRightDual instance rightRigidFunctorCategory [RightRigidCategory D] : RightRigidCategory (C β₯€ D) where @@ -60,16 +62,16 @@ instance functorHasLeftDual [LeftRigidCategory D] (F : C β₯€ D) : HasLeftDual F { app := fun X => Ξ΅_ _ _ naturality := fun X Y f => by dsimp - rw [Category.comp_id, Functor.map_inv, ← tensor_id_comp_id_tensor, Category.assoc, - leftAdjointMate_comp_evaluation, ← Category.assoc, ← comp_tensor_id, - IsIso.hom_inv_id, tensor_id, Category.id_comp] } + simp [tensorHom_def', leftAdjointMate_comp_evaluation] + simp_rw [← comp_whiskerRight_assoc] + simp only [IsIso.hom_inv_id, id_whiskerRight, Category.id_comp] } coevaluation' := { app := fun X => Ξ·_ _ _ naturality := fun X Y f => by dsimp - rw [Functor.map_inv, Category.id_comp, ← tensor_id_comp_id_tensor, ← Category.assoc, - coevaluation_comp_leftAdjointMate, Category.assoc, ← id_tensor_comp, - IsIso.inv_hom_id, tensor_id, Category.comp_id] } } + simp [tensorHom_def', coevaluation_comp_leftAdjointMate_assoc] + simp_rw [← MonoidalCategory.whiskerLeft_comp] + simp only [IsIso.inv_hom_id, MonoidalCategory.whiskerLeft_id, Category.comp_id] } } #align category_theory.monoidal.functor_has_left_dual CategoryTheory.Monoidal.functorHasLeftDual instance leftRigidFunctorCategory [LeftRigidCategory D] : LeftRigidCategory (C β₯€ D) where diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean index 47bc8cf536c05..2e92d5d5c243c 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean @@ -32,9 +32,26 @@ def exactPairingOfFaithful [Faithful F.toFunctor] {X Y : C} (eval : Y βŠ— X ⟢ evaluation' := eval coevaluation' := coeval evaluation_coevaluation' := - F.toFunctor.map_injective (by simp [map_eval, map_coeval, MonoidalFunctor.map_tensor]) + F.toFunctor.map_injective (by + simp_rw [Functor.map_comp, F.map_whiskerRight, F.map_whiskerLeft, map_eval, map_coeval] + simp only [comp_whiskerRight, Category.assoc, MonoidalCategory.whiskerLeft_comp, + LaxMonoidalFunctor.associativity_assoc, IsIso.hom_inv_id_assoc, IsIso.inv_comp_eq] + simp_rw [← whiskerLeft_comp_assoc] + simp only [IsIso.hom_inv_id_assoc, MonoidalCategory.whiskerLeft_comp, Category.assoc, + ExactPairing.evaluation_coevaluation_assoc, LaxMonoidalFunctor.left_unitality, + LaxMonoidalFunctor.right_unitality_inv] + simp_rw [← comp_whiskerRight_assoc] + simp ) coevaluation_evaluation' := - F.toFunctor.map_injective (by simp [map_eval, map_coeval, MonoidalFunctor.map_tensor]) + F.toFunctor.map_injective (by + simp_rw [Functor.map_comp, F.map_whiskerRight, F.map_whiskerLeft, map_eval, map_coeval] + simp only [MonoidalCategory.whiskerLeft_comp, Category.assoc, comp_whiskerRight, + LaxMonoidalFunctor.associativity_inv_assoc, IsIso.hom_inv_id_assoc, IsIso.inv_comp_eq] + simp_rw [← comp_whiskerRight_assoc] + simp only [IsIso.hom_inv_id_assoc, comp_whiskerRight, Category.assoc, ExactPairing.coevaluation_evaluation_assoc, + LaxMonoidalFunctor.right_unitality, LaxMonoidalFunctor.left_unitality_inv] + simp_rw [← whiskerLeft_comp_assoc] + simp ) #align category_theory.exact_pairing_of_faithful CategoryTheory.exactPairingOfFaithful /-- Given a pair of objects which are sent by a fully faithful functor to a pair of objects diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index ec34ec0ff8e5c..e53206083ffed 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -116,6 +116,7 @@ attribute [local simp] id_tensorHom tensorHom_id -- porting note: it was @[simps {attrs := [`_refl_lemma]}] /-- Transport a monoidal structure along an equivalence of (plain) categories. -/ +@[simps!] def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D := .ofTensorHom (tensorObj := fun X Y ↦ e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y)) (tensorHom := fun f g ↦ e.functor.map (e.inverse.map f βŠ— e.inverse.map g)) @@ -236,21 +237,21 @@ def Transported (_ : C β‰Œ D) := D instance (e : C β‰Œ D) : Category (Transported e) := (inferInstance : Category D) -- @[simps!] -instance transportMonoidalCategory (e : C β‰Œ D) : MonoidalCategory (Transported e) := +instance (e : C β‰Œ D) : MonoidalCategory (Transported e) := transport e instance (e : C β‰Œ D) : Inhabited (Transported e) := βŸ¨πŸ™_ _⟩ -theorem transport_tensorUnit' (e : C β‰Œ D) : πŸ™_ (Transported e) = e.functor.obj (πŸ™_ C) := rfl +-- theorem transport_tensorUnit' (e : C β‰Œ D) : πŸ™_ (Transported e) = e.functor.obj (πŸ™_ C) := rfl -theorem transport_tensorObj (e : C β‰Œ D) (X Y : Transported e) : - X βŠ— Y = e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y) := - rfl +-- theorem transport_tensorObj (e : C β‰Œ D) (X Y : Transported e) : +-- X βŠ— Y = e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y) := +-- rfl -theorem transport_tensorHom (e : C β‰Œ D) {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Transported e} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : - f βŠ— g = e.functor.map (e.inverse.map f βŠ— e.inverse.map g) := by - rfl +-- theorem transport_tensorHom (e : C β‰Œ D) {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Transported e} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : +-- f βŠ— g = e.functor.map (e.inverse.map f βŠ— e.inverse.map g) := by +-- rfl theorem transport_associator (e : C β‰Œ D) (X Y Z : Transported e) : Ξ±_ X Y Z = diff --git a/Mathlib/CategoryTheory/Monoidal/Types/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Types/Basic.lean index 989b7bce218ac..6f7e16ca0d8a2 100644 --- a/Mathlib/CategoryTheory/Monoidal/Types/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Types/Basic.lean @@ -33,6 +33,16 @@ theorem tensor_apply {W X Y Z : Type u} (f : W ⟢ X) (g : Y ⟢ Z) (p : W βŠ— Y rfl #align category_theory.tensor_apply CategoryTheory.tensor_apply +@[simp] +theorem whiskerLeft_apply (X : Type u) {Y Z : Type u} (f : Y ⟢ Z) (p : X βŠ— Y) : + (X ◁ f) p = (p.1, f p.2) := + rfl + +@[simp] +theorem whiskerRight_apply {Y Z : Type u} (f : Y ⟢ Z) (X : Type u) (p : Y βŠ— X) : + (f β–· X) p = (f p.1, p.2) := + rfl + @[simp] theorem leftUnitor_hom_apply {X : Type u} {x : X} {p : PUnit} : ((Ξ»_ X).hom : πŸ™_ (Type u) βŠ— X β†’ X) (p, x) = x := diff --git a/Mathlib/RepresentationTheory/Action.lean b/Mathlib/RepresentationTheory/Action.lean index f23629d4528d2..50cf3f8d0468e 100644 --- a/Mathlib/RepresentationTheory/Action.lean +++ b/Mathlib/RepresentationTheory/Action.lean @@ -604,7 +604,7 @@ section variable [Preadditive V] [MonoidalPreadditive V] -attribute [local simp] MonoidalPreadditive.tensor_add MonoidalPreadditive.add_tensor +attribute [local simp] MonoidalPreadditive.whiskerLeft_add MonoidalPreadditive.add_whiskerRight instance : MonoidalPreadditive (Action V G) where @@ -980,32 +980,35 @@ variable {V} variable {W : Type (u + 1)} [LargeCategory W] [MonoidalCategory V] [MonoidalCategory W] (F : MonoidalFunctor V W) (G : MonCat.{u}) +open MonoidalCategory + +attribute [local simp] id_tensorHom tensorHom_id + /-- A monoidal functor induces a monoidal functor between the categories of `G`-actions within those categories. -/ -@[simps] +@[simps!] +def mapActionAux : LaxMonoidalFunctor (Action V G) (Action W G) := .ofTensorHom + (F := F.toFunctor.mapAction G) + (Ξ΅ := + { hom := F.Ξ΅ + comm := fun g => by + dsimp [FunctorCategoryEquivalence.inverse, Functor.mapAction] + rw [Category.id_comp, F.map_id, Category.comp_id] }) + (ΞΌ := fun X Y => + { hom := F.ΞΌ X.V Y.V + comm := fun g => F.toLaxMonoidalFunctor.ΞΌ_natural (X.ρ g) (Y.ρ g) }) + (ΞΌ_natural := by intros; ext; simp) + (associativity := by intros; ext; simp) + (left_unitality := by intros; ext; simp) + (right_unitality := by intros; ext; simp) + +/-- A monoidal functor induces a monoidal functor between +the categories of `G`-actions within those categories. -/ +@[simps!] def mapAction : MonoidalFunctor (Action V G) (Action W G) := - { F.toFunctor.mapAction G with - Ξ΅ := - { hom := F.Ξ΅ - comm := fun g => by - dsimp [FunctorCategoryEquivalence.inverse, Functor.mapAction] - rw [Category.id_comp, F.map_id, Category.comp_id] } - ΞΌ := fun X Y => - { hom := F.ΞΌ X.V Y.V - comm := fun g => F.toLaxMonoidalFunctor.ΞΌ_natural (X.ρ g) (Y.ρ g) } - Ξ΅_isIso := by infer_instance - ΞΌ_isIso := by infer_instance - ΞΌ_natural := by intros; ext; simp - associativity := by intros; ext; simp - left_unitality := by intros; ext; simp - right_unitality := by - intros - ext - dsimp - simp only [MonoidalCategory.rightUnitor_conjugation, - LaxMonoidalFunctor.right_unitality, Category.id_comp, Category.assoc, - LaxMonoidalFunctor.right_unitality_inv_assoc, Category.comp_id, Iso.hom_inv_id] - rw [← F.map_comp, Iso.inv_hom_id, F.map_id, Category.comp_id] } + { mapActionAux F G with + Ξ΅_isIso := by dsimp [mapActionAux]; infer_instance + ΞΌ_isIso := by dsimp [mapActionAux]; infer_instance } set_option linter.uppercaseLean3 false in #align category_theory.monoidal_functor.map_Action CategoryTheory.MonoidalFunctor.mapAction diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index 14fd9fede21ef..c2c2b13866a5d 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -114,10 +114,15 @@ namespace MonoidalCoherence instance refl (X : C) [LiftObj X] : MonoidalCoherence X X := βŸ¨πŸ™ _⟩ @[simps] -instance tensor (X Y Z : C) [LiftObj X] [LiftObj Y] [LiftObj Z] [MonoidalCoherence Y Z] : +instance whiskerLeft (X Y Z : C) [LiftObj X] [LiftObj Y] [LiftObj Z] [MonoidalCoherence Y Z] : MonoidalCoherence (X βŠ— Y) (X βŠ— Z) := ⟨X ◁ MonoidalCoherence.hom⟩ +@[simps] +instance whiskerRight (X Y Z : C) [LiftObj X] [LiftObj Y] [LiftObj Z] [MonoidalCoherence X Y] : + MonoidalCoherence (X βŠ— Z) (Y βŠ— Z) := + ⟨MonoidalCoherence.hom β–· Z⟩ + @[simps] instance tensor_right (X Y : C) [LiftObj X] [LiftObj Y] [MonoidalCoherence (πŸ™_ C) Y] : MonoidalCoherence X (X βŠ— Y) := @@ -179,7 +184,7 @@ def monoidalComp {W X Y Z : C} [LiftObj X] [LiftObj Y] f ≫ MonoidalCoherence.hom ≫ g @[inherit_doc monoidalComp] -infixr:80 " βŠ—β‰« " => monoidalComp -- type as \ot \gg +infixr:79 " βŠ—β‰« " => monoidalComp -- type as \ot \gg /-- Compose two isomorphisms in a monoidal category, inserting unitors and associators between as necessary. -/ @@ -188,7 +193,7 @@ noncomputable def monoidalIsoComp {W X Y Z : C} [LiftObj X] [LiftObj Y] f β‰ͺ≫ asIso MonoidalCoherence.hom β‰ͺ≫ g @[inherit_doc monoidalIsoComp] -infixr:80 " β‰ͺβŠ—β‰« " => monoidalIsoComp -- type as \ot \gg +infixr:79 " β‰ͺβŠ—β‰« " => monoidalIsoComp -- type as \ot \gg example {U V W X Y : C} (f : U ⟢ V βŠ— (W βŠ— X)) (g : (V βŠ— W) βŠ— X ⟢ Y) : U ⟢ Y := f βŠ—β‰« g From 03aa21f1ce3743d847cd027e43b0c8bb7decd080 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 3 Aug 2023 21:02:52 +0900 Subject: [PATCH 05/62] Bimod done --- .../Category/ModuleCat/Monoidal/Basic.lean | 6 - Mathlib/CategoryTheory/Monad/EquivMon.lean | 6 +- Mathlib/CategoryTheory/Monoidal/Bimod.lean | 583 +++++++++--------- Mathlib/CategoryTheory/Monoidal/Braided.lean | 12 - Mathlib/CategoryTheory/Monoidal/Discrete.lean | 1 - Mathlib/CategoryTheory/Monoidal/End.lean | 2 +- .../CategoryTheory/Monoidal/Free/Basic.lean | 41 +- .../Monoidal/FunctorCategory.lean | 1 - .../Monoidal/Internal/FunctorCategory.lean | 4 + Mathlib/CategoryTheory/Monoidal/Mod_.lean | 31 +- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 33 - .../CategoryTheory/Monoidal/Preadditive.lean | 43 +- .../CategoryTheory/Monoidal/Subcategory.lean | 24 - .../CategoryTheory/Monoidal/Transport.lean | 77 --- Mathlib/RepresentationTheory/Character.lean | 4 +- Mathlib/RepresentationTheory/FdRep.lean | 5 +- Mathlib/Tactic/CategoryTheory/Coherence.lean | 8 +- 17 files changed, 356 insertions(+), 525 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean index 3eee418fe0149..9d643164107eb 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean @@ -71,12 +71,6 @@ def whiskerRight {M₁ Mβ‚‚ : ModuleCat R} (f : M₁ ⟢ Mβ‚‚) (N : ModuleCat R) tensorObj M₁ N ⟢ tensorObj Mβ‚‚ N := f.rTensor N --- theorem whiskerLeft_id (M N : ModuleCat R) : tensorHom (πŸ™ M) (πŸ™ N) = πŸ™ (ModuleCat.of R (M βŠ— N)) := by --- -- Porting note: even with high priority ext fails to find this --- apply TensorProduct.ext --- rfl - - theorem tensor_id (M N : ModuleCat R) : tensorHom (πŸ™ M) (πŸ™ N) = πŸ™ (ModuleCat.of R (M βŠ— N)) := by -- Porting note: even with high priority ext fails to find this apply TensorProduct.ext diff --git a/Mathlib/CategoryTheory/Monad/EquivMon.lean b/Mathlib/CategoryTheory/Monad/EquivMon.lean index 74a46af127c22..b24fe287b95c6 100644 --- a/Mathlib/CategoryTheory/Monad/EquivMon.lean +++ b/Mathlib/CategoryTheory/Monad/EquivMon.lean @@ -66,14 +66,14 @@ def ofMon (M : Mon_ (C β₯€ C)) : Monad C where ΞΌ' := M.mul left_unit' := fun X => by -- Porting note: now using `erw` - erw [← NatTrans.id_hcomp_app M.one, ← NatTrans.comp_app, M.mul_one] + erw [← whiskerLeft_app, ← NatTrans.comp_app, M.mul_one] rfl right_unit' := fun X => by -- Porting note: now using `erw` - erw [← NatTrans.hcomp_id_app M.one, ← NatTrans.comp_app, M.one_mul] + erw [← whiskerRight_app, ← NatTrans.comp_app, M.one_mul] rfl assoc' := fun X => by - rw [← NatTrans.hcomp_id_app, ← NatTrans.comp_app] + rw [← whiskerLeft_app, ← whiskerRight_app, ← NatTrans.comp_app] -- Porting note: had to add this step: erw [M.mul_assoc] simp diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index 62685dcafaa40..9abdf782f05d7 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -22,6 +22,8 @@ open CategoryTheory.MonoidalCategory variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] +attribute [local simp] id_tensorHom tensorHom_id + section open CategoryTheory.Limits @@ -33,19 +35,19 @@ section variable [βˆ€ X : C, PreservesColimitsOfSize.{0, 0} (tensorLeft X)] theorem id_tensor_Ο€_preserves_coequalizer_inv_desc {W X Y Z : C} (f g : X ⟢ Y) (h : Z βŠ— Y ⟢ W) - (wh : (πŸ™ Z βŠ— f) ≫ h = (πŸ™ Z βŠ— g) ≫ h) : - (πŸ™ Z βŠ— coequalizer.Ο€ f g) ≫ + (wh : (Z ◁ f) ≫ h = (Z ◁ g) ≫ h) : + (Z ◁ coequalizer.Ο€ f g) ≫ (PreservesCoequalizer.iso (tensorLeft Z) f g).inv ≫ coequalizer.desc h wh = h := map_Ο€_preserves_coequalizer_inv_desc (tensorLeft Z) f g h wh #align id_tensor_Ο€_preserves_coequalizer_inv_desc id_tensor_Ο€_preserves_coequalizer_inv_desc theorem id_tensor_Ο€_preserves_coequalizer_inv_colimMap_desc {X Y Z X' Y' Z' : C} (f g : X ⟢ Y) - (f' g' : X' ⟢ Y') (p : Z βŠ— X ⟢ X') (q : Z βŠ— Y ⟢ Y') (wf : (πŸ™ Z βŠ— f) ≫ q = p ≫ f') - (wg : (πŸ™ Z βŠ— g) ≫ q = p ≫ g') (h : Y' ⟢ Z') (wh : f' ≫ h = g' ≫ h) : - (πŸ™ Z βŠ— coequalizer.Ο€ f g) ≫ + (f' g' : X' ⟢ Y') (p : Z βŠ— X ⟢ X') (q : Z βŠ— Y ⟢ Y') (wf : (Z ◁ f) ≫ q = p ≫ f') + (wg : (Z ◁ g) ≫ q = p ≫ g') (h : Y' ⟢ Z') (wh : f' ≫ h = g' ≫ h) : + (Z ◁ coequalizer.Ο€ f g) ≫ (PreservesCoequalizer.iso (tensorLeft Z) f g).inv ≫ - colimMap (parallelPairHom (πŸ™ Z βŠ— f) (πŸ™ Z βŠ— g) f' g' p q wf wg) ≫ coequalizer.desc h wh = + colimMap (parallelPairHom (Z ◁ f) (Z ◁ g) f' g' p q wf wg) ≫ coequalizer.desc h wh = q ≫ h := map_Ο€_preserves_coequalizer_inv_colimMap_desc (tensorLeft Z) f g f' g' p q wf wg h wh #align id_tensor_Ο€_preserves_coequalizer_inv_colim_map_desc id_tensor_Ο€_preserves_coequalizer_inv_colimMap_desc @@ -57,19 +59,19 @@ section variable [βˆ€ X : C, PreservesColimitsOfSize.{0, 0} (tensorRight X)] theorem Ο€_tensor_id_preserves_coequalizer_inv_desc {W X Y Z : C} (f g : X ⟢ Y) (h : Y βŠ— Z ⟢ W) - (wh : (f βŠ— πŸ™ Z) ≫ h = (g βŠ— πŸ™ Z) ≫ h) : - (coequalizer.Ο€ f g βŠ— πŸ™ Z) ≫ + (wh : (f β–· Z) ≫ h = (g β–· Z) ≫ h) : + (coequalizer.Ο€ f g β–· Z) ≫ (PreservesCoequalizer.iso (tensorRight Z) f g).inv ≫ coequalizer.desc h wh = h := map_Ο€_preserves_coequalizer_inv_desc (tensorRight Z) f g h wh #align Ο€_tensor_id_preserves_coequalizer_inv_desc Ο€_tensor_id_preserves_coequalizer_inv_desc theorem Ο€_tensor_id_preserves_coequalizer_inv_colimMap_desc {X Y Z X' Y' Z' : C} (f g : X ⟢ Y) - (f' g' : X' ⟢ Y') (p : X βŠ— Z ⟢ X') (q : Y βŠ— Z ⟢ Y') (wf : (f βŠ— πŸ™ Z) ≫ q = p ≫ f') - (wg : (g βŠ— πŸ™ Z) ≫ q = p ≫ g') (h : Y' ⟢ Z') (wh : f' ≫ h = g' ≫ h) : - (coequalizer.Ο€ f g βŠ— πŸ™ Z) ≫ + (f' g' : X' ⟢ Y') (p : X βŠ— Z ⟢ X') (q : Y βŠ— Z ⟢ Y') (wf : (f β–· Z) ≫ q = p ≫ f') + (wg : (g β–· Z) ≫ q = p ≫ g') (h : Y' ⟢ Z') (wh : f' ≫ h = g' ≫ h) : + (coequalizer.Ο€ f g β–· Z) ≫ (PreservesCoequalizer.iso (tensorRight Z) f g).inv ≫ - colimMap (parallelPairHom (f βŠ— πŸ™ Z) (g βŠ— πŸ™ Z) f' g' p q wf wg) ≫ coequalizer.desc h wh = + colimMap (parallelPairHom (f β–· Z) (g β–· Z) f' g' p q wf wg) ≫ coequalizer.desc h wh = q ≫ h := map_Ο€_preserves_coequalizer_inv_colimMap_desc (tensorRight Z) f g f' g' p q wf wg h wh #align Ο€_tensor_id_preserves_coequalizer_inv_colim_map_desc Ο€_tensor_id_preserves_coequalizer_inv_colimMap_desc @@ -82,16 +84,16 @@ end structure Bimod (A B : Mon_ C) where X : C actLeft : A.X βŠ— X ⟢ X - one_actLeft : (A.one βŠ— πŸ™ X) ≫ actLeft = (Ξ»_ X).hom := by aesop_cat + one_actLeft : (A.one β–· X) ≫ actLeft = (Ξ»_ X).hom := by aesop_cat left_assoc : - (A.mul βŠ— πŸ™ X) ≫ actLeft = (Ξ±_ A.X A.X X).hom ≫ (πŸ™ A.X βŠ— actLeft) ≫ actLeft := by aesop_cat + (A.mul β–· X) ≫ actLeft = (Ξ±_ A.X A.X X).hom ≫ (A.X ◁ actLeft) ≫ actLeft := by aesop_cat actRight : X βŠ— B.X ⟢ X - actRight_one : (πŸ™ X βŠ— B.one) ≫ actRight = (ρ_ X).hom := by aesop_cat + actRight_one : (X ◁ B.one) ≫ actRight = (ρ_ X).hom := by aesop_cat right_assoc : - (πŸ™ X βŠ— B.mul) ≫ actRight = (Ξ±_ X B.X B.X).inv ≫ (actRight βŠ— πŸ™ B.X) ≫ actRight := by + (X ◁ B.mul) ≫ actRight = (Ξ±_ X B.X B.X).inv ≫ (actRight β–· B.X) ≫ actRight := by aesop_cat middle_assoc : - (actLeft βŠ— πŸ™ B.X) ≫ actRight = (Ξ±_ A.X X B.X).hom ≫ (πŸ™ A.X βŠ— actRight) ≫ actLeft := by + (actLeft β–· B.X) ≫ actRight = (Ξ±_ A.X X B.X).hom ≫ (A.X ◁ actRight) ≫ actLeft := by aesop_cat set_option linter.uppercaseLean3 false in #align Bimod Bimod @@ -107,8 +109,8 @@ variable {A B : Mon_ C} (M : Bimod A B) @[ext] structure Hom (M N : Bimod A B) where hom : M.X ⟢ N.X - left_act_hom : M.actLeft ≫ hom = (πŸ™ A.X βŠ— hom) ≫ N.actLeft := by aesop_cat - right_act_hom : M.actRight ≫ hom = (hom βŠ— πŸ™ B.X) ≫ N.actRight := by aesop_cat + left_act_hom : M.actLeft ≫ hom = (A.X ◁ hom) ≫ N.actLeft := by aesop_cat + right_act_hom : M.actRight ≫ hom = (hom β–· B.X) ≫ N.actRight := by aesop_cat set_option linter.uppercaseLean3 false in #align Bimod.hom Bimod.Hom @@ -159,20 +161,20 @@ and checking compatibility with left and right actions only in the forward direc -/ @[simps] def isoOfIso {X Y : Mon_ C} {P Q : Bimod X Y} (f : P.X β‰… Q.X) - (f_left_act_hom : P.actLeft ≫ f.hom = (πŸ™ X.X βŠ— f.hom) ≫ Q.actLeft) - (f_right_act_hom : P.actRight ≫ f.hom = (f.hom βŠ— πŸ™ Y.X) ≫ Q.actRight) : P β‰… Q where + (f_left_act_hom : P.actLeft ≫ f.hom = (X.X ◁ f.hom) ≫ Q.actLeft) + (f_right_act_hom : P.actRight ≫ f.hom = (f.hom β–· Y.X) ≫ Q.actRight) : P β‰… Q where hom := { hom := f.hom } inv := { hom := f.inv left_act_hom := by rw [← cancel_mono f.hom, Category.assoc, Category.assoc, Iso.inv_hom_id, Category.comp_id, - f_left_act_hom, ← Category.assoc, ← id_tensor_comp, Iso.inv_hom_id, - MonoidalCategory.tensor_id, Category.id_comp] + f_left_act_hom, ← Category.assoc, ← MonoidalCategory.whiskerLeft_comp, Iso.inv_hom_id, + MonoidalCategory.whiskerLeft_id, Category.id_comp] right_act_hom := by rw [← cancel_mono f.hom, Category.assoc, Category.assoc, Iso.inv_hom_id, Category.comp_id, - f_right_act_hom, ← Category.assoc, ← comp_tensor_id, Iso.inv_hom_id, - MonoidalCategory.tensor_id, Category.id_comp] } + f_right_act_hom, ← Category.assoc, ← comp_whiskerRight, Iso.inv_hom_id, + MonoidalCategory.id_whiskerRight, Category.id_comp] } hom_inv_id := by ext; dsimp; rw [Iso.hom_inv_id] inv_hom_id := by ext; dsimp; rw [Iso.inv_hom_id] set_option linter.uppercaseLean3 false in @@ -209,7 +211,7 @@ variable {R S T : Mon_ C} (P : Bimod R S) (Q : Bimod S T) /-- The underlying object of the tensor product of two bimodules. -/ noncomputable def X : C := - coequalizer (P.actRight βŠ— πŸ™ Q.X) ((Ξ±_ _ _ _).hom ≫ (πŸ™ P.X βŠ— Q.actLeft)) + coequalizer (P.actRight β–· Q.X) ((Ξ±_ _ _ _).hom ≫ (P.X ◁ Q.actLeft)) set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.X Bimod.TensorBimod.X @@ -222,58 +224,56 @@ noncomputable def actLeft : R.X βŠ— X P Q ⟢ X P Q := (PreservesCoequalizer.iso (tensorLeft R.X) _ _).inv ≫ colimMap (parallelPairHom _ _ _ _ - ((πŸ™ _ βŠ— (Ξ±_ _ _ _).hom) ≫ (Ξ±_ _ _ _).inv ≫ (P.actLeft βŠ— πŸ™ S.X βŠ— πŸ™ Q.X) ≫ (Ξ±_ _ _ _).inv) - ((Ξ±_ _ _ _).inv ≫ (P.actLeft βŠ— πŸ™ Q.X)) + ((Ξ±_ _ _ _).inv ≫ ((Ξ±_ _ _ _).inv β–· _) ≫ (P.actLeft β–· S.X β–· Q.X)) + ((Ξ±_ _ _ _).inv ≫ (P.actLeft β–· Q.X)) (by dsimp - slice_lhs 1 2 => rw [associator_inv_naturality] - slice_rhs 3 4 => rw [associator_inv_naturality] - slice_rhs 4 5 => rw [← tensor_comp, middle_assoc, tensor_comp, comp_tensor_id] + simp only [Category.assoc] + slice_lhs 1 2 => rw [associator_inv_naturality_middle] + slice_rhs 3 4 => rw [← comp_whiskerRight, middle_assoc, comp_whiskerRight] coherence) (by dsimp - slice_lhs 1 1 => rw [id_tensor_comp] - slice_lhs 2 3 => rw [associator_inv_naturality] - slice_lhs 3 4 => rw [tensor_id, id_tensor_comp_tensor_id] - slice_rhs 4 6 => rw [Iso.inv_hom_id_assoc] - slice_rhs 3 4 => rw [tensor_id, tensor_id_comp_id_tensor])) + slice_lhs 1 1 => rw [MonoidalCategory.whiskerLeft_comp] + slice_lhs 2 3 => rw [associator_inv_naturality_right] + slice_lhs 3 4 => rw [whisker_exchange] + coherence)) set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.act_left Bimod.TensorBimod.actLeft -theorem id_tensor_Ο€_actLeft : - (πŸ™ R.X βŠ— coequalizer.Ο€ _ _) ≫ actLeft P Q = - (Ξ±_ _ _ _).inv ≫ (P.actLeft βŠ— πŸ™ Q.X) ≫ coequalizer.Ο€ _ _ := by +theorem whiskerLeft_Ο€_actLeft : + (R.X ◁ coequalizer.Ο€ _ _) ≫ actLeft P Q = + (Ξ±_ _ _ _).inv ≫ (P.actLeft β–· Q.X) ≫ coequalizer.Ο€ _ _ := by erw [map_Ο€_preserves_coequalizer_inv_colimMap (tensorLeft _)] simp only [Category.assoc] set_option linter.uppercaseLean3 false in -#align Bimod.tensor_Bimod.id_tensor_Ο€_act_left Bimod.TensorBimod.id_tensor_Ο€_actLeft +#align Bimod.tensor_Bimod.id_tensor_Ο€_act_left Bimod.TensorBimod.whiskerLeft_Ο€_actLeft -theorem one_act_left' : (R.one βŠ— πŸ™ _) ≫ actLeft P Q = (Ξ»_ _).hom := by +theorem one_act_left' : (R.one β–· _) ≫ actLeft P Q = (Ξ»_ _).hom := by refine' (cancel_epi ((tensorLeft _).map (coequalizer.Ο€ _ _))).1 _ dsimp [X] -- porting note: had to replace `rw` by `erw` - slice_lhs 1 2 => erw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_lhs 2 3 => rw [id_tensor_Ο€_actLeft] - slice_lhs 1 2 => rw [← MonoidalCategory.tensor_id, associator_inv_naturality] - slice_lhs 2 3 => rw [← comp_tensor_id, one_actLeft] + slice_lhs 1 2 => erw [whisker_exchange] + slice_lhs 2 3 => rw [whiskerLeft_Ο€_actLeft] + slice_lhs 1 2 => rw [associator_inv_naturality_left] + slice_lhs 2 3 => rw [← comp_whiskerRight, one_actLeft] slice_rhs 1 2 => rw [leftUnitor_naturality] coherence set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.one_act_left' Bimod.TensorBimod.one_act_left' theorem left_assoc' : - (R.mul βŠ— πŸ™ _) ≫ actLeft P Q = (Ξ±_ R.X R.X _).hom ≫ (πŸ™ R.X βŠ— actLeft P Q) ≫ actLeft P Q := by + (R.mul β–· _) ≫ actLeft P Q = (Ξ±_ R.X R.X _).hom ≫ (R.X ◁ actLeft P Q) ≫ actLeft P Q := by refine' (cancel_epi ((tensorLeft _).map (coequalizer.Ο€ _ _))).1 _ dsimp [X] - -- porting note: had to replace some `rw` by `erw` - slice_lhs 1 2 => erw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_lhs 2 3 => rw [id_tensor_Ο€_actLeft] - slice_lhs 1 2 => rw [← MonoidalCategory.tensor_id, associator_inv_naturality] - slice_lhs 2 3 => rw [← comp_tensor_id, left_assoc, comp_tensor_id, comp_tensor_id] - slice_rhs 1 2 => erw [← MonoidalCategory.tensor_id, associator_naturality] - slice_rhs 2 3 => rw [← id_tensor_comp, id_tensor_Ο€_actLeft, id_tensor_comp, id_tensor_comp] - slice_rhs 4 5 => rw [id_tensor_Ο€_actLeft] - slice_rhs 3 4 => rw [associator_inv_naturality] + slice_lhs 1 2 => rw [whisker_exchange] + slice_lhs 2 3 => rw [whiskerLeft_Ο€_actLeft] + slice_lhs 1 2 => rw [associator_inv_naturality_left] + slice_lhs 2 3 => rw [← comp_whiskerRight, left_assoc, comp_whiskerRight, comp_whiskerRight] + slice_rhs 1 2 => rw [associator_naturality_right] + slice_rhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, whiskerLeft_Ο€_actLeft, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + slice_rhs 4 5 => rw [whiskerLeft_Ο€_actLeft] + slice_rhs 3 4 => rw [associator_inv_naturality_middle] coherence set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.left_assoc' Bimod.TensorBimod.left_assoc' @@ -289,62 +289,55 @@ noncomputable def actRight : X P Q βŠ— T.X ⟢ X P Q := (PreservesCoequalizer.iso (tensorRight T.X) _ _).inv ≫ colimMap (parallelPairHom _ _ _ _ - ((Ξ±_ _ _ _).hom ≫ (Ξ±_ _ _ _).hom ≫ (πŸ™ P.X βŠ— πŸ™ S.X βŠ— Q.actRight) ≫ (Ξ±_ _ _ _).inv) - ((Ξ±_ _ _ _).hom ≫ (πŸ™ P.X βŠ— Q.actRight)) + ((Ξ±_ _ _ _).hom ≫ (Ξ±_ _ _ _).hom ≫ (P.X ◁ S.X ◁ Q.actRight) ≫ (Ξ±_ _ _ _).inv) + ((Ξ±_ _ _ _).hom ≫ (P.X ◁ Q.actRight)) (by dsimp - slice_lhs 1 2 => rw [associator_naturality] - slice_lhs 2 3 => rw [tensor_id, tensor_id_comp_id_tensor] - slice_rhs 3 4 => rw [associator_inv_naturality] - slice_rhs 2 4 => rw [Iso.hom_inv_id_assoc] - slice_rhs 2 3 => rw [tensor_id, id_tensor_comp_tensor_id]) + slice_lhs 1 2 => rw [associator_naturality_left] + slice_lhs 2 3 => rw [← whisker_exchange] + simp) (by dsimp - slice_lhs 1 1 => rw [comp_tensor_id] - slice_lhs 2 3 => rw [associator_naturality] - slice_lhs 3 4 => rw [← id_tensor_comp, middle_assoc, id_tensor_comp] - slice_rhs 4 6 => rw [Iso.inv_hom_id_assoc] - slice_rhs 3 4 => rw [← id_tensor_comp] - coherence)) + simp only [comp_whiskerRight, whisker_assoc, Category.assoc, Iso.inv_hom_id_assoc] + slice_lhs 3 4 => rw [← MonoidalCategory.whiskerLeft_comp, middle_assoc, MonoidalCategory.whiskerLeft_comp] + simp)) set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.act_right Bimod.TensorBimod.actRight theorem Ο€_tensor_id_actRight : - (coequalizer.Ο€ _ _ βŠ— πŸ™ T.X) ≫ actRight P Q = - (Ξ±_ _ _ _).hom ≫ (πŸ™ P.X βŠ— Q.actRight) ≫ coequalizer.Ο€ _ _ := by + (coequalizer.Ο€ _ _ β–· T.X) ≫ actRight P Q = + (Ξ±_ _ _ _).hom ≫ (P.X ◁ Q.actRight) ≫ coequalizer.Ο€ _ _ := by erw [map_Ο€_preserves_coequalizer_inv_colimMap (tensorRight _)] simp only [Category.assoc] set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.Ο€_tensor_id_act_right Bimod.TensorBimod.Ο€_tensor_id_actRight -theorem actRight_one' : (πŸ™ _ βŠ— T.one) ≫ actRight P Q = (ρ_ _).hom := by +theorem actRight_one' : (_ ◁ T.one) ≫ actRight P Q = (ρ_ _).hom := by refine' (cancel_epi ((tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp [X] -- porting note: had to replace `rw` by `erw` - slice_lhs 1 2 => erw [tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] + slice_lhs 1 2 =>erw [← whisker_exchange] slice_lhs 2 3 => rw [Ο€_tensor_id_actRight] - slice_lhs 1 2 => rw [← MonoidalCategory.tensor_id, associator_naturality] - slice_lhs 2 3 => rw [← id_tensor_comp, actRight_one] - slice_rhs 1 2 => rw [rightUnitor_naturality] - coherence + slice_lhs 1 2 => rw [associator_naturality_right] + slice_lhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, actRight_one] + simp set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.act_right_one' Bimod.TensorBimod.actRight_one' theorem right_assoc' : - (πŸ™ _ βŠ— T.mul) ≫ actRight P Q = - (Ξ±_ _ T.X T.X).inv ≫ (actRight P Q βŠ— πŸ™ T.X) ≫ actRight P Q := by + (_ ◁ T.mul) ≫ actRight P Q = + (Ξ±_ _ T.X T.X).inv ≫ (actRight P Q β–· T.X) ≫ actRight P Q := by refine' (cancel_epi ((tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp [X] -- porting note: had to replace some `rw` by `erw` - slice_lhs 1 2 => erw [tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] + slice_lhs 1 2 => rw [← whisker_exchange] slice_lhs 2 3 => rw [Ο€_tensor_id_actRight] - slice_lhs 1 2 => rw [← MonoidalCategory.tensor_id, associator_naturality] - slice_lhs 2 3 => rw [← id_tensor_comp, right_assoc, id_tensor_comp, id_tensor_comp] - slice_rhs 1 2 => erw [← MonoidalCategory.tensor_id, associator_inv_naturality] - slice_rhs 2 3 => rw [← comp_tensor_id, Ο€_tensor_id_actRight, comp_tensor_id, comp_tensor_id] + slice_lhs 1 2 => rw [associator_naturality_right] + slice_lhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, right_assoc, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + slice_rhs 1 2 => rw [associator_inv_naturality_left] + slice_rhs 2 3 => rw [← comp_whiskerRight, Ο€_tensor_id_actRight, comp_whiskerRight, comp_whiskerRight] slice_rhs 4 5 => rw [Ο€_tensor_id_actRight] - slice_rhs 3 4 => rw [associator_naturality] - coherence + simp set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.right_assoc' Bimod.TensorBimod.right_assoc' @@ -357,21 +350,20 @@ variable [βˆ€ X : C, PreservesColimitsOfSize.{0, 0} (tensorLeft X)] variable [βˆ€ X : C, PreservesColimitsOfSize.{0, 0} (tensorRight X)] theorem middle_assoc' : - (actLeft P Q βŠ— πŸ™ T.X) ≫ actRight P Q = - (Ξ±_ R.X _ T.X).hom ≫ (πŸ™ R.X βŠ— actRight P Q) ≫ actLeft P Q := by + (actLeft P Q β–· T.X) ≫ actRight P Q = + (Ξ±_ R.X _ T.X).hom ≫ (R.X ◁ actRight P Q) ≫ actLeft P Q := by refine' (cancel_epi ((tensorLeft _ β‹™ tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp [X] - slice_lhs 1 2 => rw [← comp_tensor_id, id_tensor_Ο€_actLeft, comp_tensor_id, comp_tensor_id] + slice_lhs 1 2 => rw [← comp_whiskerRight, whiskerLeft_Ο€_actLeft, comp_whiskerRight, comp_whiskerRight] slice_lhs 3 4 => rw [Ο€_tensor_id_actRight] - slice_lhs 2 3 => rw [associator_naturality] - slice_lhs 3 4 => rw [MonoidalCategory.tensor_id, tensor_id_comp_id_tensor] + slice_lhs 2 3 => rw [associator_naturality_left] -- porting note: had to replace `rw` by `erw` - slice_rhs 1 2 => erw [associator_naturality] - slice_rhs 2 3 => rw [← id_tensor_comp, Ο€_tensor_id_actRight, id_tensor_comp, id_tensor_comp] - slice_rhs 4 5 => rw [id_tensor_Ο€_actLeft] - slice_rhs 3 4 => rw [associator_inv_naturality] - slice_rhs 4 5 => rw [MonoidalCategory.tensor_id, id_tensor_comp_tensor_id] - coherence + slice_rhs 1 2 => rw [associator_naturality_middle] + slice_rhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, Ο€_tensor_id_actRight, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + slice_rhs 4 5 => rw [whiskerLeft_Ο€_actLeft] + slice_rhs 3 4 => rw [associator_inv_naturality_right] + slice_rhs 4 5 => rw [whisker_exchange] + simp set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.middle_assoc' Bimod.TensorBimod.middle_assoc' @@ -401,60 +393,67 @@ set_option linter.uppercaseLean3 false in /-- Tensor product of two morphisms of bimodule objects. -/ @[simps] -noncomputable def tensorHom {X Y Z : Mon_ C} {M₁ Mβ‚‚ : Bimod X Y} {N₁ Nβ‚‚ : Bimod Y Z} (f : M₁ ⟢ Mβ‚‚) - (g : N₁ ⟢ Nβ‚‚) : M₁.tensorBimod N₁ ⟢ Mβ‚‚.tensorBimod Nβ‚‚ where +noncomputable def whiskerLeft {X Y Z : Mon_ C} (M : Bimod X Y) {N₁ Nβ‚‚ : Bimod Y Z} (f : N₁ ⟢ Nβ‚‚) : + M.tensorBimod N₁ ⟢ M.tensorBimod Nβ‚‚ where hom := colimMap - (parallelPairHom _ _ _ _ ((f.hom βŠ— πŸ™ Y.X) βŠ— g.hom) (f.hom βŠ— g.hom) + (parallelPairHom _ _ _ _ (_ ◁ f.hom) (_ ◁ f.hom) + (by rw [whisker_exchange]) (by - rw [← tensor_comp, ← tensor_comp, Hom.right_act_hom, Category.id_comp, Category.comp_id]) - (by - slice_lhs 2 3 => rw [← tensor_comp, Hom.left_act_hom, Category.id_comp] - slice_rhs 1 2 => rw [associator_naturality] - slice_rhs 2 3 => rw [← tensor_comp, Category.comp_id])) + simp only [Category.assoc, tensor_whiskerLeft, Iso.inv_hom_id_assoc, + Iso.cancel_iso_hom_left] + slice_lhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, Hom.left_act_hom] + simp)) left_act_hom := by refine' (cancel_epi ((tensorLeft _).map (coequalizer.Ο€ _ _))).1 _ dsimp - slice_lhs 1 2 => rw [TensorBimod.id_tensor_Ο€_actLeft] + slice_lhs 1 2 => rw [TensorBimod.whiskerLeft_Ο€_actLeft] slice_lhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] - slice_lhs 2 3 => rw [← tensor_comp, Hom.left_act_hom, Category.id_comp] - slice_rhs 1 2 => rw [← id_tensor_comp, ΞΉ_colimMap, parallelPairHom_app_one, id_tensor_comp] - slice_rhs 2 3 => rw [TensorBimod.id_tensor_Ο€_actLeft] - slice_rhs 1 2 => rw [associator_inv_naturality] - slice_rhs 2 3 => rw [← tensor_comp, Category.comp_id] + slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, ΞΉ_colimMap, parallelPairHom_app_one, MonoidalCategory.whiskerLeft_comp] + slice_rhs 2 3 => rw [TensorBimod.whiskerLeft_Ο€_actLeft] + slice_rhs 1 2 => rw [associator_inv_naturality_right] + slice_rhs 2 3 => rw [whisker_exchange] + simp right_act_hom := by refine' (cancel_epi ((tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp slice_lhs 1 2 => rw [TensorBimod.Ο€_tensor_id_actRight] slice_lhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] - slice_lhs 2 3 => rw [← tensor_comp, Category.id_comp, Hom.right_act_hom] - slice_rhs 1 2 => rw [← comp_tensor_id, ΞΉ_colimMap, parallelPairHom_app_one, comp_tensor_id] + slice_lhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, Hom.right_act_hom] + slice_rhs 1 2 => rw [← comp_whiskerRight, ΞΉ_colimMap, parallelPairHom_app_one, comp_whiskerRight] slice_rhs 2 3 => rw [TensorBimod.Ο€_tensor_id_actRight] - slice_rhs 1 2 => rw [associator_naturality] - slice_rhs 2 3 => rw [← tensor_comp, Category.comp_id] -set_option linter.uppercaseLean3 false in -#align Bimod.tensor_hom Bimod.tensorHom - -theorem tensor_id {X Y Z : Mon_ C} {M : Bimod X Y} {N : Bimod Y Z} : - tensorHom (πŸ™ M) (πŸ™ N) = πŸ™ (M.tensorBimod N) := by - ext - apply Limits.coequalizer.hom_ext - simp only [id_hom', MonoidalCategory.tensor_id, tensorHom_hom, ΞΉ_colimMap, - parallelPairHom_app_one] - dsimp; dsimp only [TensorBimod.X] - simp only [Category.id_comp, Category.comp_id] -set_option linter.uppercaseLean3 false in -#align Bimod.tensor_id Bimod.tensor_id + simp -theorem tensor_comp {X Y Z : Mon_ C} {M₁ Mβ‚‚ M₃ : Bimod X Y} {N₁ Nβ‚‚ N₃ : Bimod Y Z} (f₁ : M₁ ⟢ Mβ‚‚) - (fβ‚‚ : Mβ‚‚ ⟢ M₃) (g₁ : N₁ ⟢ Nβ‚‚) (gβ‚‚ : Nβ‚‚ ⟢ N₃) : - tensorHom (f₁ ≫ fβ‚‚) (g₁ ≫ gβ‚‚) = tensorHom f₁ g₁ ≫ tensorHom fβ‚‚ gβ‚‚ := by - ext - apply Limits.coequalizer.hom_ext - simp only [comp_hom', MonoidalCategory.tensor_comp, tensorHom_hom, - ΞΉ_colimMap, parallelPairHom_app_one, Category.assoc, ΞΉ_colimMap_assoc] -set_option linter.uppercaseLean3 false in -#align Bimod.tensor_comp Bimod.tensor_comp +/-- Tensor product of two morphisms of bimodule objects. -/ +@[simps] +noncomputable def whiskerRight {X Y Z : Mon_ C} {M₁ Mβ‚‚ : Bimod X Y} (f : M₁ ⟢ Mβ‚‚) (N : Bimod Y Z) : + M₁.tensorBimod N ⟢ Mβ‚‚.tensorBimod N where + hom := + colimMap + (parallelPairHom _ _ _ _ (f.hom β–· _ β–· _) (f.hom β–· _) + (by rw [← comp_whiskerRight, Hom.right_act_hom, comp_whiskerRight]) + (by + slice_lhs 2 3 => rw [whisker_exchange] + simp)) + left_act_hom := by + refine' (cancel_epi ((tensorLeft _).map (coequalizer.Ο€ _ _))).1 _ + dsimp + slice_lhs 1 2 => rw [TensorBimod.whiskerLeft_Ο€_actLeft] + slice_lhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] + slice_lhs 2 3 => rw [← comp_whiskerRight, Hom.left_act_hom] + slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, ΞΉ_colimMap, parallelPairHom_app_one, MonoidalCategory.whiskerLeft_comp] + slice_rhs 2 3 => rw [TensorBimod.whiskerLeft_Ο€_actLeft] + slice_rhs 1 2 => rw [associator_inv_naturality_middle] + simp + right_act_hom := by + refine' (cancel_epi ((tensorRight _).map (coequalizer.Ο€ _ _))).1 _ + dsimp + slice_lhs 1 2 => rw [TensorBimod.Ο€_tensor_id_actRight] + slice_lhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] + slice_lhs 2 3 => rw [whisker_exchange] + slice_rhs 1 2 => rw [← comp_whiskerRight, ΞΉ_colimMap, parallelPairHom_app_one, comp_whiskerRight] + slice_rhs 2 3 => rw [TensorBimod.Ο€_tensor_id_actRight] + simp end @@ -470,19 +469,15 @@ variable {R S T U : Mon_ C} (P : Bimod R S) (Q : Bimod S T) (L : Bimod T U) the associator isomorphism. -/ noncomputable def homAux : (P.tensorBimod Q).X βŠ— L.X ⟢ (P.tensorBimod (Q.tensorBimod L)).X := (PreservesCoequalizer.iso (tensorRight L.X) _ _).inv ≫ - coequalizer.desc ((Ξ±_ _ _ _).hom ≫ (πŸ™ P.X βŠ— coequalizer.Ο€ _ _) ≫ coequalizer.Ο€ _ _) + coequalizer.desc ((Ξ±_ _ _ _).hom ≫ (P.X ◁ coequalizer.Ο€ _ _) ≫ coequalizer.Ο€ _ _) (by dsimp; dsimp [TensorBimod.X] - slice_lhs 1 2 => rw [associator_naturality] - slice_lhs 2 3 => - rw [MonoidalCategory.tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] + slice_lhs 1 2 => rw [associator_naturality_left] + slice_lhs 2 3 => rw [← whisker_exchange] slice_lhs 3 4 => rw [coequalizer.condition] - slice_lhs 2 3 => rw [← MonoidalCategory.tensor_id, associator_naturality] - slice_lhs 3 4 => rw [← id_tensor_comp, TensorBimod.id_tensor_Ο€_actLeft, id_tensor_comp] - slice_rhs 1 1 => rw [comp_tensor_id] - slice_rhs 2 3 => rw [associator_naturality] - slice_rhs 3 4 => rw [← id_tensor_comp] - coherence) + slice_lhs 2 3 => rw [associator_naturality_right] + slice_lhs 3 4 => rw [← MonoidalCategory.whiskerLeft_comp, TensorBimod.whiskerLeft_Ο€_actLeft, MonoidalCategory.whiskerLeft_comp] + simp) set_option linter.uppercaseLean3 false in #align Bimod.associator_Bimod.hom_aux Bimod.AssociatorBimod.homAux @@ -495,70 +490,66 @@ noncomputable def hom : refine' (cancel_epi ((tensorRight _ β‹™ tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp [TensorBimod.X] slice_lhs 1 2 => - rw [← comp_tensor_id, TensorBimod.Ο€_tensor_id_actRight, comp_tensor_id, comp_tensor_id] + rw [← comp_whiskerRight, TensorBimod.Ο€_tensor_id_actRight, comp_whiskerRight, comp_whiskerRight] slice_lhs 3 5 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] - slice_lhs 2 3 => rw [associator_naturality] - slice_lhs 3 4 => rw [← id_tensor_comp, coequalizer.condition, id_tensor_comp, id_tensor_comp] - slice_rhs 1 2 => rw [associator_naturality] - slice_rhs 2 3 => - rw [MonoidalCategory.tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] + slice_lhs 2 3 => rw [associator_naturality_middle] + slice_lhs 3 4 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.condition, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + slice_rhs 1 2 => rw [associator_naturality_left] + slice_rhs 2 3 => rw [← whisker_exchange] slice_rhs 3 5 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] - slice_rhs 2 3 => rw [← MonoidalCategory.tensor_id, associator_naturality] - coherence) + simp) set_option linter.uppercaseLean3 false in #align Bimod.associator_Bimod.hom Bimod.AssociatorBimod.hom theorem hom_left_act_hom' : ((P.tensorBimod Q).tensorBimod L).actLeft ≫ hom P Q L = - (πŸ™ R.X βŠ— hom P Q L) ≫ (P.tensorBimod (Q.tensorBimod L)).actLeft := by + (R.X ◁ hom P Q L) ≫ (P.tensorBimod (Q.tensorBimod L)).actLeft := by dsimp; dsimp [hom, homAux] refine' (cancel_epi ((tensorLeft _).map (coequalizer.Ο€ _ _))).1 _ rw [tensorLeft_map] - slice_lhs 1 2 => rw [TensorBimod.id_tensor_Ο€_actLeft] + slice_lhs 1 2 => rw [TensorBimod.whiskerLeft_Ο€_actLeft] slice_lhs 3 4 => rw [coequalizer.Ο€_desc] - slice_rhs 1 2 => rw [← id_tensor_comp, coequalizer.Ο€_desc, id_tensor_comp] + slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.Ο€_desc, MonoidalCategory.whiskerLeft_comp] refine' (cancel_epi ((tensorRight _ β‹™ tensorLeft _).map (coequalizer.Ο€ _ _))).1 _ dsimp; dsimp [TensorBimod.X] - slice_lhs 1 2 => rw [associator_inv_naturality] + slice_lhs 1 2 => rw [associator_inv_naturality_middle] slice_lhs 2 3 => - rw [← comp_tensor_id, TensorBimod.id_tensor_Ο€_actLeft, comp_tensor_id, comp_tensor_id] + rw [← comp_whiskerRight, TensorBimod.whiskerLeft_Ο€_actLeft, comp_whiskerRight, comp_whiskerRight] slice_lhs 4 6 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] - slice_lhs 3 4 => rw [associator_naturality] - slice_lhs 4 5 => rw [MonoidalCategory.tensor_id, tensor_id_comp_id_tensor] + slice_lhs 3 4 => rw [associator_naturality_left] slice_rhs 1 3 => - rw [← id_tensor_comp, ← id_tensor_comp, Ο€_tensor_id_preserves_coequalizer_inv_desc, - id_tensor_comp, id_tensor_comp] - slice_rhs 3 4 => erw [TensorBimod.id_tensor_Ο€_actLeft P (Q.tensorBimod L)] - slice_rhs 2 3 => erw [associator_inv_naturality] - slice_rhs 3 4 => erw [MonoidalCategory.tensor_id, id_tensor_comp_tensor_id] + rw [← MonoidalCategory.whiskerLeft_comp, ← MonoidalCategory.whiskerLeft_comp, Ο€_tensor_id_preserves_coequalizer_inv_desc, + MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + slice_rhs 3 4 => erw [TensorBimod.whiskerLeft_Ο€_actLeft P (Q.tensorBimod L)] + slice_rhs 2 3 => erw [associator_inv_naturality_right] + slice_rhs 3 4 => erw [whisker_exchange] coherence set_option linter.uppercaseLean3 false in #align Bimod.associator_Bimod.hom_left_act_hom' Bimod.AssociatorBimod.hom_left_act_hom' theorem hom_right_act_hom' : ((P.tensorBimod Q).tensorBimod L).actRight ≫ hom P Q L = - (hom P Q L βŠ— πŸ™ U.X) ≫ (P.tensorBimod (Q.tensorBimod L)).actRight := by + (hom P Q L β–· U.X) ≫ (P.tensorBimod (Q.tensorBimod L)).actRight := by dsimp; dsimp [hom, homAux] refine' (cancel_epi ((tensorRight _).map (coequalizer.Ο€ _ _))).1 _ rw [tensorRight_map] slice_lhs 1 2 => rw [TensorBimod.Ο€_tensor_id_actRight] slice_lhs 3 4 => rw [coequalizer.Ο€_desc] - slice_rhs 1 2 => rw [← comp_tensor_id, coequalizer.Ο€_desc, comp_tensor_id] + slice_rhs 1 2 => rw [← comp_whiskerRight, coequalizer.Ο€_desc, comp_whiskerRight] refine' (cancel_epi ((tensorRight _ β‹™ tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp; dsimp [TensorBimod.X] - slice_lhs 1 2 => rw [associator_naturality] - slice_lhs 2 3 => - rw [MonoidalCategory.tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] + slice_lhs 1 2 => rw [associator_naturality_left] + slice_lhs 2 3 => rw [← whisker_exchange] slice_lhs 3 5 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] - slice_lhs 2 3 => rw [← MonoidalCategory.tensor_id, associator_naturality] + slice_lhs 2 3 => rw [associator_naturality_right] slice_rhs 1 3 => - rw [← comp_tensor_id, ← comp_tensor_id, Ο€_tensor_id_preserves_coequalizer_inv_desc, - comp_tensor_id, comp_tensor_id] + rw [← comp_whiskerRight, ← comp_whiskerRight, Ο€_tensor_id_preserves_coequalizer_inv_desc, + comp_whiskerRight, comp_whiskerRight] slice_rhs 3 4 => erw [TensorBimod.Ο€_tensor_id_actRight P (Q.tensorBimod L)] - slice_rhs 2 3 => erw [associator_naturality] + slice_rhs 2 3 => erw [associator_naturality_middle] dsimp slice_rhs 3 4 => - rw [← id_tensor_comp, TensorBimod.Ο€_tensor_id_actRight, id_tensor_comp, id_tensor_comp] + rw [← MonoidalCategory.whiskerLeft_comp, TensorBimod.Ο€_tensor_id_actRight, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] coherence set_option linter.uppercaseLean3 false in #align Bimod.associator_Bimod.hom_right_act_hom' Bimod.AssociatorBimod.hom_right_act_hom' @@ -567,20 +558,19 @@ set_option linter.uppercaseLean3 false in the associator isomorphism. -/ noncomputable def invAux : P.X βŠ— (Q.tensorBimod L).X ⟢ ((P.tensorBimod Q).tensorBimod L).X := (PreservesCoequalizer.iso (tensorLeft P.X) _ _).inv ≫ - coequalizer.desc ((Ξ±_ _ _ _).inv ≫ (coequalizer.Ο€ _ _ βŠ— πŸ™ L.X) ≫ coequalizer.Ο€ _ _) + coequalizer.desc ((Ξ±_ _ _ _).inv ≫ (coequalizer.Ο€ _ _ β–· L.X) ≫ coequalizer.Ο€ _ _) (by dsimp; dsimp [TensorBimod.X] - slice_lhs 1 2 => rw [associator_inv_naturality] - rw [← Iso.inv_hom_id_assoc (Ξ±_ _ _ _) (πŸ™ P.X βŠ— Q.actRight), comp_tensor_id] + slice_lhs 1 2 => rw [associator_inv_naturality_middle] + rw [← Iso.inv_hom_id_assoc (Ξ±_ _ _ _) (P.X ◁ Q.actRight), comp_whiskerRight] slice_lhs 3 4 => - rw [← comp_tensor_id, Category.assoc, ← TensorBimod.Ο€_tensor_id_actRight, - comp_tensor_id] + rw [← comp_whiskerRight, Category.assoc, ← TensorBimod.Ο€_tensor_id_actRight, + comp_whiskerRight] slice_lhs 4 5 => rw [coequalizer.condition] - slice_lhs 3 4 => rw [associator_naturality] - slice_lhs 4 5 => rw [MonoidalCategory.tensor_id, tensor_id_comp_id_tensor] - slice_rhs 1 2 => rw [id_tensor_comp] - slice_rhs 2 3 => rw [associator_inv_naturality] - slice_rhs 3 4 => rw [MonoidalCategory.tensor_id, id_tensor_comp_tensor_id] + slice_lhs 3 4 => rw [associator_naturality_left] + slice_rhs 1 2 => rw [MonoidalCategory.whiskerLeft_comp] + slice_rhs 2 3 => rw [associator_inv_naturality_right] + slice_rhs 3 4 => rw [whisker_exchange] coherence) set_option linter.uppercaseLean3 false in #align Bimod.associator_Bimod.inv_aux Bimod.AssociatorBimod.invAux @@ -593,15 +583,15 @@ noncomputable def inv : dsimp [invAux] refine' (cancel_epi ((tensorLeft _).map (coequalizer.Ο€ _ _))).1 _ dsimp [TensorBimod.X] - slice_lhs 1 2 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] + slice_lhs 1 2 => rw [whisker_exchange] slice_lhs 2 4 => rw [id_tensor_Ο€_preserves_coequalizer_inv_desc] - slice_lhs 1 2 => rw [← MonoidalCategory.tensor_id, associator_inv_naturality] - slice_lhs 2 3 => rw [← comp_tensor_id, coequalizer.condition, comp_tensor_id, comp_tensor_id] - slice_rhs 1 2 => rw [← MonoidalCategory.tensor_id, associator_naturality] + slice_lhs 1 2 => rw [associator_inv_naturality_left] + slice_lhs 2 3 => rw [← comp_whiskerRight, coequalizer.condition, comp_whiskerRight, comp_whiskerRight] + slice_rhs 1 2 => rw [associator_naturality_right] slice_rhs 2 3 => - rw [← id_tensor_comp, TensorBimod.id_tensor_Ο€_actLeft, id_tensor_comp, id_tensor_comp] + rw [← MonoidalCategory.whiskerLeft_comp, TensorBimod.whiskerLeft_Ο€_actLeft, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] slice_rhs 4 6 => rw [id_tensor_Ο€_preserves_coequalizer_inv_desc] - slice_rhs 3 4 => rw [associator_inv_naturality] + slice_rhs 3 4 => rw [associator_inv_naturality_middle] coherence) set_option linter.uppercaseLean3 false in #align Bimod.associator_Bimod.inv Bimod.AssociatorBimod.inv @@ -650,7 +640,7 @@ set_option linter.uppercaseLean3 false in /-- The underlying morphism of the inverse component of the left unitor isomorphism. -/ noncomputable def inv : P.X ⟢ TensorBimod.X (regular R) P := - (Ξ»_ P.X).inv ≫ (R.one βŠ— πŸ™ _) ≫ coequalizer.Ο€ _ _ + (Ξ»_ P.X).inv ≫ (R.one β–· _) ≫ coequalizer.Ο€ _ _ set_option linter.uppercaseLean3 false in #align Bimod.left_unitor_Bimod.inv Bimod.LeftUnitorBimod.inv @@ -659,11 +649,11 @@ theorem hom_inv_id : hom P ≫ inv P = πŸ™ _ := by ext; dsimp slice_lhs 1 2 => rw [coequalizer.Ο€_desc] slice_lhs 1 2 => rw [leftUnitor_inv_naturality] - slice_lhs 2 3 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_lhs 3 3 => rw [← Iso.inv_hom_id_assoc (Ξ±_ R.X R.X P.X) (πŸ™ R.X βŠ— P.actLeft)] + slice_lhs 2 3 => rw [whisker_exchange] + slice_lhs 3 3 => rw [← Iso.inv_hom_id_assoc (Ξ±_ R.X R.X P.X) (R.X ◁ P.actLeft)] slice_lhs 4 6 => rw [← Category.assoc, ← coequalizer.condition] - slice_lhs 2 3 => rw [← MonoidalCategory.tensor_id, associator_inv_naturality] - slice_lhs 3 4 => rw [← comp_tensor_id, Mon_.one_mul] + slice_lhs 2 3 => rw [associator_inv_naturality_left] + slice_lhs 3 4 => rw [← comp_whiskerRight, Mon_.one_mul] slice_rhs 1 2 => rw [Category.comp_id] coherence set_option linter.uppercaseLean3 false in @@ -681,24 +671,24 @@ variable [βˆ€ X : C, PreservesColimitsOfSize.{0, 0} (tensorLeft X)] variable [βˆ€ X : C, PreservesColimitsOfSize.{0, 0} (tensorRight X)] theorem hom_left_act_hom' : - ((regular R).tensorBimod P).actLeft ≫ hom P = (πŸ™ R.X βŠ— hom P) ≫ P.actLeft := by + ((regular R).tensorBimod P).actLeft ≫ hom P = (R.X ◁ hom P) ≫ P.actLeft := by dsimp; dsimp [hom, TensorBimod.actLeft, regular] refine' (cancel_epi ((tensorLeft _).map (coequalizer.Ο€ _ _))).1 _ dsimp slice_lhs 1 4 => rw [id_tensor_Ο€_preserves_coequalizer_inv_colimMap_desc] slice_lhs 2 3 => rw [left_assoc] - slice_rhs 1 2 => rw [← id_tensor_comp, coequalizer.Ο€_desc] + slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.Ο€_desc] rw [Iso.inv_hom_id_assoc] set_option linter.uppercaseLean3 false in #align Bimod.left_unitor_Bimod.hom_left_act_hom' Bimod.LeftUnitorBimod.hom_left_act_hom' theorem hom_right_act_hom' : - ((regular R).tensorBimod P).actRight ≫ hom P = (hom P βŠ— πŸ™ S.X) ≫ P.actRight := by + ((regular R).tensorBimod P).actRight ≫ hom P = (hom P β–· S.X) ≫ P.actRight := by dsimp; dsimp [hom, TensorBimod.actRight, regular] refine' (cancel_epi ((tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp slice_lhs 1 4 => rw [Ο€_tensor_id_preserves_coequalizer_inv_colimMap_desc] - slice_rhs 1 2 => rw [← comp_tensor_id, coequalizer.Ο€_desc] + slice_rhs 1 2 => rw [← comp_whiskerRight, coequalizer.Ο€_desc] slice_rhs 1 2 => rw [middle_assoc] simp only [Category.assoc] set_option linter.uppercaseLean3 false in @@ -718,7 +708,7 @@ set_option linter.uppercaseLean3 false in /-- The underlying morphism of the inverse component of the right unitor isomorphism. -/ noncomputable def inv : P.X ⟢ TensorBimod.X P (regular S) := - (ρ_ P.X).inv ≫ (πŸ™ _ βŠ— S.one) ≫ coequalizer.Ο€ _ _ + (ρ_ P.X).inv ≫ (_ ◁ S.one) ≫ coequalizer.Ο€ _ _ set_option linter.uppercaseLean3 false in #align Bimod.right_unitor_Bimod.inv Bimod.RightUnitorBimod.inv @@ -727,10 +717,10 @@ theorem hom_inv_id : hom P ≫ inv P = πŸ™ _ := by ext; dsimp slice_lhs 1 2 => rw [coequalizer.Ο€_desc] slice_lhs 1 2 => rw [rightUnitor_inv_naturality] - slice_lhs 2 3 => rw [tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] + slice_lhs 2 3 => rw [← whisker_exchange] slice_lhs 3 4 => rw [coequalizer.condition] - slice_lhs 2 3 => rw [← MonoidalCategory.tensor_id, associator_naturality] - slice_lhs 3 4 => rw [← id_tensor_comp, Mon_.mul_one] + slice_lhs 2 3 => rw [associator_naturality_right] + slice_lhs 3 4 => rw [← MonoidalCategory.whiskerLeft_comp, Mon_.mul_one] slice_rhs 1 2 => rw [Category.comp_id] coherence set_option linter.uppercaseLean3 false in @@ -748,25 +738,25 @@ variable [βˆ€ X : C, PreservesColimitsOfSize.{0, 0} (tensorLeft X)] variable [βˆ€ X : C, PreservesColimitsOfSize.{0, 0} (tensorRight X)] theorem hom_left_act_hom' : - (P.tensorBimod (regular S)).actLeft ≫ hom P = (πŸ™ R.X βŠ— hom P) ≫ P.actLeft := by + (P.tensorBimod (regular S)).actLeft ≫ hom P = (R.X ◁ hom P) ≫ P.actLeft := by dsimp; dsimp [hom, TensorBimod.actLeft, regular] refine' (cancel_epi ((tensorLeft _).map (coequalizer.Ο€ _ _))).1 _ dsimp slice_lhs 1 4 => rw [id_tensor_Ο€_preserves_coequalizer_inv_colimMap_desc] slice_lhs 2 3 => rw [middle_assoc] - slice_rhs 1 2 => rw [← id_tensor_comp, coequalizer.Ο€_desc] + slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.Ο€_desc] rw [Iso.inv_hom_id_assoc] set_option linter.uppercaseLean3 false in #align Bimod.right_unitor_Bimod.hom_left_act_hom' Bimod.RightUnitorBimod.hom_left_act_hom' theorem hom_right_act_hom' : - (P.tensorBimod (regular S)).actRight ≫ hom P = (hom P βŠ— πŸ™ S.X) ≫ P.actRight := by + (P.tensorBimod (regular S)).actRight ≫ hom P = (hom P β–· S.X) ≫ P.actRight := by dsimp; dsimp [hom, TensorBimod.actRight, regular] refine' (cancel_epi ((tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp slice_lhs 1 4 => rw [Ο€_tensor_id_preserves_coequalizer_inv_colimMap_desc] slice_lhs 2 3 => rw [right_assoc] - slice_rhs 1 2 => rw [← comp_tensor_id, coequalizer.Ο€_desc] + slice_rhs 1 2 => rw [← comp_whiskerRight, coequalizer.Ο€_desc] rw [Iso.hom_inv_id_assoc] set_option linter.uppercaseLean3 false in #align Bimod.right_unitor_Bimod.hom_right_act_hom' Bimod.RightUnitorBimod.hom_right_act_hom' @@ -811,14 +801,34 @@ noncomputable def rightUnitorBimod {X Y : Mon_ C} (M : Bimod X Y) : M.tensorBimo set_option linter.uppercaseLean3 false in #align Bimod.right_unitor_Bimod Bimod.rightUnitorBimod -theorem whisker_left_comp_bimod {X Y Z : Mon_ C} (M : Bimod X Y) {N P Q : Bimod Y Z} (f : N ⟢ P) - (g : P ⟢ Q) : tensorHom (πŸ™ M) (f ≫ g) = tensorHom (πŸ™ M) f ≫ tensorHom (πŸ™ M) g := by - rw [← tensor_comp, Category.comp_id] +theorem whiskerLeft_id_bimod {X Y Z : Mon_ C} {M : Bimod X Y} {N : Bimod Y Z} : + whiskerLeft M (πŸ™ N) = πŸ™ (M.tensorBimod N) := by + ext + apply Limits.coequalizer.hom_ext + dsimp only [tensorBimod_X, whiskerLeft_hom, id_hom'] + simp only [MonoidalCategory.whiskerLeft_id, ΞΉ_colimMap, parallelPair_obj_one, + parallelPairHom_app_one, Category.id_comp] + erw [Category.comp_id] + +theorem id_whiskerRight_bimod {X Y Z : Mon_ C} {M : Bimod X Y} {N : Bimod Y Z} : + whiskerRight (πŸ™ M) N = πŸ™ (M.tensorBimod N) := by + ext + apply Limits.coequalizer.hom_ext + dsimp only [tensorBimod_X, whiskerRight_hom, id_hom'] + simp only [MonoidalCategory.id_whiskerRight, ΞΉ_colimMap, parallelPair_obj_one, + parallelPairHom_app_one, Category.id_comp] + erw [Category.comp_id] + +theorem whiskerLeft_comp_bimod {X Y Z : Mon_ C} (M : Bimod X Y) {N P Q : Bimod Y Z} (f : N ⟢ P) + (g : P ⟢ Q) : whiskerLeft M (f ≫ g) = whiskerLeft M f ≫ whiskerLeft M g := by + ext + apply Limits.coequalizer.hom_ext + simp set_option linter.uppercaseLean3 false in -#align Bimod.whisker_left_comp_Bimod Bimod.whisker_left_comp_bimod +#align Bimod.whisker_left_comp_Bimod Bimod.whiskerLeft_comp_bimod -theorem id_whisker_left_bimod {X Y : Mon_ C} {M N : Bimod X Y} (f : M ⟢ N) : - tensorHom (πŸ™ (regular X)) f = (leftUnitorBimod M).hom ≫ f ≫ (leftUnitorBimod N).inv := by +theorem id_whiskerLeft_bimod {X Y : Mon_ C} {M N : Bimod X Y} (f : M ⟢ N) : + whiskerLeft (regular X) f = (leftUnitorBimod M).hom ≫ f ≫ (leftUnitorBimod N).inv := by dsimp [tensorHom, regular, leftUnitorBimod] ext apply coequalizer.hom_ext @@ -829,23 +839,23 @@ theorem id_whisker_left_bimod {X Y : Mon_ C} {M N : Bimod X Y} (f : M ⟢ N) : dsimp [LeftUnitorBimod.inv] slice_rhs 1 2 => rw [Hom.left_act_hom] slice_rhs 2 3 => rw [leftUnitor_inv_naturality] - slice_rhs 3 4 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_rhs 4 4 => rw [← Iso.inv_hom_id_assoc (Ξ±_ X.X X.X N.X) (πŸ™ X.X βŠ— N.actLeft)] + slice_rhs 3 4 => rw [whisker_exchange] + slice_rhs 4 4 => rw [← Iso.inv_hom_id_assoc (Ξ±_ X.X X.X N.X) (X.X ◁ N.actLeft)] slice_rhs 5 7 => rw [← Category.assoc, ← coequalizer.condition] - slice_rhs 3 4 => rw [← MonoidalCategory.tensor_id, associator_inv_naturality] - slice_rhs 4 5 => rw [← comp_tensor_id, Mon_.one_mul] - have : (Ξ»_ (X.X βŠ— N.X)).inv ≫ (Ξ±_ (πŸ™_ C) X.X N.X).inv ≫ ((Ξ»_ X.X).hom βŠ— πŸ™ N.X) = πŸ™ _ := by + slice_rhs 3 4 => rw [associator_inv_naturality_left] + slice_rhs 4 5 => rw [← comp_whiskerRight, Mon_.one_mul] + have : (Ξ»_ (X.X βŠ— N.X)).inv ≫ (Ξ±_ (πŸ™_ C) X.X N.X).inv ≫ ((Ξ»_ X.X).hom β–· N.X) = πŸ™ _ := by pure_coherence slice_rhs 2 4 => rw [this] slice_rhs 1 2 => rw [Category.comp_id] set_option linter.uppercaseLean3 false in -#align Bimod.id_whisker_left_Bimod Bimod.id_whisker_left_bimod +#align Bimod.id_whisker_left_Bimod Bimod.id_whiskerLeft_bimod -theorem comp_whisker_left_bimod {W X Y Z : Mon_ C} (M : Bimod W X) (N : Bimod X Y) +theorem comp_whiskerLeft_bimod {W X Y Z : Mon_ C} (M : Bimod W X) (N : Bimod X Y) {P P' : Bimod Y Z} (f : P ⟢ P') : - tensorHom (πŸ™ (M.tensorBimod N)) f = + whiskerLeft (M.tensorBimod N) f = (associatorBimod M N P).hom ≫ - tensorHom (πŸ™ M) (tensorHom (πŸ™ N) f) ≫ (associatorBimod M N P').inv := by + whiskerLeft M (whiskerLeft N f) ≫ (associatorBimod M N P').inv := by dsimp [tensorHom, tensorBimod, associatorBimod] ext apply coequalizer.hom_ext @@ -858,25 +868,27 @@ theorem comp_whisker_left_bimod {W X Y Z : Mon_ C} (M : Bimod W X) (N : Bimod X rw [tensorRight_map] slice_rhs 1 3 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] slice_rhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] - slice_rhs 2 3 => rw [← id_tensor_comp, ΞΉ_colimMap, parallelPairHom_app_one] + slice_rhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, ΞΉ_colimMap, parallelPairHom_app_one] slice_rhs 3 4 => rw [coequalizer.Ο€_desc] dsimp [AssociatorBimod.invAux] - slice_rhs 2 2 => rw [id_tensor_comp] + slice_rhs 2 2 => rw [MonoidalCategory.whiskerLeft_comp] slice_rhs 3 5 => rw [id_tensor_Ο€_preserves_coequalizer_inv_desc] - slice_rhs 2 3 => rw [associator_inv_naturality] - slice_rhs 1 3 => rw [Iso.hom_inv_id_assoc, MonoidalCategory.tensor_id] - slice_lhs 1 2 => rw [tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] + slice_rhs 2 3 => rw [associator_inv_naturality_right] + slice_rhs 1 3 => rw [Iso.hom_inv_id_assoc] + slice_lhs 1 2 => rw [← whisker_exchange] set_option linter.uppercaseLean3 false in -#align Bimod.comp_whisker_left_Bimod Bimod.comp_whisker_left_bimod +#align Bimod.comp_whisker_left_Bimod Bimod.comp_whiskerLeft_bimod -theorem comp_whisker_right_bimod {X Y Z : Mon_ C} {M N P : Bimod X Y} (f : M ⟢ N) (g : N ⟢ P) - (Q : Bimod Y Z) : tensorHom (f ≫ g) (πŸ™ Q) = tensorHom f (πŸ™ Q) ≫ tensorHom g (πŸ™ Q) := by - rw [← tensor_comp, Category.comp_id] +theorem comp_whiskerRight_bimod {X Y Z : Mon_ C} {M N P : Bimod X Y} (f : M ⟢ N) (g : N ⟢ P) + (Q : Bimod Y Z) : whiskerRight (f ≫ g) Q = whiskerRight f Q ≫ whiskerRight g Q := by + ext + apply Limits.coequalizer.hom_ext + simp set_option linter.uppercaseLean3 false in -#align Bimod.comp_whisker_right_Bimod Bimod.comp_whisker_right_bimod +#align Bimod.comp_whisker_right_Bimod Bimod.comp_whiskerRight_bimod theorem whisker_right_id_bimod {X Y : Mon_ C} {M N : Bimod X Y} (f : M ⟢ N) : - tensorHom f (πŸ™ (regular Y)) = (rightUnitorBimod M).hom ≫ f ≫ (rightUnitorBimod N).inv := by + whiskerRight f (regular Y) = (rightUnitorBimod M).hom ≫ f ≫ (rightUnitorBimod N).inv := by dsimp [tensorHom, regular, rightUnitorBimod] ext apply coequalizer.hom_ext @@ -887,22 +899,19 @@ theorem whisker_right_id_bimod {X Y : Mon_ C} {M N : Bimod X Y} (f : M ⟢ N) : dsimp [RightUnitorBimod.inv] slice_rhs 1 2 => rw [Hom.right_act_hom] slice_rhs 2 3 => rw [rightUnitor_inv_naturality] - slice_rhs 3 4 => rw [tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] + slice_rhs 3 4 => rw [← whisker_exchange] slice_rhs 4 5 => rw [coequalizer.condition] - slice_rhs 3 4 => rw [← MonoidalCategory.tensor_id, associator_naturality] - slice_rhs 4 5 => rw [← id_tensor_comp, Mon_.mul_one] - have : (ρ_ (N.X βŠ— Y.X)).inv ≫ (Ξ±_ N.X Y.X (πŸ™_ C)).hom ≫ (πŸ™ N.X βŠ— (ρ_ Y.X).hom) = πŸ™ _ := by - pure_coherence - slice_rhs 2 4 => rw [this] - slice_rhs 1 2 => rw [Category.comp_id] + slice_rhs 3 4 => rw [associator_naturality_right] + slice_rhs 4 5 => rw [← MonoidalCategory.whiskerLeft_comp, Mon_.mul_one] + simp set_option linter.uppercaseLean3 false in #align Bimod.whisker_right_id_Bimod Bimod.whisker_right_id_bimod theorem whisker_right_comp_bimod {W X Y Z : Mon_ C} {M M' : Bimod W X} (f : M ⟢ M') (N : Bimod X Y) (P : Bimod Y Z) : - tensorHom f (πŸ™ (N.tensorBimod P)) = + whiskerRight f (N.tensorBimod P) = (associatorBimod M N P).inv ≫ - tensorHom (tensorHom f (πŸ™ N)) (πŸ™ P) ≫ (associatorBimod M' N P).hom := by + whiskerRight (whiskerRight f N) P ≫ (associatorBimod M' N P).hom := by dsimp [tensorHom, tensorBimod, associatorBimod] ext apply coequalizer.hom_ext @@ -915,22 +924,22 @@ theorem whisker_right_comp_bimod {W X Y Z : Mon_ C} {M M' : Bimod W X} (f : M rw [tensorLeft_map] slice_rhs 1 3 => rw [id_tensor_Ο€_preserves_coequalizer_inv_desc] slice_rhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] - slice_rhs 2 3 => rw [← comp_tensor_id, ΞΉ_colimMap, parallelPairHom_app_one] + slice_rhs 2 3 => rw [← comp_whiskerRight, ΞΉ_colimMap, parallelPairHom_app_one] slice_rhs 3 4 => rw [coequalizer.Ο€_desc] dsimp [AssociatorBimod.homAux] - slice_rhs 2 2 => rw [comp_tensor_id] + slice_rhs 2 2 => rw [comp_whiskerRight] slice_rhs 3 5 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] - slice_rhs 2 3 => rw [associator_naturality] - slice_rhs 1 3 => rw [Iso.inv_hom_id_assoc, MonoidalCategory.tensor_id] - slice_lhs 1 2 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] + slice_rhs 2 3 => rw [associator_naturality_left] + slice_rhs 1 3 => rw [Iso.inv_hom_id_assoc] + slice_lhs 1 2 => rw [whisker_exchange] set_option linter.uppercaseLean3 false in #align Bimod.whisker_right_comp_Bimod Bimod.whisker_right_comp_bimod theorem whisker_assoc_bimod {W X Y Z : Mon_ C} (M : Bimod W X) {N N' : Bimod X Y} (f : N ⟢ N') (P : Bimod Y Z) : - tensorHom (tensorHom (πŸ™ M) f) (πŸ™ P) = + whiskerRight (whiskerLeft M f) P = (associatorBimod M N P).hom ≫ - tensorHom (πŸ™ M) (tensorHom f (πŸ™ P)) ≫ (associatorBimod M N' P).inv := by + whiskerLeft M (whiskerRight f P) ≫ (associatorBimod M N' P).inv := by dsimp [tensorHom, tensorBimod, associatorBimod] ext apply coequalizer.hom_ext @@ -941,46 +950,46 @@ theorem whisker_assoc_bimod {W X Y Z : Mon_ C} (M : Bimod W X) {N N' : Bimod X Y dsimp [AssociatorBimod.homAux] refine' (cancel_epi ((tensorRight _).map (coequalizer.Ο€ _ _))).1 _ rw [tensorRight_map] - slice_lhs 1 2 => rw [← comp_tensor_id, ΞΉ_colimMap, parallelPairHom_app_one] + slice_lhs 1 2 => rw [← comp_whiskerRight, ΞΉ_colimMap, parallelPairHom_app_one] slice_rhs 1 3 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] slice_rhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] - slice_rhs 2 3 => rw [← id_tensor_comp, ΞΉ_colimMap, parallelPairHom_app_one] + slice_rhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, ΞΉ_colimMap, parallelPairHom_app_one] dsimp [AssociatorBimod.inv] slice_rhs 3 4 => rw [coequalizer.Ο€_desc] dsimp [AssociatorBimod.invAux] - slice_rhs 2 2 => rw [id_tensor_comp] + slice_rhs 2 2 => rw [MonoidalCategory.whiskerLeft_comp] slice_rhs 3 5 => rw [id_tensor_Ο€_preserves_coequalizer_inv_desc] - slice_rhs 2 3 => rw [associator_inv_naturality] + slice_rhs 2 3 => rw [associator_inv_naturality_middle] slice_rhs 1 3 => rw [Iso.hom_inv_id_assoc] - slice_lhs 1 1 => rw [comp_tensor_id] + slice_lhs 1 1 => rw [comp_whiskerRight] set_option linter.uppercaseLean3 false in #align Bimod.whisker_assoc_Bimod Bimod.whisker_assoc_bimod theorem whisker_exchange_bimod {X Y Z : Mon_ C} {M N : Bimod X Y} {P Q : Bimod Y Z} (f : M ⟢ N) - (g : P ⟢ Q) : tensorHom (πŸ™ M) g ≫ tensorHom f (πŸ™ Q) = - tensorHom f (πŸ™ P) ≫ tensorHom (πŸ™ N) g := by + (g : P ⟢ Q) : whiskerLeft M g ≫ whiskerRight f Q = + whiskerRight f P ≫ whiskerLeft N g := by dsimp [tensorHom] ext apply coequalizer.hom_ext dsimp slice_lhs 1 2 => rw [ΞΉ_colimMap, parallelPairHom_app_one] slice_lhs 2 3 => rw [ΞΉ_colimMap, parallelPairHom_app_one] - slice_lhs 1 2 => rw [id_tensor_comp_tensor_id] + slice_lhs 1 2 => rw [whisker_exchange] slice_rhs 1 2 => rw [ΞΉ_colimMap, parallelPairHom_app_one] slice_rhs 2 3 => rw [ΞΉ_colimMap, parallelPairHom_app_one] - slice_rhs 1 2 => rw [tensor_id_comp_id_tensor] + simp only [Category.assoc] set_option linter.uppercaseLean3 false in #align Bimod.whisker_exchange_Bimod Bimod.whisker_exchange_bimod set_option maxHeartbeats 400000 in theorem pentagon_bimod {V W X Y Z : Mon_ C} (M : Bimod V W) (N : Bimod W X) (P : Bimod X Y) (Q : Bimod Y Z) : - tensorHom (associatorBimod M N P).hom (πŸ™ Q) ≫ + whiskerRight (associatorBimod M N P).hom Q ≫ (associatorBimod M (N.tensorBimod P) Q).hom ≫ - tensorHom (πŸ™ M) (associatorBimod N P Q).hom = + whiskerLeft M (associatorBimod N P Q).hom = (associatorBimod (M.tensorBimod N) P Q).hom ≫ (associatorBimod M N (P.tensorBimod Q)).hom := by - dsimp [tensorHom, associatorBimod] + dsimp [associatorBimod] ext apply coequalizer.hom_ext dsimp @@ -991,35 +1000,36 @@ theorem pentagon_bimod {V W X Y Z : Mon_ C} (M : Bimod V W) (N : Bimod W X) (P : dsimp [AssociatorBimod.homAux] refine' (cancel_epi ((tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp - slice_lhs 1 2 => rw [← comp_tensor_id, coequalizer.Ο€_desc] + simp only [id_tensorHom, tensorHom_id] + slice_lhs 1 2 => rw [← comp_whiskerRight, coequalizer.Ο€_desc] slice_rhs 1 3 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] slice_rhs 3 4 => rw [coequalizer.Ο€_desc] refine' (cancel_epi ((tensorRight _ β‹™ tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp slice_lhs 1 2 => - rw [← comp_tensor_id, Ο€_tensor_id_preserves_coequalizer_inv_desc, comp_tensor_id, - comp_tensor_id] + rw [← comp_whiskerRight, Ο€_tensor_id_preserves_coequalizer_inv_desc, comp_whiskerRight, + comp_whiskerRight] slice_lhs 3 5 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] dsimp only [TensorBimod.X] - slice_lhs 2 3 => rw [associator_naturality] + slice_lhs 2 3 => rw [associator_naturality_middle] slice_lhs 5 6 => rw [ΞΉ_colimMap, parallelPairHom_app_one] - slice_lhs 4 5 => rw [← id_tensor_comp, coequalizer.Ο€_desc] + slice_lhs 4 5 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.Ο€_desc] slice_lhs 3 4 => - rw [← id_tensor_comp, Ο€_tensor_id_preserves_coequalizer_inv_desc, id_tensor_comp, - id_tensor_comp] - slice_rhs 1 2 => rw [associator_naturality] + rw [← MonoidalCategory.whiskerLeft_comp, Ο€_tensor_id_preserves_coequalizer_inv_desc, MonoidalCategory.whiskerLeft_comp, + MonoidalCategory.whiskerLeft_comp] + slice_rhs 1 2 => rw [associator_naturality_left] slice_rhs 2 3 => - rw [MonoidalCategory.tensor_id, tensor_id_comp_id_tensor, ← id_tensor_comp_tensor_id] + rw [← whisker_exchange] slice_rhs 3 5 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] - slice_rhs 2 3 => rw [← MonoidalCategory.tensor_id, associator_naturality] + slice_rhs 2 3 => rw [associator_naturality_right] coherence set_option linter.uppercaseLean3 false in #align Bimod.pentagon_Bimod Bimod.pentagon_bimod theorem triangle_bimod {X Y Z : Mon_ C} (M : Bimod X Y) (N : Bimod Y Z) : - (associatorBimod M (regular Y) N).hom ≫ tensorHom (πŸ™ M) (leftUnitorBimod N).hom = - tensorHom (rightUnitorBimod M).hom (πŸ™ N) := by - dsimp [tensorHom, associatorBimod, leftUnitorBimod, rightUnitorBimod] + (associatorBimod M (regular Y) N).hom ≫ whiskerLeft M (leftUnitorBimod N).hom = + whiskerRight (rightUnitorBimod M).hom N := by + dsimp [associatorBimod, leftUnitorBimod, rightUnitorBimod] ext apply coequalizer.hom_ext dsimp @@ -1033,8 +1043,9 @@ theorem triangle_bimod {X Y Z : Mon_ C} (M : Bimod X Y) (N : Bimod Y Z) : slice_lhs 1 3 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] slice_lhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] dsimp [LeftUnitorBimod.hom] - slice_lhs 2 3 => rw [← id_tensor_comp, coequalizer.Ο€_desc] - slice_rhs 1 2 => rw [← comp_tensor_id, coequalizer.Ο€_desc] + simp only [id_tensorHom, tensorHom_id] + slice_lhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.Ο€_desc] + slice_rhs 1 2 => rw [← comp_whiskerRight, coequalizer.Ο€_desc] slice_rhs 1 2 => rw [coequalizer.condition] simp only [Category.assoc] set_option linter.uppercaseLean3 false in @@ -1046,17 +1057,17 @@ noncomputable def monBicategory : Bicategory (Mon_ C) where homCategory X Y := (inferInstance : Category (Bimod X Y)) id X := regular X comp M N := tensorBimod M N - whiskerLeft L _ _ f := tensorHom (Bimod.id' L) f - whiskerRight f N := tensorHom f (Bimod.id' N) + whiskerLeft L _ _ f := whiskerLeft L f + whiskerRight f N := whiskerRight f N associator := associatorBimod leftUnitor := leftUnitorBimod rightUnitor := rightUnitorBimod - whiskerLeft_id _ _ := tensor_id - whiskerLeft_comp M _ _ _ f g := whisker_left_comp_bimod M f g - id_whiskerLeft := id_whisker_left_bimod - comp_whiskerLeft M N _ _ f := comp_whisker_left_bimod M N f - id_whiskerRight _ _ := tensor_id - comp_whiskerRight f g Q := comp_whisker_right_bimod f g Q + whiskerLeft_id _ _ := whiskerLeft_id_bimod + whiskerLeft_comp M _ _ _ f g := whiskerLeft_comp_bimod M f g + id_whiskerLeft := id_whiskerLeft_bimod + comp_whiskerLeft M N _ _ f := comp_whiskerLeft_bimod M N f + id_whiskerRight _ _ := id_whiskerRight_bimod + comp_whiskerRight f g Q := comp_whiskerRight_bimod f g Q whiskerRight_id := whisker_right_id_bimod whiskerRight_comp := whisker_right_comp_bimod whisker_assoc M _ _ f P := whisker_assoc_bimod M f P diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 1fb0f659eabcf..93862ffea2c46 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -51,10 +51,6 @@ class BraidedCategory (C : Type u) [Category.{v} C] [MonoidalCategory.{v} C] whe βˆ€ {X Y : C} (f : X ⟢ Y) (Z : C), f β–· Z ≫ (braiding Y Z).hom = (braiding X Z).hom ≫ Z ◁ f := by aesop_cat - -- braiding_naturality : - -- βˆ€ {X X' Y Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y'), - -- (f βŠ— g) ≫ (braiding Y Y').hom = (braiding X X').hom ≫ (g βŠ— f) := by - -- aesop_cat -- hexagon identities: hexagon_forward : βˆ€ X Y Z : C, @@ -453,14 +449,6 @@ end CommMonoid section Tensor --- /-- The strength of the tensor product functor from `C Γ— C` to `C`. -/ --- def tensor_ΞΌ (X Y : C Γ— C) : (tensor C).obj X βŠ— (tensor C).obj Y ⟢ (tensor C).obj (X βŠ— Y) := --- (Ξ±_ X.1 X.2 (Y.1 βŠ— Y.2)).hom ≫ --- (X.1 ◁ (Ξ±_ X.2 Y.1 Y.2).inv) ≫ --- (X.1 ◁ (Ξ²_ X.2 Y.1).hom β–· Y.2) ≫ --- (X.1 ◁ (Ξ±_ Y.1 X.2 Y.2).hom) ≫ (Ξ±_ X.1 Y.1 (X.2 βŠ— Y.2)).inv --- #align category_theory.tensor_ΞΌ CategoryTheory.tensor_ΞΌ - /-- The strength of the tensor product functor from `C Γ— C` to `C`. -/ def tensor_ΞΌ (X Y : C Γ— C) : (X.1 βŠ— X.2) βŠ— Y.1 βŠ— Y.2 ⟢ (X.1 βŠ— Y.1) βŠ— X.2 βŠ— Y.2 := (Ξ±_ X.1 X.2 (Y.1 βŠ— Y.2)).hom ≫ diff --git a/Mathlib/CategoryTheory/Monoidal/Discrete.lean b/Mathlib/CategoryTheory/Monoidal/Discrete.lean index f303296b04ede..e94479260ef7f 100644 --- a/Mathlib/CategoryTheory/Monoidal/Discrete.lean +++ b/Mathlib/CategoryTheory/Monoidal/Discrete.lean @@ -30,7 +30,6 @@ instance Discrete.monoidal : MonoidalCategory (Discrete M) where tensorUnit' := Discrete.mk 1 tensorObj X Y := Discrete.mk (X.as * Y.as) - -- tensorHom f g := eqToHom (by dsimp; rw [eq_of_hom f, eq_of_hom g]) whiskerLeft X _ _ f := eqToHom (by dsimp; rw [eq_of_hom f]) whiskerRight f X := eqToHom (by dsimp; rw [eq_of_hom f]) leftUnitor X := Discrete.eqToIso (one_mul X.as) diff --git a/Mathlib/CategoryTheory/Monoidal/End.lean b/Mathlib/CategoryTheory/Monoidal/End.lean index 445f6ba7b89f8..ed7ccb85c5c1e 100644 --- a/Mathlib/CategoryTheory/Monoidal/End.lean +++ b/Mathlib/CategoryTheory/Monoidal/End.lean @@ -34,7 +34,7 @@ def endofunctorMonoidalCategory : MonoidalCategory (C β₯€ C) where tensorObj F G := F β‹™ G whiskerLeft X _ _ F := whiskerLeft X F whiskerRight F X := whiskerRight F X - -- tensorHom Ξ± Ξ² := Ξ± β—« Ξ² + tensorHom Ξ± Ξ² := Ξ± β—« Ξ² tensorUnit' := 𝟭 C associator F G H := Functor.associator F G H leftUnitor F := Functor.leftUnitor F diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean index 6df3588710fc9..da20ec5b5eeb0 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean @@ -71,10 +71,8 @@ inductive Hom : F C β†’ F C β†’ Type u | ρ_hom (X : F C) : Hom (X.tensor Unit) X | ρ_inv (X : F C) : Hom X (X.tensor Unit) | comp {X Y Z} (f : Hom X Y) (g : Hom Y Z) : Hom X Z - -- | tensor {W X Y Z} (f : Hom W Y) (g : Hom X Z) : Hom (W.tensor X) (Y.tensor Z) | whiskerLeft (X : F C) {Y Z : F C} (f : Hom Y Z) : Hom (X.tensor Y) (X.tensor Z) - -- `f` cannot be earlier than `Z` since it is a recursive argument. | whiskerRight {X Y : F C} (f : Hom X Y) (Z : F C) : Hom (X.tensor Z) (Y.tensor Z) #align category_theory.free_monoidal_category.hom CategoryTheory.FreeMonoidalCategory.Hom @@ -88,8 +86,6 @@ inductive HomEquiv : βˆ€ {X Y : F C}, (X ⟢ᡐ Y) β†’ (X ⟢ᡐ Y) β†’ Prop | trans {X Y} {f g h : X ⟢ᡐ Y} : HomEquiv f g β†’ HomEquiv g h β†’ HomEquiv f h | comp {X Y Z} {f f' : X ⟢ᡐ Y} {g g' : Y ⟢ᡐ Z} : HomEquiv f f' β†’ HomEquiv g g' β†’ HomEquiv (f.comp g) (f'.comp g') - -- | tensor {W X Y Z} {f f' : W ⟢ᡐ X} {g g' : Y ⟢ᡐ Z} : - -- HomEquiv f f' β†’ HomEquiv g g' β†’ HomEquiv (f.tensor g) (f'.tensor g') | whisker_left (X) {Y Z} (f f' : Y ⟢ᡐ Z) : HomEquiv f f' β†’ HomEquiv (f.whiskerLeft X) (f'.whiskerLeft X) | whisker_right {Y Z} (f f' : Y ⟢ᡐ Z) (X) : @@ -106,7 +102,6 @@ inductive HomEquiv : βˆ€ {X Y : F C}, (X ⟢ᡐ Y) β†’ (X ⟢ᡐ Y) β†’ Prop | tensor_whisker_left (X Y) {Z Z'} (f : Z ⟢ᡐ Z') : HomEquiv (f.whiskerLeft (X.tensor Y)) ((Hom.Ξ±_hom X Y Z).comp <| ((f.whiskerLeft Y).whiskerLeft X).comp <| Hom.Ξ±_inv X Y Z') - | id_whisker_right (X Y) : HomEquiv ((Hom.id X).whiskerRight Y) (Hom.id (X.tensor Y)) | comp_whisker_right {X Y Z} (W) (f : X ⟢ᡐ Y) (g : Y ⟢ᡐ Z) : HomEquiv ((f.comp g).whiskerRight W) ((f.whiskerRight W).comp <| g.whiskerRight W) @@ -120,24 +115,12 @@ inductive HomEquiv : βˆ€ {X Y : F C}, (X ⟢ᡐ Y) β†’ (X ⟢ᡐ Y) β†’ Prop ((Hom.Ξ±_hom X Y Z).comp <| ((f.whiskerRight Z).whiskerLeft X).comp <| Hom.Ξ±_inv X Y' Z) | whisker_exchange {W X Y Z} (f : W ⟢ᡐ X) (g : Y ⟢ᡐ Z) : HomEquiv ((g.whiskerLeft W).comp <| f.whiskerRight Z) ((f.whiskerRight Y).comp <| g.whiskerLeft X) - - -- | tensor_id {X Y} : HomEquiv ((Hom.id X).tensor (Hom.id Y)) (Hom.id _) - -- | tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : F C} (f₁ : X₁ ⟢ᡐ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ᡐ Yβ‚‚) (g₁ : Y₁ ⟢ᡐ Z₁) - -- (gβ‚‚ : Yβ‚‚ ⟢ᡐ Zβ‚‚) : - -- HomEquiv ((f₁.comp g₁).tensor (fβ‚‚.comp gβ‚‚)) ((f₁.tensor fβ‚‚).comp (g₁.tensor gβ‚‚)) | Ξ±_hom_inv {X Y Z} : HomEquiv ((Hom.Ξ±_hom X Y Z).comp (Hom.Ξ±_inv X Y Z)) (Hom.id _) | Ξ±_inv_hom {X Y Z} : HomEquiv ((Hom.Ξ±_inv X Y Z).comp (Hom.Ξ±_hom X Y Z)) (Hom.id _) - -- | associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃} (f₁ : X₁ ⟢ᡐ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ᡐ Yβ‚‚) (f₃ : X₃ ⟢ᡐ Y₃) : - -- HomEquiv (((f₁.tensor fβ‚‚).tensor f₃).comp (Hom.Ξ±_hom Y₁ Yβ‚‚ Y₃)) - -- ((Hom.Ξ±_hom X₁ Xβ‚‚ X₃).comp (f₁.tensor (fβ‚‚.tensor f₃))) | ρ_hom_inv {X} : HomEquiv ((Hom.ρ_hom X).comp (Hom.ρ_inv X)) (Hom.id _) | ρ_inv_hom {X} : HomEquiv ((Hom.ρ_inv X).comp (Hom.ρ_hom X)) (Hom.id _) - -- | ρ_naturality {X Y} (f : X ⟢ᡐ Y) : - -- HomEquiv ((f.tensor (Hom.id Unit)).comp (Hom.ρ_hom Y)) ((Hom.ρ_hom X).comp f) | l_hom_inv {X} : HomEquiv ((Hom.l_hom X).comp (Hom.l_inv X)) (Hom.id _) | l_inv_hom {X} : HomEquiv ((Hom.l_inv X).comp (Hom.l_hom X)) (Hom.id _) - -- | l_naturality {X Y} (f : X ⟢ᡐ Y) : - -- HomEquiv (((Hom.id Unit).tensor f).comp (Hom.l_hom Y)) ((Hom.l_hom X).comp f) | pentagon {W X Y Z} : HomEquiv (((Hom.Ξ±_hom W X Y).whiskerRight Z).comp @@ -186,18 +169,11 @@ instance categoryFreeMonoidalCategory : Category.{u} (F C) where instance : MonoidalCategory (F C) where tensorObj X Y := FreeMonoidalCategory.tensor X Y - -- tensorHom := @fun X₁ Y₁ Xβ‚‚ Yβ‚‚ => - -- Quotient.mapβ‚‚ Hom.tensor <| by - -- intro _ _ h _ _ h' - -- exact HomEquiv.tensor h h' whiskerLeft := fun X _ _ f => Quotient.map (Hom.whiskerLeft X) (HomEquiv.whisker_left X) f whiskerRight := fun f Y => Quotient.map (fun _f => Hom.whiskerRight _f Y) (fun _ _ h => HomEquiv.whisker_right _ _ _ h) f tensorUnit' := FreeMonoidalCategory.Unit associator X Y Z := ⟨⟦Hom.Ξ±_hom X Y Z⟧, ⟦Hom.Ξ±_inv X Y Z⟧, Quotient.sound Ξ±_hom_inv, Quotient.sound Ξ±_inv_hom⟩ - -- associator_naturality := @fun X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ => by - -- rintro ⟨fβ‚βŸ© ⟨fβ‚‚βŸ© ⟨fβ‚ƒβŸ© - -- exact Quotient.sound (associator_naturality _ _ _) leftUnitor X := ⟨⟦Hom.l_hom X⟧, ⟦Hom.l_inv X⟧, Quotient.sound l_hom_inv, Quotient.sound l_inv_hom⟩ rightUnitor X := ⟨⟦Hom.ρ_hom X⟧, ⟦Hom.ρ_inv X⟧, Quotient.sound ρ_hom_inv, Quotient.sound ρ_inv_hom⟩ @@ -247,11 +223,7 @@ theorem mk_comp {X Y Z : F C} (f : X ⟢ᡐ Y) (g : Y ⟢ᡐ Z) : rfl #align category_theory.free_monoidal_category.mk_comp CategoryTheory.FreeMonoidalCategory.mk_comp --- @[simp] --- theorem mk_tensor {X₁ Y₁ Xβ‚‚ Yβ‚‚ : F C} (f : X₁ ⟢ᡐ Y₁) (g : Xβ‚‚ ⟢ᡐ Yβ‚‚) : --- ⟦f.tensor g⟧ = @MonoidalCategory.tensorHom (F C) _ _ _ _ _ _ ⟦f⟧ ⟦g⟧ := --- rfl --- #align category_theory.free_monoidal_category.mk_tensor CategoryTheory.FreeMonoidalCategory.mk_tensor +#noalign category_theory.free_monoidal_category.mk_tensor @[simp] theorem mk_whiskerLeft (X : F C) {Y₁ Yβ‚‚ : F C} (f : Y₁ ⟢ᡐ Yβ‚‚) : @@ -338,7 +310,6 @@ def projectMapAux : βˆ€ {X Y : F C}, (X ⟢ᡐ Y) β†’ (projectObj f X ⟢ projec | _, _, Hom.comp f g => projectMapAux f ≫ projectMapAux g | _, _, Hom.whiskerLeft X p => projectObj f X ◁ projectMapAux p | _, _, Hom.whiskerRight p X => projectMapAux p β–· projectObj f X - -- | _, _, Hom.tensor f g => projectMapAux f βŠ— projectMapAux g #align category_theory.free_monoidal_category.project_map_aux CategoryTheory.FreeMonoidalCategory.projectMapAux -- Porting note: this declaration generates the same panic. @@ -351,27 +322,17 @@ def projectMap (X Y : F C) : (X ⟢ Y) β†’ (projectObj f X ⟢ projectObj f Y) : | symm _ _ _ hfg' => exact hfg'.symm | trans _ _ hfg hgh => exact hfg.trans hgh | comp _ _ hf hg => dsimp only [projectMapAux]; rw [hf, hg] - -- | tensor _ _ hfg hfg' => dsimp only [projectMapAux]; rw [hfg, hfg'] | whisker_left _ _ _ _ h => dsimp only [projectMapAux]; rw [h] | whisker_right _ _ _ _ h => dsimp only [projectMapAux]; rw [h] | comp_id => dsimp only [projectMapAux]; rw [Category.comp_id] | id_comp => dsimp only [projectMapAux]; rw [Category.id_comp] | assoc => dsimp only [projectMapAux]; rw [Category.assoc] - -- | tensor_id => dsimp only [projectMapAux]; rw [MonoidalCategory.tensor_id]; rfl - -- | tensor_comp => dsimp only [projectMapAux]; rw [MonoidalCategory.tensor_comp] | Ξ±_hom_inv => dsimp only [projectMapAux]; rw [Iso.hom_inv_id] | Ξ±_inv_hom => dsimp only [projectMapAux]; rw [Iso.inv_hom_id] - -- | associator_naturality => - -- dsimp only [projectMapAux]; rw [MonoidalCategory.associator_naturality] | ρ_hom_inv => dsimp only [projectMapAux]; rw [Iso.hom_inv_id] | ρ_inv_hom => dsimp only [projectMapAux]; rw [Iso.inv_hom_id] - -- | ρ_naturality => - -- dsimp only [projectMapAux, projectObj]; rw [MonoidalCategory.rightUnitor_naturality] | l_hom_inv => dsimp only [projectMapAux]; rw [Iso.hom_inv_id] | l_inv_hom => dsimp only [projectMapAux]; rw [Iso.inv_hom_id] - -- | l_naturality => - -- dsimp only [projectMapAux, projectObj] - -- exact MonoidalCategory.leftUnitor_naturality _ | pentagon => dsimp only [projectMapAux] exact MonoidalCategory.pentagon _ _ _ _ diff --git a/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean b/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean index 6de4a575549a6..3561bc83aad67 100644 --- a/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean @@ -95,7 +95,6 @@ instance functorCategoryMonoidal : MonoidalCategory (C β₯€ D) where rightUnitor F := NatIso.ofComponents fun X => ρ_ (F.obj X) associator F G H := NatIso.ofComponents fun X => Ξ±_ (F.obj X) (G.obj X) (H.obj X) whisker_exchange := by intros; ext; simp [whisker_exchange] - -- pentagon F G H K := by ext X; dsimp; rw [pentagon] #align category_theory.monoidal.functor_category_monoidal CategoryTheory.Monoidal.functorCategoryMonoidal @[simp] diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/FunctorCategory.lean b/Mathlib/CategoryTheory/Monoidal/Internal/FunctorCategory.lean index 55fbddc6b0bc9..16c5b2b7a351c 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/FunctorCategory.lean @@ -128,6 +128,8 @@ def unitIso : 𝟭 (Mon_ (C β₯€ D)) β‰… functor β‹™ inverse := set_option linter.uppercaseLean3 false in #align category_theory.monoidal.Mon_functor_category_equivalence.unit_iso CategoryTheory.Monoidal.MonFunctorCategoryEquivalence.unitIso +attribute [local simp] id_tensorHom tensorHom_id + /-- The counit for the equivalence `Mon_ (C β₯€ D) β‰Œ C β₯€ Mon_ D`. -/ @[simps!] @@ -216,6 +218,8 @@ def unitIso : 𝟭 (CommMon_ (C β₯€ D)) β‰… functor β‹™ inverse := set_option linter.uppercaseLean3 false in #align category_theory.monoidal.CommMon_functor_category_equivalence.unit_iso CategoryTheory.Monoidal.CommMonFunctorCategoryEquivalence.unitIso +attribute [local simp] id_tensorHom tensorHom_id + /-- The counit for the equivalence `CommMon_ (C β₯€ D) β‰Œ C β₯€ CommMon_ D`. -/ @[simps!] diff --git a/Mathlib/CategoryTheory/Monoidal/Mod_.lean b/Mathlib/CategoryTheory/Monoidal/Mod_.lean index 1e23e1bccecdb..8910030dea772 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mod_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mod_.lean @@ -24,8 +24,8 @@ variable {C} structure Mod_ (A : Mon_ C) where X : C act : A.X βŠ— X ⟢ X - one_act : (A.one βŠ— πŸ™ X) ≫ act = (Ξ»_ X).hom := by aesop_cat - assoc : (A.mul βŠ— πŸ™ X) ≫ act = (Ξ±_ A.X A.X X).hom ≫ (πŸ™ A.X βŠ— act) ≫ act := by aesop_cat + one_act : (A.one β–· X) ≫ act = (Ξ»_ X).hom := by aesop_cat + assoc : (A.mul β–· X) ≫ act = (Ξ±_ A.X A.X X).hom ≫ (A.X ◁ act) ≫ act := by aesop_cat set_option linter.uppercaseLean3 false in #align Mod_ Mod_ @@ -35,8 +35,10 @@ namespace Mod_ variable {A : Mon_ C} (M : Mod_ A) +attribute [local simp] id_tensorHom tensorHom_id + theorem assoc_flip : - (πŸ™ A.X βŠ— M.act) ≫ M.act = (Ξ±_ A.X A.X M.X).inv ≫ (A.mul βŠ— πŸ™ M.X) ≫ M.act := by simp + (A.X ◁ M.act) ≫ M.act = (Ξ±_ A.X A.X M.X).inv ≫ (A.mul β–· M.X) ≫ M.act := by simp set_option linter.uppercaseLean3 false in #align Mod_.assoc_flip Mod_.assoc_flip @@ -45,7 +47,7 @@ set_option linter.uppercaseLean3 false in @[ext] structure Hom (M N : Mod_ A) where hom : M.X ⟢ N.X - act_hom : M.act ≫ hom = (πŸ™ A.X βŠ— hom) ≫ N.act := by aesop_cat + act_hom : M.act ≫ hom = (A.X ◁ hom) ≫ N.act := by aesop_cat set_option linter.uppercaseLean3 false in #align Mod_.hom Mod_.Hom @@ -121,26 +123,25 @@ between the categories of module objects. def comap {A B : Mon_ C} (f : A ⟢ B) : Mod_ B β₯€ Mod_ A where obj M := { X := M.X - act := (f.hom βŠ— πŸ™ M.X) ≫ M.act + act := (f.hom β–· M.X) ≫ M.act one_act := by - slice_lhs 1 2 => rw [← comp_tensor_id] + slice_lhs 1 2 => rw [← comp_whiskerRight] rw [f.one_hom, one_act] assoc := by -- oh, for homotopy.io in a widget! - slice_rhs 2 3 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - rw [id_tensor_comp] + slice_rhs 2 3 => rw [whisker_exchange] + simp only [whiskerRight_tensor, MonoidalCategory.whiskerLeft_comp, Category.assoc, Iso.hom_inv_id_assoc] slice_rhs 4 5 => rw [Mod_.assoc_flip] - slice_rhs 3 4 => rw [associator_inv_naturality] - slice_rhs 2 3 => rw [← tensor_id, associator_inv_naturality] - slice_rhs 1 3 => rw [Iso.hom_inv_id_assoc] - slice_rhs 1 2 => rw [← comp_tensor_id, tensor_id_comp_id_tensor] - slice_rhs 1 2 => rw [← comp_tensor_id, ← f.mul_hom] - rw [comp_tensor_id, Category.assoc] } + slice_rhs 3 4 => rw [associator_inv_naturality_middle] + slice_rhs 2 4 => rw [Iso.hom_inv_id_assoc] + slice_rhs 1 2 => rw [← MonoidalCategory.comp_whiskerRight, ← whisker_exchange] + slice_rhs 1 2 => rw [← MonoidalCategory.comp_whiskerRight, ← tensorHom_def, ← f.mul_hom] + rw [comp_whiskerRight, Category.assoc] } map g := { hom := g.hom act_hom := by dsimp - slice_rhs 1 2 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] + slice_rhs 1 2 => rw [whisker_exchange] slice_rhs 2 3 => rw [← g.act_hom] rw [Category.assoc] } set_option linter.uppercaseLean3 false in diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index cac6a84e2ec5f..21c56ec613228 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -400,39 +400,6 @@ theorem one_rightUnitor {M : Mon_ C} : variable [BraidedCategory C] --- open Mathlib.Tactic.Coherence in --- example {C : Type u} [Category.{v} C] [MonoidalCategory C] (M N : C) [LiftObj M] [LiftObj N] : --- LiftObj ((tensor C).obj (M, N)) = LiftObj (M βŠ— N) := by --- rfl - --- #check _root_.id - --- open Mathlib.Tactic.Coherence in --- example {C : Type u} [Category.{v} C] [MonoidalCategory C] (M N : C) [LiftObj M] [LiftObj N] : --- FreeMonoidalCategory.projectObj _root_.id (LiftObj.lift ((tensor C).obj (M, N))) = ((tensor C).obj (M, N)) := by --- rfl - --- open Mathlib.Tactic.Coherence in --- instance {C : Type u} [Category.{v} C] [MonoidalCategory C] (M N : C) [LiftObj M] [LiftObj N] : --- MonoidalCoherence ((M βŠ— N)) ((tensor C).obj (M, N)) := by - --- infer_instance -- fails - --- /- --- To use `monoidalComp` in expressions containing `tensor_ΞΌ`, we need, for example, the following --- instance. However, it is not automatically inferred. - --- open Mathlib.Tactic.Coherence in --- instance {C : Type u} [Category.{v} C] [MonoidalCategory C] (M N : C) : --- LiftObj ((tensor C).obj (M, N)) := by --- infer_instance -- fails - --- open Mathlib.Tactic.Coherence in --- instance {C : Type u} [Category.{v} C] [MonoidalCategory C] (M N : C) : --- MonoidalCoherence ((M βŠ— N)) ((tensor C).obj (M, N)) := by --- infer_instance -- fails --- -/ - theorem Mon_tensor_one_mul (M N : Mon_ C) : (((Ξ»_ (πŸ™_ C)).inv ≫ (M.one βŠ— N.one)) β–· (M.X βŠ— N.X)) ≫ tensor_ΞΌ C (M.X, N.X) (M.X, N.X) ≫ (M.mul βŠ— N.mul) = diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index fb8110ac81be3..c6e39ebe267c0 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -40,15 +40,12 @@ class MonoidalPreadditive : Prop where add_whiskerRight : βˆ€ {X Y Z : C} (f g : Y ⟢ Z), (f + g) β–· X = f β–· X + g β–· X := by aesop_cat #align category_theory.monoidal_preadditive CategoryTheory.MonoidalPreadditive --- attribute [simp] MonoidalPreadditive.tensor_zero MonoidalPreadditive.zero_tensor attribute [simp] MonoidalPreadditive.whiskerLeft_zero MonoidalPreadditive.zero_whiskerRight attribute [simp] MonoidalPreadditive.whiskerLeft_add MonoidalPreadditive.add_whiskerRight variable {C} variable [MonoidalPreadditive C] --- attribute [local simp] MonoidalPreadditive.tensor_add MonoidalPreadditive.add_tensor - instance tensorLeft_additive (X : C) : (tensorLeft X).Additive where #align category_theory.tensor_left_additive CategoryTheory.tensorLeft_additive @@ -158,24 +155,26 @@ theorem leftDistributor_inv {J : Type} [Fintype J] (X : C) (f : J β†’ C) : @[reassoc (attr := simp)] theorem leftDistributor_hom_comp_biproduct_Ο€ {J : Type} [Fintype J] (X : C) (f : J β†’ C) (j : J) : - (leftDistributor X f).hom ≫ biproduct.Ο€ _ j = πŸ™ X βŠ— biproduct.Ο€ _ j := by + (leftDistributor X f).hom ≫ biproduct.Ο€ _ j = X ◁ biproduct.Ο€ _ j := by simp [leftDistributor_hom, Preadditive.sum_comp, biproduct.ΞΉ_Ο€, comp_dite] @[reassoc (attr := simp)] theorem biproduct_ΞΉ_comp_leftDistributor_hom {J : Type} [Fintype J] (X : C) (f : J β†’ C) (j : J) : - (πŸ™ X βŠ— biproduct.ΞΉ _ j) ≫ (leftDistributor X f).hom = biproduct.ΞΉ (fun j => X βŠ— f j) j := by + (X ◁ biproduct.ΞΉ _ j) ≫ (leftDistributor X f).hom = biproduct.ΞΉ (fun j => X βŠ— f j) j := by simp [leftDistributor_hom, Preadditive.comp_sum, ← id_tensor_comp_assoc, biproduct.ΞΉ_Ο€, tensor_dite, dite_comp] + sorry @[reassoc (attr := simp)] theorem leftDistributor_inv_comp_biproduct_Ο€ {J : Type} [Fintype J] (X : C) (f : J β†’ C) (j : J) : - (leftDistributor X f).inv ≫ (πŸ™ X βŠ— biproduct.Ο€ _ j) = biproduct.Ο€ _ j := by + (leftDistributor X f).inv ≫ (X ◁ biproduct.Ο€ _ j) = biproduct.Ο€ _ j := by simp [leftDistributor_inv, Preadditive.sum_comp, ← id_tensor_comp, biproduct.ΞΉ_Ο€, tensor_dite, comp_dite] + sorry @[reassoc (attr := simp)] theorem biproduct_ΞΉ_comp_leftDistributor_inv {J : Type} [Fintype J] (X : C) (f : J β†’ C) (j : J) : - biproduct.ΞΉ _ j ≫ (leftDistributor X f).inv = πŸ™ X βŠ— biproduct.ΞΉ _ j := by + biproduct.ΞΉ _ j ≫ (leftDistributor X f).inv = X ◁ biproduct.ΞΉ _ j := by simp [leftDistributor_inv, Preadditive.comp_sum, ← id_tensor_comp, biproduct.ΞΉ_Ο€_assoc, dite_comp] theorem leftDistributor_assoc {J : Type} [Fintype J] (X Y : C) (f : J β†’ C) : @@ -219,24 +218,26 @@ theorem rightDistributor_inv {J : Type} [Fintype J] (f : J β†’ C) (X : C) : @[reassoc (attr := simp)] theorem rightDistributor_hom_comp_biproduct_Ο€ {J : Type} [Fintype J] (f : J β†’ C) (X : C) (j : J) : - (rightDistributor f X).hom ≫ biproduct.Ο€ _ j = biproduct.Ο€ _ j βŠ— πŸ™ X := by + (rightDistributor f X).hom ≫ biproduct.Ο€ _ j = biproduct.Ο€ _ j β–· X := by simp [rightDistributor_hom, Preadditive.sum_comp, biproduct.ΞΉ_Ο€, comp_dite] @[reassoc (attr := simp)] theorem biproduct_ΞΉ_comp_rightDistributor_hom {J : Type} [Fintype J] (f : J β†’ C) (X : C) (j : J) : - (biproduct.ΞΉ _ j βŠ— πŸ™ X) ≫ (rightDistributor f X).hom = biproduct.ΞΉ (fun j => f j βŠ— X) j := by + (biproduct.ΞΉ _ j β–· X) ≫ (rightDistributor f X).hom = biproduct.ΞΉ (fun j => f j βŠ— X) j := by simp [rightDistributor_hom, Preadditive.comp_sum, ← comp_tensor_id_assoc, biproduct.ΞΉ_Ο€, dite_tensor, dite_comp] + sorry @[reassoc (attr := simp)] theorem rightDistributor_inv_comp_biproduct_Ο€ {J : Type} [Fintype J] (f : J β†’ C) (X : C) (j : J) : - (rightDistributor f X).inv ≫ (biproduct.Ο€ _ j βŠ— πŸ™ X) = biproduct.Ο€ _ j := by + (rightDistributor f X).inv ≫ (biproduct.Ο€ _ j β–· X) = biproduct.Ο€ _ j := by simp [rightDistributor_inv, Preadditive.sum_comp, ← comp_tensor_id, biproduct.ΞΉ_Ο€, dite_tensor, comp_dite] + sorry @[reassoc (attr := simp)] theorem biproduct_ΞΉ_comp_rightDistributor_inv {J : Type} [Fintype J] (f : J β†’ C) (X : C) (j : J) : - biproduct.ΞΉ _ j ≫ (rightDistributor f X).inv = biproduct.ΞΉ _ j βŠ— πŸ™ X := by + biproduct.ΞΉ _ j ≫ (rightDistributor f X).inv = biproduct.ΞΉ _ j β–· X := by simp [rightDistributor_inv, Preadditive.comp_sum, ← id_tensor_comp, biproduct.ΞΉ_Ο€_assoc, dite_comp] @@ -274,7 +275,7 @@ theorem leftDistributor_rightDistributor_assoc {J : Type _} [Fintype J] @[ext] theorem leftDistributor_ext_left {J : Type} [Fintype J] {X Y : C} {f : J β†’ C} {g h : X βŠ— ⨁ f ⟢ Y} - (w : βˆ€ j, (πŸ™ X βŠ— biproduct.ΞΉ f j) ≫ g = (πŸ™ X βŠ— biproduct.ΞΉ f j) ≫ h) : g = h := by + (w : βˆ€ j, (X ◁ biproduct.ΞΉ f j) ≫ g = (X ◁ biproduct.ΞΉ f j) ≫ h) : g = h := by apply (cancel_epi (leftDistributor X f).inv).mp ext simp? [leftDistributor_inv, Preadditive.comp_sum_assoc, biproduct.ΞΉ_Ο€_assoc, dite_comp] says @@ -285,7 +286,7 @@ theorem leftDistributor_ext_left {J : Type} [Fintype J] {X Y : C} {f : J β†’ C} @[ext] theorem leftDistributor_ext_right {J : Type} [Fintype J] {X Y : C} {f : J β†’ C} {g h : X ⟢ Y βŠ— ⨁ f} - (w : βˆ€ j, g ≫ (πŸ™ Y βŠ— biproduct.Ο€ f j) = h ≫ (πŸ™ Y βŠ— biproduct.Ο€ f j)) : g = h := by + (w : βˆ€ j, g ≫ (Y ◁ biproduct.Ο€ f j) = h ≫ (Y ◁ biproduct.Ο€ f j)) : g = h := by apply (cancel_mono (leftDistributor Y f).hom).mp ext simp? [leftDistributor_hom, Preadditive.sum_comp, Preadditive.comp_sum_assoc, biproduct.ΞΉ_Ο€, @@ -300,25 +301,27 @@ theorem leftDistributor_ext_right {J : Type} [Fintype J] {X Y : C} {f : J β†’ C} @[ext] theorem leftDistributor_extβ‚‚_left {J : Type} [Fintype J] {X Y Z : C} {f : J β†’ C} {g h : X βŠ— (Y βŠ— ⨁ f) ⟢ Z} - (w : βˆ€ j, (πŸ™ X βŠ— (πŸ™ Y βŠ— biproduct.ΞΉ f j)) ≫ g = (πŸ™ X βŠ— (πŸ™ Y βŠ— biproduct.ΞΉ f j)) ≫ h) : + (w : βˆ€ j, (X ◁ (Y ◁ biproduct.ΞΉ f j)) ≫ g = (X ◁ (Y ◁ biproduct.ΞΉ f j)) ≫ h) : g = h := by apply (cancel_epi (Ξ±_ _ _ _).hom).mp ext simp_rw [← tensor_id, associator_naturality_assoc, w] + sorry @[ext] theorem leftDistributor_extβ‚‚_right {J : Type} [Fintype J] {X Y Z : C} {f : J β†’ C} {g h : X ⟢ Y βŠ— (Z βŠ— ⨁ f)} - (w : βˆ€ j, g ≫ (πŸ™ Y βŠ— (πŸ™ Z βŠ— biproduct.Ο€ f j)) = h ≫ (πŸ™ Y βŠ— (πŸ™ Z βŠ— biproduct.Ο€ f j))) : + (w : βˆ€ j, g ≫ (Y ◁ (Z ◁ biproduct.Ο€ f j)) = h ≫ (Y ◁ (Z ◁ biproduct.Ο€ f j))) : g = h := by apply (cancel_mono (Ξ±_ _ _ _).inv).mp ext simp_rw [← tensor_id, Category.assoc, ← associator_inv_naturality, ← Category.assoc, w] + sorry @[ext] theorem rightDistributor_ext_left {J : Type} [Fintype J] {f : J β†’ C} {X Y : C} {g h : (⨁ f) βŠ— X ⟢ Y} - (w : βˆ€ j, (biproduct.ΞΉ f j βŠ— πŸ™ X) ≫ g = (biproduct.ΞΉ f j βŠ— πŸ™ X) ≫ h) : g = h := by + (w : βˆ€ j, (biproduct.ΞΉ f j β–· X) ≫ g = (biproduct.ΞΉ f j β–· X) ≫ h) : g = h := by apply (cancel_epi (rightDistributor f X).inv).mp ext simp? [rightDistributor_inv, Preadditive.comp_sum_assoc, biproduct.ΞΉ_Ο€_assoc, dite_comp] says @@ -330,7 +333,7 @@ theorem rightDistributor_ext_left {J : Type} [Fintype J] @[ext] theorem rightDistributor_ext_right {J : Type} [Fintype J] {f : J β†’ C} {X Y : C} {g h : X ⟢ (⨁ f) βŠ— Y} - (w : βˆ€ j, g ≫ (biproduct.Ο€ f j βŠ— πŸ™ Y) = h ≫ (biproduct.Ο€ f j βŠ— πŸ™ Y)) : g = h := by + (w : βˆ€ j, g ≫ (biproduct.Ο€ f j β–· Y) = h ≫ (biproduct.Ο€ f j β–· Y)) : g = h := by apply (cancel_mono (rightDistributor f Y).hom).mp ext simp? [rightDistributor_hom, Preadditive.sum_comp, Preadditive.comp_sum_assoc, biproduct.ΞΉ_Ο€, @@ -343,19 +346,21 @@ theorem rightDistributor_ext_right {J : Type} [Fintype J] @[ext] theorem rightDistributor_extβ‚‚_left {J : Type} [Fintype J] {f : J β†’ C} {X Y Z : C} {g h : ((⨁ f) βŠ— X) βŠ— Y ⟢ Z} - (w : βˆ€ j, ((biproduct.ΞΉ f j βŠ— πŸ™ X) βŠ— πŸ™ Y) ≫ g = ((biproduct.ΞΉ f j βŠ— πŸ™ X) βŠ— πŸ™ Y) ≫ h) : + (w : βˆ€ j, ((biproduct.ΞΉ f j β–· X) β–· Y) ≫ g = ((biproduct.ΞΉ f j β–· X) β–· Y) ≫ h) : g = h := by apply (cancel_epi (Ξ±_ _ _ _).inv).mp ext simp_rw [← tensor_id, associator_inv_naturality_assoc, w] + sorry @[ext] theorem rightDistributor_extβ‚‚_right {J : Type} [Fintype J] {f : J β†’ C} {X Y Z : C} {g h : X ⟢ ((⨁ f) βŠ— Y) βŠ— Z} - (w : βˆ€ j, g ≫ ((biproduct.Ο€ f j βŠ— πŸ™ Y) βŠ— πŸ™ Z) = h ≫ ((biproduct.Ο€ f j βŠ— πŸ™ Y) βŠ— πŸ™ Z)) : + (w : βˆ€ j, g ≫ ((biproduct.Ο€ f j β–· Y) β–· Z) = h ≫ ((biproduct.Ο€ f j β–· Y) β–· Z)) : g = h := by apply (cancel_mono (Ξ±_ _ _ _).hom).mp ext simp_rw [← tensor_id, Category.assoc, ← associator_naturality, ← Category.assoc, w] + sorry end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean index 20d9b10462933..66239d565da3a 100644 --- a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean @@ -75,30 +75,6 @@ instance fullMonoidalSubcategory : MonoidalCategory (FullSubcategory P) where triangle X Y := triangle X.1 Y.1 #align category_theory.monoidal_category.full_monoidal_subcategory CategoryTheory.MonoidalCategory.fullMonoidalSubcategory --- /-- --- When `P` is a monoidal predicate, the full subcategory for `P` inherits the monoidal structure of --- `C`. --- -/ --- instance fullMonoidalSubcategory''' : MonoidalCategory (FullSubcategory P) where --- tensorObj X Y := ⟨X.1 βŠ— Y.1, prop_tensor X.2 Y.2⟩ --- whiskerLeft := @fun X Y₁ Yβ‚‚ f => by --- change X₁.1 βŠ— Xβ‚‚.1 ⟢ Y₁.1 βŠ— Yβ‚‚.1 --- change X₁.1 ⟢ Y₁.1 at f; change Xβ‚‚.1 ⟢ Yβ‚‚.1 at g; exact f βŠ— g --- tensorUnit' := βŸ¨πŸ™_ C, prop_id⟩ --- associator X Y Z := --- ⟨(Ξ±_ X.1 Y.1 Z.1).hom, (Ξ±_ X.1 Y.1 Z.1).inv, hom_inv_id (Ξ±_ X.1 Y.1 Z.1), --- inv_hom_id (Ξ±_ X.1 Y.1 Z.1)⟩ --- leftUnitor X := ⟨(Ξ»_ X.1).hom, (Ξ»_ X.1).inv, hom_inv_id (Ξ»_ X.1), inv_hom_id (Ξ»_ X.1)⟩ --- rightUnitor X := ⟨(ρ_ X.1).hom, (ρ_ X.1).inv, hom_inv_id (ρ_ X.1), inv_hom_id (ρ_ X.1)⟩ --- tensor_id X Y := tensor_id X.1 Y.1 --- tensor_comp f₁ fβ‚‚ g₁ gβ‚‚ := @tensor_comp C _ _ _ _ _ _ _ _ f₁ fβ‚‚ g₁ gβ‚‚ --- associator_naturality f₁ fβ‚‚ f₃ := @associator_naturality C _ _ _ _ _ _ _ _ f₁ fβ‚‚ f₃ --- leftUnitor_naturality f := @leftUnitor_naturality C _ _ _ _ f --- rightUnitor_naturality f := @rightUnitor_naturality C _ _ _ _ f --- pentagon W X Y Z := pentagon W.1 X.1 Y.1 Z.1 --- triangle X Y := triangle X.1 Y.1 --- #align category_theory.monoidal_category.full_monoidal_subcategory CategoryTheory.MonoidalCategory.fullMonoidalSubcategory - /-- The forgetful monoidal functor from a full monoidal subcategory into the original category ("forgetting" the condition). -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index e53206083ffed..f702419f8d9fc 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -35,82 +35,6 @@ variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] variable {D : Type uβ‚‚} [Category.{vβ‚‚} D] --- -- porting note: it was @[simps {attrs := [`_refl_lemma]}] --- /-- Transport a monoidal structure along an equivalence of (plain) categories. --- -/ --- @[simps] --- def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D where --- tensorObj X Y := e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y) --- -- tensorHom f g := e.functor.map (e.inverse.map f βŠ— e.inverse.map g) --- whiskerLeft X _ _ f := e.functor.map (e.inverse.obj X ◁ e.inverse.map f) --- whiskerRight f Y := e.functor.map (e.inverse.map f β–· e.inverse.obj Y) --- -- tensorHom_def := _ --- tensorUnit' := e.functor.obj (πŸ™_ C) --- associator X Y Z := --- e.functor.mapIso --- (((e.unitIso.app _).symm βŠ— Iso.refl _) β‰ͺ≫ --- Ξ±_ (e.inverse.obj X) (e.inverse.obj Y) (e.inverse.obj Z) β‰ͺ≫ (Iso.refl _ βŠ— e.unitIso.app _)) --- leftUnitor X := --- e.functor.mapIso (((e.unitIso.app _).symm βŠ— Iso.refl _) β‰ͺ≫ Ξ»_ (e.inverse.obj X)) β‰ͺ≫ --- e.counitIso.app _ --- rightUnitor X := --- e.functor.mapIso ((Iso.refl _ βŠ— (e.unitIso.app _).symm) β‰ͺ≫ ρ_ (e.inverse.obj X)) β‰ͺ≫ --- e.counitIso.app _ --- whiskerLeft_id := _ --- whiskerLeft_comp := _ --- id_whiskerLeft := _ --- tensor_whiskerLeft := _ --- id_whiskerRight := _ --- comp_whiskerRight := _ --- whiskerRight_id := _ --- whiskerRight_tensor := _ --- whisker_assoc := _ --- whisker_exchange := _ --- pentagon W X Y Z := by --- dsimp --- simp_rw [← id_tensorHom, ← tensorHom_id] --- simp only [Iso.hom_inv_id_app_assoc, comp_tensor_id, assoc, Equivalence.inv_fun_map, --- Functor.map_comp, id_tensor_comp, e.inverse.map_id] --- simp only [← e.functor.map_comp] --- congr 2 --- slice_lhs 4 5 => --- rw [← comp_tensor_id, Iso.hom_inv_id_app] --- dsimp --- rw [tensor_id] --- simp only [Category.id_comp, Category.assoc] --- slice_lhs 5 6 => --- rw [← id_tensor_comp, Iso.hom_inv_id_app] --- dsimp --- rw [tensor_id] --- simp only [Category.id_comp, Category.assoc] --- slice_rhs 2 3 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] --- slice_rhs 1 2 => rw [← tensor_id, ← associator_naturality] --- slice_rhs 3 4 => rw [← tensor_id, associator_naturality] --- slice_rhs 2 3 => rw [← pentagon] --- simp only [Category.assoc] --- congr 2 --- simp_rw [← id_tensorHom, ← tensorHom_id] --- slice_lhs 1 2 => rw [associator_naturality] --- simp only [Category.assoc] --- congr 1 --- slice_lhs 1 2 => --- rw [← id_tensor_comp, ← comp_tensor_id, Iso.hom_inv_id_app] --- dsimp --- rw [tensor_id, tensor_id] --- simp only [Category.id_comp, Category.assoc] --- simp [id_tensorHom, tensorHom_id] --- triangle X Y := by --- dsimp --- simp_rw [← id_tensorHom, ← tensorHom_id] --- simp only [Iso.hom_inv_id_app_assoc, comp_tensor_id, Equivalence.unit_inverse_comp, assoc, --- Equivalence.inv_fun_map, comp_id, Functor.map_comp, id_tensor_comp, e.inverse.map_id] --- simp only [← e.functor.map_comp] --- congr 2 --- slice_lhs 2 3 => --- rw [← id_tensor_comp] --- simp --- simp [id_tensorHom, tensorHom_id] - attribute [local simp] id_tensorHom tensorHom_id -- porting note: it was @[simps {attrs := [`_refl_lemma]}] @@ -236,7 +160,6 @@ def Transported (_ : C β‰Œ D) := D instance (e : C β‰Œ D) : Category (Transported e) := (inferInstance : Category D) --- @[simps!] instance (e : C β‰Œ D) : MonoidalCategory (Transported e) := transport e diff --git a/Mathlib/RepresentationTheory/Character.lean b/Mathlib/RepresentationTheory/Character.lean index 3b29acb87cc96..22bff02555fa4 100644 --- a/Mathlib/RepresentationTheory/Character.lean +++ b/Mathlib/RepresentationTheory/Character.lean @@ -55,7 +55,9 @@ theorem char_one (V : FdRep k G) : V.character 1 = FiniteDimensional.finrank k V /-- The character is multiplicative under the tensor product. -/ theorem char_tensor (V W : FdRep k G) : (V βŠ— W).character = V.character * W.character := by - ext g; convert trace_tensorProduct' (V.ρ g) (W.ρ g) + ext g + convert trace_tensorProduct' (V.ρ g) (W.ρ g) + sorry #align fdRep.char_tensor FdRep.char_tensor -- Porting note: adding variant of `char_tensor` to make the simp-set confluent diff --git a/Mathlib/RepresentationTheory/FdRep.lean b/Mathlib/RepresentationTheory/FdRep.lean index 8f6ef4fe94932..06923f5abd014 100644 --- a/Mathlib/RepresentationTheory/FdRep.lean +++ b/Mathlib/RepresentationTheory/FdRep.lean @@ -181,8 +181,9 @@ noncomputable def dualTensorIsoLinHomAux : /-- When `V` and `W` are finite dimensional representations of a group `G`, the isomorphism `dualTensorHomEquiv k V W` of vector spaces induces an isomorphism of representations. -/ noncomputable def dualTensorIsoLinHom : FdRep.of ρV.dual βŠ— W β‰… FdRep.of (linHom ρV W.ρ) := by - refine Action.mkIso (dualTensorIsoLinHomAux ρV W) ?_ - convert dualTensorHom_comm ρV W.ρ + apply Action.mkIso (dualTensorIsoLinHomAux ρV W) (fun g => ?_) + convert dualTensorHom_comm ρV W.ρ g + sorry #align fdRep.dual_tensor_iso_lin_hom FdRep.dualTensorIsoLinHom @[simp] diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index c2c2b13866a5d..37e8f7bd50dd5 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -85,10 +85,6 @@ instance LiftHom_comp {X Y Z : C} [LiftObj X] [LiftObj Y] [LiftObj Z] (f : X ⟢ [LiftHom f] [LiftHom g] : LiftHom (f ≫ g) where lift := LiftHom.lift f ≫ LiftHom.lift g --- instance LiftHom_tensor {W X Y Z : C} [LiftObj W] [LiftObj X] [LiftObj Y] [LiftObj Z] --- (f : W ⟢ X) (g : Y ⟢ Z) [LiftHom f] [LiftHom g] : LiftHom (f βŠ— g) where --- lift := LiftHom.lift f βŠ— LiftHom.lift g - instance liftHom_WhiskerLeft (X : C) [LiftObj X] {Y Z : C} [LiftObj Y] [LiftObj Z] (f : Y ⟢ Z) [LiftHom f] : LiftHom (X ◁ f) where lift := LiftObj.lift X ◁ LiftHom.lift f @@ -97,6 +93,10 @@ instance liftHom_WhiskerRight {X Y : C} (f : X ⟢ Y) [LiftObj X] [LiftObj Y] [L {Z : C} [LiftObj Z] : LiftHom (f β–· Z) where lift := LiftHom.lift f β–· LiftObj.lift Z +instance LiftHom_tensor {W X Y Z : C} [LiftObj W] [LiftObj X] [LiftObj Y] [LiftObj Z] + (f : W ⟢ X) (g : Y ⟢ Z) [LiftHom f] [LiftHom g] : LiftHom (f βŠ— g) where + lift := LiftHom.lift f βŠ— LiftHom.lift g + /-- A typeclass carrying a choice of monoidal structural isomorphism between two objects. Used by the `βŠ—β‰«` monoidal composition operator, and the `coherence` tactic. From c1978269da4d93ac5da9e998859ea2c2a083e884 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 3 Aug 2023 22:06:00 +0900 Subject: [PATCH 06/62] set `id_tensorHom` and `tensorHom_id` as simp lemmas --- Mathlib/CategoryTheory/Enriched/Basic.lean | 2 - Mathlib/CategoryTheory/Monoidal/Bimod.lean | 2 - Mathlib/CategoryTheory/Monoidal/Braided.lean | 12 ++---- Mathlib/CategoryTheory/Monoidal/Category.lean | 2 + Mathlib/CategoryTheory/Monoidal/Center.lean | 2 - .../Monoidal/CoherenceLemmas.lean | 2 - Mathlib/CategoryTheory/Monoidal/CommMon_.lean | 2 - .../Monoidal/Free/Coherence.lean | 2 - .../Monoidal/FunctorCategory.lean | 4 +- .../Monoidal/Internal/FunctorCategory.lean | 4 -- Mathlib/CategoryTheory/Monoidal/Limits.lean | 39 +++++++++++-------- Mathlib/CategoryTheory/Monoidal/Mod_.lean | 2 - Mathlib/CategoryTheory/Monoidal/Mon_.lean | 6 --- .../CategoryTheory/Monoidal/Rigid/Basic.lean | 2 - .../CategoryTheory/Monoidal/Transport.lean | 2 - Mathlib/RepresentationTheory/Action.lean | 2 - Mathlib/RepresentationTheory/FdRep.lean | 2 +- 17 files changed, 30 insertions(+), 59 deletions(-) diff --git a/Mathlib/CategoryTheory/Enriched/Basic.lean b/Mathlib/CategoryTheory/Enriched/Basic.lean index 9dd6a928a0456..1a0aea108c016 100644 --- a/Mathlib/CategoryTheory/Enriched/Basic.lean +++ b/Mathlib/CategoryTheory/Enriched/Basic.lean @@ -149,8 +149,6 @@ def categoryOfEnrichedCategoryType (C : Type u₁) [π’ž : EnrichedCategory (Typ assoc f g h := (congr_fun (e_assoc (Type v) _ _ _ _) ⟨f, g, h⟩ : _) #align category_theory.category_of_enriched_category_Type CategoryTheory.categoryOfEnrichedCategoryType -attribute [local simp] id_tensorHom tensorHom_id - /-- Construct a `Type v`-enriched category from an honest category. -/ def enrichedCategoryTypeOfCategory (C : Type u₁) [π’ž : Category.{v} C] : EnrichedCategory (Type v) C diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index 9abdf782f05d7..b1f9fa6debee0 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -22,8 +22,6 @@ open CategoryTheory.MonoidalCategory variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] -attribute [local simp] id_tensorHom tensorHom_id - section open CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 93862ffea2c46..4c5ed507991ea 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -603,15 +603,11 @@ theorem associator_monoidal (X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C) : ((Ξ±_ X₁ Xβ‚‚ X₃).hom βŠ— (Ξ±_ Y₁ Yβ‚‚ Y₃).hom) ≫ tensor_ΞΌ C (X₁, Xβ‚‚ βŠ— X₃) (Y₁, Yβ‚‚ βŠ— Y₃) ≫ ((X₁ βŠ— Y₁) ◁ tensor_ΞΌ C (Xβ‚‚, X₃) (Yβ‚‚, Y₃)) := by dsimp [tensor_ΞΌ] - simp only [tensor_whiskerLeft, braiding_tensor_right, comp_whiskerRight, whisker_assoc, assoc, - MonoidalCategory.whiskerLeft_comp, Iso.inv_hom_id_assoc, whiskerRight_tensor, braiding_tensor_left] calc - _ = πŸ™ _ βŠ—β‰« - X₁ ◁ Xβ‚‚ ◁ (Ξ²_ X₃ Y₁).hom β–· Yβ‚‚ β–· Y₃ βŠ—β‰« - X₁ ◁ ((Xβ‚‚ βŠ— Y₁) ◁ (Ξ²_ X₃ Yβ‚‚).hom ≫ (Ξ²_ Xβ‚‚ Y₁).hom β–· (Yβ‚‚ βŠ— X₃)) β–· Y₃ βŠ—β‰« πŸ™ _ := ?eq1 - _ = _ := ?eq2 - case eq1 => coherence - case eq2 => rw [whisker_exchange]; coherence + _ = πŸ™ _ βŠ—β‰« X₁ ◁ Xβ‚‚ ◁ (Ξ²_ X₃ Y₁).hom β–· Yβ‚‚ β–· Y₃ βŠ—β‰« + X₁ ◁ ((Xβ‚‚ βŠ— Y₁) ◁ (Ξ²_ X₃ Yβ‚‚).hom ≫ + (Ξ²_ Xβ‚‚ Y₁).hom β–· (Yβ‚‚ βŠ— X₃)) β–· Y₃ βŠ—β‰« πŸ™ _ := by simp; coherence + _ = _ := by rw [whisker_exchange]; simp; coherence #align category_theory.associator_monoidal CategoryTheory.associator_monoidal end Tensor diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 59ffdc83e42a3..393b971dea483 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -225,10 +225,12 @@ theorem tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) ( (f₁ ≫ g₁) βŠ— (fβ‚‚ ≫ gβ‚‚) = (f₁ βŠ— fβ‚‚) ≫ (g₁ βŠ— gβ‚‚) := by simp [whisker_exchange_assoc] +@[simp] theorem id_tensorHom (X : C) {Y₁ Yβ‚‚ : C} (f : Y₁ ⟢ Yβ‚‚) : (πŸ™ X) βŠ— f = X ◁ f := by simp +@[simp] theorem tensorHom_id {X₁ Xβ‚‚ : C} (f : X₁ ⟢ Xβ‚‚) (Y : C) : f βŠ— (πŸ™ Y) = f β–· Y := by simp diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index 84eb1981adaba..a4a407aa88245 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -127,8 +127,6 @@ instance isIso_of_f_isIso {X Y : Center C} (f : X ⟢ Y) [IsIso f.f] : IsIso f : infer_instance #align category_theory.center.is_iso_of_f_is_iso CategoryTheory.Center.isIso_of_f_isIso -attribute [local simp] id_tensorHom tensorHom_id - /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ @[simps] def tensorObj (X Y : Center C) : Center C := diff --git a/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean b/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean index 6092a2dddbcba..206fd3b9460c9 100644 --- a/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean +++ b/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean @@ -24,8 +24,6 @@ namespace CategoryTheory.MonoidalCategory variable {C : Type _} [Category C] [MonoidalCategory C] -attribute [local simp] id_tensorHom tensorHom_id - -- See Proposition 2.2.4 of @[reassoc] theorem leftUnitor_tensor'' (X Y : C) : diff --git a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean index c21fb4fee1028..1850ab3cfa042 100644 --- a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean @@ -189,8 +189,6 @@ def commMonToLaxBraided : CommMon_ C β₯€ LaxBraidedFunctor (Discrete PUnit.{u + set_option linter.uppercaseLean3 false in #align CommMon_.equiv_lax_braided_functor_punit.CommMon_to_lax_braided CommMon_.EquivLaxBraidedFunctorPunit.commMonToLaxBraided -attribute [local simp] id_tensorHom tensorHom_id - /-- Implementation of `CommMon_.equivLaxBraidedFunctorPunit`. -/ @[simps!] def unitIso : diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean index bae8738e28fb3..a99cc00799983 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean @@ -300,8 +300,6 @@ theorem normalizeObj_congr (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ Y apply congr_argβ‚‚ _ rfl (congr_fun ih _) | _ => funext; rfl -attribute [local simp] id_tensorHom tensorHom_id - theorem normalize_naturality (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ Y) : inclusionObj n ◁ f ≫ (normalizeIsoApp' C Y n).hom = (normalizeIsoApp' C X n).hom ≫ diff --git a/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean b/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean index 3561bc83aad67..a6becb1d62f86 100644 --- a/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean @@ -33,8 +33,6 @@ variable {D : Type uβ‚‚} [Category.{vβ‚‚} D] [MonoidalCategory.{vβ‚‚} D] namespace FunctorCategory -attribute [local simp] id_tensorHom tensorHom_id - variable (F G F' G' : C β₯€ D) /-- (An auxiliary definition for `functorCategoryMonoidal`.) @@ -79,7 +77,7 @@ end FunctorCategory open CategoryTheory.Monoidal.FunctorCategory -attribute [local simp] id_tensorHom tensorHom_id tensorHom_def +attribute [local simp] tensorHom_def /-- When `C` is any category, and `D` is a monoidal category, the functor category `C β₯€ D` has a natural pointwise monoidal structure, diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/FunctorCategory.lean b/Mathlib/CategoryTheory/Monoidal/Internal/FunctorCategory.lean index 16c5b2b7a351c..55fbddc6b0bc9 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/FunctorCategory.lean @@ -128,8 +128,6 @@ def unitIso : 𝟭 (Mon_ (C β₯€ D)) β‰… functor β‹™ inverse := set_option linter.uppercaseLean3 false in #align category_theory.monoidal.Mon_functor_category_equivalence.unit_iso CategoryTheory.Monoidal.MonFunctorCategoryEquivalence.unitIso -attribute [local simp] id_tensorHom tensorHom_id - /-- The counit for the equivalence `Mon_ (C β₯€ D) β‰Œ C β₯€ Mon_ D`. -/ @[simps!] @@ -218,8 +216,6 @@ def unitIso : 𝟭 (CommMon_ (C β₯€ D)) β‰… functor β‹™ inverse := set_option linter.uppercaseLean3 false in #align category_theory.monoidal.CommMon_functor_category_equivalence.unit_iso CategoryTheory.Monoidal.CommMonFunctorCategoryEquivalence.unitIso -attribute [local simp] id_tensorHom tensorHom_id - /-- The counit for the equivalence `CommMon_ (C β₯€ D) β‰Œ C β₯€ CommMon_ D`. -/ @[simps!] diff --git a/Mathlib/CategoryTheory/Monoidal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Limits.lean index 6cbe45e0c14aa..d2a73f9e61858 100644 --- a/Mathlib/CategoryTheory/Monoidal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Limits.lean @@ -68,38 +68,43 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorH (associativity := fun X Y Z => by ext j; dsimp simp only [limit.lift_Ο€, Cones.postcompose_obj_Ο€, Monoidal.associator_hom_app, limit.lift_map, - NatTrans.comp_app, Category.assoc] - slice_lhs 2 2 => rw [← tensor_id_comp_id_tensor] + NatTrans.comp_app, Category.assoc, tensorHom_id, id_tensorHom] + -- simp only [tensorHom_id, Cones.postcompose_obj_pt, Functor.const_obj_obj, Monoidal.tensorObj_obj, id_tensorHom] + slice_lhs 2 2 => rw [tensorHom_def'] slice_lhs 1 2 => - rw [← comp_tensor_id, limit.lift_Ο€] + rw [← comp_whiskerRight, limit.lift_Ο€] dsimp - slice_lhs 1 2 => rw [tensor_id_comp_id_tensor] - conv_lhs => rw [associator_naturality] - conv_rhs => rw [← id_tensor_comp_tensor_id (limit.Ο€ (Y βŠ— Z) j)] - slice_rhs 2 3 => - rw [← id_tensor_comp, limit.lift_Ο€] + rw [tensorHom_def] + slice_lhs 1 2 => rw [← whisker_exchange] + simp only [tensor_whiskerLeft, comp_whiskerRight, whisker_assoc, Category.assoc, + Iso.inv_hom_id_assoc, Iso.cancel_iso_hom_left] + slice_lhs 4 5 => rw [associator_naturality_left] + conv_rhs => rw [tensorHom_def _ (limit.Ο€ (Y βŠ— Z) j)] + slice_rhs 1 2 => + rw [← MonoidalCategory.whiskerLeft_comp, limit.lift_Ο€] dsimp - dsimp; simp) + rw [tensorHom_def] + simp) (left_unitality := fun X => by ext j; dsimp simp - conv_rhs => rw [← tensor_id_comp_id_tensor (limit.Ο€ X j)] + conv_rhs => rw [tensorHom_def' _ (limit.Ο€ X j)] slice_rhs 1 2 => - rw [← comp_tensor_id] + rw [← comp_whiskerRight] erw [limit.lift_Ο€] dsimp - slice_rhs 2 3 => rw [leftUnitor_naturality'] - simp [id_tensorHom, tensorHom_id]) + slice_rhs 2 3 => rw [leftUnitor_naturality] + simp) (right_unitality := fun X => by ext j; dsimp simp - conv_rhs => rw [← id_tensor_comp_tensor_id _ (limit.Ο€ X j)] + conv_rhs => rw [tensorHom_def (limit.Ο€ X j)] slice_rhs 1 2 => - rw [← id_tensor_comp] + rw [← MonoidalCategory.whiskerLeft_comp] erw [limit.lift_Ο€] dsimp - slice_rhs 2 3 => rw [rightUnitor_naturality'] - simp [id_tensorHom, tensorHom_id]) + slice_rhs 2 3 => rw [rightUnitor_naturality] + simp) #align category_theory.limits.limit_lax_monoidal CategoryTheory.Limits.limitLaxMonoidal /-- The limit functor `F ↦ limit F` bundled as a lax monoidal functor. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Mod_.lean b/Mathlib/CategoryTheory/Monoidal/Mod_.lean index 8910030dea772..a19dfdc61cc65 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mod_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mod_.lean @@ -35,8 +35,6 @@ namespace Mod_ variable {A : Mon_ C} (M : Mod_ A) -attribute [local simp] id_tensorHom tensorHom_id - theorem assoc_flip : (A.X ◁ M.act) ≫ M.act = (Ξ±_ A.X A.X M.X).inv ≫ (A.mul β–· M.X) ≫ M.act := by simp set_option linter.uppercaseLean3 false in diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index 21c56ec613228..0ba5122dc6158 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -96,8 +96,6 @@ structure Hom (M N : Mon_ C) where attribute [reassoc (attr := simp)] Hom.one_hom Hom.mul_hom -attribute [local simp] id_tensorHom tensorHom_id - /-- The identity morphism on a monoid object. -/ @[simps] def id (M : Mon_ C) : Hom M M where @@ -283,8 +281,6 @@ attribute [local aesop safe tactic (rule_sets [CategoryTheory])] attribute [local simp] eqToIso_map -attribute [local simp] id_tensorHom tensorHom_id - /-- Implementation of `Mon_.equivLaxMonoidalFunctorPUnit`. -/ @[simps!] def unitIso : @@ -372,8 +368,6 @@ which have also been proved in `Mathlib.CategoryTheory.Monoidal.Braided`. variable {C} -attribute [local simp] id_tensorHom tensorHom_id - -- The proofs that associators and unitors preserve monoid units don't require braiding. theorem one_associator {M N P : Mon_ C} : ((Ξ»_ (πŸ™_ C)).inv ≫ ((Ξ»_ (πŸ™_ C)).inv ≫ (M.one βŠ— N.one) βŠ— P.one)) ≫ (Ξ±_ M.X N.X P.X).hom = diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean index 854c7ad7e758c..82583d6ea4a7e 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean @@ -64,8 +64,6 @@ open CategoryTheory MonoidalCategory universe v v₁ vβ‚‚ v₃ u u₁ uβ‚‚ u₃ -attribute [local simp] id_tensorHom tensorHom_id - noncomputable section namespace CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index f702419f8d9fc..daf8f80c2b8f7 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -35,8 +35,6 @@ variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] variable {D : Type uβ‚‚} [Category.{vβ‚‚} D] -attribute [local simp] id_tensorHom tensorHom_id - -- porting note: it was @[simps {attrs := [`_refl_lemma]}] /-- Transport a monoidal structure along an equivalence of (plain) categories. -/ diff --git a/Mathlib/RepresentationTheory/Action.lean b/Mathlib/RepresentationTheory/Action.lean index 50cf3f8d0468e..8fd795c9ad0f7 100644 --- a/Mathlib/RepresentationTheory/Action.lean +++ b/Mathlib/RepresentationTheory/Action.lean @@ -982,8 +982,6 @@ variable {W : Type (u + 1)} [LargeCategory W] [MonoidalCategory V] [MonoidalCate open MonoidalCategory -attribute [local simp] id_tensorHom tensorHom_id - /-- A monoidal functor induces a monoidal functor between the categories of `G`-actions within those categories. -/ @[simps!] diff --git a/Mathlib/RepresentationTheory/FdRep.lean b/Mathlib/RepresentationTheory/FdRep.lean index 06923f5abd014..ee262e2955546 100644 --- a/Mathlib/RepresentationTheory/FdRep.lean +++ b/Mathlib/RepresentationTheory/FdRep.lean @@ -182,7 +182,7 @@ noncomputable def dualTensorIsoLinHomAux : `dualTensorHomEquiv k V W` of vector spaces induces an isomorphism of representations. -/ noncomputable def dualTensorIsoLinHom : FdRep.of ρV.dual βŠ— W β‰… FdRep.of (linHom ρV W.ρ) := by apply Action.mkIso (dualTensorIsoLinHomAux ρV W) (fun g => ?_) - convert dualTensorHom_comm ρV W.ρ g + -- convert dualTensorHom_comm ρV W.ρ g sorry #align fdRep.dual_tensor_iso_lin_hom FdRep.dualTensorIsoLinHom From d8b00d54bf9e541f5dc98cc65d485ad7b1e518e6 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Fri, 4 Aug 2023 00:46:13 +0900 Subject: [PATCH 07/62] fix style --- Mathlib/CategoryTheory/Monoidal/Bimod.lean | 61 ++++++++++++------- Mathlib/CategoryTheory/Monoidal/Category.lean | 41 +++++++------ .../Monoidal/Free/Coherence.lean | 41 ++++++------- .../Monoidal/NaturalTransformation.lean | 17 +++--- .../Monoidal/Rigid/OfEquivalence.lean | 3 +- 5 files changed, 93 insertions(+), 70 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index b1f9fa6debee0..8f7dd4732f178 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -331,9 +331,11 @@ theorem right_assoc' : slice_lhs 1 2 => rw [← whisker_exchange] slice_lhs 2 3 => rw [Ο€_tensor_id_actRight] slice_lhs 1 2 => rw [associator_naturality_right] - slice_lhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, right_assoc, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + slice_lhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, right_assoc, + MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] slice_rhs 1 2 => rw [associator_inv_naturality_left] - slice_rhs 2 3 => rw [← comp_whiskerRight, Ο€_tensor_id_actRight, comp_whiskerRight, comp_whiskerRight] + slice_rhs 2 3 => rw [← comp_whiskerRight, Ο€_tensor_id_actRight, comp_whiskerRight, + comp_whiskerRight] slice_rhs 4 5 => rw [Ο€_tensor_id_actRight] simp set_option linter.uppercaseLean3 false in @@ -352,12 +354,14 @@ theorem middle_assoc' : (Ξ±_ R.X _ T.X).hom ≫ (R.X ◁ actRight P Q) ≫ actLeft P Q := by refine' (cancel_epi ((tensorLeft _ β‹™ tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp [X] - slice_lhs 1 2 => rw [← comp_whiskerRight, whiskerLeft_Ο€_actLeft, comp_whiskerRight, comp_whiskerRight] + slice_lhs 1 2 => rw [← comp_whiskerRight, whiskerLeft_Ο€_actLeft, comp_whiskerRight, + comp_whiskerRight] slice_lhs 3 4 => rw [Ο€_tensor_id_actRight] slice_lhs 2 3 => rw [associator_naturality_left] -- porting note: had to replace `rw` by `erw` slice_rhs 1 2 => rw [associator_naturality_middle] - slice_rhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, Ο€_tensor_id_actRight, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + slice_rhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, Ο€_tensor_id_actRight, + MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] slice_rhs 4 5 => rw [whiskerLeft_Ο€_actLeft] slice_rhs 3 4 => rw [associator_inv_naturality_right] slice_rhs 4 5 => rw [whisker_exchange] @@ -389,7 +393,7 @@ noncomputable def tensorBimod {X Y Z : Mon_ C} (M : Bimod X Y) (N : Bimod Y Z) : set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod Bimod.tensorBimod -/-- Tensor product of two morphisms of bimodule objects. -/ +/-- Left whiskering for morphisms of bimodule objects. -/ @[simps] noncomputable def whiskerLeft {X Y Z : Mon_ C} (M : Bimod X Y) {N₁ Nβ‚‚ : Bimod Y Z} (f : N₁ ⟢ Nβ‚‚) : M.tensorBimod N₁ ⟢ M.tensorBimod Nβ‚‚ where @@ -407,7 +411,8 @@ noncomputable def whiskerLeft {X Y Z : Mon_ C} (M : Bimod X Y) {N₁ Nβ‚‚ : Bimo dsimp slice_lhs 1 2 => rw [TensorBimod.whiskerLeft_Ο€_actLeft] slice_lhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] - slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, ΞΉ_colimMap, parallelPairHom_app_one, MonoidalCategory.whiskerLeft_comp] + slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, ΞΉ_colimMap, parallelPairHom_app_one, + MonoidalCategory.whiskerLeft_comp] slice_rhs 2 3 => rw [TensorBimod.whiskerLeft_Ο€_actLeft] slice_rhs 1 2 => rw [associator_inv_naturality_right] slice_rhs 2 3 => rw [whisker_exchange] @@ -422,7 +427,7 @@ noncomputable def whiskerLeft {X Y Z : Mon_ C} (M : Bimod X Y) {N₁ Nβ‚‚ : Bimo slice_rhs 2 3 => rw [TensorBimod.Ο€_tensor_id_actRight] simp -/-- Tensor product of two morphisms of bimodule objects. -/ +/-- Right whiskering for morphisms of bimodule objects. -/ @[simps] noncomputable def whiskerRight {X Y Z : Mon_ C} {M₁ Mβ‚‚ : Bimod X Y} (f : M₁ ⟢ Mβ‚‚) (N : Bimod Y Z) : M₁.tensorBimod N ⟢ Mβ‚‚.tensorBimod N where @@ -439,7 +444,8 @@ noncomputable def whiskerRight {X Y Z : Mon_ C} {M₁ Mβ‚‚ : Bimod X Y} (f : M slice_lhs 1 2 => rw [TensorBimod.whiskerLeft_Ο€_actLeft] slice_lhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] slice_lhs 2 3 => rw [← comp_whiskerRight, Hom.left_act_hom] - slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, ΞΉ_colimMap, parallelPairHom_app_one, MonoidalCategory.whiskerLeft_comp] + slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, ΞΉ_colimMap, parallelPairHom_app_one, + MonoidalCategory.whiskerLeft_comp] slice_rhs 2 3 => rw [TensorBimod.whiskerLeft_Ο€_actLeft] slice_rhs 1 2 => rw [associator_inv_naturality_middle] simp @@ -449,7 +455,8 @@ noncomputable def whiskerRight {X Y Z : Mon_ C} {M₁ Mβ‚‚ : Bimod X Y} (f : M slice_lhs 1 2 => rw [TensorBimod.Ο€_tensor_id_actRight] slice_lhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] slice_lhs 2 3 => rw [whisker_exchange] - slice_rhs 1 2 => rw [← comp_whiskerRight, ΞΉ_colimMap, parallelPairHom_app_one, comp_whiskerRight] + slice_rhs 1 2 => rw [← comp_whiskerRight, ΞΉ_colimMap, parallelPairHom_app_one, + comp_whiskerRight] slice_rhs 2 3 => rw [TensorBimod.Ο€_tensor_id_actRight] simp @@ -474,7 +481,9 @@ noncomputable def homAux : (P.tensorBimod Q).X βŠ— L.X ⟢ (P.tensorBimod (Q.ten slice_lhs 2 3 => rw [← whisker_exchange] slice_lhs 3 4 => rw [coequalizer.condition] slice_lhs 2 3 => rw [associator_naturality_right] - slice_lhs 3 4 => rw [← MonoidalCategory.whiskerLeft_comp, TensorBimod.whiskerLeft_Ο€_actLeft, MonoidalCategory.whiskerLeft_comp] + slice_lhs 3 4 => + rw [← MonoidalCategory.whiskerLeft_comp, + TensorBimod.whiskerLeft_Ο€_actLeft, MonoidalCategory.whiskerLeft_comp] simp) set_option linter.uppercaseLean3 false in #align Bimod.associator_Bimod.hom_aux Bimod.AssociatorBimod.homAux @@ -487,11 +496,13 @@ noncomputable def hom : dsimp [homAux] refine' (cancel_epi ((tensorRight _ β‹™ tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp [TensorBimod.X] - slice_lhs 1 2 => - rw [← comp_whiskerRight, TensorBimod.Ο€_tensor_id_actRight, comp_whiskerRight, comp_whiskerRight] + slice_lhs 1 2 => rw [← comp_whiskerRight, TensorBimod.Ο€_tensor_id_actRight, + comp_whiskerRight, comp_whiskerRight] slice_lhs 3 5 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] slice_lhs 2 3 => rw [associator_naturality_middle] - slice_lhs 3 4 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.condition, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + slice_lhs 3 4 => + rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.condition, + MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] slice_rhs 1 2 => rw [associator_naturality_left] slice_rhs 2 3 => rw [← whisker_exchange] slice_rhs 3 5 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] @@ -507,17 +518,20 @@ theorem hom_left_act_hom' : rw [tensorLeft_map] slice_lhs 1 2 => rw [TensorBimod.whiskerLeft_Ο€_actLeft] slice_lhs 3 4 => rw [coequalizer.Ο€_desc] - slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.Ο€_desc, MonoidalCategory.whiskerLeft_comp] + slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.Ο€_desc, + MonoidalCategory.whiskerLeft_comp] refine' (cancel_epi ((tensorRight _ β‹™ tensorLeft _).map (coequalizer.Ο€ _ _))).1 _ dsimp; dsimp [TensorBimod.X] slice_lhs 1 2 => rw [associator_inv_naturality_middle] slice_lhs 2 3 => - rw [← comp_whiskerRight, TensorBimod.whiskerLeft_Ο€_actLeft, comp_whiskerRight, comp_whiskerRight] + rw [← comp_whiskerRight, TensorBimod.whiskerLeft_Ο€_actLeft, + comp_whiskerRight, comp_whiskerRight] slice_lhs 4 6 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] slice_lhs 3 4 => rw [associator_naturality_left] slice_rhs 1 3 => - rw [← MonoidalCategory.whiskerLeft_comp, ← MonoidalCategory.whiskerLeft_comp, Ο€_tensor_id_preserves_coequalizer_inv_desc, - MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + rw [← MonoidalCategory.whiskerLeft_comp, ← MonoidalCategory.whiskerLeft_comp, + Ο€_tensor_id_preserves_coequalizer_inv_desc, MonoidalCategory.whiskerLeft_comp, + MonoidalCategory.whiskerLeft_comp] slice_rhs 3 4 => erw [TensorBimod.whiskerLeft_Ο€_actLeft P (Q.tensorBimod L)] slice_rhs 2 3 => erw [associator_inv_naturality_right] slice_rhs 3 4 => erw [whisker_exchange] @@ -547,7 +561,8 @@ theorem hom_right_act_hom' : slice_rhs 2 3 => erw [associator_naturality_middle] dsimp slice_rhs 3 4 => - rw [← MonoidalCategory.whiskerLeft_comp, TensorBimod.Ο€_tensor_id_actRight, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + rw [← MonoidalCategory.whiskerLeft_comp, TensorBimod.Ο€_tensor_id_actRight, + MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] coherence set_option linter.uppercaseLean3 false in #align Bimod.associator_Bimod.hom_right_act_hom' Bimod.AssociatorBimod.hom_right_act_hom' @@ -584,10 +599,12 @@ noncomputable def inv : slice_lhs 1 2 => rw [whisker_exchange] slice_lhs 2 4 => rw [id_tensor_Ο€_preserves_coequalizer_inv_desc] slice_lhs 1 2 => rw [associator_inv_naturality_left] - slice_lhs 2 3 => rw [← comp_whiskerRight, coequalizer.condition, comp_whiskerRight, comp_whiskerRight] + slice_lhs 2 3 => + rw [← comp_whiskerRight, coequalizer.condition, comp_whiskerRight, comp_whiskerRight] slice_rhs 1 2 => rw [associator_naturality_right] slice_rhs 2 3 => - rw [← MonoidalCategory.whiskerLeft_comp, TensorBimod.whiskerLeft_Ο€_actLeft, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + rw [← MonoidalCategory.whiskerLeft_comp, TensorBimod.whiskerLeft_Ο€_actLeft, + MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] slice_rhs 4 6 => rw [id_tensor_Ο€_preserves_coequalizer_inv_desc] slice_rhs 3 4 => rw [associator_inv_naturality_middle] coherence) @@ -1013,8 +1030,8 @@ theorem pentagon_bimod {V W X Y Z : Mon_ C} (M : Bimod V W) (N : Bimod W X) (P : slice_lhs 5 6 => rw [ΞΉ_colimMap, parallelPairHom_app_one] slice_lhs 4 5 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.Ο€_desc] slice_lhs 3 4 => - rw [← MonoidalCategory.whiskerLeft_comp, Ο€_tensor_id_preserves_coequalizer_inv_desc, MonoidalCategory.whiskerLeft_comp, - MonoidalCategory.whiskerLeft_comp] + rw [← MonoidalCategory.whiskerLeft_comp, Ο€_tensor_id_preserves_coequalizer_inv_desc, + MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] slice_rhs 1 2 => rw [associator_naturality_left] slice_rhs 2 3 => rw [← whisker_exchange] diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 393b971dea483..6ebf7e0549ee3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -149,17 +149,9 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] where aesop_cat #align category_theory.monoidal_category CategoryTheory.MonoidalCategory --- attribute [simp] MonoidalCategory.tensor_id --- attribute [reassoc] MonoidalCategory.tensor_comp --- attribute [simp] MonoidalCategory.tensor_comp --- attribute [reassoc] MonoidalCategory.associator_naturality --- attribute [reassoc] MonoidalCategory.leftUnitor_naturality --- attribute [reassoc] MonoidalCategory.rightUnitor_naturality attribute [reassoc (attr := simp)] MonoidalCategory.pentagon attribute [reassoc (attr := simp)] MonoidalCategory.triangle --- attribute [simp] MonoidalCategory.whiskerLeft_eq_tensorHom_id MonoidalCategory.whiskerRight_eq_id_tensorHom - namespace MonoidalCategory attribute [reassoc] @@ -167,8 +159,10 @@ attribute [reassoc] whiskerRight_tensor whisker_assoc whisker_exchange attribute [simp] - whiskerLeft_id whiskerLeft_comp id_whiskerLeft tensor_whiskerLeft id_whiskerRight comp_whiskerRight - whiskerRight_id whiskerRight_tensor whisker_assoc + whiskerLeft_id whiskerRight_id + whiskerLeft_comp id_whiskerLeft tensor_whiskerLeft comp_whiskerRight id_whiskerRight + whiskerRight_tensor whisker_assoc + end MonoidalCategory -- Porting Note: This is here to make `tensorUnit` explicitly depend on `C`, which was done in @@ -221,8 +215,9 @@ Composition of tensor products is tensor product of compositions: `(f₁ βŠ— g₁) ∘ (fβ‚‚ βŠ— gβ‚‚) = (f₁ ∘ fβ‚‚) βŠ— (g₁ βŠ— gβ‚‚)` -/ @[reassoc, simp] -theorem tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : Y₁ ⟢ Z₁) (gβ‚‚ : Yβ‚‚ ⟢ Zβ‚‚) : - (f₁ ≫ g₁) βŠ— (fβ‚‚ ≫ gβ‚‚) = (f₁ βŠ— fβ‚‚) ≫ (g₁ βŠ— gβ‚‚) := by +theorem tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} + (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : Y₁ ⟢ Z₁) (gβ‚‚ : Yβ‚‚ ⟢ Zβ‚‚) : + (f₁ ≫ g₁) βŠ— (fβ‚‚ ≫ gβ‚‚) = (f₁ βŠ— fβ‚‚) ≫ (g₁ βŠ— gβ‚‚) := by simp [whisker_exchange_assoc] @[simp] @@ -585,9 +580,10 @@ theorem rightUnitor_tensor_inv (X Y : C) : (ρ_ (X βŠ— Y)).inv = X ◁ (ρ_ Y).inv ≫ (Ξ±_ X Y (πŸ™_ C)).inv := by simp @[reassoc] -theorem associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (f₃ : X₃ ⟢ Y₃) : - tensorHom (tensorHom f₁ fβ‚‚) f₃ ≫ (associator Y₁ Yβ‚‚ Y₃).hom = - (associator X₁ Xβ‚‚ X₃).hom ≫ tensorHom f₁ (tensorHom fβ‚‚ f₃) := by +theorem associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C} + (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (f₃ : X₃ ⟢ Y₃) : + tensorHom (tensorHom f₁ fβ‚‚) f₃ ≫ (associator Y₁ Yβ‚‚ Y₃).hom = + (associator X₁ Xβ‚‚ X₃).hom ≫ tensorHom f₁ (tensorHom fβ‚‚ f₃) := by simp @[reassoc] @@ -598,7 +594,8 @@ theorem triangle' (X Y : C) : @[reassoc] theorem pentagon' (W X Y Z : C) : - ((Ξ±_ W X Y).hom βŠ— πŸ™ Z) ≫ (Ξ±_ W (X βŠ— Y) Z).hom ≫ (πŸ™ W βŠ— (Ξ±_ X Y Z).hom) = (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom := by + ((Ξ±_ W X Y).hom βŠ— πŸ™ Z) ≫ (Ξ±_ W (X βŠ— Y) Z).hom ≫ (πŸ™ W βŠ— (Ξ±_ X Y Z).hom) = + (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom := by simp #align category_theory.monoidal_category.pentagon CategoryTheory.MonoidalCategory.pentagon' @@ -631,7 +628,9 @@ theorem rightUnitor_conjugation {X Y : C} (f : X ⟢ Y) : simp #align category_theory.monoidal_category.right_unitor_conjugation CategoryTheory.MonoidalCategory.rightUnitor_conjugation -theorem leftUnitor_conjugation {X Y : C} (f : X ⟢ Y) : πŸ™ (πŸ™_ C) βŠ— f = (Ξ»_ X).hom ≫ f ≫ (Ξ»_ Y).inv := by simp +theorem leftUnitor_conjugation {X Y : C} (f : X ⟢ Y) : + πŸ™ (πŸ™_ C) βŠ— f = (Ξ»_ X).hom ≫ f ≫ (Ξ»_ Y).inv := by + simp #align category_theory.monoidal_category.left_unitor_conjugation CategoryTheory.MonoidalCategory.leftUnitor_conjugation @[reassoc] @@ -645,11 +644,15 @@ theorem rightUnitor_naturality' {X Y : C} (f : X ⟢ Y) : simp @[reassoc] -theorem leftUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : f ≫ (Ξ»_ X').inv = (Ξ»_ X).inv ≫ (πŸ™ _ βŠ— f) := by simp +theorem leftUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : + f ≫ (Ξ»_ X').inv = (Ξ»_ X).inv ≫ (πŸ™ _ βŠ— f) := by + simp #align category_theory.monoidal_category.left_unitor_inv_naturality CategoryTheory.MonoidalCategory.leftUnitor_inv_naturality' @[reassoc] -theorem rightUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : f ≫ (ρ_ X').inv = (ρ_ X).inv ≫ (f βŠ— πŸ™ _) := by simp +theorem rightUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : + f ≫ (ρ_ X').inv = (ρ_ X).inv ≫ (f βŠ— πŸ™ _) := by + simp #align category_theory.monoidal_category.right_unitor_inv_naturality CategoryTheory.MonoidalCategory.rightUnitor_inv_naturality' end diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean index a99cc00799983..aa4b533ef919b 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean @@ -143,7 +143,8 @@ open Hom @[simp] def normalizeMapAux : βˆ€ {X Y : F C}, (X ⟢ᡐ Y) β†’ - ((Discrete.functor (fun n ↦ ⟨normalizeObj X n⟩) : N C β₯€ N C) ⟢ Discrete.functor fun n ↦ ⟨normalizeObj Y n⟩) + ((Discrete.functor (fun n ↦ ⟨normalizeObj X n⟩) : N C β₯€ N C) ⟢ + Discrete.functor fun n ↦ ⟨normalizeObj Y n⟩) | _, _, Hom.id _ => by dsimp; exact πŸ™ _ | _, _, Ξ±_hom X Y Z => by dsimp; exact Discrete.natTrans (fun _ => πŸ™ _) | _, _, Ξ±_inv _ _ _ => by dsimp; exact Discrete.natTrans (fun _ => πŸ™ _) @@ -219,8 +220,8 @@ def normalizeIsoApp : βˆ€ (X : F C) (n : N C), ((tensorFunc C).obj X).obj n β‰… ((normalize' C).obj X).obj n | of _, _ => Iso.refl _ | Unit, _ => ρ_ _ - | tensor X _, n => - (Ξ±_ _ _ _).symm β‰ͺ≫ tensorIso (normalizeIsoApp X n) (Iso.refl _) β‰ͺ≫ normalizeIsoApp _ _ + | tensor X a, n => + (Ξ±_ _ _ _).symm β‰ͺ≫ whiskerRightIso (normalizeIsoApp X n) a β‰ͺ≫ normalizeIsoApp _ _ #align category_theory.free_monoidal_category.normalize_iso_app CategoryTheory.FreeMonoidalCategory.normalizeIsoApp /-- Almost non-definitionally equall to `normalizeIsoApp`, but has a better definitional property @@ -230,10 +231,11 @@ def normalizeIsoApp' : βˆ€ (X : F C) (n : NormalMonoidalObject C), inclusionObj n βŠ— X β‰… inclusionObj (normalizeObj X n) | of _, _ => Iso.refl _ | Unit, _ => ρ_ _ - | tensor X _, n => - (Ξ±_ _ _ _).symm β‰ͺ≫ tensorIso (normalizeIsoApp' X n) (Iso.refl _) β‰ͺ≫ normalizeIsoApp' _ _ + | tensor X Y, n => + (Ξ±_ _ _ _).symm β‰ͺ≫ whiskerRightIso (normalizeIsoApp' X n) Y β‰ͺ≫ normalizeIsoApp' _ _ -theorem normalizeIsoApp_eq : βˆ€ (X : F C) (n : N C), normalizeIsoApp C X n = normalizeIsoApp' C X n.as +theorem normalizeIsoApp_eq : + βˆ€ (X : F C) (n : N C), normalizeIsoApp C X n = normalizeIsoApp' C X n.as | of X, _ => rfl | Unit, _ => rfl | tensor X Y, n => by @@ -245,7 +247,7 @@ theorem normalizeIsoApp_eq : βˆ€ (X : F C) (n : N C), normalizeIsoApp C X n = no @[simp] theorem normalizeIsoApp_tensor (X Y : F C) (n : N C) : normalizeIsoApp C (X βŠ— Y) n = - (Ξ±_ _ _ _).symm β‰ͺ≫ tensorIso (normalizeIsoApp C X n) (Iso.refl _) β‰ͺ≫ normalizeIsoApp _ _ _ := + (Ξ±_ _ _ _).symm β‰ͺ≫ whiskerRightIso (normalizeIsoApp C X n) Y β‰ͺ≫ normalizeIsoApp _ _ _ := rfl #align category_theory.free_monoidal_category.normalize_iso_app_tensor CategoryTheory.FreeMonoidalCategory.normalizeIsoApp_tensor @@ -294,10 +296,8 @@ theorem normalizeObj_congr (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ Y clear n f induction f' with | comp _ _ _ _ => apply Eq.trans <;> assumption - | whiskerLeft _ _ ih => funext; apply congr_fun ih - | whiskerRight _ _ ih => - funext - apply congr_argβ‚‚ _ rfl (congr_fun ih _) + | whiskerLeft _ _ ih => funext; apply congr_fun ih + | whiskerRight _ _ ih => funext; apply congr_argβ‚‚ _ rfl (congr_fun ih _) | _ => funext; rfl theorem normalize_naturality (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ Y) : @@ -325,19 +325,18 @@ theorem normalize_naturality (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ | whiskerLeft X f ih => intro n erw [mk_whiskerLeft] - dsimp only [tensor_eq_tensor, normalizeObj_tensor, normalizeIsoApp', Iso.trans_hom, Iso.symm_hom, tensorIso_hom, - Iso.refl_hom, Function.comp_apply, inclusion_obj] - simp only [MonoidalCategory.whiskerLeft_id, Category.assoc, Category.id_comp, tensorHom_id] + simp only [tensor_eq_tensor, normalizeObj_tensor, normalizeIsoApp', Iso.trans_hom, + Iso.symm_hom, whiskerRightIso_hom, Function.comp_apply, inclusion_obj, inclusion_map, + Category.assoc] rw [associator_inv_naturality_right_assoc, whisker_exchange_assoc, ih] + simp | @whiskerRight X Y h Ξ·' ih => intro n erw [mk_whiskerRight] - dsimp - simp only [MonoidalCategory.whiskerLeft_id, Category.assoc, Category.id_comp, tensorHom_id] - rw [associator_inv_naturality_middle_assoc, ← comp_whiskerRight_assoc] - erw [ih] - dsimp - simp + simp only [tensor_eq_tensor, normalizeObj_tensor, normalizeIsoApp', Iso.trans_hom, + Iso.symm_hom, whiskerRightIso_hom, Function.comp_apply, inclusion_obj, inclusion_map, + Category.assoc] + rw [associator_inv_naturality_middle_assoc, ← comp_whiskerRight_assoc, ih] have := dcongr_arg (fun x => (normalizeIsoApp' C Ξ·' x).hom) (normalizeObj_congr n (Quotient.mk (setoidHom X Y) h)) dsimp at this; simp [this] @@ -352,7 +351,7 @@ def normalizeIso : tensorFunc C β‰… normalize' C := intro X Y f ext ⟨n⟩ convert normalize_naturality n f using 1 - any_goals dsimp [NatIso.ofComponents]; simp; congr; apply normalizeIsoApp_eq + any_goals dsimp [NatIso.ofComponents]; congr; apply normalizeIsoApp_eq #align category_theory.free_monoidal_category.normalize_iso CategoryTheory.FreeMonoidalCategory.normalizeIso /-- The isomorphism between an object and its normal form is natural. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index 8abcfd6644237..fec02a0b4d3bf 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -118,12 +118,14 @@ def hcomp {F G : LaxMonoidalFunctor C D} {H K : LaxMonoidalFunctor D E} (Ξ± : Mo dsimp; simp conv_lhs => rw [← K.toFunctor.map_comp, Ξ±.unit] tensor := fun X Y => by - simp only [LaxMonoidalFunctor.comp_toFunctor, comp_obj, LaxMonoidalFunctor.comp_ΞΌ, NatTrans.hcomp_app, assoc, - NatTrans.naturality_assoc, tensor_assoc, tensorHom_def, whisker_exchange_assoc, MonoidalCategory.whiskerLeft_comp, - comp_whiskerRight, LaxMonoidalFunctor.ΞΌ_natural_left_assoc, LaxMonoidalFunctor.ΞΌ_natural_right_assoc] + simp only [LaxMonoidalFunctor.comp_toFunctor, comp_obj, + LaxMonoidalFunctor.comp_ΞΌ, NatTrans.hcomp_app, assoc, + NatTrans.naturality_assoc, tensor_assoc, tensorHom_def, + whisker_exchange_assoc, MonoidalCategory.whiskerLeft_comp, + comp_whiskerRight, LaxMonoidalFunctor.ΞΌ_natural_left_assoc, + LaxMonoidalFunctor.ΞΌ_natural_right_assoc] simp_rw [← K.toFunctor.map_comp] - congr 4 - simp + simp only [tensor, tensorHom_def, assoc, map_comp] } #align category_theory.monoidal_nat_trans.hcomp CategoryTheory.MonoidalNatTrans.hcomp @@ -200,8 +202,9 @@ def monoidalUnit (F : MonoidalFunctor C D) [IsEquivalence F.toFunctor] : id_comp, assoc] simp only [← Functor.map_comp, assoc] erw [e.counit_app_functor, e.counit_app_functor] - simp only [comp_obj, asEquivalence_functor, asEquivalence_inverse, id_obj, LaxMonoidalFunctor.ΞΌ_natural_left, - LaxMonoidalFunctor.ΞΌ_natural_right_assoc, IsIso.inv_hom_id_assoc, map_comp, IsEquivalence.inv_fun_map, assoc, + simp only [comp_obj, asEquivalence_functor, asEquivalence_inverse, id_obj, + LaxMonoidalFunctor.ΞΌ_natural_left, LaxMonoidalFunctor.ΞΌ_natural_right_assoc, + IsIso.inv_hom_id_assoc, map_comp, IsEquivalence.inv_fun_map, assoc, Iso.hom_inv_id_app_assoc, tensorHom_def] slice_rhs 3 4 => erw [Iso.hom_inv_id_app] simp only [id_obj, id_comp, assoc, whisker_exchange_assoc] diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean index 2e92d5d5c243c..2a96d08165e06 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean @@ -48,7 +48,8 @@ def exactPairingOfFaithful [Faithful F.toFunctor] {X Y : C} (eval : Y βŠ— X ⟢ simp only [MonoidalCategory.whiskerLeft_comp, Category.assoc, comp_whiskerRight, LaxMonoidalFunctor.associativity_inv_assoc, IsIso.hom_inv_id_assoc, IsIso.inv_comp_eq] simp_rw [← comp_whiskerRight_assoc] - simp only [IsIso.hom_inv_id_assoc, comp_whiskerRight, Category.assoc, ExactPairing.coevaluation_evaluation_assoc, + simp only [IsIso.hom_inv_id_assoc, comp_whiskerRight, Category.assoc, + ExactPairing.coevaluation_evaluation_assoc, LaxMonoidalFunctor.right_unitality, LaxMonoidalFunctor.left_unitality_inv] simp_rw [← whiskerLeft_comp_assoc] simp ) From d56ded3590f37337c98284c25a5911e75c2308dc Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Fri, 4 Aug 2023 01:10:57 +0900 Subject: [PATCH 08/62] fix style --- Mathlib/CategoryTheory/Monoidal/Braided.lean | 5 +++-- Mathlib/CategoryTheory/Monoidal/Category.lean | 5 ----- .../CategoryTheory/Monoidal/Free/Basic.lean | 6 ++++-- Mathlib/CategoryTheory/Monoidal/Functor.lean | 18 +++++++++++------- Mathlib/CategoryTheory/Monoidal/Limits.lean | 1 - Mathlib/CategoryTheory/Monoidal/Mod_.lean | 3 ++- .../CategoryTheory/Monoidal/Rigid/Basic.lean | 4 ++-- Mathlib/Tactic/CategoryTheory/Coherence.lean | 6 ++++-- 8 files changed, 26 insertions(+), 22 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 4c5ed507991ea..d88cd369d2cfe 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -543,8 +543,9 @@ theorem tensor_associativity (X₁ Xβ‚‚ Y₁ Yβ‚‚ Z₁ Zβ‚‚ : C) : (Ξ±_ (X₁ βŠ— Xβ‚‚) (Y₁ βŠ— Yβ‚‚) (Z₁ βŠ— Zβ‚‚)).hom ≫ ((X₁ βŠ— Xβ‚‚) ◁ tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚)) ≫ tensor_ΞΌ C (X₁, Xβ‚‚) (Y₁ βŠ— Z₁, Yβ‚‚ βŠ— Zβ‚‚) := by dsimp only [tensor_obj, prodMonoidal_tensorObj, tensor_ΞΌ] - simp only [whiskerRight_tensor, comp_whiskerRight, whisker_assoc, assoc, Iso.inv_hom_id_assoc, tensor_whiskerLeft, - braiding_tensor_left, MonoidalCategory.whiskerLeft_comp, braiding_tensor_right] + simp only [whiskerRight_tensor, comp_whiskerRight, whisker_assoc, assoc, Iso.inv_hom_id_assoc, + tensor_whiskerLeft, braiding_tensor_left, MonoidalCategory.whiskerLeft_comp, + braiding_tensor_right] calc _ = πŸ™ _ βŠ—β‰« X₁ ◁ ((Ξ²_ Xβ‚‚ Y₁).hom β–· (Yβ‚‚ βŠ— Z₁) ≫ (Y₁ βŠ— Xβ‚‚) ◁ (Ξ²_ Yβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 6ebf7e0549ee3..1aac620ecf8ea 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -239,11 +239,6 @@ variable {C : Type u} [π’ž : Category.{v} C] [MonoidalCategory C] namespace MonoidalCategory --- @[simp] --- theorem tensorHom_def {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : --- f βŠ— g = X₁ ◁ g ≫ f β–· Yβ‚‚ := by --- rw [← whiskerLeft_eq_tensorHom_id, ← whiskerRight_eq_id_tensorHom, ← tensor_comp, id_comp, comp_id] - @[reassoc (attr := simp)] theorem hom_inv_whiskerLeft (X : C) {Y Z : C} (f : Y β‰… Z) : X ◁ f.hom ≫ X ◁ f.inv = πŸ™ (X βŠ— Y) := by rw [← whiskerLeft_comp, hom_inv_id, whiskerLeft_id] diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean index da20ec5b5eeb0..8fc0e07f98fa2 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean @@ -114,7 +114,8 @@ inductive HomEquiv : βˆ€ {X Y : F C}, (X ⟢ᡐ Y) β†’ (X ⟢ᡐ Y) β†’ Prop HomEquiv ((f.whiskerLeft X).whiskerRight Z) ((Hom.Ξ±_hom X Y Z).comp <| ((f.whiskerRight Z).whiskerLeft X).comp <| Hom.Ξ±_inv X Y' Z) | whisker_exchange {W X Y Z} (f : W ⟢ᡐ X) (g : Y ⟢ᡐ Z) : - HomEquiv ((g.whiskerLeft W).comp <| f.whiskerRight Z) ((f.whiskerRight Y).comp <| g.whiskerLeft X) + HomEquiv ((g.whiskerLeft W).comp <| f.whiskerRight Z) + ((f.whiskerRight Y).comp <| g.whiskerLeft X) | Ξ±_hom_inv {X Y Z} : HomEquiv ((Hom.Ξ±_hom X Y Z).comp (Hom.Ξ±_inv X Y Z)) (Hom.id _) | Ξ±_inv_hom {X Y Z} : HomEquiv ((Hom.Ξ±_inv X Y Z).comp (Hom.Ξ±_hom X Y Z)) (Hom.id _) | ρ_hom_inv {X} : HomEquiv ((Hom.ρ_hom X).comp (Hom.ρ_inv X)) (Hom.id _) @@ -170,7 +171,8 @@ instance categoryFreeMonoidalCategory : Category.{u} (F C) where instance : MonoidalCategory (F C) where tensorObj X Y := FreeMonoidalCategory.tensor X Y whiskerLeft := fun X _ _ f => Quotient.map (Hom.whiskerLeft X) (HomEquiv.whisker_left X) f - whiskerRight := fun f Y => Quotient.map (fun _f => Hom.whiskerRight _f Y) (fun _ _ h => HomEquiv.whisker_right _ _ _ h) f + whiskerRight := fun f Y => + Quotient.map (fun f' => Hom.whiskerRight f' Y) (fun _ _ h => HomEquiv.whisker_right _ _ _ h) f tensorUnit' := FreeMonoidalCategory.Unit associator X Y Z := ⟨⟦Hom.Ξ±_hom X Y Z⟧, ⟦Hom.Ξ±_inv X Y Z⟧, Quotient.sound Ξ±_hom_inv, Quotient.sound Ξ±_inv_hom⟩ diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index e13974b784759..f9e3994608c99 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -118,8 +118,9 @@ variable {C D} attribute [local simp] tensorHom_def @[reassoc (attr := simp)] -theorem LaxMonoidalFunctor.ΞΌ_natural (F : LaxMonoidalFunctor C D) {X Y X' Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y') : - (F.map f βŠ— F.map g) ≫ F.ΞΌ Y Y' = F.ΞΌ X X' ≫ F.map (f βŠ— g) := by +theorem LaxMonoidalFunctor.ΞΌ_natural (F : LaxMonoidalFunctor C D) {X Y X' Y' : C} + (f : X ⟢ Y) (g : X' ⟢ Y') : + (F.map f βŠ— F.map g) ≫ F.ΞΌ Y Y' = F.ΞΌ X X' ≫ F.map (f βŠ— g) := by simp only [tensorHom_def, assoc, F.ΞΌ_natural_left, F.ΞΌ_natural_right_assoc, map_comp] @[reassoc] @@ -155,10 +156,12 @@ def LaxMonoidalFunctor.ofTensorHom (F : C β₯€ D) (Ξ±_ (F.obj X) (F.obj Y) (F.obj Z)).hom ≫ (πŸ™ (F.obj X) βŠ— ΞΌ Y Z) ≫ ΞΌ X (Y βŠ— Z) := by aesop_cat) /- unitality -/ - (left_unitality : βˆ€ X : C, (Ξ»_ (F.obj X)).hom = (Ξ΅ βŠ— πŸ™ (F.obj X)) ≫ ΞΌ (πŸ™_ C) X ≫ F.map (Ξ»_ X).hom := - by aesop_cat) - (right_unitality : βˆ€ X : C, (ρ_ (F.obj X)).hom = (πŸ™ (F.obj X) βŠ— Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ F.map (ρ_ X).hom := - by aesop_cat) : + (left_unitality : + βˆ€ X : C, (Ξ»_ (F.obj X)).hom = (Ξ΅ βŠ— πŸ™ (F.obj X)) ≫ ΞΌ (πŸ™_ C) X ≫ F.map (Ξ»_ X).hom := + by aesop_cat) + (right_unitality : + βˆ€ X : C, (ρ_ (F.obj X)).hom = (πŸ™ (F.obj X) βŠ— Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ F.map (ρ_ X).hom := + by aesop_cat) : LaxMonoidalFunctor C D where obj := F.obj map := F.map @@ -560,7 +563,8 @@ noncomputable def monoidalAdjoint (F : MonoidalFunctor C D) {G : D β₯€ C} (h : F LaxMonoidalFunctor D C := LaxMonoidalFunctor.ofTensorHom (F := G) (Ξ΅ := h.homEquiv _ _ (inv F.Ξ΅)) - (ΞΌ := fun X Y ↦ h.homEquiv _ (X βŠ— Y) (inv (F.ΞΌ (G.obj X) (G.obj Y)) ≫ (h.counit.app X βŠ— h.counit.app Y))) + (ΞΌ := fun X Y ↦ + h.homEquiv _ (X βŠ— Y) (inv (F.ΞΌ (G.obj X) (G.obj Y)) ≫ (h.counit.app X βŠ— h.counit.app Y))) (ΞΌ_natural := @fun X Y X' Y' f g => by rw [← h.homEquiv_naturality_left, ← h.homEquiv_naturality_right, Equiv.apply_eq_iff_eq, assoc, IsIso.eq_inv_comp, ← F.toLaxMonoidalFunctor.ΞΌ_natural_assoc, IsIso.hom_inv_id_assoc, ← diff --git a/Mathlib/CategoryTheory/Monoidal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Limits.lean index d2a73f9e61858..c26c58e3b3a8e 100644 --- a/Mathlib/CategoryTheory/Monoidal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Limits.lean @@ -69,7 +69,6 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorH ext j; dsimp simp only [limit.lift_Ο€, Cones.postcompose_obj_Ο€, Monoidal.associator_hom_app, limit.lift_map, NatTrans.comp_app, Category.assoc, tensorHom_id, id_tensorHom] - -- simp only [tensorHom_id, Cones.postcompose_obj_pt, Functor.const_obj_obj, Monoidal.tensorObj_obj, id_tensorHom] slice_lhs 2 2 => rw [tensorHom_def'] slice_lhs 1 2 => rw [← comp_whiskerRight, limit.lift_Ο€] diff --git a/Mathlib/CategoryTheory/Monoidal/Mod_.lean b/Mathlib/CategoryTheory/Monoidal/Mod_.lean index a19dfdc61cc65..59361b04c246c 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mod_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mod_.lean @@ -128,7 +128,8 @@ def comap {A B : Mon_ C} (f : A ⟢ B) : Mod_ B β₯€ Mod_ A where assoc := by -- oh, for homotopy.io in a widget! slice_rhs 2 3 => rw [whisker_exchange] - simp only [whiskerRight_tensor, MonoidalCategory.whiskerLeft_comp, Category.assoc, Iso.hom_inv_id_assoc] + simp only [whiskerRight_tensor, MonoidalCategory.whiskerLeft_comp, Category.assoc, + Iso.hom_inv_id_assoc] slice_rhs 4 5 => rw [Mod_.assoc_flip] slice_rhs 3 4 => rw [associator_inv_naturality_middle] slice_rhs 2 4 => rw [Iso.hom_inv_id_assoc] diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean index 82583d6ea4a7e..3de03843b95d2 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean @@ -251,8 +251,8 @@ theorem comp_rightAdjointMate {X Y Z : C} [HasRightDual X] [HasRightDual Y] [Has slice_rhs 2 3 => rw [← tensor_comp, tensor_id, Category.comp_id, ← Category.id_comp (Ξ·_ Y (Yᘁ)), tensor_comp] rw [← id_tensor_comp_tensor_id _ (Ξ·_ Y (Yᘁ)), ← tensor_id] - simp only [tensorHom_id, id_tensorHom, id_whiskerLeft, id_whiskerRight, whiskerRight_tensor, Category.assoc, - Iso.inv_hom_id_assoc, pentagon_hom_hom_inv_hom_hom] + simp only [tensorHom_id, id_tensorHom, id_whiskerLeft, id_whiskerRight, whiskerRight_tensor, + Category.assoc, Iso.inv_hom_id_assoc, pentagon_hom_hom_inv_hom_hom] rw [← associator_naturality_middle_assoc] simp_rw [← comp_whiskerRight_assoc, evaluation_coevaluation] coherence diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index 37e8f7bd50dd5..53894ce3f45cf 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -364,10 +364,12 @@ syntax (name := monoidal_simps) "monoidal_simps" : tactic elab_rules : tactic | `(tactic| monoidal_simps) => do evalTactic (← `(tactic| - simp only [Category.assoc, MonoidalCategory.tensor_whiskerLeft, MonoidalCategory.id_whiskerLeft, + simp only [ + Category.assoc, MonoidalCategory.tensor_whiskerLeft, MonoidalCategory.id_whiskerLeft, MonoidalCategory.whiskerRight_tensor, MonoidalCategory.whiskerRight_id, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_id, - MonoidalCategory.comp_whiskerRight, MonoidalCategory.id_whiskerRight, MonoidalCategory.whisker_assoc, + MonoidalCategory.comp_whiskerRight, MonoidalCategory.id_whiskerRight, + MonoidalCategory.whisker_assoc, MonoidalCategory.id_tensorHom, MonoidalCategory.tensorHom_id]; try simp only [MonoidalCategory.tensorHom_def] )) From 4cd06b718a2beaf1186bf6046a3148c98ab211f8 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Fri, 4 Aug 2023 17:10:49 +0900 Subject: [PATCH 09/62] use another def of `tensorHom` as default --- .../Category/ModuleCat/Monoidal/Basic.lean | 24 +-- .../Closed/FunctorCategory.lean | 11 +- Mathlib/CategoryTheory/Enriched/Basic.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Braided.lean | 6 +- Mathlib/CategoryTheory/Monoidal/Category.lean | 155 ++++++++---------- Mathlib/CategoryTheory/Monoidal/Center.lean | 95 ++++++----- Mathlib/CategoryTheory/Monoidal/Functor.lean | 4 +- Mathlib/CategoryTheory/Monoidal/Limits.lean | 12 +- Mathlib/CategoryTheory/Monoidal/Mod_.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 4 +- .../Monoidal/NaturalTransformation.lean | 8 +- .../CategoryTheory/Monoidal/Preadditive.lean | 6 +- .../CategoryTheory/Monoidal/Rigid/Basic.lean | 2 +- .../Monoidal/Rigid/FunctorCategory.lean | 8 +- .../CategoryTheory/Monoidal/Transport.lean | 2 +- Mathlib/Tactic/CategoryTheory/Coherence.lean | 1 + 16 files changed, 155 insertions(+), 187 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean index 9d643164107eb..7f31f17e0e76e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean @@ -140,8 +140,7 @@ theorem pentagon (W X Y Z : ModuleCat R) : whiskerRight (associator W X Y).hom Z ≫ (associator W (tensorObj X Y) Z).hom ≫ whiskerLeft W (associator X Y Z).hom = (associator (tensorObj W X) Y Z).hom ≫ (associator W X (tensorObj Y Z)).hom := by - sorry - -- convert pentagon_aux R W X Y Z using 1 + convert pentagon_aux R W X Y Z using 1 #align Module.monoidal_category.pentagon ModuleCat.MonoidalCategory.pentagon /-- (implementation) the left unitor for R-modules -/ @@ -194,27 +193,6 @@ theorem triangle (M N : ModuleCat.{u} R) : exact (TensorProduct.smul_tmul _ _ _).symm #align Module.monoidal_category.triangle ModuleCat.MonoidalCategory.triangle --- theorem id_whiskerLeft_aux {M N : ModuleCat R} (f : M ⟢ N) : --- whiskerLeft (of R R) f = TensorProduct.lift (LinearMap.complβ‚‚ (TensorProduct.mk _ _ _) f) := by --- sorry - - --- theorem id_whiskerLeft {M N : ModuleCat R} (f : M ⟢ N) : --- whiskerLeft (of R R) f = (leftUnitor M).hom ≫ f ≫ (leftUnitor N).inv := by --- -- rw [id_whiskerLeft_aux] --- apply TensorProduct.ext --- apply LinearMap.ext_ring --- apply LinearMap.ext; intro x --- dsimp --- -- dsimp [LinearMap.comprβ‚‚, whiskerLeft] --- change ((leftUnitor N).hom) ((tensorHom (πŸ™ (of R R)) f) ((1 : R) βŠ—β‚œ[R] x)) = --- f (((leftUnitor M).hom) (1 βŠ—β‚œ[R] x)) --- change _ = _ ≫ f _ ≫ _ --- change ((leftUnitor N).hom) _ = f _ --- erw [TensorProduct.lid_tmul, TensorProduct.lid_tmul] --- rw [LinearMap.map_smul] --- rfl - end MonoidalCategory open MonoidalCategory diff --git a/Mathlib/CategoryTheory/Closed/FunctorCategory.lean b/Mathlib/CategoryTheory/Closed/FunctorCategory.lean index 10114ac1ede7c..8e8f10cbd53ad 100644 --- a/Mathlib/CategoryTheory/Closed/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Closed/FunctorCategory.lean @@ -41,9 +41,9 @@ def closedUnit (F : D β₯€ C) : 𝟭 (D β₯€ C) ⟢ tensorLeft F β‹™ closedIhom F dsimp simp only [ihom.coev_naturality, closedIhom_obj_map, Monoidal.tensorObj_map] dsimp - rw [coev_app_comp_pre_app_assoc, ← Functor.map_comp] + rw [coev_app_comp_pre_app_assoc, ← Functor.map_comp, + tensorHom_def, ← comp_whiskerRight_assoc, IsIso.inv_hom_id] simp - sorry } #align category_theory.functor.closed_unit CategoryTheory.Functor.closedUnit @@ -57,9 +57,10 @@ def closedCounit (F : D β₯€ C) : closedIhom F β‹™ tensorLeft F ⟢ 𝟭 (D β₯€ C intro X Y f dsimp simp only [closedIhom_obj_map, pre_comm_ihom_map] - rw [← tensor_id_comp_id_tensor, id_tensor_comp] - simp - sorry + rw [tensorHom_def] + simp only [NatTrans.naturality, MonoidalCategory.whiskerLeft_comp, Category.assoc, + ihom.ev_naturality, comp_obj, tensorLeft_obj, id_obj, id_tensor_pre_app_comp_ev_assoc] + simp [← comp_whiskerRight_assoc] } #align category_theory.functor.closed_counit CategoryTheory.Functor.closedCounit diff --git a/Mathlib/CategoryTheory/Enriched/Basic.lean b/Mathlib/CategoryTheory/Enriched/Basic.lean index 1a0aea108c016..619fb2f1d6949 100644 --- a/Mathlib/CategoryTheory/Enriched/Basic.lean +++ b/Mathlib/CategoryTheory/Enriched/Basic.lean @@ -133,7 +133,7 @@ instance (F : LaxMonoidalFunctor V W) : EnrichedCategory W (TransportEnrichment F.toFunctor.map_comp, id_tensorHom, tensorHom_id, e_assoc, id_tensor_comp, Category.assoc, ← F.toFunctor.map_id, F.ΞΌ_natural_assoc, F.toFunctor.map_comp] - simp [id_tensorHom, tensorHom_id] + simp end diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index d88cd369d2cfe..85340d91183f6 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -112,7 +112,7 @@ theorem braiding_inv_tensor_right (X Y Z : C) : @[reassoc (attr := simp)] theorem braiding_naturality {X X' Y Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y') : (f βŠ— g) ≫ (braiding Y Y').hom = (braiding X X').hom ≫ (g βŠ— f) := by - rw [tensorHom_def f g, tensorHom_def' g f] + rw [tensorHom_def' f g, tensorHom_def g f] simp_rw [Category.assoc, braiding_naturality_left, braiding_naturality_right_assoc] end BraidedCategory @@ -495,13 +495,13 @@ theorem tensor_ΞΌ_natural {X₁ Xβ‚‚ Y₁ Yβ‚‚ U₁ Uβ‚‚ V₁ Vβ‚‚ : C} (f₁ : theorem tensor_ΞΌ_natural_left {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁: X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (Z₁ Zβ‚‚ : C) : (f₁ βŠ— fβ‚‚) β–· (Z₁ βŠ— Zβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚) = tensor_ΞΌ C (X₁, Xβ‚‚) (Z₁, Zβ‚‚) ≫ (f₁ β–· Z₁ βŠ— fβ‚‚ β–· Zβ‚‚) := by - convert tensor_ΞΌ_natural C f₁ fβ‚‚ (πŸ™ Z₁) (πŸ™ Zβ‚‚) using 1 <;> simp [id_tensorHom, tensorHom_id] + convert tensor_ΞΌ_natural C f₁ fβ‚‚ (πŸ™ Z₁) (πŸ™ Zβ‚‚) using 1 <;> simp @[reassoc] theorem tensor_ΞΌ_natural_right (Z₁ Zβ‚‚ : C) {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) : (Z₁ βŠ— Zβ‚‚) ◁ (f₁ βŠ— fβ‚‚) ≫ tensor_ΞΌ C (Z₁, Zβ‚‚) (Y₁, Yβ‚‚) = tensor_ΞΌ C (Z₁, Zβ‚‚) (X₁, Xβ‚‚) ≫ (Z₁ ◁ f₁ βŠ— Zβ‚‚ ◁ fβ‚‚) := by - convert tensor_ΞΌ_natural C (πŸ™ Z₁) (πŸ™ Zβ‚‚) f₁ fβ‚‚ using 1 <;> simp [id_tensorHom, tensorHom_id] + convert tensor_ΞΌ_natural C (πŸ™ Z₁) (πŸ™ Zβ‚‚) f₁ fβ‚‚ using 1 <;> simp theorem tensor_left_unitality (X₁ Xβ‚‚ : C) : (Ξ»_ (X₁ βŠ— Xβ‚‚)).hom = diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 1aac620ecf8ea..4390757b7e5ea 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -60,14 +60,6 @@ open CategoryTheory.Iso namespace CategoryTheory -def tensorHomDefault {C : Type u} [Category.{v} C] - (tensorObj : C β†’ C β†’ C) - (whiskerLeft : βˆ€ (X : C) {Y₁ Yβ‚‚ : C} (_f : Y₁ ⟢ Yβ‚‚), tensorObj X Y₁ ⟢ tensorObj X Yβ‚‚) - (whiskerRight : βˆ€ {X₁ Xβ‚‚ : C} (_f : X₁ ⟢ Xβ‚‚) (Y : C), tensorObj X₁ Y ⟢ tensorObj Xβ‚‚ Y) - ⦃X₁ Y₁ Xβ‚‚ Yβ‚‚ : C⦄ (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : - (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚) := - whiskerLeft X₁ g ≫ whiskerRight f Yβ‚‚ - /-- In a monoidal category, we can take the tensor product of objects, `X βŠ— Y` and of morphisms `f βŠ— g`. Tensor product does not need to be strictly associative on objects, but there is a @@ -79,13 +71,18 @@ See . -/ -- Porting note: The Mathport did not translate the temporary notation class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] where + /-- curried tensor product of objects -/ tensorObj : C β†’ C β†’ C + /-- left whiskering for morphisms -/ whiskerLeft (X : C) {Y₁ Yβ‚‚ : C} (f : Y₁ ⟢ Yβ‚‚) : tensorObj X Y₁ ⟢ tensorObj X Yβ‚‚ + /-- right whiskering for morphisms -/ whiskerRight {X₁ Xβ‚‚ : C} (f : X₁ ⟢ Xβ‚‚) (Y : C) : tensorObj X₁ Y ⟢ tensorObj Xβ‚‚ Y + /-- Tensor product of identity maps is the identity: `(πŸ™ X₁ βŠ— πŸ™ Xβ‚‚) = πŸ™ (X₁ βŠ— Xβ‚‚)` -/ + -- By default, it is defined in terms of whiskerings. tensorHom {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚) := - whiskerLeft X₁ g ≫ whiskerRight f Yβ‚‚ + whiskerRight f Xβ‚‚ ≫ whiskerLeft Y₁ g tensorHom_def {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : - tensorHom f g = whiskerLeft X₁ g ≫ whiskerRight f Yβ‚‚ := by + tensorHom f g = whiskerRight f Xβ‚‚ ≫ whiskerLeft Y₁ g := by aesop_cat -- Porting note: Adding a prime here, so I can later define `tensorUnit` unprimed with explicit -- argument `C` @@ -132,16 +129,23 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] where whiskerRight (whiskerLeft X f) Z = (associator X Y Z).hom ≫ whiskerLeft X (whiskerRight f Z) ≫ (associator X Y' Z).inv := by aesop_cat + /-- The exchnage identity for the left and right whiskerings -/ whisker_exchange : βˆ€ {W X Y Z : C} (f : W ⟢ X) (g : Y ⟢ Z), whiskerLeft W g ≫ whiskerRight f Z = whiskerRight f Y ≫ whiskerLeft X g := by aesop_cat + /-- + The pentagon identity relating the isomorphism between `X βŠ— (Y βŠ— (Z βŠ— W))` and `((X βŠ— Y) βŠ— Z) βŠ— W` + -/ pentagon : βˆ€ W X Y Z : C, whiskerRight (associator W X Y).hom Z ≫ (associator W (tensorObj X Y) Z).hom ≫ whiskerLeft W (associator X Y Z).hom = (associator (tensorObj W X) Y Z).hom ≫ (associator W X (tensorObj Y Z)).hom := by aesop_cat + /-- + The identity relating the isomorphisms between `X βŠ— (πŸ™_C βŠ— Y)`, `(X βŠ— πŸ™_C) βŠ— Y` and `X βŠ— Y` + -/ triangle : βˆ€ X Y : C, (associator X tensorUnit' Y).hom ≫ whiskerLeft X (leftUnitor Y).hom = @@ -156,7 +160,7 @@ namespace MonoidalCategory attribute [reassoc] whiskerLeft_comp id_whiskerLeft tensor_whiskerLeft comp_whiskerRight whiskerRight_id - whiskerRight_tensor whisker_assoc whisker_exchange + whiskerRight_tensor whisker_assoc whisker_exchange tensorHom_def attribute [simp] whiskerLeft_id whiskerRight_id @@ -197,38 +201,34 @@ variable {C : Type u} [π’ž : Category.{v} C] [MonoidalCategory C] /-- Notation for `tensorHom`, the tensor product of morphisms in a monoidal category -/ scoped infixr:70 " βŠ— " => MonoidalCategory.tensorHom -attribute [reassoc] tensorHom_def -attribute [local simp] tensorHom_def - @[reassoc] theorem tensorHom_def' {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : - f βŠ— g = f β–· Xβ‚‚ ≫ Y₁ ◁ g := by - rw [← whisker_exchange] - apply tensorHom_def - + f βŠ— g = X₁ ◁ g ≫ f β–· Yβ‚‚ := + whisker_exchange f g β–Έ tensorHom_def f g /-- Tensor product of identity maps is the identity: `(πŸ™ X₁ βŠ— πŸ™ Xβ‚‚) = πŸ™ (X₁ βŠ— Xβ‚‚)` -/ -theorem tensor_id (X₁ Xβ‚‚ : C) : tensorHom (πŸ™ X₁) (πŸ™ Xβ‚‚) = πŸ™ (tensorObj X₁ Xβ‚‚) := by simp +theorem tensor_id (X₁ Xβ‚‚ : C) : tensorHom (πŸ™ X₁) (πŸ™ Xβ‚‚) = πŸ™ (tensorObj X₁ Xβ‚‚) := by + simp [tensorHom_def] /-- Composition of tensor products is tensor product of compositions: -`(f₁ βŠ— g₁) ∘ (fβ‚‚ βŠ— gβ‚‚) = (f₁ ∘ fβ‚‚) βŠ— (g₁ βŠ— gβ‚‚)` +`(f₁ ≫ g₁) βŠ— (fβ‚‚ ≫ gβ‚‚) = (f₁ βŠ— fβ‚‚) ≫ (g₁ βŠ— gβ‚‚)` -/ @[reassoc, simp] theorem tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : Y₁ ⟢ Z₁) (gβ‚‚ : Yβ‚‚ ⟢ Zβ‚‚) : (f₁ ≫ g₁) βŠ— (fβ‚‚ ≫ gβ‚‚) = (f₁ βŠ— fβ‚‚) ≫ (g₁ βŠ— gβ‚‚) := by - simp [whisker_exchange_assoc] + simp [tensorHom_def, whisker_exchange_assoc] @[simp] theorem id_tensorHom (X : C) {Y₁ Yβ‚‚ : C} (f : Y₁ ⟢ Yβ‚‚) : (πŸ™ X) βŠ— f = X ◁ f := by - simp + simp [tensorHom_def] @[simp] theorem tensorHom_id {X₁ Xβ‚‚ : C} (f : X₁ ⟢ Xβ‚‚) (Y : C) : f βŠ— (πŸ™ Y) = f β–· Y := by - simp + simp [tensorHom_def] end MonoidalCategory @@ -255,6 +255,26 @@ theorem inv_hom_whiskerLeft (X : C) {Y Z : C} (f : Y β‰… Z) : theorem inv_hom_whiskerRight {X Y : C} (f : X β‰… Y) (Z : C) : f.inv β–· Z ≫ f.hom β–· Z = πŸ™ (Y βŠ— Z) := by rw [← comp_whiskerRight, inv_hom_id, id_whiskerRight] +@[reassoc (attr := simp)] +theorem hom_inv_whiskerLeft' (X : C) {Y Z : C} (f : Y ⟢ Z) [IsIso f] : + X ◁ f ≫ X ◁ inv f = πŸ™ (X βŠ— Y) := by + rw [← whiskerLeft_comp, IsIso.hom_inv_id, whiskerLeft_id] + +@[reassoc (attr := simp)] +theorem hom_inv_whiskerRight' {X Y : C} (f : X ⟢ Y) [IsIso f] (Z : C) : + f β–· Z ≫ inv f β–· Z = πŸ™ (X βŠ— Z) := by + rw [← comp_whiskerRight, IsIso.hom_inv_id, id_whiskerRight] + +@[reassoc (attr := simp)] +theorem inv_hom_whiskerLeft' (X : C) {Y Z : C} (f : Y ⟢ Z) [IsIso f] : + X ◁ inv f ≫ X ◁ f = πŸ™ (X βŠ— Z) := by + rw [← whiskerLeft_comp, IsIso.inv_hom_id, whiskerLeft_id] + +@[reassoc (attr := simp)] +theorem inv_hom_whiskerRight' {X Y : C} (f : X ⟢ Y) [IsIso f] (Z : C) : + inv f β–· Z ≫ f β–· Z = πŸ™ (Y βŠ— Z) := by + rw [← comp_whiskerRight, IsIso.inv_hom_id, id_whiskerRight] + /-- The left whiskering of a 2-isomorphism is a 2-isomorphism. -/ @[simps] def whiskerLeftIso (X : C) {Y Z : C} (f : Y β‰… Z) : X βŠ— Y β‰… X βŠ— Z where @@ -266,9 +286,8 @@ instance whiskerLeft_isIso (X : C) {Y Z : C} (f : Y ⟢ Z) [IsIso f] : IsIso (X @[simp] theorem inv_whiskerLeft (X : C) {Y Z : C} (f : Y ⟢ Z) [IsIso f] : - inv (X ◁ f) = X ◁ inv f := by - aesop_cat_nonterminal - simp only [← whiskerLeft_comp, whiskerLeft_id, IsIso.hom_inv_id] + inv (X ◁ f) = X ◁ inv f := by + aesop_cat /-- The right whiskering of a 2-isomorphism is a 2-isomorphism. -/ @[simps!] @@ -282,8 +301,7 @@ instance whiskerRight_isIso {X Y : C} (f : X ⟢ Y) (Z : C) [IsIso f] : IsIso (f @[simp] theorem inv_whiskerRight {X Y : C} (f : X ⟢ Y) (Z : C) [IsIso f] : inv (f β–· Z) = inv f β–· Z := by - aesop_cat_nonterminal - simp only [← comp_whiskerRight, id_whiskerRight, IsIso.hom_inv_id] + aesop_cat end MonoidalCategory @@ -306,8 +324,6 @@ section variable {C : Type u} [Category.{v} C] [MonoidalCategory.{v} C] -attribute [local simp] tensorHom_def - instance tensor_isIso {W X Y Z : C} (f : W ⟢ X) [IsIso f] (g : Y ⟢ Z) [IsIso g] : IsIso (f βŠ— g) := IsIso.of_iso (asIso f βŠ— asIso g) #align category_theory.monoidal_category.tensor_is_iso CategoryTheory.MonoidalCategory.tensor_isIso @@ -315,9 +331,7 @@ instance tensor_isIso {W X Y Z : C} (f : W ⟢ X) [IsIso f] (g : Y ⟢ Z) [IsIso @[simp] theorem inv_tensor {W X Y Z : C} (f : W ⟢ X) [IsIso f] (g : Y ⟢ Z) [IsIso g] : inv (f βŠ— g) = inv f βŠ— inv g := by - -- Porting note: Replaced `ext` with `aesop_cat_nonterminal` - aesop_cat_nonterminal - simp [whisker_exchange] + simp [tensorHom_def ,whisker_exchange] #align category_theory.monoidal_category.inv_tensor CategoryTheory.MonoidalCategory.inv_tensor section pentagon @@ -401,41 +415,28 @@ theorem dite_tensor {P : Prop} [Decidable P] {W X Y Z : C} (f : W ⟢ X) (g : P by split_ifs <;> rfl #align category_theory.monoidal_category.dite_tensor CategoryTheory.MonoidalCategory.dite_tensor -@[reassoc, simp] +@[reassoc] theorem comp_tensor_id (f : W ⟢ X) (g : X ⟢ Y) : f ≫ g βŠ— πŸ™ Z = (f βŠ— πŸ™ Z) ≫ (g βŠ— πŸ™ Z) := by - rw [← tensor_comp] simp #align category_theory.monoidal_category.comp_tensor_id CategoryTheory.MonoidalCategory.comp_tensor_id -@[reassoc, simp] +@[reassoc] theorem id_tensor_comp (f : W ⟢ X) (g : X ⟢ Y) : πŸ™ Z βŠ— f ≫ g = (πŸ™ Z βŠ— f) ≫ (πŸ™ Z βŠ— g) := by - rw [← tensor_comp] simp #align category_theory.monoidal_category.id_tensor_comp CategoryTheory.MonoidalCategory.id_tensor_comp -@[reassoc (attr := simp)] +@[reassoc] theorem id_tensor_comp_tensor_id (f : W ⟢ X) (g : Y ⟢ Z) : (πŸ™ Y βŠ— f) ≫ (g βŠ— πŸ™ X) = g βŠ— f := by rw [← tensor_comp] simp #align category_theory.monoidal_category.id_tensor_comp_tensor_id CategoryTheory.MonoidalCategory.id_tensor_comp_tensor_id -@[reassoc (attr := simp)] +@[reassoc] theorem tensor_id_comp_id_tensor (f : W ⟢ X) (g : Y ⟢ Z) : (g βŠ— πŸ™ W) ≫ (πŸ™ Z βŠ— f) = g βŠ— f := by rw [← tensor_comp] simp #align category_theory.monoidal_category.tensor_id_comp_id_tensor CategoryTheory.MonoidalCategory.tensor_id_comp_id_tensor --- @[simp] --- theorem rightUnitor_conjugation {X Y : C} (f : X ⟢ Y) : --- f βŠ— πŸ™ (πŸ™_ C) = (ρ_ X).hom ≫ f ≫ (ρ_ Y).inv := by --- rw [← rightUnitor_naturality_assoc, Iso.hom_inv_id, Category.comp_id] --- #align category_theory.monoidal_category.right_unitor_conjugation CategoryTheory.MonoidalCategory.rightUnitor_conjugation - --- @[simp] --- theorem leftUnitor_conjugation {X Y : C} (f : X ⟢ Y) : πŸ™ (πŸ™_ C) βŠ— f = (Ξ»_ X).hom ≫ f ≫ (Ξ»_ Y).inv --- := by rw [← leftUnitor_naturality_assoc, Iso.hom_inv_id, Category.comp_id] --- #align category_theory.monoidal_category.left_unitor_conjugation CategoryTheory.MonoidalCategory.leftUnitor_conjugation - theorem tensor_left_iff {X Y : C} (f g : X ⟢ Y) : πŸ™ (πŸ™_ C) βŠ— f = πŸ™ (πŸ™_ C) βŠ— g ↔ f = g := by simp #align category_theory.monoidal_category.tensor_left_iff CategoryTheory.MonoidalCategory.tensor_left_iff @@ -577,9 +578,8 @@ theorem rightUnitor_tensor_inv (X Y : C) : @[reassoc] theorem associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (f₃ : X₃ ⟢ Y₃) : - tensorHom (tensorHom f₁ fβ‚‚) f₃ ≫ (associator Y₁ Yβ‚‚ Y₃).hom = - (associator X₁ Xβ‚‚ X₃).hom ≫ tensorHom f₁ (tensorHom fβ‚‚ f₃) := by - simp + ((f₁ βŠ— fβ‚‚) βŠ— f₃) ≫ (Ξ±_ Y₁ Yβ‚‚ Y₃).hom = (Ξ±_ X₁ Xβ‚‚ X₃).hom ≫ (f₁ βŠ— (fβ‚‚ βŠ— f₃)) := by + simp [tensorHom_def] @[reassoc] theorem triangle' (X Y : C) : @@ -655,7 +655,7 @@ end @[reassoc] theorem associator_inv_naturality {X Y Z X' Y' Z' : C} (f : X ⟢ X') (g : Y ⟢ Y') (h : Z ⟢ Z') : (f βŠ— g βŠ— h) ≫ (Ξ±_ X' Y' Z').inv = (Ξ±_ X Y Z).inv ≫ ((f βŠ— g) βŠ— h) := by - simp + simp [tensorHom_def] #align category_theory.monoidal_category.associator_inv_naturality CategoryTheory.MonoidalCategory.associator_inv_naturality @[reassoc, simp] @@ -670,66 +670,53 @@ theorem associator_inv_conjugation {X X' Y Y' Z Z' : C} (f : X ⟢ X') (g : Y rw [associator_naturality, inv_hom_id_assoc] #align category_theory.monoidal_category.associator_inv_conjugation CategoryTheory.MonoidalCategory.associator_inv_conjugation --- TODO these next two lemmas aren't so fundamental, and perhaps could be removed --- (replacing their usages by their proofs). -@[reassoc] -theorem id_tensor_associator_naturality {X Y Z Z' : C} (h : Z ⟢ Z') : - (πŸ™ (X βŠ— Y) βŠ— h) ≫ (Ξ±_ X Y Z').hom = (Ξ±_ X Y Z).hom ≫ (πŸ™ X βŠ— πŸ™ Y βŠ— h) := by - rw [← tensor_id, associator_naturality] -#align category_theory.monoidal_category.id_tensor_associator_naturality CategoryTheory.MonoidalCategory.id_tensor_associator_naturality - -@[reassoc] -theorem id_tensor_associator_inv_naturality {X Y Z X' : C} (f : X ⟢ X') : - (f βŠ— πŸ™ (Y βŠ— Z)) ≫ (Ξ±_ X' Y Z).inv = (Ξ±_ X Y Z).inv ≫ ((f βŠ— πŸ™ Y) βŠ— πŸ™ Z) := by - rw [← tensor_id, associator_inv_naturality] -#align category_theory.monoidal_category.id_tensor_associator_inv_naturality CategoryTheory.MonoidalCategory.id_tensor_associator_inv_naturality - @[reassoc (attr := simp)] theorem hom_inv_id_tensor {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : - (f.hom βŠ— g) ≫ (f.inv βŠ— h) = (πŸ™ V βŠ— g) ≫ (πŸ™ V βŠ— h) := by - rw [← tensor_comp, f.hom_inv_id, id_tensor_comp] + (f.hom βŠ— g) ≫ (f.inv βŠ— h) = (V ◁ g) ≫ (V ◁ h) := by + rw [← tensor_comp, f.hom_inv_id, id_tensorHom, whiskerLeft_comp] + #align category_theory.monoidal_category.hom_inv_id_tensor CategoryTheory.MonoidalCategory.hom_inv_id_tensor @[reassoc (attr := simp)] theorem inv_hom_id_tensor {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : - (f.inv βŠ— g) ≫ (f.hom βŠ— h) = (πŸ™ W βŠ— g) ≫ (πŸ™ W βŠ— h) := by - rw [← tensor_comp, f.inv_hom_id, id_tensor_comp] + (f.inv βŠ— g) ≫ (f.hom βŠ— h) = (W ◁ g) ≫ (W ◁ h) := by + rw [← tensor_comp, f.inv_hom_id, id_tensorHom, whiskerLeft_comp] #align category_theory.monoidal_category.inv_hom_id_tensor CategoryTheory.MonoidalCategory.inv_hom_id_tensor @[reassoc (attr := simp)] theorem tensorHom_inv_id {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : - (g βŠ— f.hom) ≫ (h βŠ— f.inv) = (g βŠ— πŸ™ V) ≫ (h βŠ— πŸ™ V) := by - rw [← tensor_comp, f.hom_inv_id, comp_tensor_id] + (g βŠ— f.hom) ≫ (h βŠ— f.inv) = (g β–· V) ≫ (h β–· V) := by + rw [← tensor_comp, f.hom_inv_id, tensorHom_id, comp_whiskerRight] #align category_theory.monoidal_category.tensor_hom_inv_id CategoryTheory.MonoidalCategory.tensorHom_inv_id @[reassoc (attr := simp)] theorem tensor_inv_hom_id {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : - (g βŠ— f.inv) ≫ (h βŠ— f.hom) = (g βŠ— πŸ™ W) ≫ (h βŠ— πŸ™ W) := by - rw [← tensor_comp, f.inv_hom_id, comp_tensor_id] + (g βŠ— f.inv) ≫ (h βŠ— f.hom) = (g β–· W) ≫ (h β–· W) := by + rw [← tensor_comp, f.inv_hom_id, tensorHom_id, comp_whiskerRight] #align category_theory.monoidal_category.tensor_inv_hom_id CategoryTheory.MonoidalCategory.tensor_inv_hom_id @[reassoc (attr := simp)] theorem hom_inv_id_tensor' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : - (f βŠ— g) ≫ (inv f βŠ— h) = (πŸ™ V βŠ— g) ≫ (πŸ™ V βŠ— h) := by - rw [← tensor_comp, IsIso.hom_inv_id, id_tensor_comp] + (f βŠ— g) ≫ (inv f βŠ— h) = (V ◁ g) ≫ (V ◁ h) := by + rw [← tensor_comp, IsIso.hom_inv_id, id_tensorHom, whiskerLeft_comp] #align category_theory.monoidal_category.hom_inv_id_tensor' CategoryTheory.MonoidalCategory.hom_inv_id_tensor' @[reassoc (attr := simp)] theorem inv_hom_id_tensor' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : - (inv f βŠ— g) ≫ (f βŠ— h) = (πŸ™ W βŠ— g) ≫ (πŸ™ W βŠ— h) := by - rw [← tensor_comp, IsIso.inv_hom_id, id_tensor_comp] + (inv f βŠ— g) ≫ (f βŠ— h) = (W ◁ g) ≫ (W ◁ h) := by + rw [← tensor_comp, IsIso.inv_hom_id, id_tensorHom, whiskerLeft_comp] #align category_theory.monoidal_category.inv_hom_id_tensor' CategoryTheory.MonoidalCategory.inv_hom_id_tensor' @[reassoc (attr := simp)] theorem tensorHom_inv_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : - (g βŠ— f) ≫ (h βŠ— inv f) = (g βŠ— πŸ™ V) ≫ (h βŠ— πŸ™ V) := by - rw [← tensor_comp, IsIso.hom_inv_id, comp_tensor_id] + (g βŠ— f) ≫ (h βŠ— inv f) = (g β–· V) ≫ (h β–· V) := by + rw [← tensor_comp, IsIso.hom_inv_id, tensorHom_id, comp_whiskerRight] #align category_theory.monoidal_category.tensor_hom_inv_id' CategoryTheory.MonoidalCategory.tensorHom_inv_id' @[reassoc (attr := simp)] theorem tensor_inv_hom_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : - (g βŠ— inv f) ≫ (h βŠ— f) = (g βŠ— πŸ™ W) ≫ (h βŠ— πŸ™ W) := by - rw [← tensor_comp, IsIso.inv_hom_id, comp_tensor_id] + (g βŠ— inv f) ≫ (h βŠ— f) = (g β–· W) ≫ (h β–· W) := by + rw [← tensor_comp, IsIso.inv_hom_id, tensorHom_id, comp_whiskerRight] #align category_theory.monoidal_category.tensor_inv_hom_id' CategoryTheory.MonoidalCategory.tensor_inv_hom_id' def ofTensorHom diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index a4a407aa88245..6c59ab0b8e963 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -164,52 +164,63 @@ def tensorObj (X Y : Center C) : Center C := case eq3 => rw [HalfBraiding.naturality]; coherence }⟩ #align category_theory.center.tensor_obj CategoryTheory.Center.tensorObj -/-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ -@[simps] -def tensorHom {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Center C} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : - tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚ where - f := f.f βŠ— g.f - comm U := by - sorry -#align category_theory.center.tensor_hom CategoryTheory.Center.tensorHom +@[reassoc] +theorem whiskerLeft_comm (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) (U : C) : + (X.1 ◁ f.f) β–· U ≫ (HalfBraiding.Ξ² (tensorObj X Yβ‚‚).2 U).hom = + (HalfBraiding.Ξ² (tensorObj X Y₁).2 U).hom ≫ U ◁ X.1 ◁ f.f := by + dsimp only [tensorObj_fst, tensorObj_snd_Ξ², Iso.trans_hom, whiskerLeftIso_hom, + Iso.symm_hom, whiskerRightIso_hom] + calc + _ = πŸ™ _ βŠ—β‰« + X.fst ◁ (f.f β–· U ≫ (HalfBraiding.Ξ² Yβ‚‚.snd U).hom) βŠ—β‰« + (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst βŠ—β‰« πŸ™ _ := ?eq1 + _ = πŸ™ _ βŠ—β‰« + X.fst ◁ (HalfBraiding.Ξ² Y₁.snd U).hom βŠ—β‰« + ((X.fst βŠ— U) ◁ f.f ≫ (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst) βŠ—β‰« πŸ™ _ := ?eq2 + _ = _ := ?eq3 + case eq1 => coherence + case eq2 => rw [f.comm]; coherence + case eq3 => rw [whisker_exchange]; coherence /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ def whiskerLeft (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) : tensorObj X Y₁ ⟢ tensorObj X Yβ‚‚ where f := X.1 ◁ f.f - comm U := by - dsimp only [tensorObj_fst, tensorObj_snd_Ξ², Iso.trans_hom, whiskerLeftIso_hom, - Iso.symm_hom, whiskerRightIso_hom] - calc - _ = πŸ™ _ βŠ—β‰« - X.fst ◁ (f.f β–· U ≫ (HalfBraiding.Ξ² Yβ‚‚.snd U).hom) βŠ—β‰« - (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst βŠ—β‰« πŸ™ _ := ?eq1 - _ = πŸ™ _ βŠ—β‰« - X.fst ◁ (HalfBraiding.Ξ² Y₁.snd U).hom βŠ—β‰« - ((X.fst βŠ— U) ◁ f.f ≫ (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst) βŠ—β‰« πŸ™ _ := ?eq2 - _ = _ := ?eq3 - case eq1 => coherence - case eq2 => rw [f.comm]; coherence - case eq3 => rw [whisker_exchange]; coherence + comm U := whiskerLeft_comm X f U + +@[reassoc] +theorem whiskerRight_comm {X₁ Xβ‚‚: Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) (U : C) : + f.f β–· Y.1 β–· U ≫ (HalfBraiding.Ξ² (tensorObj Xβ‚‚ Y).2 U).hom = + (HalfBraiding.Ξ² (tensorObj X₁ Y).snd U).hom ≫ U ◁ f.f β–· Y.1 := by + dsimp only [tensorObj_fst, tensorObj_snd_Ξ², Iso.trans_hom, whiskerLeftIso_hom, + Iso.symm_hom, whiskerRightIso_hom] + calc + _ = πŸ™ _ βŠ—β‰« + (f.f β–· (Y.fst βŠ— U) ≫ Xβ‚‚.fst ◁ (HalfBraiding.Ξ² Y.snd U).hom) βŠ—β‰« + (HalfBraiding.Ξ² Xβ‚‚.snd U).hom β–· Y.fst βŠ—β‰« πŸ™ _ := ?eq1 + _ = πŸ™ _ βŠ—β‰« + X₁.fst ◁ (HalfBraiding.Ξ² Y.snd U).hom βŠ—β‰« + (f.f β–· U ≫ (HalfBraiding.Ξ² Xβ‚‚.snd U).hom) β–· Y.fst βŠ—β‰« πŸ™ _ := ?eq2 + _ = _ := ?eq3 + case eq1 => coherence + case eq2 => rw [← whisker_exchange]; coherence + case eq3 => rw [f.comm]; coherence /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ def whiskerRight {X₁ Xβ‚‚ : Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) : tensorObj X₁ Y ⟢ tensorObj Xβ‚‚ Y where f := f.f β–· Y.1 + comm U := whiskerRight_comm f Y U + +/-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ +@[simps] +def tensorHom {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Center C} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : + tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚ where + f := f.f βŠ— g.f comm U := by - dsimp only [tensorObj_fst, tensorObj_snd_Ξ², Iso.trans_hom, whiskerLeftIso_hom, - Iso.symm_hom, whiskerRightIso_hom] - calc - _ = πŸ™ _ βŠ—β‰« - (f.f β–· (Y.fst βŠ— U) ≫ Xβ‚‚.fst ◁ (HalfBraiding.Ξ² Y.snd U).hom) βŠ—β‰« - (HalfBraiding.Ξ² Xβ‚‚.snd U).hom β–· Y.fst βŠ—β‰« πŸ™ _ := ?eq1 - _ = πŸ™ _ βŠ—β‰« - X₁.fst ◁ (HalfBraiding.Ξ² Y.snd U).hom βŠ—β‰« - (f.f β–· U ≫ (HalfBraiding.Ξ² Xβ‚‚.snd U).hom) β–· Y.fst βŠ—β‰« πŸ™ _ := ?eq2 - _ = _ := ?eq3 - case eq1 => coherence - case eq2 => rw [← whisker_exchange]; coherence - case eq3 => rw [f.comm]; coherence + rw [tensorHom_def, comp_whiskerRight_assoc, whiskerLeft_comm, whiskerRight_comm_assoc, + MonoidalCategory.whiskerLeft_comp] +#align category_theory.center.tensor_hom CategoryTheory.Center.tensorHom /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ @[simps] @@ -238,10 +249,12 @@ attribute [local simp] associator_naturality leftUnitor_naturality rightUnitor_n attribute [local simp] Center.associator Center.leftUnitor Center.rightUnitor -attribute [local simp] Center.whiskerLeft Center.whiskerRight +attribute [local simp] Center.whiskerLeft Center.whiskerRight Center.tensorHom instance : MonoidalCategory (Center C) where tensorObj X Y := tensorObj X Y + tensorHom f g := tensorHom f g + tensorHom_def := by intros; ext; simp [tensorHom_def] whiskerLeft X _ _ f := whiskerLeft X f whiskerRight f X := whiskerRight f X tensorUnit' := tensorUnit @@ -265,18 +278,16 @@ theorem tensor_Ξ² (X Y : Center C) (U : C) : #align category_theory.center.tensor_Ξ² CategoryTheory.Center.tensor_Ξ² @[simp] -theorem whiskerLeft_f (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) : - (X ◁ f).f = X.1 ◁ f.f:= by +theorem whiskerLeft_f (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) : (X ◁ f).f = X.1 ◁ f.f:= rfl @[simp] -theorem whiskerRight_f {X₁ Xβ‚‚ : Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) : - (f β–· Y).f = f.f β–· Y.1 := by +theorem whiskerRight_f {X₁ Xβ‚‚ : Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) : (f β–· Y).f = f.f β–· Y.1 := rfl @[simp] -theorem tensor_f {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Center C} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : (f βŠ— g).f = f.f βŠ— g.f := by - rw [tensorHom_def, tensorHom_def, comp_f, whiskerLeft_f, whiskerRight_f] +theorem tensor_f {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Center C} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : (f βŠ— g).f = f.f βŠ— g.f := + rfl #align category_theory.center.tensor_f CategoryTheory.Center.tensor_f @[simp] diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index f9e3994608c99..4280b1eefc1a4 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -115,13 +115,11 @@ section variable {C D} -attribute [local simp] tensorHom_def - @[reassoc (attr := simp)] theorem LaxMonoidalFunctor.ΞΌ_natural (F : LaxMonoidalFunctor C D) {X Y X' Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y') : (F.map f βŠ— F.map g) ≫ F.ΞΌ Y Y' = F.ΞΌ X X' ≫ F.map (f βŠ— g) := by - simp only [tensorHom_def, assoc, F.ΞΌ_natural_left, F.ΞΌ_natural_right_assoc, map_comp] + simp [tensorHom_def] @[reassoc] theorem LaxMonoidalFunctor.associativity' (F : LaxMonoidalFunctor C D) (X Y Z : C) : diff --git a/Mathlib/CategoryTheory/Monoidal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Limits.lean index c26c58e3b3a8e..0a26fb47f4f51 100644 --- a/Mathlib/CategoryTheory/Monoidal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Limits.lean @@ -69,25 +69,25 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorH ext j; dsimp simp only [limit.lift_Ο€, Cones.postcompose_obj_Ο€, Monoidal.associator_hom_app, limit.lift_map, NatTrans.comp_app, Category.assoc, tensorHom_id, id_tensorHom] - slice_lhs 2 2 => rw [tensorHom_def'] + slice_lhs 2 2 => rw [tensorHom_def] slice_lhs 1 2 => rw [← comp_whiskerRight, limit.lift_Ο€] dsimp - rw [tensorHom_def] + rw [tensorHom_def'] slice_lhs 1 2 => rw [← whisker_exchange] simp only [tensor_whiskerLeft, comp_whiskerRight, whisker_assoc, Category.assoc, Iso.inv_hom_id_assoc, Iso.cancel_iso_hom_left] slice_lhs 4 5 => rw [associator_naturality_left] - conv_rhs => rw [tensorHom_def _ (limit.Ο€ (Y βŠ— Z) j)] + conv_rhs => rw [tensorHom_def' _ (limit.Ο€ (Y βŠ— Z) j)] slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp, limit.lift_Ο€] dsimp - rw [tensorHom_def] + rw [tensorHom_def'] simp) (left_unitality := fun X => by ext j; dsimp simp - conv_rhs => rw [tensorHom_def' _ (limit.Ο€ X j)] + conv_rhs => rw [tensorHom_def _ (limit.Ο€ X j)] slice_rhs 1 2 => rw [← comp_whiskerRight] erw [limit.lift_Ο€] @@ -97,7 +97,7 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorH (right_unitality := fun X => by ext j; dsimp simp - conv_rhs => rw [tensorHom_def (limit.Ο€ X j)] + conv_rhs => rw [tensorHom_def' (limit.Ο€ X j)] slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp] erw [limit.lift_Ο€] diff --git a/Mathlib/CategoryTheory/Monoidal/Mod_.lean b/Mathlib/CategoryTheory/Monoidal/Mod_.lean index 59361b04c246c..676dac683e395 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mod_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mod_.lean @@ -134,7 +134,7 @@ def comap {A B : Mon_ C} (f : A ⟢ B) : Mod_ B β₯€ Mod_ A where slice_rhs 3 4 => rw [associator_inv_naturality_middle] slice_rhs 2 4 => rw [Iso.hom_inv_id_assoc] slice_rhs 1 2 => rw [← MonoidalCategory.comp_whiskerRight, ← whisker_exchange] - slice_rhs 1 2 => rw [← MonoidalCategory.comp_whiskerRight, ← tensorHom_def, ← f.mul_hom] + slice_rhs 1 2 => rw [← MonoidalCategory.comp_whiskerRight, ← tensorHom_def', ← f.mul_hom] rw [comp_whiskerRight, Category.assoc] } map g := { hom := g.hom diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index 0ba5122dc6158..17e5dc4494369 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -74,12 +74,12 @@ variable {M : Mon_ C} @[simp] theorem one_mul_hom {Z : C} (f : Z ⟢ M.X) : (M.one βŠ— f) ≫ M.mul = (Ξ»_ Z).hom ≫ f := by - rw [tensorHom_def_assoc, M.one_mul, leftUnitor_naturality] + rw [tensorHom_def'_assoc, M.one_mul, leftUnitor_naturality] #align Mon_.one_mul_hom Mon_.one_mul_hom @[simp] theorem mul_one_hom {Z : C} (f : Z ⟢ M.X) : (f βŠ— M.one) ≫ M.mul = (ρ_ Z).hom ≫ f := by - rw [tensorHom_def'_assoc, M.mul_one, rightUnitor_naturality] + rw [tensorHom_def_assoc, M.mul_one, rightUnitor_naturality] #align Mon_.mul_one_hom Mon_.mul_one_hom theorem assoc_flip : diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index fec02a0b4d3bf..c2cc79a1d1d95 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -31,8 +31,6 @@ namespace CategoryTheory open MonoidalCategory -attribute [local simp] tensorHom_def - variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] {D : Type uβ‚‚} [Category.{vβ‚‚} D] [MonoidalCategory.{vβ‚‚} D] @@ -120,12 +118,12 @@ def hcomp {F G : LaxMonoidalFunctor C D} {H K : LaxMonoidalFunctor D E} (Ξ± : Mo tensor := fun X Y => by simp only [LaxMonoidalFunctor.comp_toFunctor, comp_obj, LaxMonoidalFunctor.comp_ΞΌ, NatTrans.hcomp_app, assoc, - NatTrans.naturality_assoc, tensor_assoc, tensorHom_def, + NatTrans.naturality_assoc, tensor_assoc, tensorHom_def', whisker_exchange_assoc, MonoidalCategory.whiskerLeft_comp, comp_whiskerRight, LaxMonoidalFunctor.ΞΌ_natural_left_assoc, LaxMonoidalFunctor.ΞΌ_natural_right_assoc] simp_rw [← K.toFunctor.map_comp] - simp only [tensor, tensorHom_def, assoc, map_comp] + simp only [tensor, tensorHom_def', assoc, map_comp] } #align category_theory.monoidal_nat_trans.hcomp CategoryTheory.MonoidalNatTrans.hcomp @@ -205,7 +203,7 @@ def monoidalUnit (F : MonoidalFunctor C D) [IsEquivalence F.toFunctor] : simp only [comp_obj, asEquivalence_functor, asEquivalence_inverse, id_obj, LaxMonoidalFunctor.ΞΌ_natural_left, LaxMonoidalFunctor.ΞΌ_natural_right_assoc, IsIso.inv_hom_id_assoc, map_comp, IsEquivalence.inv_fun_map, assoc, - Iso.hom_inv_id_app_assoc, tensorHom_def] + Iso.hom_inv_id_app_assoc, tensorHom_def'] slice_rhs 3 4 => erw [Iso.hom_inv_id_app] simp only [id_obj, id_comp, assoc, whisker_exchange_assoc] simp only [← whiskerLeft_comp_assoc, ← comp_whiskerRight_assoc] diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index c6e39ebe267c0..e852d3c4df078 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -85,8 +85,6 @@ theorem monoidalPreadditive_of_faithful {D} [Category D] [Preadditive D] [Monoid open BigOperators -attribute [local simp] tensorHom_def - theorem whiskerLeft_sum (P : C) {Q R : C} {J : Type _} (s : Finset J) (g : J β†’ (Q ⟢ R)) : P ◁ βˆ‘ j in s, g j = βˆ‘ j in s, P ◁ g j := map_sum ((tensoringLeft C).obj P).mapAddHom g s @@ -97,12 +95,12 @@ theorem sum_whiskerRight {Q R : C} {J : Type _} (s : Finset J) (g : J β†’ (Q ⟢ theorem tensor_sum {P Q R S : C} {J : Type _} (s : Finset J) (f : P ⟢ Q) (g : J β†’ (R ⟢ S)) : (f βŠ— βˆ‘ j in s, g j) = βˆ‘ j in s, f βŠ— g j := by - simp [whiskerLeft_sum, sum_whiskerRight, Preadditive.sum_comp] + simp [whiskerLeft_sum, sum_whiskerRight, Preadditive.comp_sum, tensorHom_def] #align category_theory.tensor_sum CategoryTheory.tensor_sum theorem sum_tensor {P Q R S : C} {J : Type _} (s : Finset J) (f : P ⟢ Q) (g : J β†’ (R ⟢ S)) : (βˆ‘ j in s, g j) βŠ— f = βˆ‘ j in s, g j βŠ— f := by - simp [whiskerLeft_sum, sum_whiskerRight, Preadditive.comp_sum] + simp [whiskerLeft_sum, sum_whiskerRight, Preadditive.sum_comp, tensorHom_def] #align category_theory.sum_tensor CategoryTheory.sum_tensor -- In a closed monoidal category, this would hold because diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean index 3de03843b95d2..149c21c566a12 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean @@ -445,7 +445,7 @@ theorem tensorLeftHomEquiv_tensor {X X' Y Y' Z Z' : C} [ExactPairing Y Y'] (f : (g : X' ⟢ Z') : (tensorLeftHomEquiv (X βŠ— X') Y Y' (Z βŠ— Z')).symm ((f βŠ— g) ≫ (Ξ±_ _ _ _).hom) = (Ξ±_ _ _ _).inv ≫ ((tensorLeftHomEquiv X Y Y' Z).symm f βŠ— g) := by - simp [tensorLeftHomEquiv, tensorHom_def] + simp [tensorLeftHomEquiv, tensorHom_def'] #align category_theory.tensor_left_hom_equiv_tensor CategoryTheory.tensorLeftHomEquiv_tensor /-- `tensorRightHomEquiv` commutes with tensoring on the left -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/FunctorCategory.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/FunctorCategory.lean index bbeb2d8fc4cce..f1a9512a51a2f 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/FunctorCategory.lean @@ -62,16 +62,12 @@ instance functorHasLeftDual [LeftRigidCategory D] (F : C β₯€ D) : HasLeftDual F { app := fun X => Ξ΅_ _ _ naturality := fun X Y f => by dsimp - simp [tensorHom_def', leftAdjointMate_comp_evaluation] - simp_rw [← comp_whiskerRight_assoc] - simp only [IsIso.hom_inv_id, id_whiskerRight, Category.id_comp] } + simp [tensorHom_def, leftAdjointMate_comp_evaluation] } coevaluation' := { app := fun X => Ξ·_ _ _ naturality := fun X Y f => by dsimp - simp [tensorHom_def', coevaluation_comp_leftAdjointMate_assoc] - simp_rw [← MonoidalCategory.whiskerLeft_comp] - simp only [IsIso.inv_hom_id, MonoidalCategory.whiskerLeft_id, Category.comp_id] } } + simp [tensorHom_def, coevaluation_comp_leftAdjointMate_assoc] } } #align category_theory.monoidal.functor_has_left_dual CategoryTheory.Monoidal.functorHasLeftDual instance leftRigidFunctorCategory [LeftRigidCategory D] : LeftRigidCategory (C β₯€ D) where diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index daf8f80c2b8f7..6cfbe5c3140d9 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -68,7 +68,7 @@ def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D := .ofTensorHom slice_lhs 2 3 => rw [← id_tensor_comp] simp - simp [id_tensorHom, tensorHom_id]) + simp) (pentagon := fun W X Y Z ↦ by dsimp simp only [Iso.hom_inv_id_app_assoc, comp_tensor_id, assoc, Equivalence.inv_fun_map, diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index 53894ce3f45cf..214c2f7010be8 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -371,6 +371,7 @@ elab_rules : tactic MonoidalCategory.comp_whiskerRight, MonoidalCategory.id_whiskerRight, MonoidalCategory.whisker_assoc, MonoidalCategory.id_tensorHom, MonoidalCategory.tensorHom_id]; + -- I'm not sure if `tensorHom` should be expanded. try simp only [MonoidalCategory.tensorHom_def] )) From 2993b1efded9b1941af76e68823695b8145d39f0 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sat, 5 Aug 2023 17:56:11 +0900 Subject: [PATCH 10/62] set default `tensorHom` --- .../Category/ModuleCat/Monoidal/Basic.lean | 14 +++ .../ModuleCat/Monoidal/Symmetric.lean | 4 +- Mathlib/CategoryTheory/Monoidal/Category.lean | 50 +++++--- Mathlib/CategoryTheory/Monoidal/Center.lean | 8 +- .../CategoryTheory/Monoidal/Functorial.lean | 10 +- .../Monoidal/Internal/Module.lean | 6 +- .../CategoryTheory/Monoidal/Subcategory.lean | 2 + .../CategoryTheory/Monoidal/Transport.lean | 113 ++++++------------ Mathlib/RepresentationTheory/Action.lean | 10 ++ Mathlib/RepresentationTheory/Character.lean | 4 +- Mathlib/RepresentationTheory/FdRep.lean | 7 +- 11 files changed, 117 insertions(+), 111 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean index 7f31f17e0e76e..d6cdeddfc544e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean @@ -201,6 +201,8 @@ instance monoidalCategory : MonoidalCategory (ModuleCat.{u} R) := MonoidalCatego -- data (tensorObj := MonoidalCategory.tensorObj) (tensorHom := @tensorHom _ _) + (whiskerLeft := @whiskerLeft _ _) + (whiskerRight := @whiskerRight _ _) (tensorUnit' := ModuleCat.of R R) (associator := associator) (leftUnitor := leftUnitor) @@ -227,6 +229,18 @@ theorem hom_apply {K L M N : ModuleCat.{u} R} (f : K ⟢ L) (g : M ⟢ N) (k : K rfl #align Module.monoidal_category.hom_apply ModuleCat.MonoidalCategory.hom_apply +@[simp] +theorem whiskerLeft_apply (L : ModuleCat.{u} R) {M N : ModuleCat.{u} R} (f : M ⟢ N) + (l : L) (m : M) : + (L ◁ f) (l βŠ—β‚œ m) = l βŠ—β‚œ f m := + rfl + +@[simp] +theorem whiskerRight_apply {L M : ModuleCat.{u} R} (f : L ⟢ M) (N : ModuleCat.{u} R) + (l : L) (n : N) : + (f β–· N) (l βŠ—β‚œ n) = f l βŠ—β‚œ n := + rfl + @[simp] theorem leftUnitor_hom_apply {M : ModuleCat.{u} R} (r : R) (m : M) : ((Ξ»_ M).hom : πŸ™_ (ModuleCat R) βŠ— M ⟢ M) (r βŠ—β‚œ[R] m) = r β€’ m := diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean index ca76373143f9f..924c42cb3858b 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean @@ -53,7 +53,7 @@ theorem braiding_naturality_right (X : ModuleCat R) {Y Z : ModuleCat R} (f : Y @[simp] theorem hexagon_forward (X Y Z : ModuleCat.{u} R) : (Ξ±_ X Y Z).hom ≫ (braiding X _).hom ≫ (Ξ±_ Y Z X).hom = - ((braiding X Y).hom βŠ— πŸ™ Z) ≫ (Ξ±_ Y X Z).hom ≫ (πŸ™ Y βŠ— (braiding X Z).hom) := by + ((braiding X Y).hom β–· Z) ≫ (Ξ±_ Y X Z).hom ≫ (Y ◁ (braiding X Z).hom) := by apply TensorProduct.ext_threefold intro x y z rfl @@ -63,7 +63,7 @@ set_option linter.uppercaseLean3 false in @[simp] theorem hexagon_reverse (X Y Z : ModuleCat.{u} R) : (Ξ±_ X Y Z).inv ≫ (braiding _ Z).hom ≫ (Ξ±_ Z X Y).inv = - (πŸ™ X βŠ— (Y.braiding Z).hom) ≫ (Ξ±_ X Z Y).inv ≫ ((X.braiding Z).hom βŠ— πŸ™ Y) := by + (X ◁ (Y.braiding Z).hom) ≫ (Ξ±_ X Z Y).inv ≫ ((X.braiding Z).hom β–· Y) := by apply (cancel_epi (Ξ±_ X Y Z).hom).1 apply TensorProduct.ext_threefold intro x y z diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 4390757b7e5ea..8f2eb477ede70 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -722,8 +722,16 @@ theorem tensor_inv_hom_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) def ofTensorHom (tensorObj : C β†’ C β†’ C) (tensorHom : βˆ€ {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C}, (X₁ ⟢ Y₁) β†’ (Xβ‚‚ ⟢ Yβ‚‚) β†’ (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚)) + (whiskerLeft : βˆ€ (X : C) {Y₁ Yβ‚‚ : C} (_f : Y₁ ⟢ Yβ‚‚), tensorObj X Y₁ ⟢ tensorObj X Yβ‚‚ := + fun X _ _ f ↦ tensorHom (πŸ™ X) f) + (whiskerRight : βˆ€ {X₁ Xβ‚‚ : C} (_f : X₁ ⟢ Xβ‚‚) (Y : C), tensorObj X₁ Y ⟢ tensorObj Xβ‚‚ Y := + fun f Y ↦ tensorHom f (πŸ™ Y)) (tensor_id : βˆ€ X₁ Xβ‚‚ : C, tensorHom (πŸ™ X₁) (πŸ™ Xβ‚‚) = πŸ™ (tensorObj X₁ Xβ‚‚) := by aesop_cat) + (id_tensorHom : βˆ€ (X : C) {Y₁ Yβ‚‚ : C} (f : Y₁ ⟢ Yβ‚‚), tensorHom (πŸ™ X) f = whiskerLeft X f := by + aesop_cat) + (tensorHom_id : βˆ€ {X₁ Xβ‚‚ : C} (f : X₁ ⟢ Xβ‚‚) (Y : C), tensorHom f (πŸ™ Y) = whiskerRight f Y := by + aesop_cat) (tensor_comp : βˆ€ {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : Y₁ ⟢ Z₁) (gβ‚‚ : Yβ‚‚ ⟢ Zβ‚‚), tensorHom (f₁ ≫ g₁) (fβ‚‚ ≫ gβ‚‚) = tensorHom f₁ fβ‚‚ ≫ tensorHom g₁ gβ‚‚ := by @@ -759,25 +767,37 @@ def ofTensorHom MonoidalCategory C where tensorObj := tensorObj tensorHom := tensorHom - whiskerLeft X _ _ f := tensorHom (πŸ™ X) f - whiskerRight f X := tensorHom f (πŸ™ X) - tensorHom_def f g := by rw [← tensor_comp, id_comp, comp_id] + whiskerLeft X _ _ f := whiskerLeft X f + whiskerRight f X := whiskerRight f X + tensorHom_def := by intros; simp [← id_tensorHom, ←tensorHom_id, ← tensor_comp] tensorUnit' := tensorUnit' leftUnitor := leftUnitor rightUnitor := rightUnitor associator := associator - whiskerLeft_id := by intros; simp [tensor_id] - whiskerLeft_comp := by intros; simp [← tensor_comp] - id_whiskerLeft := by intros; rw [← assoc, ← leftUnitor_naturality]; simp - tensor_whiskerLeft := by intros; rw [← assoc, ← associator_naturality]; simp [tensor_id] - id_whiskerRight := by intros; simp [tensor_id] - comp_whiskerRight := by intros; simp [← tensor_comp] - whiskerRight_id := by intros; rw [← assoc, ← rightUnitor_naturality]; simp - whiskerRight_tensor := by intros; rw [associator_naturality]; simp [tensor_id] - whisker_assoc := by intros; rw [← assoc, ← associator_naturality]; simp - whisker_exchange := by intros; dsimp; rw [← tensor_comp, ← tensor_comp]; simp - pentagon := pentagon - triangle := triangle + whiskerLeft_id := by intros; simp [← id_tensorHom, ← tensor_id] + whiskerLeft_comp := by intros; simp [← id_tensorHom, ← tensor_comp] + id_whiskerLeft := by intros; rw [← assoc, ← leftUnitor_naturality]; simp [← id_tensorHom] + tensor_whiskerLeft := by + intros + simp only [← id_tensorHom, ← tensorHom_id] + rw [← assoc, ← associator_naturality] + simp [tensor_id] + id_whiskerRight := by intros; simp [← tensorHom_id, tensor_id] + comp_whiskerRight := by intros; simp [← tensorHom_id, ← tensor_comp] + whiskerRight_id := by intros; rw [← assoc, ← rightUnitor_naturality]; simp [← tensorHom_id] + whiskerRight_tensor := by + intros + simp only [← id_tensorHom, ← tensorHom_id] + rw [associator_naturality] + simp [tensor_id] + whisker_assoc := by + intros + simp only [← id_tensorHom, ← tensorHom_id] + rw [← assoc, ← associator_naturality] + simp + whisker_exchange := by simp [← id_tensorHom, ← tensorHom_id, ← tensor_comp] + pentagon := by intros; simp [← id_tensorHom, ← tensorHom_id, pentagon] + triangle := by intros; simp [← id_tensorHom, ← tensorHom_id, triangle] end diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index 6c59ab0b8e963..67170c93ea0df 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -166,8 +166,8 @@ def tensorObj (X Y : Center C) : Center C := @[reassoc] theorem whiskerLeft_comm (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) (U : C) : - (X.1 ◁ f.f) β–· U ≫ (HalfBraiding.Ξ² (tensorObj X Yβ‚‚).2 U).hom = - (HalfBraiding.Ξ² (tensorObj X Y₁).2 U).hom ≫ U ◁ X.1 ◁ f.f := by + (X.1 ◁ f.f) β–· U ≫ ((tensorObj X Yβ‚‚).2.Ξ² U).hom = + ((tensorObj X Y₁).2.Ξ² U).hom ≫ U ◁ X.1 ◁ f.f := by dsimp only [tensorObj_fst, tensorObj_snd_Ξ², Iso.trans_hom, whiskerLeftIso_hom, Iso.symm_hom, whiskerRightIso_hom] calc @@ -190,8 +190,8 @@ def whiskerLeft (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) : @[reassoc] theorem whiskerRight_comm {X₁ Xβ‚‚: Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) (U : C) : - f.f β–· Y.1 β–· U ≫ (HalfBraiding.Ξ² (tensorObj Xβ‚‚ Y).2 U).hom = - (HalfBraiding.Ξ² (tensorObj X₁ Y).snd U).hom ≫ U ◁ f.f β–· Y.1 := by + f.f β–· Y.1 β–· U ≫ ((tensorObj Xβ‚‚ Y).2.Ξ² U).hom = + ((tensorObj X₁ Y).2.Ξ² U).hom ≫ U ◁ f.f β–· Y.1 := by dsimp only [tensorObj_fst, tensorObj_snd_Ξ², Iso.trans_hom, whiskerLeftIso_hom, Iso.symm_hom, whiskerRightIso_hom] calc diff --git a/Mathlib/CategoryTheory/Monoidal/Functorial.lean b/Mathlib/CategoryTheory/Monoidal/Functorial.lean index 5582488929ed8..8be20eaae585c 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functorial.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functorial.lean @@ -102,11 +102,11 @@ def LaxMonoidal.ofTensorHom (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] [LaxMonoid LaxMonoidal.{v₁, vβ‚‚} F where Ξ΅ := Ξ΅ ΞΌ := ΞΌ - ΞΌ_natural_left := sorry - ΞΌ_natural_right := sorry - associativity := sorry - left_unitality := sorry - right_unitality := sorry + ΞΌ_natural_left f X := by intros; simpa using ΞΌ_natural f (πŸ™ X) + ΞΌ_natural_right X f := by intros; simpa using ΞΌ_natural (πŸ™ X) f + associativity X Y Z := by intros; simpa using associativity X Y Z + left_unitality X := by intros; simpa using left_unitality X + right_unitality X := by intros; simpa using right_unitality X attribute [simp] LaxMonoidal.ΞΌ_natural_left LaxMonoidal.ΞΌ_natural_right diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean index 2a1c0d440a5a0..a008b261e975b 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean @@ -121,7 +121,7 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where -- ModuleCat.MonoidalCategory.leftUnitor_hom_apply] -- Porting note : because `dsimp` is not effective, `rw` needs to be changed to `erw` erw [LinearMap.mul'_apply, MonoidalCategory.leftUnitor_hom_apply, ← Algebra.smul_def] - rw [id_apply] + erw [id_apply] mul_one := by -- Porting note : `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext_ring ?_ @@ -134,7 +134,7 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where rw [CategoryTheory.comp_apply] erw [LinearMap.mul'_apply, ModuleCat.MonoidalCategory.rightUnitor_hom_apply, ← Algebra.commutes, ← Algebra.smul_def] - rw [id_apply] + erw [id_apply] mul_assoc := by -- Porting note : `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext fun y => @@ -147,7 +147,7 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where rw [comprβ‚‚_apply, comprβ‚‚_apply, comprβ‚‚_apply, comprβ‚‚_apply, CategoryTheory.comp_apply, CategoryTheory.comp_apply, CategoryTheory.comp_apply] erw [LinearMap.mul'_apply, LinearMap.mul'_apply] - rw [id_apply, TensorProduct.mk_apply] + erw [id_apply, TensorProduct.mk_apply] erw [TensorProduct.mk_apply, TensorProduct.mk_apply, id_apply, LinearMap.mul'_apply, LinearMap.mul'_apply] simp only [LinearMap.mul'_apply, mul_assoc] diff --git a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean index 66239d565da3a..623bfb10d9acf 100644 --- a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean @@ -53,6 +53,8 @@ When `P` is a monoidal predicate, the full subcategory for `P` inherits the mono -/ instance fullMonoidalSubcategory : MonoidalCategory (FullSubcategory P) where tensorObj X Y := ⟨X.1 βŠ— Y.1, prop_tensor X.2 Y.2⟩ + tensorHom f g := f βŠ— g + tensorHom_def f g := tensorHom_def (C := C) f g whiskerLeft := fun X _ _ f ↦ X.1 ◁ f whiskerRight := fun f Y ↦ (fun f ↦ f β–· Y.1) f tensorUnit' := βŸ¨πŸ™_ C, prop_id⟩ diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index 6cfbe5c3140d9..a54880af3adeb 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -38,89 +38,84 @@ variable {D : Type uβ‚‚} [Category.{vβ‚‚} D] -- porting note: it was @[simps {attrs := [`_refl_lemma]}] /-- Transport a monoidal structure along an equivalence of (plain) categories. -/ -@[simps!] +@[simps (config := {rhsMd := .default}) + tensorObj tensorHom tensorUnit' associator leftUnitor rightUnitor] def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D := .ofTensorHom (tensorObj := fun X Y ↦ e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y)) (tensorHom := fun f g ↦ e.functor.map (e.inverse.map f βŠ— e.inverse.map g)) - (tensor_comp := by - intro X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ f₁ fβ‚‚ g₁ gβ‚‚ - dsimp - simp only [← e.functor.map_comp] - congr 1 - simp [tensorHom_def, whisker_exchange_assoc]) + (whiskerLeft := fun X _ _ f ↦ e.functor.map (e.inverse.obj X ◁ e.inverse.map f)) + (whiskerRight := fun f X ↦ e.functor.map (e.inverse.map f β–· e.inverse.obj X)) (tensorUnit' := e.functor.obj (πŸ™_ C)) (associator := fun X Y Z ↦ e.functor.mapIso - (((e.unitIso.app _).symm βŠ— Iso.refl _) β‰ͺ≫ - Ξ±_ (e.inverse.obj X) (e.inverse.obj Y) (e.inverse.obj Z) β‰ͺ≫ (Iso.refl _ βŠ— e.unitIso.app _))) + ((whiskerRightIso (e.unitIso.app _).symm _) β‰ͺ≫ + Ξ±_ (e.inverse.obj X) (e.inverse.obj Y) (e.inverse.obj Z) β‰ͺ≫ + (whiskerLeftIso _ (e.unitIso.app _)))) (leftUnitor := fun X ↦ - e.functor.mapIso (((e.unitIso.app _).symm βŠ— Iso.refl _) β‰ͺ≫ Ξ»_ (e.inverse.obj X)) β‰ͺ≫ + e.functor.mapIso ((whiskerRightIso (e.unitIso.app _).symm _) β‰ͺ≫ Ξ»_ (e.inverse.obj X)) β‰ͺ≫ e.counitIso.app _) (rightUnitor := fun X ↦ - e.functor.mapIso ((Iso.refl _ βŠ— (e.unitIso.app _).symm) β‰ͺ≫ ρ_ (e.inverse.obj X)) β‰ͺ≫ + e.functor.mapIso ((whiskerLeftIso _ (e.unitIso.app _).symm) β‰ͺ≫ ρ_ (e.inverse.obj X)) β‰ͺ≫ e.counitIso.app _) (triangle := fun X Y ↦ by dsimp - simp only [Iso.hom_inv_id_app_assoc, comp_tensor_id, Equivalence.unit_inverse_comp, assoc, - Equivalence.inv_fun_map, comp_id, Functor.map_comp, id_tensor_comp, e.inverse.map_id] + simp only [Functor.map_comp, Functor.map_id, assoc, Equivalence.inv_fun_map, + Functor.comp_obj, Functor.id_obj, + Equivalence.unit_inverse_comp, comp_id, Iso.hom_inv_id_app_assoc, + id_tensorHom, MonoidalCategory.whiskerLeft_comp, + tensorHom_id, comp_whiskerRight, whisker_assoc, triangle_assoc_comp_right] simp only [← e.functor.map_comp] congr 2 slice_lhs 2 3 => - rw [← id_tensor_comp] + rw [← MonoidalCategory.whiskerLeft_comp] simp simp) (pentagon := fun W X Y Z ↦ by dsimp - simp only [Iso.hom_inv_id_app_assoc, comp_tensor_id, assoc, Equivalence.inv_fun_map, - Functor.map_comp, id_tensor_comp, e.inverse.map_id] + simp only [Functor.map_comp, Equivalence.inv_fun_map, Functor.comp_obj, Functor.id_obj, assoc, + Iso.hom_inv_id_app_assoc, Functor.map_id, tensorHom_id, comp_whiskerRight, whisker_assoc, + id_tensorHom, MonoidalCategory.whiskerLeft_comp] simp only [← e.functor.map_comp] congr 2 - slice_lhs 4 5 => - rw [← comp_tensor_id, Iso.hom_inv_id_app] + slice_lhs 6 7 => + rw [← comp_whiskerRight, Iso.hom_inv_id_app] dsimp - rw [tensor_id] - simp only [Category.id_comp, Category.assoc] + rw [id_whiskerRight] + simp only [id_comp, assoc, Iso.inv_hom_id_assoc] slice_lhs 5 6 => - rw [← id_tensor_comp, Iso.hom_inv_id_app] + rw [← MonoidalCategory.whiskerLeft_comp, Iso.hom_inv_id_app] dsimp - rw [tensor_id] + rw [MonoidalCategory.whiskerLeft_id] simp only [Category.id_comp, Category.assoc] - slice_rhs 2 3 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_rhs 1 2 => rw [← tensor_id, ← associator_naturality] - slice_rhs 3 4 => rw [← tensor_id, associator_naturality] - slice_rhs 2 3 => rw [← pentagon'] - simp only [Category.assoc] - congr 2 - slice_lhs 1 2 => rw [associator_naturality] - simp only [Category.assoc] - congr 1 - slice_lhs 1 2 => - rw [← id_tensor_comp, ← comp_tensor_id, Iso.hom_inv_id_app] + slice_rhs 2 3 => rw [whisker_exchange] + slice_lhs 4 5 => + rw [← MonoidalCategory.whiskerLeft_comp, ← comp_whiskerRight, Iso.hom_inv_id_app] dsimp - rw [tensor_id, tensor_id] - simp only [Category.id_comp, Category.assoc]) + rw [id_whiskerRight, MonoidalCategory.whiskerLeft_id] + simp) (leftUnitor_naturality := fun f ↦ by dsimp simp only [Functor.map_comp, Functor.map_id, Category.assoc] erw [← e.counitIso.hom.naturality] simp only [Functor.comp_map, ← e.functor.map_comp_assoc] congr 2 - rw [id_tensor_comp_tensor_id_assoc, ← tensor_id_comp_id_tensor_assoc, - leftUnitor_naturality']) + simp only [id_tensorHom] + rw [whisker_exchange_assoc, leftUnitor_naturality]) (rightUnitor_naturality := fun f ↦ by dsimp simp only [Functor.map_comp, Functor.map_id, Category.assoc] erw [← e.counitIso.hom.naturality] simp only [Functor.comp_map, ← e.functor.map_comp_assoc] congr 2 - erw [tensor_id_comp_id_tensor_assoc, ← id_tensor_comp_tensor_id_assoc, - rightUnitor_naturality']) + simp only [tensorHom_id] + erw [← whisker_exchange_assoc, rightUnitor_naturality]) (associator_naturality := fun f₁ fβ‚‚ f₃ ↦ by dsimp simp only [Equivalence.inv_fun_map, Functor.map_comp, Category.assoc] simp only [← e.functor.map_comp] congr 1 conv_lhs => rw [← tensor_id_comp_id_tensor] + simp only [← id_tensorHom, ← tensorHom_id] slice_lhs 2 3 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor, ← tensor_id] simp only [Category.assoc] slice_lhs 3 4 => rw [associator_naturality] @@ -164,51 +159,16 @@ instance (e : C β‰Œ D) : MonoidalCategory (Transported e) := instance (e : C β‰Œ D) : Inhabited (Transported e) := βŸ¨πŸ™_ _⟩ --- theorem transport_tensorUnit' (e : C β‰Œ D) : πŸ™_ (Transported e) = e.functor.obj (πŸ™_ C) := rfl - --- theorem transport_tensorObj (e : C β‰Œ D) (X Y : Transported e) : --- X βŠ— Y = e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y) := --- rfl - --- theorem transport_tensorHom (e : C β‰Œ D) {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Transported e} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : --- f βŠ— g = e.functor.map (e.inverse.map f βŠ— e.inverse.map g) := by --- rfl - -theorem transport_associator (e : C β‰Œ D) (X Y Z : Transported e) : - Ξ±_ X Y Z = - e.functor.mapIso - (((e.unitIso.app (e.inverse.obj X βŠ— e.inverse.obj Y)).symm βŠ— - Iso.refl (e.inverse.obj Z)) β‰ͺ≫ - Ξ±_ (e.inverse.obj X) (e.inverse.obj Y) (e.inverse.obj Z) β‰ͺ≫ - (Iso.refl (e.inverse.obj X) βŠ— e.unitIso.app (e.inverse.obj Y βŠ— e.inverse.obj Z))) := - rfl - -theorem transport_leftUnitor (e : C β‰Œ D) (X : Transported e) : - Ξ»_ X = - e.functor.mapIso (((e.unitIso.app (πŸ™_ C)).symm βŠ— Iso.refl (e.inverse.obj X)) β‰ͺ≫ - Ξ»_ (e.inverse.obj X)) β‰ͺ≫ e.counitIso.app X := - rfl - -theorem transport_rightUnitor (e : C β‰Œ D) (X : Transported e) : - ρ_ X = - e.functor.mapIso ((Iso.refl (e.inverse.obj X) βŠ— (e.unitIso.app (πŸ™_ C)).symm) β‰ͺ≫ - ρ_ (e.inverse.obj X)) β‰ͺ≫ e.counitIso.app X := - rfl - section attribute [local simp] transport_tensorUnit' section -attribute [local simp] - transport_tensorObj transport_tensorHom transport_associator - transport_leftUnitor transport_rightUnitor - /-- We can upgrade `e.functor` to a lax monoidal functor from `C` to `D` with the transported structure. -/ -@[simp] +@[simps!] def laxToTransported (e : C β‰Œ D) : LaxMonoidalFunctor C (Transported e) := .ofTensorHom (F := e.functor) (Ξ΅ := πŸ™ (e.functor.obj (πŸ™_ C))) @@ -222,6 +182,7 @@ def laxToTransported (e : C β‰Œ D) : LaxMonoidalFunctor C (Transported e) := .of rw [comp_id, comp_id]) (associativity := fun X Y Z ↦ by dsimp + simp only [← id_tensorHom, ← tensorHom_id] rw [Equivalence.inv_fun_map, Equivalence.inv_fun_map, Functor.map_comp, Functor.map_comp, assoc, assoc, e.inverse.map_id, e.inverse.map_id, comp_tensor_id, id_tensor_comp, Functor.map_comp, assoc, id_tensor_comp, @@ -249,6 +210,7 @@ def laxToTransported (e : C β‰Œ D) : LaxMonoidalFunctor C (Transported e) := .of Iso.inv_hom_id_assoc, assoc, id_comp, comp_id]) (left_unitality := fun X ↦ by dsimp + simp only [← id_tensorHom, ← tensorHom_id] rw [e.inverse.map_id, e.inverse.map_id, tensor_id, Functor.map_comp, assoc, Equivalence.counit_app_functor, ← e.functor.map_comp, ← e.functor.map_comp, ← e.functor.map_comp, ← e.functor.map_comp, ← leftUnitor_naturality', @@ -256,6 +218,7 @@ def laxToTransported (e : C β‰Œ D) : LaxMonoidalFunctor C (Transported e) := .of rfl) (right_unitality := fun X ↦ by dsimp + simp only [← id_tensorHom, ← tensorHom_id] rw [Functor.map_comp, assoc, e.inverse.map_id, e.inverse.map_id, tensor_id, Functor.map_id, id_comp, Equivalence.counit_app_functor, ← e.functor.map_comp, ← e.functor.map_comp, ← e.functor.map_comp, ← rightUnitor_naturality', diff --git a/Mathlib/RepresentationTheory/Action.lean b/Mathlib/RepresentationTheory/Action.lean index 8fd795c9ad0f7..3abb5706e0574 100644 --- a/Mathlib/RepresentationTheory/Action.lean +++ b/Mathlib/RepresentationTheory/Action.lean @@ -506,6 +506,16 @@ theorem tensorHom {W X Y Z : Action V G} (f : W ⟢ X) (g : Y ⟢ Z) : (f βŠ— g) set_option linter.uppercaseLean3 false in #align Action.tensor_hom Action.tensorHom +@[simp] +theorem whiskerLeft_v (X : Action V G) {Y Z : Action V G} (f : Y ⟢ Z) : + (X ◁ f).hom = X.V ◁ f.hom := + rfl + +@[simp] +theorem whiskerRight_v {X Y : Action V G} (f : X ⟢ Y) (Z : Action V G) : + (f β–· Z).hom = f.hom β–· Z.V := + rfl + -- porting note: removed @[simp] as the simpNF linter complains theorem associator_hom_hom {X Y Z : Action V G} : Hom.hom (Ξ±_ X Y Z).hom = (Ξ±_ X.V Y.V Z.V).hom := by diff --git a/Mathlib/RepresentationTheory/Character.lean b/Mathlib/RepresentationTheory/Character.lean index 22bff02555fa4..3b29acb87cc96 100644 --- a/Mathlib/RepresentationTheory/Character.lean +++ b/Mathlib/RepresentationTheory/Character.lean @@ -55,9 +55,7 @@ theorem char_one (V : FdRep k G) : V.character 1 = FiniteDimensional.finrank k V /-- The character is multiplicative under the tensor product. -/ theorem char_tensor (V W : FdRep k G) : (V βŠ— W).character = V.character * W.character := by - ext g - convert trace_tensorProduct' (V.ρ g) (W.ρ g) - sorry + ext g; convert trace_tensorProduct' (V.ρ g) (W.ρ g) #align fdRep.char_tensor FdRep.char_tensor -- Porting note: adding variant of `char_tensor` to make the simp-set confluent diff --git a/Mathlib/RepresentationTheory/FdRep.lean b/Mathlib/RepresentationTheory/FdRep.lean index ee262e2955546..8eb46ff0ae633 100644 --- a/Mathlib/RepresentationTheory/FdRep.lean +++ b/Mathlib/RepresentationTheory/FdRep.lean @@ -175,15 +175,14 @@ noncomputable def dualTensorIsoLinHomAux : (FdRep.of ρV.dual βŠ— W).V β‰… (FdRep.of (linHom ρV W.ρ)).V := -- Porting note: had to make all types explicit arguments @LinearEquiv.toFGModuleCatIso k _ (FdRep.of ρV.dual βŠ— W).V (V β†’β‚—[k] W) - _ _ _ _ _ _ (dualTensorHomEquiv k V W) + _ _ _ _ _ _ (dualTensorHomEquiv k V W) #align fdRep.dual_tensor_iso_lin_hom_aux FdRep.dualTensorIsoLinHomAux /-- When `V` and `W` are finite dimensional representations of a group `G`, the isomorphism `dualTensorHomEquiv k V W` of vector spaces induces an isomorphism of representations. -/ noncomputable def dualTensorIsoLinHom : FdRep.of ρV.dual βŠ— W β‰… FdRep.of (linHom ρV W.ρ) := by - apply Action.mkIso (dualTensorIsoLinHomAux ρV W) (fun g => ?_) - -- convert dualTensorHom_comm ρV W.ρ g - sorry + refine Action.mkIso (dualTensorIsoLinHomAux ρV W) ?_ + convert dualTensorHom_comm ρV W.ρ #align fdRep.dual_tensor_iso_lin_hom FdRep.dualTensorIsoLinHom @[simp] From 936abcf8ff48222c4d926f2df9ce96b974351b38 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sat, 5 Aug 2023 18:27:20 +0900 Subject: [PATCH 11/62] `Monoidal.Preadditive` --- Mathlib/CategoryTheory/Monoidal/Bimod.lean | 10 ++++-- Mathlib/CategoryTheory/Monoidal/Category.lean | 10 ++++++ .../CategoryTheory/Monoidal/Preadditive.lean | 32 +++++++------------ 3 files changed, 29 insertions(+), 23 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index 8f7dd4732f178..bc7f2c88c3fb0 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -269,7 +269,9 @@ theorem left_assoc' : slice_lhs 1 2 => rw [associator_inv_naturality_left] slice_lhs 2 3 => rw [← comp_whiskerRight, left_assoc, comp_whiskerRight, comp_whiskerRight] slice_rhs 1 2 => rw [associator_naturality_right] - slice_rhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, whiskerLeft_Ο€_actLeft, MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + slice_rhs 2 3 => + rw [← MonoidalCategory.whiskerLeft_comp, whiskerLeft_Ο€_actLeft, + MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] slice_rhs 4 5 => rw [whiskerLeft_Ο€_actLeft] slice_rhs 3 4 => rw [associator_inv_naturality_middle] coherence @@ -297,7 +299,8 @@ noncomputable def actRight : X P Q βŠ— T.X ⟢ X P Q := (by dsimp simp only [comp_whiskerRight, whisker_assoc, Category.assoc, Iso.inv_hom_id_assoc] - slice_lhs 3 4 => rw [← MonoidalCategory.whiskerLeft_comp, middle_assoc, MonoidalCategory.whiskerLeft_comp] + slice_lhs 3 4 => + rw [← MonoidalCategory.whiskerLeft_comp, middle_assoc, MonoidalCategory.whiskerLeft_comp] simp)) set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.act_right Bimod.TensorBimod.actRight @@ -423,7 +426,8 @@ noncomputable def whiskerLeft {X Y Z : Mon_ C} (M : Bimod X Y) {N₁ Nβ‚‚ : Bimo slice_lhs 1 2 => rw [TensorBimod.Ο€_tensor_id_actRight] slice_lhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] slice_lhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, Hom.right_act_hom] - slice_rhs 1 2 => rw [← comp_whiskerRight, ΞΉ_colimMap, parallelPairHom_app_one, comp_whiskerRight] + slice_rhs 1 2 => + rw [← comp_whiskerRight, ΞΉ_colimMap, parallelPairHom_app_one, comp_whiskerRight] slice_rhs 2 3 => rw [TensorBimod.Ο€_tensor_id_actRight] simp diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 8f2eb477ede70..4ac80d751f32c 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -405,6 +405,16 @@ end pentagon variable {U V W X Y Z : C} +theorem whiskerLeft_dite {P : Prop} [Decidable P] + (X : C) {Y Z : C} (f : P β†’ (Y ⟢ Z)) (f' : Β¬P β†’ (Y ⟢ Z)) : + X ◁ (if h : P then f h else f' h) = if h : P then X ◁ f h else X ◁ f' h := by + split_ifs <;> rfl + +theorem dite_whiskerRight {P : Prop} [Decidable P] + {X Y : C} (f : P β†’ (X ⟢ Y)) (f' : Β¬P β†’ (X ⟢ Y)) (Z : C): + (if h : P then f h else f' h) β–· Z = if h : P then f h β–· Z else f' h β–· Z := by + split_ifs <;> rfl + theorem tensor_dite {P : Prop} [Decidable P] {W X Y Z : C} (f : W ⟢ X) (g : P β†’ (Y ⟢ Z)) (g' : Β¬P β†’ (Y ⟢ Z)) : (f βŠ— if h : P then g h else g' h) = if h : P then f βŠ— g h else f βŠ— g' h := by split_ifs <;> rfl diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index e852d3c4df078..dffb33a3ac339 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -159,16 +159,14 @@ theorem leftDistributor_hom_comp_biproduct_Ο€ {J : Type} [Fintype J] (X : C) (f @[reassoc (attr := simp)] theorem biproduct_ΞΉ_comp_leftDistributor_hom {J : Type} [Fintype J] (X : C) (f : J β†’ C) (j : J) : (X ◁ biproduct.ΞΉ _ j) ≫ (leftDistributor X f).hom = biproduct.ΞΉ (fun j => X βŠ— f j) j := by - simp [leftDistributor_hom, Preadditive.comp_sum, ← id_tensor_comp_assoc, biproduct.ΞΉ_Ο€, - tensor_dite, dite_comp] - sorry + simp [leftDistributor_hom, Preadditive.comp_sum, ← whiskerLeft_comp_assoc, biproduct.ΞΉ_Ο€, + whiskerLeft_dite, dite_comp] @[reassoc (attr := simp)] theorem leftDistributor_inv_comp_biproduct_Ο€ {J : Type} [Fintype J] (X : C) (f : J β†’ C) (j : J) : (leftDistributor X f).inv ≫ (X ◁ biproduct.Ο€ _ j) = biproduct.Ο€ _ j := by - simp [leftDistributor_inv, Preadditive.sum_comp, ← id_tensor_comp, biproduct.ΞΉ_Ο€, tensor_dite, - comp_dite] - sorry + simp [leftDistributor_inv, Preadditive.sum_comp, ← MonoidalCategory.whiskerLeft_comp, + biproduct.ΞΉ_Ο€, whiskerLeft_dite, comp_dite] @[reassoc (attr := simp)] theorem biproduct_ΞΉ_comp_leftDistributor_inv {J : Type} [Fintype J] (X : C) (f : J β†’ C) (j : J) : @@ -222,16 +220,14 @@ theorem rightDistributor_hom_comp_biproduct_Ο€ {J : Type} [Fintype J] (f : J β†’ @[reassoc (attr := simp)] theorem biproduct_ΞΉ_comp_rightDistributor_hom {J : Type} [Fintype J] (f : J β†’ C) (X : C) (j : J) : (biproduct.ΞΉ _ j β–· X) ≫ (rightDistributor f X).hom = biproduct.ΞΉ (fun j => f j βŠ— X) j := by - simp [rightDistributor_hom, Preadditive.comp_sum, ← comp_tensor_id_assoc, biproduct.ΞΉ_Ο€, - dite_tensor, dite_comp] - sorry + simp [rightDistributor_hom, Preadditive.comp_sum, ← comp_whiskerRight_assoc, biproduct.ΞΉ_Ο€, + dite_whiskerRight, dite_comp] @[reassoc (attr := simp)] theorem rightDistributor_inv_comp_biproduct_Ο€ {J : Type} [Fintype J] (f : J β†’ C) (X : C) (j : J) : (rightDistributor f X).inv ≫ (biproduct.Ο€ _ j β–· X) = biproduct.Ο€ _ j := by - simp [rightDistributor_inv, Preadditive.sum_comp, ← comp_tensor_id, biproduct.ΞΉ_Ο€, dite_tensor, - comp_dite] - sorry + simp [rightDistributor_inv, Preadditive.sum_comp, ← MonoidalCategory.comp_whiskerRight, + biproduct.ΞΉ_Ο€, dite_whiskerRight, comp_dite] @[reassoc (attr := simp)] theorem biproduct_ΞΉ_comp_rightDistributor_inv {J : Type} [Fintype J] (f : J β†’ C) (X : C) (j : J) : @@ -303,8 +299,7 @@ theorem leftDistributor_extβ‚‚_left {J : Type} [Fintype J] g = h := by apply (cancel_epi (Ξ±_ _ _ _).hom).mp ext - simp_rw [← tensor_id, associator_naturality_assoc, w] - sorry + simp [w] @[ext] theorem leftDistributor_extβ‚‚_right {J : Type} [Fintype J] @@ -313,8 +308,7 @@ theorem leftDistributor_extβ‚‚_right {J : Type} [Fintype J] g = h := by apply (cancel_mono (Ξ±_ _ _ _).inv).mp ext - simp_rw [← tensor_id, Category.assoc, ← associator_inv_naturality, ← Category.assoc, w] - sorry + simp [w] @[ext] theorem rightDistributor_ext_left {J : Type} [Fintype J] @@ -348,8 +342,7 @@ theorem rightDistributor_extβ‚‚_left {J : Type} [Fintype J] g = h := by apply (cancel_epi (Ξ±_ _ _ _).inv).mp ext - simp_rw [← tensor_id, associator_inv_naturality_assoc, w] - sorry + simp [w] @[ext] theorem rightDistributor_extβ‚‚_right {J : Type} [Fintype J] @@ -358,7 +351,6 @@ theorem rightDistributor_extβ‚‚_right {J : Type} [Fintype J] g = h := by apply (cancel_mono (Ξ±_ _ _ _).hom).mp ext - simp_rw [← tensor_id, Category.assoc, ← associator_naturality, ← Category.assoc, w] - sorry + simp [w] end CategoryTheory From edb01a51988423eb6e3534e5afbf93964aef212d Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 5 Aug 2023 20:27:08 +1000 Subject: [PATCH 12/62] linting errors --- Mathlib/CategoryTheory/Monoidal/Braided.lean | 1 - Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean | 1 - Mathlib/CategoryTheory/Monoidal/Functorial.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Preadditive.lean | 2 -- 4 files changed, 1 insertion(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 85340d91183f6..5e0dab43ec74e 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -224,7 +224,6 @@ theorem braiding_leftUnitor_auxβ‚‚ (X : C) : #align category_theory.braiding_left_unitor_auxβ‚‚ CategoryTheory.braiding_leftUnitor_auxβ‚‚ -@[simp] theorem braiding_leftUnitor (X : C) : (Ξ²_ X (πŸ™_ C)).hom ≫ (Ξ»_ X).hom = (ρ_ X).hom := by rw [← whiskerRight_iff, comp_whiskerRight, braiding_leftUnitor_auxβ‚‚] #align category_theory.braiding_left_unitor CategoryTheory.braiding_leftUnitor diff --git a/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean b/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean index 206fd3b9460c9..c26c281827cdb 100644 --- a/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean +++ b/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean @@ -59,7 +59,6 @@ theorem pentagon_inv_inv_hom (W X Y Z : C) : coherence #align category_theory.monoidal_category.pentagon_inv_inv_hom CategoryTheory.MonoidalCategory.pentagon_inv_inv_hom -@[reassoc (attr := simp)] theorem triangle_assoc_comp_right_inv' (X Y : C) : ((ρ_ X).inv βŠ— πŸ™ Y) ≫ (Ξ±_ X (πŸ™_ C) Y).hom = πŸ™ X βŠ— (Ξ»_ Y).inv := by coherence diff --git a/Mathlib/CategoryTheory/Monoidal/Functorial.lean b/Mathlib/CategoryTheory/Monoidal/Functorial.lean index 8be20eaae585c..1814b5061fdcb 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functorial.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functorial.lean @@ -108,7 +108,7 @@ def LaxMonoidal.ofTensorHom (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] [LaxMonoid left_unitality X := by intros; simpa using left_unitality X right_unitality X := by intros; simpa using right_unitality X -attribute [simp] LaxMonoidal.ΞΌ_natural_left LaxMonoidal.ΞΌ_natural_right +attribute [simp, nolint simpNF] LaxMonoidal.ΞΌ_natural_left LaxMonoidal.ΞΌ_natural_right -- The unitality axioms cannot be used as simp lemmas because they require -- higher-order matching to figure out the `F` and `X` from `F X`. diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index dffb33a3ac339..fd835c8c37999 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -192,7 +192,6 @@ def rightDistributor {J : Type} [Fintype J] (f : J β†’ C) (X : C) : (⨁ f) βŠ— (tensorRight X).mapBiproduct f #align category_theory.right_distributor CategoryTheory.rightDistributor -@[simp] theorem rightDistributor_hom {J : Type} [Fintype J] (f : J β†’ C) (X : C) : (rightDistributor f X).hom = βˆ‘ j : J, (biproduct.Ο€ f j β–· X) ≫ biproduct.ΞΉ (fun j => f j βŠ— X) j := by @@ -203,7 +202,6 @@ theorem rightDistributor_hom {J : Type} [Fintype J] (f : J β†’ C) (X : C) : Finset.sum_dite_eq', Finset.mem_univ, eqToHom_refl, Category.comp_id, ite_true] #align category_theory.right_distributor_hom CategoryTheory.rightDistributor_hom -@[simp] theorem rightDistributor_inv {J : Type} [Fintype J] (f : J β†’ C) (X : C) : (rightDistributor f X).inv = βˆ‘ j : J, biproduct.Ο€ _ j ≫ (biproduct.ΞΉ f j β–· X) := by ext From 6cbee423e908175cec7c81c4d260c48298af3319 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sat, 5 Aug 2023 22:04:10 +0900 Subject: [PATCH 13/62] update docs --- Mathlib/CategoryTheory/Monoidal/Category.lean | 467 ++++++++++-------- Mathlib/CategoryTheory/Monoidal/Center.lean | 16 +- .../Monoidal/Free/Coherence.lean | 18 - Mathlib/CategoryTheory/Monoidal/Functor.lean | 1 - .../Monoidal/Rigid/OfEquivalence.lean | 26 +- Mathlib/RepresentationTheory/FdRep.lean | 2 +- Mathlib/Tactic/CategoryTheory/Coherence.lean | 3 +- 7 files changed, 265 insertions(+), 268 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 4ac80d751f32c..4c94032bcf674 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -21,25 +21,41 @@ The tensor product can be expressed as a functor via `tensor : C Γ— C β₯€ C`. The unitors and associator are gathered together as natural isomorphisms in `leftUnitor_nat_iso`, `rightUnitor_nat_iso` and `associator_nat_iso`. -Some consequences of the definition are proved in other files, -e.g. `(Ξ»_ (πŸ™_ C)).hom = (ρ_ (πŸ™_ C)).hom` in `CategoryTheory.Monoidal.UnitorsEqual`. +Some consequences of the definition are proved in other files after proving the coherence theorem, +e.g. `(Ξ»_ (πŸ™_ C)).hom = (ρ_ (πŸ™_ C)).hom` in `CategoryTheory.Monoidal.CoherenceLemmas`. -## Implementation -Dealing with unitors and associators is painful, and at this stage we do not have a useful -implementation of coherence for monoidal categories. +## Implementation notes -In an effort to lessen the pain, we put some effort into choosing the right `simp` lemmas. -Generally, the rule is that the component index of a natural transformation "weighs more" -in considering the complexity of an expression than does a structural isomorphism (associator, etc). +In the definition of monoidal categories, we also provide the whiskering operators `whiskerLeft` +and `whiskerRight`. These are products of an object and a morphism (the terminology "whiskering" +is borrowed from the 2-category theory). The tensor product of morphisms `tensorHom` can be defined +in terms of the whiskerings. There are two possible such definitions, which are related by +the exchange property of the whiskerings. These two definitions are accessed by `tensorHom_def` +and `tensorHom_def'`. By default, `tensorHom` is defined so that `tensorHom_def` holds +definitionally. -As an example when we prove Proposition 2.2.4 of - -we state it as a `@[simp]` lemma as -``` -(Ξ»_ (X βŠ— Y)).hom = (Ξ±_ (πŸ™_ C) X Y).inv ≫ (Ξ»_ X).hom βŠ— (πŸ™ Y) -``` +If you want to provide `tensorHom` and define `whiskerLeft` and `whiskerRight` in terms of it, +you can use the alternative constructor `CategoryTheory.MonoidalCategory.ofTensorHom`. -This is far from completely effective, but seems to prove a useful principle. +The whiskerings are useful when considering simp-normal forms of morphisms in monoidal categories. + +### Simp-normal form for morphisms + +Rewriting involving associators and unitors could be very complicated. We try to ease this +complexity by putting carefully chosen simp lemmas that rewrite any morphisms into simp-normal +form defined below. Rewriting into simp-normal form is especially useful in preprocessing +performed by the `coherence` tactic. + +The simp-normal form of morphisms is defined to be an expression that has the minimal number of +parentheses. More precisely, +1. it is a composition of morphisms like `f₁ ≫ fβ‚‚ ≫ f₃ ≫ fβ‚„ ≫ fβ‚…` such that each `fα΅’` is + either a structural morphisms (morphisms made up only of identities, associators, unitors) + or non-structural morphisms, and +2. each non-structural morphism in the composition is of the form `X₁ ◁ Xβ‚‚ ◁ X₃ ◁ f β–· Xβ‚„ β–· Xβ‚…`, + where each `Xα΅’` is a object that is not the identity or a tensor and `f` is a non-structural + morphisms that is not the identity or a composite. + +Note that `X₁ ◁ Xβ‚‚ ◁ X₃ ◁ f β–· Xβ‚„ β–· Xβ‚…` is actually `X₁ ◁ (Xβ‚‚ ◁ (X₃ ◁ ((f β–· Xβ‚„) β–· Xβ‚…)))`. ## References * Tensor categories, Etingof, Gelaki, Nikshych, Ostrik, @@ -86,11 +102,11 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] where aesop_cat -- Porting note: Adding a prime here, so I can later define `tensorUnit` unprimed with explicit -- argument `C` - /-- The tensor unity in the monoidal structure `πŸ™_C` -/ + /-- The tensor unity in the monoidal structure `πŸ™_ C` -/ tensorUnit' : C - /-- The left unitor: `πŸ™_C βŠ— X ≃ X` -/ + /-- The left unitor: `πŸ™_ C βŠ— X ≃ X` -/ leftUnitor : βˆ€ X : C, tensorObj tensorUnit' X β‰… X - /-- The right unitor: `X βŠ— πŸ™_C ≃ X` -/ + /-- The right unitor: `X βŠ— πŸ™_ C ≃ X` -/ rightUnitor : βˆ€ X : C, tensorObj X tensorUnit' β‰… X /-- The associator isomorphism `(X βŠ— Y) βŠ— Z ≃ X βŠ— (Y βŠ— Z)` -/ associator : βˆ€ X Y Z : C, tensorObj (tensorObj X Y) Z β‰… tensorObj X (tensorObj Y Z) @@ -144,7 +160,7 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] where (associator (tensorObj W X) Y Z).hom ≫ (associator W X (tensorObj Y Z)).hom := by aesop_cat /-- - The identity relating the isomorphisms between `X βŠ— (πŸ™_C βŠ— Y)`, `(X βŠ— πŸ™_C) βŠ— Y` and `X βŠ— Y` + The identity relating the isomorphisms between `X βŠ— (πŸ™_ C βŠ— Y)`, `(X βŠ— πŸ™_ C) βŠ— Y` and `X βŠ— Y` -/ triangle : βˆ€ X Y : C, @@ -167,20 +183,24 @@ attribute [simp] whiskerLeft_comp id_whiskerLeft tensor_whiskerLeft comp_whiskerRight id_whiskerRight whiskerRight_tensor whisker_assoc -end MonoidalCategory - -- Porting Note: This is here to make `tensorUnit` explicitly depend on `C`, which was done in -- Lean 3 using the `[]` notation in the `tensorUnit'` field. -open CategoryTheory.MonoidalCategory in /-- The tensor unity in the monoidal structure `πŸ™_C` -/ -abbrev MonoidalCategory.tensorUnit (C : Type u) [Category.{v} C] [MonoidalCategory C] : C := +abbrev tensorUnit (C : Type u) [Category.{v} C] [MonoidalCategory C] : C := tensorUnit' (C := C) -namespace MonoidalCategory - /-- Notation for `tensorObj`, the tensor product of objects in a monoidal category -/ scoped infixr:70 " βŠ— " => MonoidalCategory.tensorObj +/-- Notation for the `whiskerLeft` -/ +scoped infixr:81 " ◁ " => whiskerLeft + +/-- Notation for the `whiskerRight` -/ +scoped infixl:81 " β–· " => whiskerRight + +/-- Notation for `tensorHom`, the tensor product of morphisms in a monoidal category -/ +scoped infixr:70 " βŠ— " => MonoidalCategory.tensorHom + /-- Notation for `tensorUnit`, the two-sided identity of `βŠ—` -/ scoped notation "πŸ™_" => tensorUnit @@ -193,14 +213,8 @@ scoped notation "Ξ»_" => leftUnitor /-- Notation for the `rightUnitor`: `X βŠ— πŸ™_C ≃ X` -/ scoped notation "ρ_" => rightUnitor -scoped infixr:81 " ◁ " => whiskerLeft -scoped infixl:81 " β–· " => whiskerRight - variable {C : Type u} [π’ž : Category.{v} C] [MonoidalCategory C] -/-- Notation for `tensorHom`, the tensor product of morphisms in a monoidal category -/ -scoped infixr:70 " βŠ— " => MonoidalCategory.tensorHom - @[reassoc] theorem tensorHom_def' {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : f βŠ— g = X₁ ◁ g ≫ f β–· Yβ‚‚ := @@ -334,75 +348,6 @@ theorem inv_tensor {W X Y Z : C} (f : W ⟢ X) [IsIso f] (g : Y ⟢ Z) [IsIso g] simp [tensorHom_def ,whisker_exchange] #align category_theory.monoidal_category.inv_tensor CategoryTheory.MonoidalCategory.inv_tensor -section pentagon - -variable (W X Y Z : C) - -@[reassoc (attr := simp)] -theorem pentagon_inv : - W ◁ (Ξ±_ X Y Z).inv ≫ (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z = - (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv := - eq_of_inv_eq_inv (by simp) - -@[reassoc] -theorem pentagon_inv' (W X Y Z : C) : - (πŸ™ W βŠ— (Ξ±_ X Y Z).inv) ≫ (Ξ±_ W (X βŠ— Y) Z).inv ≫ ((Ξ±_ W X Y).inv βŠ— πŸ™ Z) = - (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv := - CategoryTheory.eq_of_inv_eq_inv (by simp [pentagon]) -#align category_theory.monoidal_category.pentagon_inv CategoryTheory.MonoidalCategory.pentagon_inv' - -@[reassoc (attr := simp)] -theorem pentagon_inv_inv_hom_hom_inv : - (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z ≫ (Ξ±_ (W βŠ— X) Y Z).hom = - W ◁ (Ξ±_ X Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).inv := by - rw [← cancel_epi (W ◁ (Ξ±_ X Y Z).inv), ← cancel_mono (Ξ±_ (W βŠ— X) Y Z).inv] - simp - -@[reassoc (attr := simp)] -theorem pentagon_inv_hom_hom_hom_inv : - (Ξ±_ (W βŠ— X) Y Z).inv ≫ (Ξ±_ W X Y).hom β–· Z ≫ (Ξ±_ W (X βŠ— Y) Z).hom = - (Ξ±_ W X (Y βŠ— Z)).hom ≫ W ◁ (Ξ±_ X Y Z).inv := - eq_of_inv_eq_inv (by simp) - -@[reassoc (attr := simp)] -theorem pentagon_hom_inv_inv_inv_inv : - W ◁ (Ξ±_ X Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv = - (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z := - by simp [← cancel_epi (W ◁ (Ξ±_ X Y Z).inv)] - -@[reassoc (attr := simp)] -theorem pentagon_hom_hom_inv_hom_hom : - (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom ≫ W ◁ (Ξ±_ X Y Z).inv = - (Ξ±_ W X Y).hom β–· Z ≫ (Ξ±_ W (X βŠ— Y) Z).hom := - eq_of_inv_eq_inv (by simp) - -@[reassoc (attr := simp)] -theorem pentagon_hom_inv_inv_inv_hom : - (Ξ±_ W X (Y βŠ— Z)).hom ≫ W ◁ (Ξ±_ X Y Z).inv ≫ (Ξ±_ W (X βŠ— Y) Z).inv = - (Ξ±_ (W βŠ— X) Y Z).inv ≫ (Ξ±_ W X Y).hom β–· Z := by - rw [← cancel_epi (Ξ±_ W X (Y βŠ— Z)).inv, ← cancel_mono ((Ξ±_ W X Y).inv β–· Z)] - simp - -@[reassoc (attr := simp)] -theorem pentagon_hom_hom_inv_inv_hom : - (Ξ±_ W (X βŠ— Y) Z).hom ≫ W ◁ (Ξ±_ X Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).inv = - (Ξ±_ W X Y).inv β–· Z ≫ (Ξ±_ (W βŠ— X) Y Z).hom := - eq_of_inv_eq_inv (by simp) - -@[reassoc (attr := simp)] -theorem pentagon_inv_hom_hom_hom_hom : - (Ξ±_ W X Y).inv β–· Z ≫ (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom = - (Ξ±_ W (X βŠ— Y) Z).hom ≫ W ◁ (Ξ±_ X Y Z).hom := - by simp [← cancel_epi ((Ξ±_ W X Y).hom β–· Z)] - -@[reassoc (attr := simp)] -theorem pentagon_inv_inv_hom_inv_inv : - (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv ≫ (Ξ±_ W X Y).hom β–· Z = - W ◁ (Ξ±_ X Y Z).inv ≫ (Ξ±_ W (X βŠ— Y) Z).inv := - eq_of_inv_eq_inv (by simp) - -end pentagon - variable {U V W X Y Z : C} theorem whiskerLeft_dite {P : Prop} [Decidable P] @@ -425,54 +370,17 @@ theorem dite_tensor {P : Prop} [Decidable P] {W X Y Z : C} (f : W ⟢ X) (g : P by split_ifs <;> rfl #align category_theory.monoidal_category.dite_tensor CategoryTheory.MonoidalCategory.dite_tensor -@[reassoc] -theorem comp_tensor_id (f : W ⟢ X) (g : X ⟢ Y) : f ≫ g βŠ— πŸ™ Z = (f βŠ— πŸ™ Z) ≫ (g βŠ— πŸ™ Z) := by - simp -#align category_theory.monoidal_category.comp_tensor_id CategoryTheory.MonoidalCategory.comp_tensor_id - -@[reassoc] -theorem id_tensor_comp (f : W ⟢ X) (g : X ⟢ Y) : πŸ™ Z βŠ— f ≫ g = (πŸ™ Z βŠ— f) ≫ (πŸ™ Z βŠ— g) := by - simp -#align category_theory.monoidal_category.id_tensor_comp CategoryTheory.MonoidalCategory.id_tensor_comp - -@[reassoc] -theorem id_tensor_comp_tensor_id (f : W ⟢ X) (g : Y ⟢ Z) : (πŸ™ Y βŠ— f) ≫ (g βŠ— πŸ™ X) = g βŠ— f := by - rw [← tensor_comp] - simp -#align category_theory.monoidal_category.id_tensor_comp_tensor_id CategoryTheory.MonoidalCategory.id_tensor_comp_tensor_id - -@[reassoc] -theorem tensor_id_comp_id_tensor (f : W ⟢ X) (g : Y ⟢ Z) : (g βŠ— πŸ™ W) ≫ (πŸ™ Z βŠ— f) = g βŠ— f := by - rw [← tensor_comp] - simp -#align category_theory.monoidal_category.tensor_id_comp_id_tensor CategoryTheory.MonoidalCategory.tensor_id_comp_id_tensor - -theorem tensor_left_iff {X Y : C} (f g : X ⟢ Y) : πŸ™ (πŸ™_ C) βŠ— f = πŸ™ (πŸ™_ C) βŠ— g ↔ f = g := by simp -#align category_theory.monoidal_category.tensor_left_iff CategoryTheory.MonoidalCategory.tensor_left_iff - -theorem tensor_right_iff {X Y : C} (f g : X ⟢ Y) : f βŠ— πŸ™ (πŸ™_ C) = g βŠ— πŸ™ (πŸ™_ C) ↔ f = g := by simp -#align category_theory.monoidal_category.tensor_right_iff CategoryTheory.MonoidalCategory.tensor_right_iff - -/-! The lemmas in the next section are true by coherence, -but we prove them directly as they are used in proving the coherence theorem. -/ - - -section - -@[reassoc (attr := simp)] -theorem triangle_assoc_comp_right (X Y : C) : - (Ξ±_ X (πŸ™_ C) Y).inv ≫ ((ρ_ X).hom β–· Y) = X ◁ (Ξ»_ Y).hom := by - rw [← triangle, Iso.inv_hom_id_assoc] - -@[reassoc (attr := simp)] -theorem triangle_assoc_comp_right_inv (X Y : C) : - (ρ_ X).inv β–· Y ≫ (Ξ±_ X (πŸ™_ C) Y).hom = X ◁ (Ξ»_ Y).inv := by - simp [← cancel_mono (X ◁ (Ξ»_ Y).hom)] +@[simp] +theorem whiskerLeft_eqToHom (X : C) {Y Z : C} (f : Y = Z) : + X ◁ eqToHom f = eqToHom (congr_argβ‚‚ tensorObj rfl f) := by + cases f + simp only [whiskerLeft_id, eqToHom_refl] -@[reassoc (attr := simp)] -theorem triangle_assoc_comp_left_inv (X Y : C) : - (X ◁ (Ξ»_ Y).inv) ≫ (Ξ±_ X (πŸ™_ C) Y).inv = (ρ_ X).inv β–· Y := by - simp [← cancel_mono ((ρ_ X).hom β–· Y)] +@[simp] +theorem eqToHom_whiskerRight {X Y : C} (f : X = Y) (Z : C) : + eqToHom f β–· Z = eqToHom (congr_argβ‚‚ tensorObj f rfl) := by + cases f + simp only [id_whiskerRight, eqToHom_refl] @[reassoc] theorem associator_naturality_left {X X' : C} (f : X ⟢ X') (Y Z : C) : @@ -541,6 +449,89 @@ theorem whiskerLeft_iff {X Y : C} (f g : X ⟢ Y) : πŸ™_ C ◁ f = πŸ™_ C ◁ theorem whiskerRight_iff {X Y : C} (f g : X ⟢ Y) : f β–· πŸ™_ C = g β–· πŸ™_ C ↔ f = g := by simp +/-! The lemmas in the next section are true by coherence, +but we prove them directly as they are used in proving the coherence theorem. -/ + +section + +@[reassoc (attr := simp)] +theorem pentagon_inv : + W ◁ (Ξ±_ X Y Z).inv ≫ (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z = + (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv := + eq_of_inv_eq_inv (by simp) + +@[reassoc] +theorem pentagon_inv' (W X Y Z : C) : + (πŸ™ W βŠ— (Ξ±_ X Y Z).inv) ≫ (Ξ±_ W (X βŠ— Y) Z).inv ≫ ((Ξ±_ W X Y).inv βŠ— πŸ™ Z) = + (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv := + CategoryTheory.eq_of_inv_eq_inv (by simp [pentagon]) +#align category_theory.monoidal_category.pentagon_inv CategoryTheory.MonoidalCategory.pentagon_inv' + +@[reassoc (attr := simp)] +theorem pentagon_inv_inv_hom_hom_inv : + (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z ≫ (Ξ±_ (W βŠ— X) Y Z).hom = + W ◁ (Ξ±_ X Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).inv := by + rw [← cancel_epi (W ◁ (Ξ±_ X Y Z).inv), ← cancel_mono (Ξ±_ (W βŠ— X) Y Z).inv] + simp + +@[reassoc (attr := simp)] +theorem pentagon_inv_hom_hom_hom_inv : + (Ξ±_ (W βŠ— X) Y Z).inv ≫ (Ξ±_ W X Y).hom β–· Z ≫ (Ξ±_ W (X βŠ— Y) Z).hom = + (Ξ±_ W X (Y βŠ— Z)).hom ≫ W ◁ (Ξ±_ X Y Z).inv := + eq_of_inv_eq_inv (by simp) + +@[reassoc (attr := simp)] +theorem pentagon_hom_inv_inv_inv_inv : + W ◁ (Ξ±_ X Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv = + (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z := + by simp [← cancel_epi (W ◁ (Ξ±_ X Y Z).inv)] + +@[reassoc (attr := simp)] +theorem pentagon_hom_hom_inv_hom_hom : + (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom ≫ W ◁ (Ξ±_ X Y Z).inv = + (Ξ±_ W X Y).hom β–· Z ≫ (Ξ±_ W (X βŠ— Y) Z).hom := + eq_of_inv_eq_inv (by simp) + +@[reassoc (attr := simp)] +theorem pentagon_hom_inv_inv_inv_hom : + (Ξ±_ W X (Y βŠ— Z)).hom ≫ W ◁ (Ξ±_ X Y Z).inv ≫ (Ξ±_ W (X βŠ— Y) Z).inv = + (Ξ±_ (W βŠ— X) Y Z).inv ≫ (Ξ±_ W X Y).hom β–· Z := by + rw [← cancel_epi (Ξ±_ W X (Y βŠ— Z)).inv, ← cancel_mono ((Ξ±_ W X Y).inv β–· Z)] + simp + +@[reassoc (attr := simp)] +theorem pentagon_hom_hom_inv_inv_hom : + (Ξ±_ W (X βŠ— Y) Z).hom ≫ W ◁ (Ξ±_ X Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).inv = + (Ξ±_ W X Y).inv β–· Z ≫ (Ξ±_ (W βŠ— X) Y Z).hom := + eq_of_inv_eq_inv (by simp) + +@[reassoc (attr := simp)] +theorem pentagon_inv_hom_hom_hom_hom : + (Ξ±_ W X Y).inv β–· Z ≫ (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom = + (Ξ±_ W (X βŠ— Y) Z).hom ≫ W ◁ (Ξ±_ X Y Z).hom := + by simp [← cancel_epi ((Ξ±_ W X Y).hom β–· Z)] + +@[reassoc (attr := simp)] +theorem pentagon_inv_inv_hom_inv_inv : + (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv ≫ (Ξ±_ W X Y).hom β–· Z = + W ◁ (Ξ±_ X Y Z).inv ≫ (Ξ±_ W (X βŠ— Y) Z).inv := + eq_of_inv_eq_inv (by simp) + +@[reassoc (attr := simp)] +theorem triangle_assoc_comp_right (X Y : C) : + (Ξ±_ X (πŸ™_ C) Y).inv ≫ ((ρ_ X).hom β–· Y) = X ◁ (Ξ»_ Y).hom := by + rw [← triangle, Iso.inv_hom_id_assoc] + +@[reassoc (attr := simp)] +theorem triangle_assoc_comp_right_inv (X Y : C) : + (ρ_ X).inv β–· Y ≫ (Ξ±_ X (πŸ™_ C) Y).hom = X ◁ (Ξ»_ Y).inv := by + simp [← cancel_mono (X ◁ (Ξ»_ Y).hom)] + +@[reassoc (attr := simp)] +theorem triangle_assoc_comp_left_inv (X Y : C) : + (X ◁ (Ξ»_ Y).inv) ≫ (Ξ±_ X (πŸ™_ C) Y).inv = (ρ_ X).inv β–· Y := by + simp [← cancel_mono ((ρ_ X).hom β–· Y)] + /-- We state it as a simp lemma, which is regarded as an involved version of `id_whiskerRight X Y : πŸ™ X β–· Y = πŸ™ (X βŠ— Y)`. -/ @@ -585,83 +576,14 @@ theorem rightUnitor_tensor (X Y : C) : theorem rightUnitor_tensor_inv (X Y : C) : (ρ_ (X βŠ— Y)).inv = X ◁ (ρ_ Y).inv ≫ (Ξ±_ X Y (πŸ™_ C)).inv := by simp +end + @[reassoc] theorem associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (f₃ : X₃ ⟢ Y₃) : ((f₁ βŠ— fβ‚‚) βŠ— f₃) ≫ (Ξ±_ Y₁ Yβ‚‚ Y₃).hom = (Ξ±_ X₁ Xβ‚‚ X₃).hom ≫ (f₁ βŠ— (fβ‚‚ βŠ— f₃)) := by simp [tensorHom_def] -@[reassoc] -theorem triangle' (X Y : C) : - (Ξ±_ X (πŸ™_ C) Y).hom ≫ (πŸ™ X βŠ— (Ξ»_ Y).hom) = (ρ_ X).hom βŠ— πŸ™ Y := by - simp -#align category_theory.monoidal_category.triangle CategoryTheory.MonoidalCategory.triangle' - -@[reassoc] -theorem pentagon' (W X Y Z : C) : - ((Ξ±_ W X Y).hom βŠ— πŸ™ Z) ≫ (Ξ±_ W (X βŠ— Y) Z).hom ≫ (πŸ™ W βŠ— (Ξ±_ X Y Z).hom) = - (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom := by - simp -#align category_theory.monoidal_category.pentagon CategoryTheory.MonoidalCategory.pentagon' - -@[reassoc] -theorem rightUnitor_tensor' (X Y : C) : - (ρ_ (X βŠ— Y)).hom = (Ξ±_ X Y (πŸ™_ C)).hom ≫ (πŸ™ X βŠ— (ρ_ Y).hom) := by - simp -#align category_theory.monoidal_category.right_unitor_tensor CategoryTheory.MonoidalCategory.rightUnitor_tensor' - -@[reassoc] -theorem rightUnitor_tensor_inv' (X Y : C) : - (ρ_ (X βŠ— Y)).inv = (πŸ™ X βŠ— (ρ_ Y).inv) ≫ (Ξ±_ X Y (πŸ™_ C)).inv := - eq_of_inv_eq_inv (by simp) -#align category_theory.monoidal_category.right_unitor_tensor_inv CategoryTheory.MonoidalCategory.rightUnitor_tensor_inv' - -@[reassoc] -theorem triangle_assoc_comp_right' (X Y : C) : - (Ξ±_ X (πŸ™_ C) Y).inv ≫ ((ρ_ X).hom βŠ— πŸ™ Y) = πŸ™ X βŠ— (Ξ»_ Y).hom := by - simp -#align category_theory.monoidal_category.triangle_assoc_comp_right CategoryTheory.MonoidalCategory.triangle_assoc_comp_right' - -@[reassoc] -theorem triangle_assoc_comp_left_inv' (X Y : C) : - (πŸ™ X βŠ— (Ξ»_ Y).inv) ≫ (Ξ±_ X (πŸ™_ C) Y).inv = (ρ_ X).inv βŠ— πŸ™ Y := by - simp -#align category_theory.monoidal_category.triangle_assoc_comp_left_inv CategoryTheory.MonoidalCategory.triangle_assoc_comp_left_inv' - -theorem rightUnitor_conjugation {X Y : C} (f : X ⟢ Y) : - f βŠ— πŸ™ (πŸ™_ C) = (ρ_ X).hom ≫ f ≫ (ρ_ Y).inv := by - simp -#align category_theory.monoidal_category.right_unitor_conjugation CategoryTheory.MonoidalCategory.rightUnitor_conjugation - -theorem leftUnitor_conjugation {X Y : C} (f : X ⟢ Y) : - πŸ™ (πŸ™_ C) βŠ— f = (Ξ»_ X).hom ≫ f ≫ (Ξ»_ Y).inv := by - simp -#align category_theory.monoidal_category.left_unitor_conjugation CategoryTheory.MonoidalCategory.leftUnitor_conjugation - -@[reassoc] -theorem leftUnitor_naturality' {X Y : C} (f : X ⟢ Y) : - (πŸ™ (πŸ™_ C) βŠ— f) ≫ (Ξ»_ Y).hom = (Ξ»_ X).hom ≫ f := - by simp - -@[reassoc] -theorem rightUnitor_naturality' {X Y : C} (f : X ⟢ Y) : - (f βŠ— πŸ™ (πŸ™_ C)) ≫ (ρ_ Y).hom = (ρ_ X).hom ≫ f := by - simp - -@[reassoc] -theorem leftUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : - f ≫ (Ξ»_ X').inv = (Ξ»_ X).inv ≫ (πŸ™ _ βŠ— f) := by - simp -#align category_theory.monoidal_category.left_unitor_inv_naturality CategoryTheory.MonoidalCategory.leftUnitor_inv_naturality' - -@[reassoc] -theorem rightUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : - f ≫ (ρ_ X').inv = (ρ_ X).inv ≫ (f βŠ— πŸ™ _) := by - simp -#align category_theory.monoidal_category.right_unitor_inv_naturality CategoryTheory.MonoidalCategory.rightUnitor_inv_naturality' - -end - @[reassoc] theorem associator_inv_naturality {X Y Z X' Y' Z' : C} (f : X ⟢ X') (g : Y ⟢ Y') (h : Z ⟢ Z') : (f βŠ— g βŠ— h) ≫ (Ξ±_ X' Y' Z').inv = (Ξ±_ X Y Z).inv ≫ ((f βŠ— g) βŠ— h) := by @@ -729,6 +651,10 @@ theorem tensor_inv_hom_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) rw [← tensor_comp, IsIso.inv_hom_id, tensorHom_id, comp_whiskerRight] #align category_theory.monoidal_category.tensor_inv_hom_id' CategoryTheory.MonoidalCategory.tensor_inv_hom_id' +/-- +A constructor for monoidal categaories that requires `tensorHom` instead of `whiskerLeft` and +`whiskerRight`. +-/ def ofTensorHom (tensorObj : C β†’ C β†’ C) (tensorHom : βˆ€ {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C}, (X₁ ⟢ Y₁) β†’ (Xβ‚‚ ⟢ Yβ‚‚) β†’ (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚)) @@ -809,6 +735,111 @@ def ofTensorHom pentagon := by intros; simp [← id_tensorHom, ← tensorHom_id, pentagon] triangle := by intros; simp [← id_tensorHom, ← tensorHom_id, triangle] +section + +/- The lemmas of this section might be redundant because they should be stated in terms of the +whiskering operators. We leave them in order not to break proofs that existed before we +have the whiskering operators. -/ + +@[reassoc] +theorem comp_tensor_id (f : W ⟢ X) (g : X ⟢ Y) : f ≫ g βŠ— πŸ™ Z = (f βŠ— πŸ™ Z) ≫ (g βŠ— πŸ™ Z) := by + simp +#align category_theory.monoidal_category.comp_tensor_id CategoryTheory.MonoidalCategory.comp_tensor_id + +@[reassoc] +theorem id_tensor_comp (f : W ⟢ X) (g : X ⟢ Y) : πŸ™ Z βŠ— f ≫ g = (πŸ™ Z βŠ— f) ≫ (πŸ™ Z βŠ— g) := by + simp +#align category_theory.monoidal_category.id_tensor_comp CategoryTheory.MonoidalCategory.id_tensor_comp + +@[reassoc] +theorem id_tensor_comp_tensor_id (f : W ⟢ X) (g : Y ⟢ Z) : (πŸ™ Y βŠ— f) ≫ (g βŠ— πŸ™ X) = g βŠ— f := by + rw [← tensor_comp] + simp +#align category_theory.monoidal_category.id_tensor_comp_tensor_id CategoryTheory.MonoidalCategory.id_tensor_comp_tensor_id + +@[reassoc] +theorem tensor_id_comp_id_tensor (f : W ⟢ X) (g : Y ⟢ Z) : (g βŠ— πŸ™ W) ≫ (πŸ™ Z βŠ— f) = g βŠ— f := by + rw [← tensor_comp] + simp +#align category_theory.monoidal_category.tensor_id_comp_id_tensor CategoryTheory.MonoidalCategory.tensor_id_comp_id_tensor + +theorem tensor_left_iff {X Y : C} (f g : X ⟢ Y) : πŸ™ (πŸ™_ C) βŠ— f = πŸ™ (πŸ™_ C) βŠ— g ↔ f = g := by simp +#align category_theory.monoidal_category.tensor_left_iff CategoryTheory.MonoidalCategory.tensor_left_iff + +theorem tensor_right_iff {X Y : C} (f g : X ⟢ Y) : f βŠ— πŸ™ (πŸ™_ C) = g βŠ— πŸ™ (πŸ™_ C) ↔ f = g := by simp +#align category_theory.monoidal_category.tensor_right_iff CategoryTheory.MonoidalCategory.tensor_right_iff + +@[reassoc] +theorem triangle' (X Y : C) : + (Ξ±_ X (πŸ™_ C) Y).hom ≫ (πŸ™ X βŠ— (Ξ»_ Y).hom) = (ρ_ X).hom βŠ— πŸ™ Y := by + simp +#align category_theory.monoidal_category.triangle CategoryTheory.MonoidalCategory.triangle' + +@[reassoc] +theorem pentagon' (W X Y Z : C) : + ((Ξ±_ W X Y).hom βŠ— πŸ™ Z) ≫ (Ξ±_ W (X βŠ— Y) Z).hom ≫ (πŸ™ W βŠ— (Ξ±_ X Y Z).hom) = + (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom := by + simp +#align category_theory.monoidal_category.pentagon CategoryTheory.MonoidalCategory.pentagon' + +@[reassoc] +theorem rightUnitor_tensor' (X Y : C) : + (ρ_ (X βŠ— Y)).hom = (Ξ±_ X Y (πŸ™_ C)).hom ≫ (πŸ™ X βŠ— (ρ_ Y).hom) := by + simp +#align category_theory.monoidal_category.right_unitor_tensor CategoryTheory.MonoidalCategory.rightUnitor_tensor' + +@[reassoc] +theorem rightUnitor_tensor_inv' (X Y : C) : + (ρ_ (X βŠ— Y)).inv = (πŸ™ X βŠ— (ρ_ Y).inv) ≫ (Ξ±_ X Y (πŸ™_ C)).inv := + eq_of_inv_eq_inv (by simp) +#align category_theory.monoidal_category.right_unitor_tensor_inv CategoryTheory.MonoidalCategory.rightUnitor_tensor_inv' + +@[reassoc] +theorem triangle_assoc_comp_right' (X Y : C) : + (Ξ±_ X (πŸ™_ C) Y).inv ≫ ((ρ_ X).hom βŠ— πŸ™ Y) = πŸ™ X βŠ— (Ξ»_ Y).hom := by + simp +#align category_theory.monoidal_category.triangle_assoc_comp_right CategoryTheory.MonoidalCategory.triangle_assoc_comp_right' + +@[reassoc] +theorem triangle_assoc_comp_left_inv' (X Y : C) : + (πŸ™ X βŠ— (Ξ»_ Y).inv) ≫ (Ξ±_ X (πŸ™_ C) Y).inv = (ρ_ X).inv βŠ— πŸ™ Y := by + simp +#align category_theory.monoidal_category.triangle_assoc_comp_left_inv CategoryTheory.MonoidalCategory.triangle_assoc_comp_left_inv' + +theorem rightUnitor_conjugation {X Y : C} (f : X ⟢ Y) : + f βŠ— πŸ™ (πŸ™_ C) = (ρ_ X).hom ≫ f ≫ (ρ_ Y).inv := by + simp +#align category_theory.monoidal_category.right_unitor_conjugation CategoryTheory.MonoidalCategory.rightUnitor_conjugation + +theorem leftUnitor_conjugation {X Y : C} (f : X ⟢ Y) : + πŸ™ (πŸ™_ C) βŠ— f = (Ξ»_ X).hom ≫ f ≫ (Ξ»_ Y).inv := by + simp +#align category_theory.monoidal_category.left_unitor_conjugation CategoryTheory.MonoidalCategory.leftUnitor_conjugation + +@[reassoc] +theorem leftUnitor_naturality' {X Y : C} (f : X ⟢ Y) : + (πŸ™ (πŸ™_ C) βŠ— f) ≫ (Ξ»_ Y).hom = (Ξ»_ X).hom ≫ f := + by simp + +@[reassoc] +theorem rightUnitor_naturality' {X Y : C} (f : X ⟢ Y) : + (f βŠ— πŸ™ (πŸ™_ C)) ≫ (ρ_ Y).hom = (ρ_ X).hom ≫ f := by + simp + +@[reassoc] +theorem leftUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : + f ≫ (Ξ»_ X').inv = (Ξ»_ X).inv ≫ (πŸ™ _ βŠ— f) := by + simp +#align category_theory.monoidal_category.left_unitor_inv_naturality CategoryTheory.MonoidalCategory.leftUnitor_inv_naturality' + +@[reassoc] +theorem rightUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : + f ≫ (ρ_ X').inv = (ρ_ X).inv ≫ (f βŠ— πŸ™ _) := by + simp +#align category_theory.monoidal_category.right_unitor_inv_naturality CategoryTheory.MonoidalCategory.rightUnitor_inv_naturality' + +end + end section diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index 67170c93ea0df..9e2287b224649 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -17,18 +17,19 @@ We define `Center C` to be pairs `⟨X, b⟩`, where `X : C` and `b` is a half-b We show that `Center C` is braided monoidal, and provide the monoidal functor `Center.forget` from `Center C` back to `C`. -## Future work +## Implementation notes -Verifying the various axioms here is done by tedious rewriting. +Verifying the various axioms directly requires tedious rewriting. Using the `slice` tactic may make the proofs marginally more readable. More exciting, however, would be to make possible one of the following options: 1. Integration with homotopy.io / globular to give "picture proofs". 2. The monoidal coherence theorem, so we can ignore associators - (after which most of these proofs are trivial; - I'm unsure if the monoidal coherence theorem is even usable in dependent type theory). + (after which most of these proofs are trivial). 3. Automating these proofs using `rewrite_search` or some relative. +In this file, we take the second approach using the monoidal composition `βŠ—β‰«` and the +`coherence` tactic. -/ @@ -140,12 +141,13 @@ def tensorObj (X Y : Center C) : Center C := simp only [HalfBraiding.monoidal] -- We'd like to commute `X.1 ◁ U ◁ (HalfBraiding.Ξ² Y.2 U').hom` -- and `((HalfBraiding.Ξ² X.2 U).hom β–· U' β–· Y.1)` past each other. - -- We do this with the help of the monoidal composition `βŠ—β‰«`. + -- We do this with the help of the monoidal composition `βŠ—β‰«` and the `coherence` tactic. calc _ = πŸ™ _ βŠ—β‰« X.1 ◁ (HalfBraiding.Ξ² Y.2 U).hom β–· U' βŠ—β‰« - (_ ◁ (HalfBraiding.Ξ² Y.2 U').hom ≫ (HalfBraiding.Ξ² X.2 U).hom β–· _) βŠ—β‰« - U ◁ (HalfBraiding.Ξ² X.2 U').hom β–· Y.1 βŠ—β‰« πŸ™ _ := ?eq1 + _ ◁ (HalfBraiding.Ξ² Y.2 U').hom ≫ + (HalfBraiding.Ξ² X.2 U).hom β–· _ βŠ—β‰« + U ◁ (HalfBraiding.Ξ² X.2 U').hom β–· Y.1 βŠ—β‰« πŸ™ _ := ?eq1 _ = _ := ?eq2 case eq1 => coherence case eq2 => rw [whisker_exchange]; coherence diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean index aa4b533ef919b..543f8997edb1b 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean @@ -41,24 +41,6 @@ universe u namespace CategoryTheory -namespace MonoidalCategory - -variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] - -@[simp] -theorem whiskerLeft_eqToHom (X : C) {Y Z : C} (f : Y = Z) : - X ◁ eqToHom f = eqToHom (congr_argβ‚‚ tensorObj rfl f) := by - cases f - simp only [whiskerLeft_id, eqToHom_refl] - -@[simp] -theorem eqToHom_whiskerRight {X Y : C} (f : X = Y) (Z : C) : - eqToHom f β–· Z = eqToHom (congr_argβ‚‚ tensorObj f rfl) := by - cases f - simp only [id_whiskerRight, eqToHom_refl] - -end MonoidalCategory - open MonoidalCategory namespace FreeMonoidalCategory diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index 4280b1eefc1a4..f7c60052ae393 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -97,7 +97,6 @@ structure LaxMonoidalFunctor extends C β₯€ D where initialize_simps_projections LaxMonoidalFunctor (+toFunctor, -obj, -map) --Porting note: was `[simp, reassoc.1]` --- attribute [reassoc (attr := simp)] LaxMonoidalFunctor.ΞΌ_natural attribute [reassoc (attr := simp)] LaxMonoidalFunctor.ΞΌ_natural_left attribute [reassoc (attr := simp)] LaxMonoidalFunctor.ΞΌ_natural_right diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean index 2a96d08165e06..011d618df61b4 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean @@ -31,28 +31,10 @@ def exactPairingOfFaithful [Faithful F.toFunctor] {X Y : C} (eval : Y βŠ— X ⟢ (map_coeval : F.map coeval = inv F.Ξ΅ ≫ Ξ·_ _ _ ≫ F.ΞΌ _ _) : ExactPairing X Y where evaluation' := eval coevaluation' := coeval - evaluation_coevaluation' := - F.toFunctor.map_injective (by - simp_rw [Functor.map_comp, F.map_whiskerRight, F.map_whiskerLeft, map_eval, map_coeval] - simp only [comp_whiskerRight, Category.assoc, MonoidalCategory.whiskerLeft_comp, - LaxMonoidalFunctor.associativity_assoc, IsIso.hom_inv_id_assoc, IsIso.inv_comp_eq] - simp_rw [← whiskerLeft_comp_assoc] - simp only [IsIso.hom_inv_id_assoc, MonoidalCategory.whiskerLeft_comp, Category.assoc, - ExactPairing.evaluation_coevaluation_assoc, LaxMonoidalFunctor.left_unitality, - LaxMonoidalFunctor.right_unitality_inv] - simp_rw [← comp_whiskerRight_assoc] - simp ) - coevaluation_evaluation' := - F.toFunctor.map_injective (by - simp_rw [Functor.map_comp, F.map_whiskerRight, F.map_whiskerLeft, map_eval, map_coeval] - simp only [MonoidalCategory.whiskerLeft_comp, Category.assoc, comp_whiskerRight, - LaxMonoidalFunctor.associativity_inv_assoc, IsIso.hom_inv_id_assoc, IsIso.inv_comp_eq] - simp_rw [← comp_whiskerRight_assoc] - simp only [IsIso.hom_inv_id_assoc, comp_whiskerRight, Category.assoc, - ExactPairing.coevaluation_evaluation_assoc, - LaxMonoidalFunctor.right_unitality, LaxMonoidalFunctor.left_unitality_inv] - simp_rw [← whiskerLeft_comp_assoc] - simp ) + evaluation_coevaluation' := F.toFunctor.map_injective <| by + simp [map_eval, map_coeval, F.map_whiskerRight, F.map_whiskerLeft] + coevaluation_evaluation' := F.toFunctor.map_injective <| by + simp [map_eval, map_coeval, F.map_whiskerRight, F.map_whiskerLeft] #align category_theory.exact_pairing_of_faithful CategoryTheory.exactPairingOfFaithful /-- Given a pair of objects which are sent by a fully faithful functor to a pair of objects diff --git a/Mathlib/RepresentationTheory/FdRep.lean b/Mathlib/RepresentationTheory/FdRep.lean index 8eb46ff0ae633..8f6ef4fe94932 100644 --- a/Mathlib/RepresentationTheory/FdRep.lean +++ b/Mathlib/RepresentationTheory/FdRep.lean @@ -175,7 +175,7 @@ noncomputable def dualTensorIsoLinHomAux : (FdRep.of ρV.dual βŠ— W).V β‰… (FdRep.of (linHom ρV W.ρ)).V := -- Porting note: had to make all types explicit arguments @LinearEquiv.toFGModuleCatIso k _ (FdRep.of ρV.dual βŠ— W).V (V β†’β‚—[k] W) - _ _ _ _ _ _ (dualTensorHomEquiv k V W) + _ _ _ _ _ _ (dualTensorHomEquiv k V W) #align fdRep.dual_tensor_iso_lin_hom_aux FdRep.dualTensorIsoLinHomAux /-- When `V` and `W` are finite dimensional representations of a group `G`, the isomorphism diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index 214c2f7010be8..8586932ae91cc 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -397,6 +397,7 @@ elab_rules : tactic evalTactic (← `(tactic| simp only [bicategoricalComp]; simp only [monoidalComp]; - try (first | whisker_simps | monoidal_simps); + (try whisker_simps); + (try monoidal_simps); )) coherence_loop From 59cd705a64ebc31f0bce67888a9fb1ec80331a5d Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sat, 5 Aug 2023 22:52:09 +0900 Subject: [PATCH 14/62] Update Mathlib/CategoryTheory/Monoidal/Bimod.lean Co-authored-by: Scott Morrison --- Mathlib/CategoryTheory/Monoidal/Bimod.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index bc7f2c88c3fb0..552fb40544333 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -300,7 +300,8 @@ noncomputable def actRight : X P Q βŠ— T.X ⟢ X P Q := dsimp simp only [comp_whiskerRight, whisker_assoc, Category.assoc, Iso.inv_hom_id_assoc] slice_lhs 3 4 => - rw [← MonoidalCategory.whiskerLeft_comp, middle_assoc, MonoidalCategory.whiskerLeft_comp] + rw [← MonoidalCategory.whiskerLeft_comp, middle_assoc, + MonoidalCategory.whiskerLeft_comp] simp)) set_option linter.uppercaseLean3 false in #align Bimod.tensor_Bimod.act_right Bimod.TensorBimod.actRight From 6e0afadbd759f575ceb9bcaed6547003b00e5884 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 6 Aug 2023 17:06:46 +0900 Subject: [PATCH 15/62] Inline one-line proofs --- Mathlib/CategoryTheory/Monoidal/Braided.lean | 6 ++-- Mathlib/CategoryTheory/Monoidal/Center.lean | 36 ++++++++------------ 2 files changed, 16 insertions(+), 26 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 5e0dab43ec74e..5bd592f6126fe 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -548,10 +548,8 @@ theorem tensor_associativity (X₁ Xβ‚‚ Y₁ Yβ‚‚ Z₁ Zβ‚‚ : C) : calc _ = πŸ™ _ βŠ—β‰« X₁ ◁ ((Ξ²_ Xβ‚‚ Y₁).hom β–· (Yβ‚‚ βŠ— Z₁) ≫ (Y₁ βŠ— Xβ‚‚) ◁ (Ξ²_ Yβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« - X₁ ◁ Y₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom β–· Yβ‚‚ β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := ?eq1 - _ = _ := ?eq2 - case eq1 => coherence - case eq2 => rw [← whisker_exchange]; coherence + X₁ ◁ Y₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom β–· Yβ‚‚ β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := by coherence + _ = _ := by rw [← whisker_exchange]; coherence #align category_theory.tensor_associativity CategoryTheory.tensor_associativity /-- The tensor product functor from `C Γ— C` to `C` as a monoidal functor. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index 9e2287b224649..3a9dfc9ed73bb 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -147,23 +147,19 @@ def tensorObj (X Y : Center C) : Center C := X.1 ◁ (HalfBraiding.Ξ² Y.2 U).hom β–· U' βŠ—β‰« _ ◁ (HalfBraiding.Ξ² Y.2 U').hom ≫ (HalfBraiding.Ξ² X.2 U).hom β–· _ βŠ—β‰« - U ◁ (HalfBraiding.Ξ² X.2 U').hom β–· Y.1 βŠ—β‰« πŸ™ _ := ?eq1 - _ = _ := ?eq2 - case eq1 => coherence - case eq2 => rw [whisker_exchange]; coherence + U ◁ (HalfBraiding.Ξ² X.2 U').hom β–· Y.1 βŠ—β‰« πŸ™ _ := by coherence + _ = _ := by rw [whisker_exchange]; coherence naturality := fun {U U'} f => by dsimp only [Iso.trans_hom, whiskerLeftIso_hom, Iso.symm_hom, whiskerRightIso_hom] calc _ = πŸ™ _ βŠ—β‰« (X.1 ◁ (Y.1 ◁ f ≫ (HalfBraiding.Ξ² Y.2 U').hom)) βŠ—β‰« - (HalfBraiding.Ξ² X.2 U').hom β–· Y.1 βŠ—β‰« πŸ™ _ := ?eq1 + (HalfBraiding.Ξ² X.2 U').hom β–· Y.1 βŠ—β‰« πŸ™ _ := by coherence _ = πŸ™ _ βŠ—β‰« X.1 ◁ (HalfBraiding.Ξ² Y.2 U).hom βŠ—β‰« - (X.1 ◁ f ≫ (HalfBraiding.Ξ² X.2 U').hom) β–· Y.1 βŠ—β‰« πŸ™ _ := ?eq2 - _ = _ := ?eq3 - case eq1 => coherence - case eq2 => rw [HalfBraiding.naturality]; coherence - case eq3 => rw [HalfBraiding.naturality]; coherence }⟩ + (X.1 ◁ f ≫ (HalfBraiding.Ξ² X.2 U').hom) β–· Y.1 βŠ—β‰« πŸ™ _ := by + rw [HalfBraiding.naturality]; coherence + _ = _ := by rw [HalfBraiding.naturality]; coherence }⟩ #align category_theory.center.tensor_obj CategoryTheory.Center.tensorObj @[reassoc] @@ -175,14 +171,12 @@ theorem whiskerLeft_comm (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Y calc _ = πŸ™ _ βŠ—β‰« X.fst ◁ (f.f β–· U ≫ (HalfBraiding.Ξ² Yβ‚‚.snd U).hom) βŠ—β‰« - (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst βŠ—β‰« πŸ™ _ := ?eq1 + (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst βŠ—β‰« πŸ™ _ := by coherence _ = πŸ™ _ βŠ—β‰« X.fst ◁ (HalfBraiding.Ξ² Y₁.snd U).hom βŠ—β‰« - ((X.fst βŠ— U) ◁ f.f ≫ (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst) βŠ—β‰« πŸ™ _ := ?eq2 - _ = _ := ?eq3 - case eq1 => coherence - case eq2 => rw [f.comm]; coherence - case eq3 => rw [whisker_exchange]; coherence + ((X.fst βŠ— U) ◁ f.f ≫ (HalfBraiding.Ξ² X.snd U).hom β–· Yβ‚‚.fst) βŠ—β‰« πŸ™ _ := by + rw [f.comm]; coherence + _ = _ := by rw [whisker_exchange]; coherence /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ def whiskerLeft (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) : @@ -199,14 +193,12 @@ theorem whiskerRight_comm {X₁ Xβ‚‚: Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center calc _ = πŸ™ _ βŠ—β‰« (f.f β–· (Y.fst βŠ— U) ≫ Xβ‚‚.fst ◁ (HalfBraiding.Ξ² Y.snd U).hom) βŠ—β‰« - (HalfBraiding.Ξ² Xβ‚‚.snd U).hom β–· Y.fst βŠ—β‰« πŸ™ _ := ?eq1 + (HalfBraiding.Ξ² Xβ‚‚.snd U).hom β–· Y.fst βŠ—β‰« πŸ™ _ := by coherence _ = πŸ™ _ βŠ—β‰« X₁.fst ◁ (HalfBraiding.Ξ² Y.snd U).hom βŠ—β‰« - (f.f β–· U ≫ (HalfBraiding.Ξ² Xβ‚‚.snd U).hom) β–· Y.fst βŠ—β‰« πŸ™ _ := ?eq2 - _ = _ := ?eq3 - case eq1 => coherence - case eq2 => rw [← whisker_exchange]; coherence - case eq3 => rw [f.comm]; coherence + (f.f β–· U ≫ (HalfBraiding.Ξ² Xβ‚‚.snd U).hom) β–· Y.fst βŠ—β‰« πŸ™ _ := by + rw [← whisker_exchange]; coherence + _ = _ := by rw [f.comm]; coherence /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ def whiskerRight {X₁ Xβ‚‚ : Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) : From 37c7ddabdb276d636ee811b32a0c1a7ccec455ba Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 6 Aug 2023 17:07:59 +0900 Subject: [PATCH 16/62] delete a draft proof --- Mathlib/CategoryTheory/Monoidal/Braided.lean | 17 ----------------- 1 file changed, 17 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 5bd592f6126fe..3884b4222145c 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -456,23 +456,6 @@ def tensor_ΞΌ (X Y : C Γ— C) : (X.1 βŠ— X.2) βŠ— Y.1 βŠ— Y.2 ⟢ (X.1 βŠ— Y.1) (X.1 ◁ (Ξ±_ Y.1 X.2 Y.2).hom) ≫ (Ξ±_ X.1 Y.1 (X.2 βŠ— Y.2)).inv #align category_theory.tensor_ΞΌ CategoryTheory.tensor_ΞΌ --- theorem tensor_ΞΌ_natural_left {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁: X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (Z₁ Zβ‚‚ : C) : --- (f₁ βŠ— fβ‚‚) β–· (Z₁ βŠ— Zβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚) = --- tensor_ΞΌ C (X₁, Xβ‚‚) (Z₁, Zβ‚‚) ≫ (f₁ β–· Z₁ βŠ— fβ‚‚ β–· Zβ‚‚) := by --- dsimp only [tensor_ΞΌ, prodMonoidal_tensorObj, tensor_obj] --- calc --- _ = πŸ™ _ βŠ—β‰« --- f₁ β–· Xβ‚‚ β–· Z₁ β–· Zβ‚‚ βŠ—β‰« Y₁ ◁ (fβ‚‚ β–· Z₁ ≫ (Ξ²_ Yβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := ?eq1 --- _ = πŸ™ _ βŠ—β‰« --- (f₁ β–· (Xβ‚‚ βŠ— Z₁) ≫ Y₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom) β–· Zβ‚‚ βŠ—β‰« Y₁ ◁ Z₁ ◁ fβ‚‚ β–· Zβ‚‚ βŠ—β‰« πŸ™ _ := ?eq2 --- _ = πŸ™ _ βŠ—β‰« --- X₁ ◁ (Ξ²_ Xβ‚‚ Z₁).hom β–· Zβ‚‚ βŠ—β‰« (f₁ β–· Z₁ β–· (Xβ‚‚ βŠ— Zβ‚‚) ≫ (Y₁ βŠ— Z₁) ◁ fβ‚‚ β–· Zβ‚‚) βŠ—β‰« πŸ™ _ := ?eq3 --- _ = _ := ?eq4 --- case eq1 => rw [tensorHom_def']; coherence --- case eq2 => rw [braiding_naturality_left]; coherence --- case eq3 => rw [← whisker_exchange]; coherence --- case eq4 => rw [tensorHom_def']; coherence - theorem tensor_ΞΌ_natural {X₁ Xβ‚‚ Y₁ Yβ‚‚ U₁ Uβ‚‚ V₁ Vβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : U₁ ⟢ V₁) (gβ‚‚ : Uβ‚‚ ⟢ Vβ‚‚) : ((f₁ βŠ— fβ‚‚) βŠ— g₁ βŠ— gβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (V₁, Vβ‚‚) = From 4ac65e4b02b5446a5ce037b558bd1fc8b2296a92 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 6 Aug 2023 17:12:12 +0900 Subject: [PATCH 17/62] fix typos Co-authored-by: Scott Morrison --- Mathlib/CategoryTheory/Monoidal/Category.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 4c94032bcf674..dae21098de8bb 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -303,7 +303,7 @@ theorem inv_whiskerLeft (X : C) {Y Z : C} (f : Y ⟢ Z) [IsIso f] : inv (X ◁ f) = X ◁ inv f := by aesop_cat -/-- The right whiskering of a 2-isomorphism is a 2-isomorphism. -/ +/-- The right whiskering of an isomorphism is an isomorphism. -/ @[simps!] def whiskerRightIso {X Y : C} (f : X β‰… Y) (Z : C) : X βŠ— Z β‰… Y βŠ— Z where hom := f.hom β–· Z From 2642550c86136e7ec7a41494b0f800fe47c89235 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 6 Aug 2023 17:12:28 +0900 Subject: [PATCH 18/62] Update Mathlib/CategoryTheory/Monoidal/Category.lean fix typos Co-authored-by: Scott Morrison --- Mathlib/CategoryTheory/Monoidal/Category.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index dae21098de8bb..f3fe105ea6bf3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -145,7 +145,7 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] where whiskerRight (whiskerLeft X f) Z = (associator X Y Z).hom ≫ whiskerLeft X (whiskerRight f Z) ≫ (associator X Y' Z).inv := by aesop_cat - /-- The exchnage identity for the left and right whiskerings -/ + /-- The exchange identity for the left and right whiskerings -/ whisker_exchange : βˆ€ {W X Y Z : C} (f : W ⟢ X) (g : Y ⟢ Z), whiskerLeft W g ≫ whiskerRight f Z = whiskerRight f Y ≫ whiskerLeft X g := by From 029b506a6f059e3ec5eec10d924933ded3f2b5ca Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 6 Aug 2023 17:12:44 +0900 Subject: [PATCH 19/62] fix typos Co-authored-by: Scott Morrison --- Mathlib/CategoryTheory/Monoidal/Category.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index f3fe105ea6bf3..3e78754bf40bd 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -289,7 +289,7 @@ theorem inv_hom_whiskerRight' {X Y : C} (f : X ⟢ Y) [IsIso f] (Z : C) : inv f β–· Z ≫ f β–· Z = πŸ™ (Y βŠ— Z) := by rw [← comp_whiskerRight, IsIso.inv_hom_id, id_whiskerRight] -/-- The left whiskering of a 2-isomorphism is a 2-isomorphism. -/ +/-- The left whiskering of an isomorphism is an isomorphism. -/ @[simps] def whiskerLeftIso (X : C) {Y Z : C} (f : Y β‰… Z) : X βŠ— Y β‰… X βŠ— Z where hom := X ◁ f.hom From fd4e5cca6a7af8ec42d3393f18801acef615c4b5 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 7 Aug 2023 20:06:06 +0900 Subject: [PATCH 20/62] docs: notation for whiskerings --- Mathlib/CategoryTheory/Monoidal/Category.lean | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 3e78754bf40bd..592963e6092ad 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -26,9 +26,11 @@ e.g. `(Ξ»_ (πŸ™_ C)).hom = (ρ_ (πŸ™_ C)).hom` in `CategoryTheory.Monoidal.Coh ## Implementation notes -In the definition of monoidal categories, we also provide the whiskering operators `whiskerLeft` -and `whiskerRight`. These are products of an object and a morphism (the terminology "whiskering" -is borrowed from the 2-category theory). The tensor product of morphisms `tensorHom` can be defined +In the definition of monoidal categories, we also provide the whiskering operators: +* `whiskerLeft (X : C) {Y₁ Yβ‚‚ : C} (f : Y₁ ⟢ Yβ‚‚) : X βŠ— Y₁ ⟢ X βŠ— Yβ‚‚`, denoted by `X ◁ f`, +* `whiskerRight {X₁ Xβ‚‚ : C} (f : X₁ ⟢ Xβ‚‚) (Y : C) : X₁ βŠ— Y ⟢ Xβ‚‚ βŠ— Y`, denoted by `f β–· Y`. +These are products of an object and a morphism (the terminology "whiskering" +is borrowed from 2-category theory). The tensor product of morphisms `tensorHom` can be defined in terms of the whiskerings. There are two possible such definitions, which are related by the exchange property of the whiskerings. These two definitions are accessed by `tensorHom_def` and `tensorHom_def'`. By default, `tensorHom` is defined so that `tensorHom_def` holds From 35c69a8ee4086d6f643794d67ee21ccdf2466e9e Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 7 Aug 2023 20:07:08 +0900 Subject: [PATCH 21/62] fix docs Co-authored-by: Johan Commelin --- Mathlib/CategoryTheory/Monoidal/Category.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 592963e6092ad..e833c0c0de721 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -197,7 +197,7 @@ scoped infixr:70 " βŠ— " => MonoidalCategory.tensorObj /-- Notation for the `whiskerLeft` -/ scoped infixr:81 " ◁ " => whiskerLeft -/-- Notation for the `whiskerRight` -/ +/-- Notation for the `whiskerRight` operator of monoidal categories -/ scoped infixl:81 " β–· " => whiskerRight /-- Notation for `tensorHom`, the tensor product of morphisms in a monoidal category -/ From faf6921b259b07f72bde4a0f976d9a4545886cbe Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 7 Aug 2023 20:07:35 +0900 Subject: [PATCH 22/62] fix docs Co-authored-by: Johan Commelin --- Mathlib/CategoryTheory/Monoidal/Category.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index e833c0c0de721..37891f41ad0c5 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -44,7 +44,7 @@ The whiskerings are useful when considering simp-normal forms of morphisms in mo ### Simp-normal form for morphisms Rewriting involving associators and unitors could be very complicated. We try to ease this -complexity by putting carefully chosen simp lemmas that rewrite any morphisms into simp-normal +complexity by putting carefully chosen simp lemmas that rewrite any morphisms into the simp-normal form defined below. Rewriting into simp-normal form is especially useful in preprocessing performed by the `coherence` tactic. From a158ba79bd696287a37d3dcbbe48ba5c4554eb8d Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 7 Aug 2023 20:07:47 +0900 Subject: [PATCH 23/62] fix docs Co-authored-by: Johan Commelin --- Mathlib/CategoryTheory/Monoidal/Category.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 37891f41ad0c5..23463182f3d46 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -654,7 +654,7 @@ theorem tensor_inv_hom_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) #align category_theory.monoidal_category.tensor_inv_hom_id' CategoryTheory.MonoidalCategory.tensor_inv_hom_id' /-- -A constructor for monoidal categaories that requires `tensorHom` instead of `whiskerLeft` and +A constructor for monoidal categories that requires `tensorHom` instead of `whiskerLeft` and `whiskerRight`. -/ def ofTensorHom From ccfa744b730658d5bd35f9c55a230ff79b5b3f9e Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 7 Aug 2023 20:08:33 +0900 Subject: [PATCH 24/62] fix docs Co-authored-by: Johan Commelin --- Mathlib/CategoryTheory/Monoidal/Category.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 23463182f3d46..4852d8844511a 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -194,7 +194,7 @@ abbrev tensorUnit (C : Type u) [Category.{v} C] [MonoidalCategory C] : C := /-- Notation for `tensorObj`, the tensor product of objects in a monoidal category -/ scoped infixr:70 " βŠ— " => MonoidalCategory.tensorObj -/-- Notation for the `whiskerLeft` -/ +/-- Notation for the `whiskerLeft` operator of monoidal categories -/ scoped infixr:81 " ◁ " => whiskerLeft /-- Notation for the `whiskerRight` operator of monoidal categories -/ From 1089cdf0e4b6601fe6d73cdadf7a0d5b24133dfe Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 7 Aug 2023 20:09:31 +0900 Subject: [PATCH 25/62] fix docs --- Mathlib/CategoryTheory/Monoidal/Category.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 592963e6092ad..b229ac0d87624 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -223,7 +223,7 @@ theorem tensorHom_def' {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ whisker_exchange f g β–Έ tensorHom_def f g /-- Tensor product of identity maps is the identity: `(πŸ™ X₁ βŠ— πŸ™ Xβ‚‚) = πŸ™ (X₁ βŠ— Xβ‚‚)` -/ -theorem tensor_id (X₁ Xβ‚‚ : C) : tensorHom (πŸ™ X₁) (πŸ™ Xβ‚‚) = πŸ™ (tensorObj X₁ Xβ‚‚) := by +theorem tensor_id (X₁ Xβ‚‚ : C) : (πŸ™ X₁) βŠ— (πŸ™ Xβ‚‚) = πŸ™ (X₁ βŠ— Xβ‚‚) := by simp [tensorHom_def] /-- From e432ac4cc5cac410654e174151e4f90ddf84c3ff Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 7 Aug 2023 23:30:06 +0900 Subject: [PATCH 26/62] fix --- Mathlib/CategoryTheory/Monoidal/Category.lean | 94 ++----------------- 1 file changed, 8 insertions(+), 86 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index fdd3a0a7c17f1..6615a244bc35f 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -96,17 +96,11 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] where /-- right whiskering for morphisms -/ whiskerRight {X₁ Xβ‚‚ : C} (f : X₁ ⟢ Xβ‚‚) (Y : C) : tensorObj X₁ Y ⟢ tensorObj Xβ‚‚ Y /-- curried tensor product of morphisms -/ - tensorHom : βˆ€ {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C}, (X₁ ⟢ Y₁) β†’ (Xβ‚‚ ⟢ Yβ‚‚) β†’ (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚) - /-- Tensor product of identity maps is the identity: `(πŸ™ X₁ βŠ— πŸ™ Xβ‚‚) = πŸ™ (X₁ βŠ— Xβ‚‚)` -/ - tensor_id : βˆ€ X₁ Xβ‚‚ : C, tensorHom (πŸ™ X₁) (πŸ™ Xβ‚‚) = πŸ™ (tensorObj X₁ Xβ‚‚) := by aesop_cat - /-- - Composition of tensor products is tensor product of compositions: - `(f₁ βŠ— g₁) ∘ (fβ‚‚ βŠ— gβ‚‚) = (f₁ ∘ fβ‚‚) βŠ— (g₁ βŠ— gβ‚‚)` - -/ - tensor_comp : - βˆ€ {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : Y₁ ⟢ Z₁) (gβ‚‚ : Yβ‚‚ ⟢ Zβ‚‚), - tensorHom (f₁ ≫ g₁) (fβ‚‚ ≫ gβ‚‚) = tensorHom f₁ fβ‚‚ ≫ tensorHom g₁ gβ‚‚ := by - aesop_cat + tensorHom {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚) := + whiskerRight f Xβ‚‚ ≫ whiskerLeft Y₁ g + tensorHom_def {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : + tensorHom f g = whiskerRight f Xβ‚‚ ≫ whiskerLeft Y₁ g := by + aesop_cat -- Porting note: Adding a prime here, so I can later define `tensorUnit` unprimed with explicit -- argument `C` /-- The tensor unity in the monoidal structure `πŸ™_ C` -/ @@ -192,18 +186,12 @@ attribute [simp] -- Porting Note: This is here to make `tensorUnit` explicitly depend on `C`, which was done in -- Lean 3 using the `[]` notation in the `tensorUnit'` field. -/-- The tensor unity in the monoidal structure `πŸ™_C` -/ +/-- The tensor unity in the monoidal structure `πŸ™_ C` -/ abbrev tensorUnit (C : Type u) [Category.{v} C] [MonoidalCategory C] : C := tensorUnit' (C := C) /-- Notation for `tensorObj`, the tensor product of objects in a monoidal category -/ -scoped infixr:70 " βŠ— " => MonoidalCategory.tensorObj - -/-- Notation for the `whiskerLeft` operator of monoidal categories -/ -scoped infixr:81 " ◁ " => whiskerLeft - -/-- Notation for the `whiskerRight` operator of monoidal categories -/ -scoped infixl:81 " β–· " => whiskerRight +scoped infixr:70 " βŠ— " => tensorObj /-- Notation for the `whiskerLeft` operator of monoidal categories -/ scoped infixr:81 " ◁ " => whiskerLeft @@ -212,7 +200,7 @@ scoped infixr:81 " ◁ " => whiskerLeft scoped infixl:81 " β–· " => whiskerRight /-- Notation for `tensorHom`, the tensor product of morphisms in a monoidal category -/ -scoped infixr:70 " βŠ— " => MonoidalCategory.tensorHom +scoped infixr:70 " βŠ— " => tensorHom /-- Notation for `tensorUnit`, the two-sided identity of `βŠ—` -/ scoped notation "πŸ™_" => tensorUnit @@ -853,70 +841,6 @@ theorem rightUnitor_inv_naturality' {X X' : C} (f : X ⟢ X') : end -/-- -A constructor for monoidal categories that requires `tensorHom` instead of `whiskerLeft` and -`whiskerRight`. --/ -def ofTensorHom - (tensorObj : C β†’ C β†’ C) - (tensorHom : βˆ€ {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C}, (X₁ ⟢ Y₁) β†’ (Xβ‚‚ ⟢ Yβ‚‚) β†’ (tensorObj X₁ Xβ‚‚ ⟢ tensorObj Y₁ Yβ‚‚)) - (whiskerLeft : βˆ€ (X : C) {Y₁ Yβ‚‚ : C} (_f : Y₁ ⟢ Yβ‚‚), tensorObj X Y₁ ⟢ tensorObj X Yβ‚‚ := - fun X _ _ f ↦ tensorHom (πŸ™ X) f) - (whiskerRight : βˆ€ {X₁ Xβ‚‚ : C} (_f : X₁ ⟢ Xβ‚‚) (Y : C), tensorObj X₁ Y ⟢ tensorObj Xβ‚‚ Y := - fun f Y ↦ tensorHom f (πŸ™ Y)) - (tensor_id : βˆ€ X₁ Xβ‚‚ : C, tensorHom (πŸ™ X₁) (πŸ™ Xβ‚‚) = πŸ™ (tensorObj X₁ Xβ‚‚) := by - aesop_cat) - (id_tensorHom : βˆ€ (X : C) {Y₁ Yβ‚‚ : C} (f : Y₁ ⟢ Yβ‚‚), tensorHom (πŸ™ X) f = whiskerLeft X f := by - aesop_cat) - (tensorHom_id : βˆ€ {X₁ Xβ‚‚ : C} (f : X₁ ⟢ Xβ‚‚) (Y : C), tensorHom f (πŸ™ Y) = whiskerRight f Y := by - aesop_cat) - (tensor_comp : - βˆ€ {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (g₁ : Y₁ ⟢ Z₁) (gβ‚‚ : Yβ‚‚ ⟢ Zβ‚‚), - tensorHom (f₁ ≫ g₁) (fβ‚‚ ≫ gβ‚‚) = tensorHom f₁ fβ‚‚ ≫ tensorHom g₁ gβ‚‚ := by - aesop_cat) - (tensorUnit' : C) - (associator : βˆ€ X Y Z : C, tensorObj (tensorObj X Y) Z β‰… tensorObj X (tensorObj Y Z)) - (associator_naturality : - βˆ€ {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃ : C} (f₁ : X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (f₃ : X₃ ⟢ Y₃), - tensorHom (tensorHom f₁ fβ‚‚) f₃ ≫ (associator Y₁ Yβ‚‚ Y₃).hom = - (associator X₁ Xβ‚‚ X₃).hom ≫ tensorHom f₁ (tensorHom fβ‚‚ f₃) := by - aesop_cat) - (leftUnitor : βˆ€ X : C, tensorObj tensorUnit' X β‰… X) - (leftUnitor_naturality : - βˆ€ {X Y : C} (f : X ⟢ Y), - tensorHom (πŸ™ tensorUnit') f ≫ (leftUnitor Y).hom = (leftUnitor X).hom ≫ f := by - aesop_cat) - (rightUnitor : βˆ€ X : C, tensorObj X tensorUnit' β‰… X) - (rightUnitor_naturality : - βˆ€ {X Y : C} (f : X ⟢ Y), - tensorHom f (πŸ™ tensorUnit') ≫ (rightUnitor Y).hom = (rightUnitor X).hom ≫ f := by - aesop_cat) - (pentagon : - βˆ€ W X Y Z : C, - tensorHom (associator W X Y).hom (πŸ™ Z) ≫ - (associator W (tensorObj X Y) Z).hom ≫ tensorHom (πŸ™ W) (associator X Y Z).hom = - (associator (tensorObj W X) Y Z).hom ≫ (associator W X (tensorObj Y Z)).hom := by - aesop_cat) - (triangle : - βˆ€ X Y : C, - (associator X tensorUnit' Y).hom ≫ tensorHom (πŸ™ X) (leftUnitor Y).hom = - tensorHom (rightUnitor X).hom (πŸ™ Y) := by - aesop_cat) : - MonoidalCategory C where - tensorObj := tensorObj - tensorHom := tensorHom - whiskerLeft X _ _ f := whiskerLeft X f - whiskerRight f X := whiskerRight f X - tensorHom_def := by intros; simp [← id_tensorHom, ←tensorHom_id, ← tensor_comp] - tensorUnit' := tensorUnit' - leftUnitor := leftUnitor - rightUnitor := rightUnitor - associator := associator - whiskerLeft_id := by intros; simp [← id_tensorHom, ← tensor_id] - id_whiskerRight := by intros; simp [← tensorHom_id, tensor_id] - pentagon := by intros; simp [← id_tensorHom, ← tensorHom_id, pentagon] - triangle := by intros; simp [← id_tensorHom, ← tensorHom_id, triangle] - end section @@ -1152,8 +1076,6 @@ instance prodMonoidal : MonoidalCategory (C₁ Γ— Cβ‚‚) where whiskerLeft X _ _ f := (whiskerLeft X.1 f.1, whiskerLeft X.2 f.2) whiskerRight f X := (whiskerRight f.1 X.1, whiskerRight f.2 X.2) whisker_exchange := by simp [whisker_exchange] - whiskerLeft X _ _ f := (whiskerLeft X.1 f.1, whiskerLeft X.2 f.2) - whiskerRight f X := (whiskerRight f.1 X.1, whiskerRight f.2 X.2) tensorHom_def := by simp [tensorHom_def] tensorUnit' := (πŸ™_ C₁, πŸ™_ Cβ‚‚) associator X Y Z := (Ξ±_ X.1 Y.1 Z.1).prod (Ξ±_ X.2 Y.2 Z.2) From 41b22ddfd2702f56c4c04a2e834c5db91a3256c1 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 7 Aug 2023 23:34:59 +0900 Subject: [PATCH 27/62] golf --- Mathlib/CategoryTheory/Closed/FunctorCategory.lean | 11 +++-------- 1 file changed, 3 insertions(+), 8 deletions(-) diff --git a/Mathlib/CategoryTheory/Closed/FunctorCategory.lean b/Mathlib/CategoryTheory/Closed/FunctorCategory.lean index 8e8f10cbd53ad..05442d5c4452a 100644 --- a/Mathlib/CategoryTheory/Closed/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Closed/FunctorCategory.lean @@ -41,10 +41,8 @@ def closedUnit (F : D β₯€ C) : 𝟭 (D β₯€ C) ⟢ tensorLeft F β‹™ closedIhom F dsimp simp only [ihom.coev_naturality, closedIhom_obj_map, Monoidal.tensorObj_map] dsimp - rw [coev_app_comp_pre_app_assoc, ← Functor.map_comp, - tensorHom_def, ← comp_whiskerRight_assoc, IsIso.inv_hom_id] - simp - } + rw [coev_app_comp_pre_app_assoc, ← Functor.map_comp, tensorHom_def] + simp } #align category_theory.functor.closed_unit CategoryTheory.Functor.closedUnit /-- Auxiliary definition for `CategoryTheory.Functor.closed`. @@ -58,10 +56,7 @@ def closedCounit (F : D β₯€ C) : closedIhom F β‹™ tensorLeft F ⟢ 𝟭 (D β₯€ C dsimp simp only [closedIhom_obj_map, pre_comm_ihom_map] rw [tensorHom_def] - simp only [NatTrans.naturality, MonoidalCategory.whiskerLeft_comp, Category.assoc, - ihom.ev_naturality, comp_obj, tensorLeft_obj, id_obj, id_tensor_pre_app_comp_ev_assoc] - simp [← comp_whiskerRight_assoc] - } + simp } #align category_theory.functor.closed_counit CategoryTheory.Functor.closedCounit /-- If `C` is a monoidal closed category and `D` is a groupoid, then every functor `F : D β₯€ C` is From 1fc3593d877c673a631ae400d3256fbe91b1f192 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 7 Aug 2023 23:41:12 +0900 Subject: [PATCH 28/62] fix --- Mathlib/CategoryTheory/Monoidal/Discrete.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib/CategoryTheory/Monoidal/Discrete.lean b/Mathlib/CategoryTheory/Monoidal/Discrete.lean index e94479260ef7f..b69000f60d044 100644 --- a/Mathlib/CategoryTheory/Monoidal/Discrete.lean +++ b/Mathlib/CategoryTheory/Monoidal/Discrete.lean @@ -32,6 +32,7 @@ instance Discrete.monoidal : MonoidalCategory (Discrete M) tensorObj X Y := Discrete.mk (X.as * Y.as) whiskerLeft X _ _ f := eqToHom (by dsimp; rw [eq_of_hom f]) whiskerRight f X := eqToHom (by dsimp; rw [eq_of_hom f]) + tensorHom f g := eqToHom (by dsimp; rw [eq_of_hom f, eq_of_hom g]) leftUnitor X := Discrete.eqToIso (one_mul X.as) rightUnitor X := Discrete.eqToIso (mul_one X.as) associator X Y Z := Discrete.eqToIso (mul_assoc _ _ _) From 1bf248a1ef0ee049f7a38726deaded492a5f7c86 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 21 Sep 2023 16:35:26 +0900 Subject: [PATCH 29/62] fix some errors --- Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Functor.lean | 4 ++-- Mathlib/CategoryTheory/Monoidal/Preadditive.lean | 9 ++++----- Mathlib/CategoryTheory/Monoidal/Transport.lean | 5 ++--- Mathlib/Tactic/CategoryTheory/Coherence.lean | 2 ++ 5 files changed, 11 insertions(+), 11 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean index 543f8997edb1b..d9eb7b813739b 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean @@ -321,7 +321,7 @@ theorem normalize_naturality (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ rw [associator_inv_naturality_middle_assoc, ← comp_whiskerRight_assoc, ih] have := dcongr_arg (fun x => (normalizeIsoApp' C Ξ·' x).hom) (normalizeObj_congr n (Quotient.mk (setoidHom X Y) h)) - dsimp at this; simp [this] + simp [this] end diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index f7c60052ae393..692f88dd829ad 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -173,9 +173,9 @@ def LaxMonoidalFunctor.ofTensorHom (F : C β₯€ D) associativity := fun X Y Z => by simp_rw [← tensorHom_id, ← id_tensorHom, associativity] left_unitality := fun X => by - simp_rw [← tensorHom_id, ← id_tensorHom, left_unitality] + simp_rw [← tensorHom_id, left_unitality] right_unitality := fun X => by - simp_rw [← tensorHom_id, ← id_tensorHom, right_unitality] + simp_rw [← id_tensorHom, right_unitality] --Porting note: was `[simp, reassoc.1]` @[reassoc (attr := simp)] diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index cd1d410b8de37..f01aa226de459 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -111,8 +111,7 @@ instance (X : C) : PreservesFiniteBiproducts (tensorLeft X) where { preserves := fun {f} => { preserves := fun {b} i => isBilimitOfTotal _ (by dsimp - simp_rw [← whiskerLeft_comp] - simp_rw [← tensorHom_id, ← id_tensorHom] + simp_rw [← id_tensorHom] simp only [← tensor_comp, Category.comp_id, ← tensor_sum, ← tensor_id, IsBilimit.total i]) } } @@ -121,7 +120,7 @@ instance (X : C) : PreservesFiniteBiproducts (tensorRight X) where { preserves := fun {f} => { preserves := fun {b} i => isBilimitOfTotal _ (by dsimp - simp_rw [← tensorHom_id, ← id_tensorHom] + simp_rw [← tensorHom_id] simp only [← tensor_comp, Category.comp_id, ← sum_tensor, ← tensor_id, IsBilimit.total i]) } } @@ -181,7 +180,7 @@ theorem leftDistributor_assoc {J : Type} [Fintype J] (X Y : C) (f : J β†’ C) : asIso_hom, comp_zero, comp_dite, Preadditive.sum_comp, Preadditive.comp_sum, tensor_sum, id_tensor_comp, tensorIso_hom, leftDistributor_hom, biproduct.mapIso_hom, biproduct.ΞΉ_map, biproduct.ΞΉ_Ο€, Finset.sum_dite_irrel, Finset.sum_dite_eq', Finset.sum_const_zero] - simp_rw [← tensorHom_id, ← id_tensorHom] + simp_rw [← id_tensorHom] simp only [← id_tensor_comp, biproduct.ΞΉ_Ο€] simp only [id_tensor_comp, tensor_dite, comp_dite] simp @@ -242,7 +241,7 @@ theorem rightDistributor_assoc {J : Type} [Fintype J] (f : J β†’ C) (X Y : C) : comp_tensor_id, tensorIso_hom, rightDistributor_hom, biproduct.mapIso_hom, biproduct.ΞΉ_map, biproduct.ΞΉ_Ο€, Finset.sum_dite_irrel, Finset.sum_dite_eq', Finset.sum_const_zero, Finset.mem_univ, if_true] - simp_rw [← tensorHom_id, ← id_tensorHom] + simp_rw [← tensorHom_id] simp only [← comp_tensor_id, biproduct.ΞΉ_Ο€, dite_tensor, comp_dite] simp #align category_theory.right_distributor_assoc CategoryTheory.rightDistributor_assoc diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index 90df47af004af..be8402103c400 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -254,9 +254,8 @@ def fromTransported (e : C β‰Œ D) : MonoidalFunctor (Transported e) C := #align category_theory.monoidal.from_transported CategoryTheory.Monoidal.fromTransported instance instIsEquivalence_fromTransported (e : C β‰Œ D) : - IsEquivalence (fromTransported e).toFunctor := by - dsimp [fromTransported] - infer_instance + IsEquivalence (fromTransported e).toFunctor := + inferInstanceAs <| IsEquivalence <| (e.functor.asEquivalence).inverse /-- The unit isomorphism upgrades to a monoidal isomorphism. -/ @[simps! hom inv] diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index a190e623a8aef..be34eeab9987a 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -361,6 +361,8 @@ def coherence_loop (maxSteps := 47) : TacticM Unit := -- and whose second terms can be identified by recursively called `coherence`. coherence_loop maxSteps' +open Lean.Parser.Tactic + /-- Simp lemmas for rewriting a hom in monoical categories into a normal form. -/ From 22014bbd0121a304e092193329c0012c20197e3f Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 21 Sep 2023 16:37:27 +0900 Subject: [PATCH 30/62] revert precedence --- Mathlib/Tactic/CategoryTheory/Coherence.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index be34eeab9987a..359405908de60 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -186,7 +186,7 @@ def monoidalComp {W X Y Z : C} [LiftObj X] [LiftObj Y] f ≫ MonoidalCoherence.hom ≫ g @[inherit_doc Mathlib.Tactic.Coherence.monoidalComp] -scoped[CategoryTheory.MonoidalCategory] infixr:79 " βŠ—β‰« " => +scoped[CategoryTheory.MonoidalCategory] infixr:80 " βŠ—β‰« " => Mathlib.Tactic.Coherence.monoidalComp -- type as \ot \gg /-- Compose two isomorphisms in a monoidal category, @@ -196,7 +196,7 @@ noncomputable def monoidalIsoComp {W X Y Z : C} [LiftObj X] [LiftObj Y] f β‰ͺ≫ asIso MonoidalCoherence.hom β‰ͺ≫ g @[inherit_doc Mathlib.Tactic.Coherence.monoidalIsoComp] -scoped[CategoryTheory.MonoidalCategory] infixr:79 " β‰ͺβŠ—β‰« " => +scoped[CategoryTheory.MonoidalCategory] infixr:80 " β‰ͺβŠ—β‰« " => Mathlib.Tactic.Coherence.monoidalIsoComp -- type as \ll \ot \gg example {U V W X Y : C} (f : U ⟢ V βŠ— (W βŠ— X)) (g : (V βŠ— W) βŠ— X ⟢ Y) : U ⟢ Y := f βŠ—β‰« g From c1bce84c5ba2714d3215f03485a3f410211fef0a Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 21 Sep 2023 17:00:16 +0900 Subject: [PATCH 31/62] add parentheses --- Mathlib/CategoryTheory/Monoidal/Center.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index 3a9dfc9ed73bb..1a3897219f6fa 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -145,8 +145,8 @@ def tensorObj (X Y : Center C) : Center C := calc _ = πŸ™ _ βŠ—β‰« X.1 ◁ (HalfBraiding.Ξ² Y.2 U).hom β–· U' βŠ—β‰« - _ ◁ (HalfBraiding.Ξ² Y.2 U').hom ≫ - (HalfBraiding.Ξ² X.2 U).hom β–· _ βŠ—β‰« + (_ ◁ (HalfBraiding.Ξ² Y.2 U').hom ≫ + (HalfBraiding.Ξ² X.2 U).hom β–· _) βŠ—β‰« U ◁ (HalfBraiding.Ξ² X.2 U').hom β–· Y.1 βŠ—β‰« πŸ™ _ := by coherence _ = _ := by rw [whisker_exchange]; coherence naturality := fun {U U'} f => by From 2dc4ce3ab7e93b348ff80900b81edab08fe4b429 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 21 Sep 2023 17:01:29 +0900 Subject: [PATCH 32/62] fix no progress simp --- Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean | 4 ++-- Mathlib/CategoryTheory/Monoidal/Bimod.lean | 5 +---- 2 files changed, 3 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean index 924c42cb3858b..87be76429e8c9 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean @@ -41,13 +41,13 @@ set_option linter.uppercaseLean3 false in @[simp] theorem braiding_naturality_left {X Y : ModuleCat R} (f : X ⟢ Y) (Z : ModuleCat R) : f β–· Z ≫ (braiding Y Z).hom = (braiding X Z).hom ≫ Z ◁ f := by - simp_rw [← id_tensorHom, tensorHom_id] + simp_rw [← id_tensorHom] apply braiding_naturality @[simp] theorem braiding_naturality_right (X : ModuleCat R) {Y Z : ModuleCat R} (f : Y ⟢ Z) : X ◁ f ≫ (braiding X Z).hom = (braiding X Y).hom ≫ f β–· X := by - simp_rw [← id_tensorHom, tensorHom_id] + simp_rw [← id_tensorHom] apply braiding_naturality @[simp] diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index 552fb40544333..e7e26cdee0654 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -300,7 +300,7 @@ noncomputable def actRight : X P Q βŠ— T.X ⟢ X P Q := dsimp simp only [comp_whiskerRight, whisker_assoc, Category.assoc, Iso.inv_hom_id_assoc] slice_lhs 3 4 => - rw [← MonoidalCategory.whiskerLeft_comp, middle_assoc, + rw [← MonoidalCategory.whiskerLeft_comp, middle_assoc, MonoidalCategory.whiskerLeft_comp] simp)) set_option linter.uppercaseLean3 false in @@ -988,7 +988,6 @@ set_option linter.uppercaseLean3 false in theorem whisker_exchange_bimod {X Y Z : Mon_ C} {M N : Bimod X Y} {P Q : Bimod Y Z} (f : M ⟢ N) (g : P ⟢ Q) : whiskerLeft M g ≫ whiskerRight f Q = whiskerRight f P ≫ whiskerLeft N g := by - dsimp [tensorHom] ext apply coequalizer.hom_ext dsimp @@ -1020,7 +1019,6 @@ theorem pentagon_bimod {V W X Y Z : Mon_ C} (M : Bimod V W) (N : Bimod W X) (P : dsimp [AssociatorBimod.homAux] refine' (cancel_epi ((tensorRight _).map (coequalizer.Ο€ _ _))).1 _ dsimp - simp only [id_tensorHom, tensorHom_id] slice_lhs 1 2 => rw [← comp_whiskerRight, coequalizer.Ο€_desc] slice_rhs 1 3 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] slice_rhs 3 4 => rw [coequalizer.Ο€_desc] @@ -1063,7 +1061,6 @@ theorem triangle_bimod {X Y Z : Mon_ C} (M : Bimod X Y) (N : Bimod Y Z) : slice_lhs 1 3 => rw [Ο€_tensor_id_preserves_coequalizer_inv_desc] slice_lhs 3 4 => rw [ΞΉ_colimMap, parallelPairHom_app_one] dsimp [LeftUnitorBimod.hom] - simp only [id_tensorHom, tensorHom_id] slice_lhs 2 3 => rw [← MonoidalCategory.whiskerLeft_comp, coequalizer.Ο€_desc] slice_rhs 1 2 => rw [← comp_whiskerRight, coequalizer.Ο€_desc] slice_rhs 1 2 => rw [coequalizer.condition] From a5083a288b9c1f1ac8395da17bc3a4585e282d20 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 21 Sep 2023 20:34:44 +0900 Subject: [PATCH 33/62] docBlame linter --- Mathlib/CategoryTheory/Monoidal/Braided.lean | 2 ++ Mathlib/CategoryTheory/Monoidal/Category.lean | 9 +++++++++ Mathlib/CategoryTheory/Monoidal/Functor.lean | 6 ++++++ 3 files changed, 17 insertions(+) diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 76e36cc546be3..85b02ff94365a 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -42,10 +42,12 @@ and also satisfies the two hexagon identities. class BraidedCategory (C : Type u) [Category.{v} C] [MonoidalCategory.{v} C] where -- braiding natural iso: braiding : βˆ€ X Y : C, X βŠ— Y β‰… Y βŠ— X + /-- naturality of the braiding -/ braiding_naturality_right : βˆ€ (X : C) {Y Z : C} (f : Y ⟢ Z), X ◁ f ≫ (braiding X Z).hom = (braiding X Y).hom ≫ f β–· X := by aesop_cat + /-- naturality of the braiding -/ braiding_naturality_left : βˆ€ {X Y : C} (f : X ⟢ Y) (Z : C), f β–· Z ≫ (braiding Y Z).hom = (braiding X Z).hom ≫ Z ◁ f := by diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 6615a244bc35f..fbf6d000740ea 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -111,36 +111,45 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] where rightUnitor : βˆ€ X : C, tensorObj X tensorUnit' β‰… X /-- The associator isomorphism `(X βŠ— Y) βŠ— Z ≃ X βŠ— (Y βŠ— Z)` -/ associator : βˆ€ X Y Z : C, tensorObj (tensorObj X Y) Z β‰… tensorObj X (tensorObj Y Z) + /-- The left whiskering preserves the identity. -/ whiskerLeft_id : βˆ€ (X Y : C), whiskerLeft X (πŸ™ Y) = πŸ™ (tensorObj X Y) := by aesop_cat + /-- The left whiskering commutes with the composition. -/ whiskerLeft_comp : βˆ€ (W : C) {X Y Z : C} (f : X ⟢ Y) (g : Y ⟢ Z), whiskerLeft W (f ≫ g) = whiskerLeft W f ≫ whiskerLeft W g := by aesop_cat + /-- `πŸ™_ C ◁ f` is equal to `f` up to unitors. -/ id_whiskerLeft : βˆ€ {X Y : C} (f : X ⟢ Y), whiskerLeft tensorUnit' f = (leftUnitor X).hom ≫ f ≫ (leftUnitor Y).inv := by aesop_cat + /-- `(X βŠ— Y) ◁ f` is equal to `X ◁ Y ◁ f` up to associators. -/ tensor_whiskerLeft : βˆ€ (X Y : C) {Z Z' : C} (f : Z ⟢ Z'), whiskerLeft (tensorObj X Y) f = (associator X Y Z).hom ≫ whiskerLeft X (whiskerLeft Y f) ≫ (associator X Y Z').inv := by aesop_cat + /-- The right whiskering preserves the identity. -/ id_whiskerRight : βˆ€ (X Y : C), whiskerRight (πŸ™ X) Y = πŸ™ (tensorObj X Y) := by aesop_cat + /-- The right whiskering commutes with the composition. -/ comp_whiskerRight : βˆ€ {W X Y : C} (f : W ⟢ X) (g : X ⟢ Y) (Z : C), whiskerRight (f ≫ g) Z = whiskerRight f Z ≫ whiskerRight g Z := by aesop_cat + /-- `f β–· πŸ™_ C` is equal to `f` up to unitors. -/ whiskerRight_id : βˆ€ {X Y : C} (f : X ⟢ Y), whiskerRight f tensorUnit' = (rightUnitor X).hom ≫ f ≫ (rightUnitor Y).inv := by aesop_cat + /-- `f β–· (Y βŠ— Z)` is equal to `f β–· Y β–· Z` up to associators. -/ whiskerRight_tensor : βˆ€ {X X' : C} (f : X ⟢ X') (Y Z : C), whiskerRight f (tensorObj Y Z) = (associator X Y Z).inv ≫ whiskerRight (whiskerRight f Y) Z ≫ (associator X' Y Z).hom := by aesop_cat + /-- `(X ◁ f) β–· Z` is equal to `X ◁ (f β–· Z)` up to associators. -/ whisker_assoc : βˆ€ (X : C) {Y Y' : C} (f : Y ⟢ Y') (Z : C), whiskerRight (whiskerLeft X f) Z = diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index 692f88dd829ad..26982bd109307 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -69,10 +69,12 @@ structure LaxMonoidalFunctor extends C β₯€ D where Ξ΅ : πŸ™_ D ⟢ obj (πŸ™_ C) /-- tensorator -/ ΞΌ : βˆ€ X Y : C, obj X βŠ— obj Y ⟢ obj (X βŠ— Y) + /-- naturality of the tensorator -/ ΞΌ_natural_left : βˆ€ {X Y : C} (f : X ⟢ Y) (X' : C), (map f β–· obj X') ≫ ΞΌ Y X' = ΞΌ X X' ≫ map (f β–· X') := by aesop_cat + /-- naturality of the tensorator -/ ΞΌ_natural_right : βˆ€ {X Y : C} (X' : C) (f : X ⟢ Y) , (obj X' ◁ map f) ≫ ΞΌ X' Y = ΞΌ X' X ≫ map (X' ◁ f) := by @@ -136,6 +138,10 @@ theorem LaxMonoidalFunctor.right_unitality' (F : LaxMonoidalFunctor C D) (X : C (ρ_ (F.obj X)).hom = (πŸ™ (F.obj X) βŠ— F.Ξ΅) ≫ F.ΞΌ X (πŸ™_ C) ≫ F.map (ρ_ X).hom := by simp +/-- +A constructor for lax monoidal functors whose axioms are described by `tensorHom` instead of +`whiskerLeft` and `whiskerRight`. +-/ @[simps] def LaxMonoidalFunctor.ofTensorHom (F : C β₯€ D) /- unit morphism -/ From 286c179d18d99d5c71c20b5ab893d2904b6d89c2 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 21 Sep 2023 21:53:21 +0900 Subject: [PATCH 34/62] docBlame linter --- Mathlib/CategoryTheory/Monoidal/Functorial.lean | 7 ++++++- Mathlib/CategoryTheory/Monoidal/Linear.lean | 2 ++ Mathlib/CategoryTheory/Monoidal/Preadditive.lean | 4 ++++ 3 files changed, 12 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Functorial.lean b/Mathlib/CategoryTheory/Monoidal/Functorial.lean index 1814b5061fdcb..1f5ed07b96884 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functorial.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functorial.lean @@ -47,6 +47,8 @@ open MonoidalCategory variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] {D : Type uβ‚‚} [Category.{vβ‚‚} D] [MonoidalCategory.{vβ‚‚} D] +/-- An unbundled description of lax monoidal functors without axioms. See `LaxMonoidal` for +the full description. -/ class LaxMonoidalStruct (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] where /-- unit morphism -/ Ξ΅ : πŸ™_ D ⟢ F (πŸ™_ C) @@ -57,10 +59,12 @@ class LaxMonoidalStruct (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] where -- but that isn't the immediate plan. /-- An unbundled description of lax monoidal functors. -/ class LaxMonoidal (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] extends LaxMonoidalStruct F where + /-- naturality of the tensorator -/ ΞΌ_natural_left : βˆ€ {X Y : C} (f : X ⟢ Y) (X' : C), (map F f β–· F X') ≫ ΞΌ Y X' = ΞΌ X X' ≫ map F (f β–· X') := by aesop_cat + /-- naturality of the tensorator -/ ΞΌ_natural_right : βˆ€ {X Y : C} (X' : C) (f : X ⟢ Y) , (F X' ◁ map F f) ≫ ΞΌ X' Y = ΞΌ X' X ≫ map F (X' ◁ f) := by @@ -71,9 +75,10 @@ class LaxMonoidal (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] extends LaxMonoidalS (ΞΌ X Y β–· F Z) ≫ ΞΌ (X βŠ— Y) Z ≫ map F (Ξ±_ X Y Z).hom = (Ξ±_ (F X) (F Y) (F Z)).hom ≫ (F X ◁ ΞΌ Y Z) ≫ ΞΌ X (Y βŠ— Z) := by aesop_cat - -- unitality + /-- unitality of lax monoidal functors -/ left_unitality : βˆ€ X : C, (Ξ»_ (F X)).hom = (Ξ΅ β–· F X) ≫ ΞΌ (πŸ™_ C) X ≫ map F (Ξ»_ X).hom := by aesop_cat + /-- unitality of lax monoidal functors -/ right_unitality : βˆ€ X : C, (ρ_ (F X)).hom = (F X ◁ Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ map F (ρ_ X).hom := by aesop_cat diff --git a/Mathlib/CategoryTheory/Monoidal/Linear.lean b/Mathlib/CategoryTheory/Monoidal/Linear.lean index c8a2e24043adb..ca3c0b11c5d8c 100644 --- a/Mathlib/CategoryTheory/Monoidal/Linear.lean +++ b/Mathlib/CategoryTheory/Monoidal/Linear.lean @@ -32,8 +32,10 @@ variable [MonoidalCategory C] /-- A category is `MonoidalLinear R` if tensoring is `R`-linear in both factors. -/ class MonoidalLinear [MonoidalPreadditive C] : Prop where + /-- The left whiskering commutes with the scalar product. -/ whiskerLeft_smul : βˆ€ (X : C) {Y Z : C} (r : R) (f : Y ⟢ Z) , X ◁ (r β€’ f) = r β€’ (X ◁ f) := by aesop_cat + /-- The right whiskering commutes with the scalar product. -/ smul_whiskerRight : βˆ€ (r : R) {Y Z : C} (f : Y ⟢ Z) (X : C), (r β€’ f) β–· X = r β€’ (f β–· X) := by aesop_cat #align category_theory.monoidal_linear CategoryTheory.MonoidalLinear diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index f01aa226de459..ba832b5400f82 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -34,9 +34,13 @@ Note we don't `extend Preadditive C` here, as `Abelian C` already extends it, and we'll need to have both typeclasses sometimes. -/ class MonoidalPreadditive : Prop where + /-- The left whiskering preserves the zero. -/ whiskerLeft_zero : βˆ€ {X Y Z : C}, X ◁ (0 : Y ⟢ Z) = 0 := by aesop_cat + /-- The right whiskering preserves the zero. -/ zero_whiskerRight : βˆ€ {X Y Z : C}, (0 : Y ⟢ Z) β–· X = 0 := by aesop_cat + /-- The left whiskering commutes with the addition. -/ whiskerLeft_add : βˆ€ {X Y Z : C} (f g : Y ⟢ Z), X ◁ (f + g) = X ◁ f + X ◁ g := by aesop_cat + /-- The right whiskering commutes with the addition. -/ add_whiskerRight : βˆ€ {X Y Z : C} (f g : Y ⟢ Z), (f + g) β–· X = f β–· X + g β–· X := by aesop_cat #align category_theory.monoidal_preadditive CategoryTheory.MonoidalPreadditive From 74f0a6f85a981e35362f3ce63bb36768c88bc7d4 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 5 Nov 2023 07:47:45 +0900 Subject: [PATCH 35/62] Revert "docBlame linter" thanks to leanprover/std4#269 --- Mathlib/CategoryTheory/Monoidal/Braided.lean | 2 -- Mathlib/CategoryTheory/Monoidal/Category.lean | 9 --------- Mathlib/CategoryTheory/Monoidal/Functor.lean | 2 -- Mathlib/CategoryTheory/Monoidal/Functorial.lean | 4 ---- Mathlib/CategoryTheory/Monoidal/Linear.lean | 2 -- Mathlib/CategoryTheory/Monoidal/Preadditive.lean | 4 ---- 6 files changed, 23 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 85b02ff94365a..76e36cc546be3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -42,12 +42,10 @@ and also satisfies the two hexagon identities. class BraidedCategory (C : Type u) [Category.{v} C] [MonoidalCategory.{v} C] where -- braiding natural iso: braiding : βˆ€ X Y : C, X βŠ— Y β‰… Y βŠ— X - /-- naturality of the braiding -/ braiding_naturality_right : βˆ€ (X : C) {Y Z : C} (f : Y ⟢ Z), X ◁ f ≫ (braiding X Z).hom = (braiding X Y).hom ≫ f β–· X := by aesop_cat - /-- naturality of the braiding -/ braiding_naturality_left : βˆ€ {X Y : C} (f : X ⟢ Y) (Z : C), f β–· Z ≫ (braiding Y Z).hom = (braiding X Z).hom ≫ Z ◁ f := by diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index fbf6d000740ea..6615a244bc35f 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -111,45 +111,36 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] where rightUnitor : βˆ€ X : C, tensorObj X tensorUnit' β‰… X /-- The associator isomorphism `(X βŠ— Y) βŠ— Z ≃ X βŠ— (Y βŠ— Z)` -/ associator : βˆ€ X Y Z : C, tensorObj (tensorObj X Y) Z β‰… tensorObj X (tensorObj Y Z) - /-- The left whiskering preserves the identity. -/ whiskerLeft_id : βˆ€ (X Y : C), whiskerLeft X (πŸ™ Y) = πŸ™ (tensorObj X Y) := by aesop_cat - /-- The left whiskering commutes with the composition. -/ whiskerLeft_comp : βˆ€ (W : C) {X Y Z : C} (f : X ⟢ Y) (g : Y ⟢ Z), whiskerLeft W (f ≫ g) = whiskerLeft W f ≫ whiskerLeft W g := by aesop_cat - /-- `πŸ™_ C ◁ f` is equal to `f` up to unitors. -/ id_whiskerLeft : βˆ€ {X Y : C} (f : X ⟢ Y), whiskerLeft tensorUnit' f = (leftUnitor X).hom ≫ f ≫ (leftUnitor Y).inv := by aesop_cat - /-- `(X βŠ— Y) ◁ f` is equal to `X ◁ Y ◁ f` up to associators. -/ tensor_whiskerLeft : βˆ€ (X Y : C) {Z Z' : C} (f : Z ⟢ Z'), whiskerLeft (tensorObj X Y) f = (associator X Y Z).hom ≫ whiskerLeft X (whiskerLeft Y f) ≫ (associator X Y Z').inv := by aesop_cat - /-- The right whiskering preserves the identity. -/ id_whiskerRight : βˆ€ (X Y : C), whiskerRight (πŸ™ X) Y = πŸ™ (tensorObj X Y) := by aesop_cat - /-- The right whiskering commutes with the composition. -/ comp_whiskerRight : βˆ€ {W X Y : C} (f : W ⟢ X) (g : X ⟢ Y) (Z : C), whiskerRight (f ≫ g) Z = whiskerRight f Z ≫ whiskerRight g Z := by aesop_cat - /-- `f β–· πŸ™_ C` is equal to `f` up to unitors. -/ whiskerRight_id : βˆ€ {X Y : C} (f : X ⟢ Y), whiskerRight f tensorUnit' = (rightUnitor X).hom ≫ f ≫ (rightUnitor Y).inv := by aesop_cat - /-- `f β–· (Y βŠ— Z)` is equal to `f β–· Y β–· Z` up to associators. -/ whiskerRight_tensor : βˆ€ {X X' : C} (f : X ⟢ X') (Y Z : C), whiskerRight f (tensorObj Y Z) = (associator X Y Z).inv ≫ whiskerRight (whiskerRight f Y) Z ≫ (associator X' Y Z).hom := by aesop_cat - /-- `(X ◁ f) β–· Z` is equal to `X ◁ (f β–· Z)` up to associators. -/ whisker_assoc : βˆ€ (X : C) {Y Y' : C} (f : Y ⟢ Y') (Z : C), whiskerRight (whiskerLeft X f) Z = diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index 26982bd109307..fa3fc75827466 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -69,12 +69,10 @@ structure LaxMonoidalFunctor extends C β₯€ D where Ξ΅ : πŸ™_ D ⟢ obj (πŸ™_ C) /-- tensorator -/ ΞΌ : βˆ€ X Y : C, obj X βŠ— obj Y ⟢ obj (X βŠ— Y) - /-- naturality of the tensorator -/ ΞΌ_natural_left : βˆ€ {X Y : C} (f : X ⟢ Y) (X' : C), (map f β–· obj X') ≫ ΞΌ Y X' = ΞΌ X X' ≫ map (f β–· X') := by aesop_cat - /-- naturality of the tensorator -/ ΞΌ_natural_right : βˆ€ {X Y : C} (X' : C) (f : X ⟢ Y) , (obj X' ◁ map f) ≫ ΞΌ X' Y = ΞΌ X' X ≫ map (X' ◁ f) := by diff --git a/Mathlib/CategoryTheory/Monoidal/Functorial.lean b/Mathlib/CategoryTheory/Monoidal/Functorial.lean index 1f5ed07b96884..62dd50047f7de 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functorial.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functorial.lean @@ -59,12 +59,10 @@ class LaxMonoidalStruct (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] where -- but that isn't the immediate plan. /-- An unbundled description of lax monoidal functors. -/ class LaxMonoidal (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] extends LaxMonoidalStruct F where - /-- naturality of the tensorator -/ ΞΌ_natural_left : βˆ€ {X Y : C} (f : X ⟢ Y) (X' : C), (map F f β–· F X') ≫ ΞΌ Y X' = ΞΌ X X' ≫ map F (f β–· X') := by aesop_cat - /-- naturality of the tensorator -/ ΞΌ_natural_right : βˆ€ {X Y : C} (X' : C) (f : X ⟢ Y) , (F X' ◁ map F f) ≫ ΞΌ X' Y = ΞΌ X' X ≫ map F (X' ◁ f) := by @@ -75,10 +73,8 @@ class LaxMonoidal (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] extends LaxMonoidalS (ΞΌ X Y β–· F Z) ≫ ΞΌ (X βŠ— Y) Z ≫ map F (Ξ±_ X Y Z).hom = (Ξ±_ (F X) (F Y) (F Z)).hom ≫ (F X ◁ ΞΌ Y Z) ≫ ΞΌ X (Y βŠ— Z) := by aesop_cat - /-- unitality of lax monoidal functors -/ left_unitality : βˆ€ X : C, (Ξ»_ (F X)).hom = (Ξ΅ β–· F X) ≫ ΞΌ (πŸ™_ C) X ≫ map F (Ξ»_ X).hom := by aesop_cat - /-- unitality of lax monoidal functors -/ right_unitality : βˆ€ X : C, (ρ_ (F X)).hom = (F X ◁ Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ map F (ρ_ X).hom := by aesop_cat diff --git a/Mathlib/CategoryTheory/Monoidal/Linear.lean b/Mathlib/CategoryTheory/Monoidal/Linear.lean index ca3c0b11c5d8c..c8a2e24043adb 100644 --- a/Mathlib/CategoryTheory/Monoidal/Linear.lean +++ b/Mathlib/CategoryTheory/Monoidal/Linear.lean @@ -32,10 +32,8 @@ variable [MonoidalCategory C] /-- A category is `MonoidalLinear R` if tensoring is `R`-linear in both factors. -/ class MonoidalLinear [MonoidalPreadditive C] : Prop where - /-- The left whiskering commutes with the scalar product. -/ whiskerLeft_smul : βˆ€ (X : C) {Y Z : C} (r : R) (f : Y ⟢ Z) , X ◁ (r β€’ f) = r β€’ (X ◁ f) := by aesop_cat - /-- The right whiskering commutes with the scalar product. -/ smul_whiskerRight : βˆ€ (r : R) {Y Z : C} (f : Y ⟢ Z) (X : C), (r β€’ f) β–· X = r β€’ (f β–· X) := by aesop_cat #align category_theory.monoidal_linear CategoryTheory.MonoidalLinear diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index ba832b5400f82..f01aa226de459 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -34,13 +34,9 @@ Note we don't `extend Preadditive C` here, as `Abelian C` already extends it, and we'll need to have both typeclasses sometimes. -/ class MonoidalPreadditive : Prop where - /-- The left whiskering preserves the zero. -/ whiskerLeft_zero : βˆ€ {X Y Z : C}, X ◁ (0 : Y ⟢ Z) = 0 := by aesop_cat - /-- The right whiskering preserves the zero. -/ zero_whiskerRight : βˆ€ {X Y Z : C}, (0 : Y ⟢ Z) β–· X = 0 := by aesop_cat - /-- The left whiskering commutes with the addition. -/ whiskerLeft_add : βˆ€ {X Y Z : C} (f g : Y ⟢ Z), X ◁ (f + g) = X ◁ f + X ◁ g := by aesop_cat - /-- The right whiskering commutes with the addition. -/ add_whiskerRight : βˆ€ {X Y Z : C} (f g : Y ⟢ Z), (f + g) β–· X = f β–· X + g β–· X := by aesop_cat #align category_theory.monoidal_preadditive CategoryTheory.MonoidalPreadditive From dfeb3cf16406a6019a02199a940a7884f7ebbd96 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 5 Nov 2023 09:40:35 +0900 Subject: [PATCH 36/62] resolve conflists --- Mathlib/CategoryTheory/Closed/Monoidal.lean | 9 +- Mathlib/CategoryTheory/Monoidal/Category.lean | 50 +-- Mathlib/CategoryTheory/Monoidal/Center.lean | 2 +- Mathlib/CategoryTheory/Monoidal/End.lean | 4 +- .../CategoryTheory/Monoidal/Free/Basic.lean | 2 +- .../Monoidal/FunctorCategory.lean | 5 +- .../CategoryTheory/Monoidal/Transport.lean | 300 ++++-------------- 7 files changed, 81 insertions(+), 291 deletions(-) diff --git a/Mathlib/CategoryTheory/Closed/Monoidal.lean b/Mathlib/CategoryTheory/Closed/Monoidal.lean index 19a7af00b1244..5c1183e2a03f4 100644 --- a/Mathlib/CategoryTheory/Closed/Monoidal.lean +++ b/Mathlib/CategoryTheory/Closed/Monoidal.lean @@ -70,7 +70,14 @@ def unitClosed : Closed (πŸ™_ C) where { toFun := fun a => (leftUnitor X).inv ≫ a invFun := fun a => (leftUnitor X).hom ≫ a left_inv := by aesop_cat - right_inv := by aesop_cat } } } + right_inv := by aesop_cat } + homEquiv_naturality_left_symm := fun f g => by + dsimp + rw [leftUnitor_naturality_assoc] + -- This used to be automatic before leanprover/lean4#2644 + homEquiv_naturality_right := by -- aesop failure + dsimp + simp }} #align category_theory.unit_closed CategoryTheory.unitClosed variable (A B : C) {X X' Y Y' Z : C} diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 7a9110452ab51..fcd9c69c8d9fb 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -155,16 +155,6 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] extends MonoidalCate tensorHom_def {X₁ Y₁ Xβ‚‚ Yβ‚‚ : C} (f : X₁ ⟢ Y₁) (g: Xβ‚‚ ⟢ Yβ‚‚) : f βŠ— g = (f β–· Xβ‚‚) ≫ (Y₁ ◁ g) := by aesop_cat - -- Porting note: Adding a prime here, so I can later define `tensorUnit` unprimed with explicit - -- argument `C` - /-- The tensor unity in the monoidal structure `πŸ™_ C` -/ - tensorUnit' : C - /-- The left unitor: `πŸ™_ C βŠ— X ≃ X` -/ - leftUnitor : βˆ€ X : C, tensorObj tensorUnit' X β‰… X - /-- The right unitor: `X βŠ— πŸ™_ C ≃ X` -/ - rightUnitor : βˆ€ X : C, tensorObj X tensorUnit' β‰… X - /-- The associator isomorphism `(X βŠ— Y) βŠ— Z ≃ X βŠ— (Y βŠ— Z)` -/ - associator : βˆ€ X Y Z : C, tensorObj (tensorObj X Y) Z β‰… tensorObj X (tensorObj Y Z) whiskerLeft_id : βˆ€ (X Y : C), whiskerLeft X (πŸ™ Y) = πŸ™ (tensorObj X Y) := by aesop_cat whiskerLeft_comp : @@ -173,7 +163,7 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] extends MonoidalCate aesop_cat id_whiskerLeft : βˆ€ {X Y : C} (f : X ⟢ Y), - whiskerLeft tensorUnit' f = (leftUnitor X).hom ≫ f ≫ (leftUnitor Y).inv := by + whiskerLeft tensorUnit f = (leftUnitor X).hom ≫ f ≫ (leftUnitor Y).inv := by aesop_cat tensor_whiskerLeft : βˆ€ (X Y : C) {Z Z' : C} (f : Z ⟢ Z'), @@ -188,7 +178,7 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] extends MonoidalCate aesop_cat whiskerRight_id : βˆ€ {X Y : C} (f : X ⟢ Y), - whiskerRight f tensorUnit' = (rightUnitor X).hom ≫ f ≫ (rightUnitor Y).inv := by + whiskerRight f tensorUnit = (rightUnitor X).hom ≫ f ≫ (rightUnitor Y).inv := by aesop_cat whiskerRight_tensor : βˆ€ {X X' : C} (f : X ⟢ X') (Y Z : C), @@ -210,14 +200,14 @@ class MonoidalCategory (C : Type u) [π’ž : Category.{v} C] extends MonoidalCate -/ pentagon : βˆ€ W X Y Z : C, - ((Ξ±_ W X Y).hom βŠ— πŸ™ Z) ≫ (Ξ±_ W (X βŠ— Y) Z).hom ≫ (πŸ™ W βŠ— (Ξ±_ X Y Z).hom) = + ((Ξ±_ W X Y).hom β–· Z) ≫ (Ξ±_ W (X βŠ— Y) Z).hom ≫ (W ◁ (Ξ±_ X Y Z).hom) = (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom := by aesop_cat /-- The identity relating the isomorphisms between `X βŠ— (πŸ™_ C βŠ— Y)`, `(X βŠ— πŸ™_ C) βŠ— Y` and `X βŠ— Y` -/ triangle : - βˆ€ X Y : C, (Ξ±_ X (πŸ™_ _) Y).hom ≫ (πŸ™ X βŠ— (Ξ»_ Y).hom) = ((ρ_ X).hom βŠ— πŸ™ Y) := by + βˆ€ X Y : C, (Ξ±_ X (πŸ™_ _) Y).hom ≫ (X ◁ (Ξ»_ Y).hom) = ((ρ_ X).hom β–· Y) := by aesop_cat #align category_theory.monoidal_category CategoryTheory.MonoidalCategory @@ -235,36 +225,6 @@ attribute [simp] whiskerLeft_comp id_whiskerLeft tensor_whiskerLeft comp_whiskerRight id_whiskerRight whiskerRight_tensor whisker_assoc --- Porting Note: This is here to make `tensorUnit` explicitly depend on `C`, which was done in --- Lean 3 using the `[]` notation in the `tensorUnit'` field. -/-- The tensor unity in the monoidal structure `πŸ™_ C` -/ -abbrev tensorUnit (C : Type u) [Category.{v} C] [MonoidalCategory C] : C := - tensorUnit' (C := C) - -/-- Notation for `tensorObj`, the tensor product of objects in a monoidal category -/ -scoped infixr:70 " βŠ— " => tensorObj - -/-- Notation for the `whiskerLeft` operator of monoidal categories -/ -scoped infixr:81 " ◁ " => whiskerLeft - -/-- Notation for the `whiskerRight` operator of monoidal categories -/ -scoped infixl:81 " β–· " => whiskerRight - -/-- Notation for `tensorHom`, the tensor product of morphisms in a monoidal category -/ -scoped infixr:70 " βŠ— " => tensorHom - -/-- Notation for `tensorUnit`, the two-sided identity of `βŠ—` -/ -scoped notation "πŸ™_" => tensorUnit - -/-- Notation for the monoidal `associator`: `(X βŠ— Y) βŠ— Z) ≃ X βŠ— (Y βŠ— Z)` -/ -scoped notation "Ξ±_" => associator - -/-- Notation for the `leftUnitor`: `πŸ™_C βŠ— X ≃ X` -/ -scoped notation "Ξ»_" => leftUnitor - -/-- Notation for the `rightUnitor`: `X βŠ— πŸ™_C ≃ X` -/ -scoped notation "ρ_" => rightUnitor - variable {C : Type u} [π’ž : Category.{v} C] [MonoidalCategory C] @[reassoc] @@ -1102,7 +1062,7 @@ variable (Cβ‚‚ : Type uβ‚‚) [Category.{vβ‚‚} Cβ‚‚] [MonoidalCategory.{vβ‚‚} Cβ‚‚ attribute [local simp] associator_naturality leftUnitor_naturality rightUnitor_naturality pentagon attribute [local simp] tensorHom_def -@[simps! tensorObj tensorUnit' whiskerLeft whiskerRight associator] +@[simps! tensorObj tensorUnit whiskerLeft whiskerRight associator] instance prodMonoidal : MonoidalCategory (C₁ Γ— Cβ‚‚) where tensorObj X Y := (X.1 βŠ— Y.1, X.2 βŠ— Y.2) tensorHom f g := (f.1 βŠ— g.1, f.2 βŠ— g.2) diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index 1a3897219f6fa..d69e980feab6b 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -251,7 +251,7 @@ instance : MonoidalCategory (Center C) where tensorHom_def := by intros; ext; simp [tensorHom_def] whiskerLeft X _ _ f := whiskerLeft X f whiskerRight f X := whiskerRight f X - tensorUnit' := tensorUnit + tensorUnit := tensorUnit associator := associator leftUnitor := leftUnitor rightUnitor := rightUnitor diff --git a/Mathlib/CategoryTheory/Monoidal/End.lean b/Mathlib/CategoryTheory/Monoidal/End.lean index a10e21b45f391..c66a95eae46b1 100644 --- a/Mathlib/CategoryTheory/Monoidal/End.lean +++ b/Mathlib/CategoryTheory/Monoidal/End.lean @@ -62,8 +62,8 @@ attribute [local instance] endofunctorMonoidalCategory (Ξ± βŠ— Ξ²).app X = Ξ².app (F.obj X) ≫ K.map (Ξ±.app X) := rfl @[simp] theorem endofunctorMonoidalCategory_whiskerLeft_app - {F G H : C β₯€ C} {Ξ± : F ⟢ G} (X : C) : - (H ◁ Ξ±).app X = Ξ±.app (H.obj X) := rfl + {F H K : C β₯€ C} {Ξ² : H ⟢ K} (X : C) : + (F ◁ Ξ²).app X = Ξ².app (F.obj X) := rfl @[simp] theorem endofunctorMonoidalCategory_whiskerRight_app {F G H : C β₯€ C} {Ξ± : F ⟢ G} (X : C) : diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean index 8fc0e07f98fa2..28c8196869c92 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean @@ -173,7 +173,7 @@ instance : MonoidalCategory (F C) where whiskerLeft := fun X _ _ f => Quotient.map (Hom.whiskerLeft X) (HomEquiv.whisker_left X) f whiskerRight := fun f Y => Quotient.map (fun f' => Hom.whiskerRight f' Y) (fun _ _ h => HomEquiv.whisker_right _ _ _ h) f - tensorUnit' := FreeMonoidalCategory.Unit + tensorUnit := FreeMonoidalCategory.Unit associator X Y Z := ⟨⟦Hom.Ξ±_hom X Y Z⟧, ⟦Hom.Ξ±_inv X Y Z⟧, Quotient.sound Ξ±_hom_inv, Quotient.sound Ξ±_inv_hom⟩ leftUnitor X := ⟨⟦Hom.l_hom X⟧, ⟦Hom.l_inv X⟧, Quotient.sound l_hom_inv, Quotient.sound l_inv_hom⟩ diff --git a/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean b/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean index 23145d8320753..429b5701b8a6d 100644 --- a/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean @@ -88,12 +88,10 @@ instance functorCategoryMonoidalStruct : MonoidalCategoryStruct (C β₯€ D) where tensorHom Ξ± Ξ² := tensorHom Ξ± Ξ² whiskerLeft F _ _ Ξ± := FunctorCategory.whiskerLeft F Ξ± whiskerRight Ξ± F := FunctorCategory.whiskerRight Ξ± F - tensorUnit' := (CategoryTheory.Functor.const C).obj (πŸ™_ D) + tensorUnit := (CategoryTheory.Functor.const C).obj (πŸ™_ D) leftUnitor F := NatIso.ofComponents fun X => Ξ»_ (F.obj X) rightUnitor F := NatIso.ofComponents fun X => ρ_ (F.obj X) associator F G H := NatIso.ofComponents fun X => Ξ±_ (F.obj X) (G.obj X) (H.obj X) - whisker_exchange := by intros; ext; simp [whisker_exchange] -#align category_theory.monoidal.functor_category_monoidal CategoryTheory.Monoidal.functorCategoryMonoidal @[simp] theorem tensorUnit_obj {X} : (πŸ™_ (C β₯€ D)).obj X = πŸ™_ D := @@ -173,6 +171,7 @@ where `(F βŠ— G).obj X = F.obj X βŠ— G.obj X`. -/ instance functorCategoryMonoidal : MonoidalCategory (C β₯€ D) where tensorHom_def := by intros; ext; simp [tensorHom_def] + whisker_exchange := by intros; ext; simp [whisker_exchange] pentagon F G H K := by ext X; dsimp; rw [pentagon] #align category_theory.monoidal.functor_category_monoidal CategoryTheory.Monoidal.functorCategoryMonoidal diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index 719c2eb378216..524eaf4ce210f 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -82,6 +82,11 @@ attribute [nolint docBlame] InducingFunctorData.leftUnitor_eq InducingFunctorData.rightUnitor_eq +theorem InducingFunctorData.leftUnitor_inv_eq [MonoidalCategoryStruct D] {F : D β₯€ C} (fData : InducingFunctorData F) (X : D) : + F.map (Ξ»_ X).inv = (((fData.ΞΌIso (πŸ™_ D) X).symm β‰ͺ≫ (fData.Ξ΅Iso.symm βŠ— Iso.refl (F.obj X))) β‰ͺ≫ Ξ»_ (F.obj X)).inv := by + rw [← Functor.mapIso_inv, Iso.inv_eq_inv] + simp [fData.leftUnitor_eq] + /-- Induce the lawfulness of the monoidal structure along an faithful functor of (plain) categories, where the operations are already defined on the destination type `D`. @@ -96,46 +101,59 @@ abbrev induced [MonoidalCategoryStruct D] (F : D β₯€ C) [Faithful F] tensorHom_def {X₁ Y₁ Xβ‚‚ Yβ‚‚} f g := F.map_injective <| by rw [fData.tensorHom_eq, Functor.map_comp, fData.whiskerRight_eq, fData.whiskerLeft_eq] simp only [tensorHom_def, assoc, Iso.hom_inv_id_assoc] - tensor_id X₁ Xβ‚‚ := F.map_injective <| by cases fData; aesop_cat - tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚} f₁ fβ‚‚ g₁ gβ‚‚ := F.map_injective <| by cases fData; aesop_cat + -- tensor_id X₁ Xβ‚‚ := F.map_injective <| by cases fData; aesop_cat + -- tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚} f₁ fβ‚‚ g₁ gβ‚‚ := F.map_injective <| by cases fData; aesop_cat whiskerLeft_id X Y := F.map_injective <| by simp [fData.whiskerLeft_eq] + whiskerLeft_comp W X Y Z f g := F.map_injective <| by cases fData; aesop_cat + id_whiskerLeft {X Y} f := by + rw [← Category.assoc, Iso.eq_comp_inv] + apply F.map_injective + simp [fData.leftUnitor_eq, fData.whiskerLeft_eq, whisker_exchange_assoc] id_whiskerRight X Y := F.map_injective <| by simp [fData.whiskerRight_eq] + tensor_whiskerLeft X Y Z Z' f := by + rw [← Category.assoc, Iso.eq_comp_inv] + apply F.map_injective + simp [fData.associator_eq, fData.whiskerLeft_eq, whisker_exchange_assoc] + comp_whiskerRight := sorry + whiskerRight_id := sorry + whiskerRight_tensor :=sorry + whisker_assoc := sorry + whisker_exchange := sorry triangle X Y := F.map_injective <| by cases fData; aesop_cat pentagon W X Y Z := F.map_injective <| by - simp only [Functor.map_comp, fData.tensorHom_eq, fData.associator_eq, Iso.trans_assoc, - Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, Functor.map_id, comp_tensor_id, - associator_conjugation, tensor_id, assoc, id_tensor_comp, Iso.hom_inv_id_assoc, - tensor_hom_inv_id_assoc, id_comp, hom_inv_id_tensor_assoc, Iso.inv_hom_id_assoc, - id_tensor_comp_tensor_id_assoc, Iso.cancel_iso_inv_left] - slice_lhs 6 8 => - rw [← id_tensor_comp, hom_inv_id_tensor, tensor_id, comp_id, - tensor_id] - simp only [comp_id, assoc, pentagon_assoc, Iso.inv_hom_id_assoc, - ← associator_naturality_assoc, tensor_id, tensor_id_comp_id_tensor_assoc] - leftUnitor_naturality {X Y : D} f := F.map_injective <| by - have := leftUnitor_naturality (F.map f) - simp only [Functor.map_comp, fData.tensorHom_eq, Functor.map_id, fData.leftUnitor_eq, - Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, assoc, - Iso.hom_inv_id_assoc, id_tensor_comp_tensor_id_assoc, Iso.cancel_iso_inv_left] - rw [←this, ←assoc, ←tensor_comp, id_comp, comp_id] - rightUnitor_naturality {X Y : D} f := F.map_injective <| by - have := rightUnitor_naturality (F.map f) - simp only [Functor.map_comp, fData.tensorHom_eq, Functor.map_id, fData.rightUnitor_eq, - Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, assoc, - Iso.hom_inv_id_assoc, tensor_id_comp_id_tensor_assoc, Iso.cancel_iso_inv_left] - rw [←this, ←assoc, ←tensor_comp, id_comp, comp_id] - associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃} f₁ fβ‚‚ f₃ := F.map_injective <| by - have := associator_naturality (F.map f₁) (F.map fβ‚‚) (F.map f₃) - simp [fData.associator_eq, fData.tensorHom_eq] - simp_rw [←assoc, ←tensor_comp, assoc, Iso.hom_inv_id, ←assoc] - congr 1 - conv_rhs => rw [←comp_id (F.map f₁), ←id_comp (F.map f₁)] - simp only [tensor_comp] - simp only [tensor_id, comp_id, assoc, tensor_hom_inv_id_assoc, id_comp] - slice_rhs 2 3 => rw [←this] - simp only [← assoc, Iso.inv_hom_id, comp_id] - congr 2 - simp_rw [←tensor_comp, id_comp] + simp only [Functor.map_comp, fData.whiskerRight_eq, fData.associator_eq, Iso.trans_assoc, + Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, tensorHom_id, id_tensorHom, + comp_whiskerRight, whisker_assoc, assoc, fData.whiskerLeft_eq, + MonoidalCategory.whiskerLeft_comp, Iso.hom_inv_id_assoc, hom_inv_whiskerLeft_assoc, + hom_inv_whiskerRight_assoc, Iso.inv_hom_id_assoc, Iso.cancel_iso_inv_left] + slice_lhs 5 6 => + rw [← MonoidalCategory.whiskerLeft_comp, hom_inv_whiskerRight] + rw [whisker_exchange_assoc] + simp + -- leftUnitor_naturality {X Y : D} f := F.map_injective <| by + -- have := leftUnitor_naturality (F.map f) + -- simp only [Functor.map_comp, fData.tensorHom_eq, Functor.map_id, fData.leftUnitor_eq, + -- Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, assoc, + -- Iso.hom_inv_id_assoc, id_tensor_comp_tensor_id_assoc, Iso.cancel_iso_inv_left] + -- rw [←this, ←assoc, ←tensor_comp, id_comp, comp_id] + -- rightUnitor_naturality {X Y : D} f := F.map_injective <| by + -- have := rightUnitor_naturality (F.map f) + -- simp only [Functor.map_comp, fData.tensorHom_eq, Functor.map_id, fData.rightUnitor_eq, + -- Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, assoc, + -- Iso.hom_inv_id_assoc, tensor_id_comp_id_tensor_assoc, Iso.cancel_iso_inv_left] + -- rw [←this, ←assoc, ←tensor_comp, id_comp, comp_id] + -- associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃} f₁ fβ‚‚ f₃ := F.map_injective <| by + -- have := associator_naturality (F.map f₁) (F.map fβ‚‚) (F.map f₃) + -- simp [fData.associator_eq, fData.tensorHom_eq] + -- simp_rw [←assoc, ←tensor_comp, assoc, Iso.hom_inv_id, ←assoc] + -- congr 1 + -- conv_rhs => rw [←comp_id (F.map f₁), ←id_comp (F.map f₁)] + -- simp only [tensor_comp] + -- simp only [tensor_id, comp_id, assoc, tensor_hom_inv_id_assoc, id_comp] + -- slice_rhs 2 3 => rw [←this] + -- simp only [← assoc, Iso.inv_hom_id, comp_id] + -- congr 2 + -- simp_rw [←tensor_comp, id_comp] /-- @@ -150,119 +168,12 @@ def fromInduced [MonoidalCategoryStruct D] (F : D β₯€ C) [Faithful F] { toFunctor := F Ξ΅ := fData.Ξ΅Iso.hom ΞΌ := fun X Y => (fData.ΞΌIso X Y).hom - ΞΌ_natural := by cases fData; aesop_cat + ΞΌ_natural_left := by cases fData; aesop_cat + ΞΌ_natural_right := by cases fData; aesop_cat associativity := by cases fData; aesop_cat left_unitality := by cases fData; aesop_cat right_unitality := by cases fData; aesop_cat } -/-- Transport a monoidal structure along an equivalence of (plain) categories. --/ -@[simps (config := {rhsMd := .default}) - tensorObj tensorHom tensorUnit' associator leftUnitor rightUnitor] -def transport (e : C β‰Œ D) : MonoidalCategory.{vβ‚‚} D := .ofTensorHom - (tensorObj := fun X Y ↦ e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y)) - (tensorHom := fun f g ↦ e.functor.map (e.inverse.map f βŠ— e.inverse.map g)) - (whiskerLeft := fun X _ _ f ↦ e.functor.map (e.inverse.obj X ◁ e.inverse.map f)) - (whiskerRight := fun f X ↦ e.functor.map (e.inverse.map f β–· e.inverse.obj X)) - (tensorUnit' := e.functor.obj (πŸ™_ C)) - (associator := fun X Y Z ↦ - e.functor.mapIso - ((whiskerRightIso (e.unitIso.app _).symm _) β‰ͺ≫ - Ξ±_ (e.inverse.obj X) (e.inverse.obj Y) (e.inverse.obj Z) β‰ͺ≫ - (whiskerLeftIso _ (e.unitIso.app _)))) - (leftUnitor := fun X ↦ - e.functor.mapIso ((whiskerRightIso (e.unitIso.app _).symm _) β‰ͺ≫ Ξ»_ (e.inverse.obj X)) β‰ͺ≫ - e.counitIso.app _) - (rightUnitor := fun X ↦ - e.functor.mapIso ((whiskerLeftIso _ (e.unitIso.app _).symm) β‰ͺ≫ ρ_ (e.inverse.obj X)) β‰ͺ≫ - e.counitIso.app _) - (triangle := fun X Y ↦ by - dsimp - simp only [Functor.map_comp, Functor.map_id, assoc, Equivalence.inv_fun_map, - Functor.comp_obj, Functor.id_obj, - Equivalence.unit_inverse_comp, comp_id, Iso.hom_inv_id_app_assoc, - id_tensorHom, MonoidalCategory.whiskerLeft_comp, - tensorHom_id, comp_whiskerRight, whisker_assoc, triangle_assoc_comp_right] - simp only [← e.functor.map_comp] - congr 2 - slice_lhs 2 3 => - rw [← MonoidalCategory.whiskerLeft_comp] - simp - simp) - (pentagon := fun W X Y Z ↦ by - dsimp - simp only [Functor.map_comp, Equivalence.inv_fun_map, Functor.comp_obj, Functor.id_obj, assoc, - Iso.hom_inv_id_app_assoc, Functor.map_id, tensorHom_id, comp_whiskerRight, whisker_assoc, - id_tensorHom, MonoidalCategory.whiskerLeft_comp] - simp only [← e.functor.map_comp] - congr 2 - slice_lhs 6 7 => - rw [← comp_whiskerRight, Iso.hom_inv_id_app] - dsimp - rw [id_whiskerRight] - simp only [id_comp, assoc, Iso.inv_hom_id_assoc] - slice_lhs 5 6 => - rw [← MonoidalCategory.whiskerLeft_comp, Iso.hom_inv_id_app] - dsimp - rw [MonoidalCategory.whiskerLeft_id] - simp only [Category.id_comp, Category.assoc] - slice_rhs 2 3 => rw [whisker_exchange] - slice_lhs 4 5 => - rw [← MonoidalCategory.whiskerLeft_comp, ← comp_whiskerRight, Iso.hom_inv_id_app] - dsimp - rw [id_whiskerRight, MonoidalCategory.whiskerLeft_id] - simp) - (leftUnitor_naturality := fun f ↦ by - dsimp - simp only [Functor.map_comp, Functor.map_id, Category.assoc] - erw [← e.counitIso.hom.naturality] - simp only [Functor.comp_map, ← e.functor.map_comp_assoc] - congr 2 - simp only [id_tensorHom] - rw [whisker_exchange_assoc, leftUnitor_naturality]) - (rightUnitor_naturality := fun f ↦ by - dsimp - simp only [Functor.map_comp, Functor.map_id, Category.assoc] - erw [← e.counitIso.hom.naturality] - simp only [Functor.comp_map, ← e.functor.map_comp_assoc] - congr 2 - simp only [tensorHom_id] - erw [← whisker_exchange_assoc, rightUnitor_naturality]) - (associator_naturality := fun f₁ fβ‚‚ f₃ ↦ by - dsimp - simp only [Equivalence.inv_fun_map, Functor.map_comp, Category.assoc] - simp only [← e.functor.map_comp] - congr 1 - conv_lhs => rw [← tensor_id_comp_id_tensor] - simp only [← id_tensorHom, ← tensorHom_id] - slice_lhs 2 3 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor, ← tensor_id] - simp only [Category.assoc] - slice_lhs 3 4 => rw [associator_naturality] - conv_lhs => simp only [comp_tensor_id] - slice_lhs 3 4 => - rw [← comp_tensor_id, Iso.hom_inv_id_app] - dsimp - rw [tensor_id] - simp only [Category.id_comp, Category.assoc] - slice_lhs 2 3 => rw [associator_naturality] - simp only [Category.assoc] - congr 2 - slice_lhs 1 1 => rw [← tensor_id_comp_id_tensor] - slice_lhs 2 3 => rw [← id_tensor_comp, tensor_id_comp_id_tensor] - slice_lhs 1 2 => rw [tensor_id_comp_id_tensor] - conv_rhs => - congr - Β· skip - Β· rw [← id_tensor_comp_tensor_id, id_tensor_comp] - simp only [Category.assoc] - slice_rhs 1 2 => - rw [← id_tensor_comp, Iso.hom_inv_id_app] - dsimp - rw [tensor_id] - simp only [Category.id_comp, Category.assoc] - conv_rhs => rw [id_tensor_comp] - slice_rhs 2 3 => rw [id_tensor_comp_tensor_id, ← tensor_id_comp_id_tensor] - slice_rhs 1 2 => rw [id_tensor_comp_tensor_id]) @[simps] def transportStruct (e : C β‰Œ D) : MonoidalCategoryStruct.{vβ‚‚} D where tensorObj X Y := e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y) @@ -308,93 +219,6 @@ instance Transported.instMonoidalCategory (e : C β‰Œ D) : MonoidalCategory (Tran instance (e : C β‰Œ D) : Inhabited (Transported e) := βŸ¨πŸ™_ _⟩ -section - -attribute [local simp] transport_tensorUnit' - -section - -/-- -We can upgrade `e.functor` to a lax monoidal functor from `C` to `D` with the transported structure. --/ -@[simps!] -def laxToTransported (e : C β‰Œ D) : LaxMonoidalFunctor C (Transported e) := .ofTensorHom - (F := e.functor) - (Ξ΅ := πŸ™ (e.functor.obj (πŸ™_ C))) - (ΞΌ := fun X Y ↦ e.functor.map (e.unitInv.app X βŠ— e.unitInv.app Y)) - (ΞΌ_natural := fun f g ↦ by - dsimp - rw [Equivalence.inv_fun_map, Equivalence.inv_fun_map, tensor_comp, Functor.map_comp, - tensor_comp, ← e.functor.map_comp, ← e.functor.map_comp, ← e.functor.map_comp, - assoc, assoc, ← tensor_comp, Iso.hom_inv_id_app, Iso.hom_inv_id_app, ← tensor_comp] - dsimp - rw [comp_id, comp_id]) - (associativity := fun X Y Z ↦ by - dsimp - simp only [← id_tensorHom, ← tensorHom_id] - rw [Equivalence.inv_fun_map, Equivalence.inv_fun_map, Functor.map_comp, - Functor.map_comp, assoc, assoc, e.inverse.map_id, e.inverse.map_id, - comp_tensor_id, id_tensor_comp, Functor.map_comp, assoc, id_tensor_comp, - comp_tensor_id, ← e.functor.map_comp, ← e.functor.map_comp, ← e.functor.map_comp, - ← e.functor.map_comp, ← e.functor.map_comp, ← e.functor.map_comp, ← e.functor.map_comp] - congr 2 - slice_lhs 3 3 => rw [← tensor_id_comp_id_tensor] - slice_lhs 2 3 => - rw [← comp_tensor_id, Iso.hom_inv_id_app] - dsimp - rw [tensor_id] - rw [id_comp] - slice_rhs 2 3 => - rw [←id_tensor_comp, Iso.hom_inv_id_app] - dsimp - rw [tensor_id] - rw [id_comp] - conv_rhs => rw [← id_tensor_comp_tensor_id _ (e.unitInv.app X)] - dsimp only [Functor.comp_obj] - slice_rhs 3 4 => - rw [← id_tensor_comp, Iso.hom_inv_id_app] - dsimp - rw [tensor_id] - simp only [associator_conjugation, ←tensor_id, ←tensor_comp, Iso.inv_hom_id, - Iso.inv_hom_id_assoc, assoc, id_comp, comp_id]) - (left_unitality := fun X ↦ by - dsimp - simp only [← id_tensorHom, ← tensorHom_id] - rw [e.inverse.map_id, e.inverse.map_id, tensor_id, Functor.map_comp, assoc, - Equivalence.counit_app_functor, ← e.functor.map_comp, ← e.functor.map_comp, - ← e.functor.map_comp, ← e.functor.map_comp, ← leftUnitor_naturality', - ← tensor_comp_assoc, comp_id, id_comp, id_comp] - rfl) - (right_unitality := fun X ↦ by - dsimp - simp only [← id_tensorHom, ← tensorHom_id] - rw [Functor.map_comp, assoc, e.inverse.map_id, e.inverse.map_id, tensor_id, - Functor.map_id, id_comp, Equivalence.counit_app_functor, ← e.functor.map_comp, - ← e.functor.map_comp, ← e.functor.map_comp, ← rightUnitor_naturality', - ← tensor_comp_assoc, id_comp, comp_id] - rfl) -#align category_theory.monoidal.lax_to_transported CategoryTheory.Monoidal.laxToTransported - -end - -/-- We can upgrade `e.functor` to a monoidal functor from `C` to `D` with the transported structure. --/ -@[simps] -def toTransported (e : C β‰Œ D) : MonoidalFunctor C (Transported e) where - toLaxMonoidalFunctor := laxToTransported e - Ξ΅_isIso := by - dsimp - infer_instance - ΞΌ_isIso X Y := by - dsimp - infer_instance -#align category_theory.monoidal.to_transported CategoryTheory.Monoidal.toTransported - -end - -instance (e : C β‰Œ D) : IsEquivalence (toTransported e).toFunctor := - inferInstanceAs (IsEquivalence e.functor) - /-- We can upgrade `e.inverse` to a monoidal functor from `D` with the transported structure to `C`. -/ @[simps!] @@ -404,21 +228,21 @@ def fromTransported (e : C β‰Œ D) : MonoidalFunctor (Transported e) C := by #align category_theory.monoidal.from_transported CategoryTheory.Monoidal.fromTransported instance instIsEquivalence_fromTransported (e : C β‰Œ D) : - IsEquivalence (fromTransported e).toFunctor := - inferInstanceAs <| IsEquivalence <| (e.functor.asEquivalence).inverse + IsEquivalence (fromTransported e).toFunctor := by + dsimp [fromTransported] + infer_instance #noalign category_theory.monoidal.lax_to_transported /-- We can upgrade `e.functor` to a monoidal functor from `C` to `D` with the transported structure. -/ -@[simps!] +@[simps] def toTransported (e : C β‰Œ D) : MonoidalFunctor C (Transported e) := monoidalInverse (fromTransported e) #align category_theory.monoidal.to_transported CategoryTheory.Monoidal.toTransported -instance (e : C β‰Œ D) : IsEquivalence (toTransported e).toFunctor := by - dsimp [toTransported] - infer_instance +instance (e : C β‰Œ D) : IsEquivalence (toTransported e).toFunctor := + inferInstanceAs (IsEquivalence e.functor) /-- The unit isomorphism upgrades to a monoidal isomorphism. -/ @[simps! hom inv] From 020d085a8894c23fd93081cceb379b4d9a36d1c7 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 5 Nov 2023 10:18:57 +0900 Subject: [PATCH 37/62] fix --- Mathlib/CategoryTheory/Monoidal/Braided.lean | 6 +-- Mathlib/CategoryTheory/Monoidal/Category.lean | 9 ++-- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 2 - .../CategoryTheory/Monoidal/Transport.lean | 47 +++++++------------ 4 files changed, 23 insertions(+), 41 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index 76e36cc546be3..93c8366e33b2e 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -103,9 +103,9 @@ theorem braiding_inv_tensor_left (X Y Z : C) : @[simp] theorem braiding_inv_tensor_right (X Y Z : C) : - (Ξ²_ X (Y βŠ— Z)).inv = - (Ξ±_ Y Z X).hom ≫ Y ◁ (Ξ²_ X Z).inv ≫ (Ξ±_ Y X Z).inv ≫ - (Ξ²_ X Y).inv β–· Z ≫ (Ξ±_ X Y Z).hom := + (Ξ²_ X (Y βŠ— Z)).inv = + (Ξ±_ Y Z X).hom ≫ Y ◁ (Ξ²_ X Z).inv ≫ (Ξ±_ Y X Z).inv ≫ + (Ξ²_ X Y).inv β–· Z ≫ (Ξ±_ X Y Z).hom := eq_of_inv_eq_inv (by simp) @[reassoc (attr := simp)] diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index fcd9c69c8d9fb..ee3d959ac1a01 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -618,7 +618,6 @@ theorem associator_inv_conjugation {X X' Y Y' Z Z' : C} (f : X ⟢ X') (g : Y theorem hom_inv_id_tensor {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : (f.hom βŠ— g) ≫ (f.inv βŠ— h) = (V ◁ g) ≫ (V ◁ h) := by rw [← tensor_comp, f.hom_inv_id, id_tensorHom, whiskerLeft_comp] - #align category_theory.monoidal_category.hom_inv_id_tensor CategoryTheory.MonoidalCategory.hom_inv_id_tensor @[reassoc (attr := simp)] @@ -628,10 +627,10 @@ theorem inv_hom_id_tensor {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ #align category_theory.monoidal_category.inv_hom_id_tensor CategoryTheory.MonoidalCategory.inv_hom_id_tensor @[reassoc (attr := simp)] -theorem tensorHom_inv_id {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : +theorem tensor_hom_inv_id {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : (g βŠ— f.hom) ≫ (h βŠ— f.inv) = (g β–· V) ≫ (h β–· V) := by rw [← tensor_comp, f.hom_inv_id, tensorHom_id, comp_whiskerRight] -#align category_theory.monoidal_category.tensor_hom_inv_id CategoryTheory.MonoidalCategory.tensorHom_inv_id +#align category_theory.monoidal_category.tensor_hom_inv_id CategoryTheory.MonoidalCategory.tensor_hom_inv_id @[reassoc (attr := simp)] theorem tensor_inv_hom_id {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : @@ -652,10 +651,10 @@ theorem inv_hom_id_tensor' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) #align category_theory.monoidal_category.inv_hom_id_tensor' CategoryTheory.MonoidalCategory.inv_hom_id_tensor' @[reassoc (attr := simp)] -theorem tensorHom_inv_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : +theorem tensor_hom_inv_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : (g βŠ— f) ≫ (h βŠ— inv f) = (g β–· V) ≫ (h β–· V) := by rw [← tensor_comp, IsIso.hom_inv_id, tensorHom_id, comp_whiskerRight] -#align category_theory.monoidal_category.tensor_hom_inv_id' CategoryTheory.MonoidalCategory.tensorHom_inv_id' +#align category_theory.monoidal_category.tensor_hom_inv_id' CategoryTheory.MonoidalCategory.tensor_hom_inv_id' @[reassoc (attr := simp)] theorem tensor_inv_hom_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index fe2ec965dba9f..f559eb48a0212 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -498,9 +498,7 @@ instance monMonoidal : MonoidalCategory (Mon_ C) := .ofTensorHom (tensor_id := by intros; ext; apply tensor_id) (tensor_comp := by intros; ext; apply tensor_comp) (associator_naturality := by intros; ext; dsimp; apply associator_naturality) - (leftUnitor := fun M ↦ isoOfIso (Ξ»_ M.X) one_leftUnitor mul_leftUnitor) (leftUnitor_naturality := by intros; ext; dsimp; apply leftUnitor_naturality') - (rightUnitor := fun M ↦ isoOfIso (ρ_ M.X) one_rightUnitor mul_rightUnitor) (rightUnitor_naturality := by intros; ext; dsimp; apply rightUnitor_naturality') (pentagon := by intros; ext; dsimp; apply pentagon') (triangle := by intros; ext; dsimp; apply triangle') diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index 524eaf4ce210f..1fc2577dc6678 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -114,11 +114,21 @@ abbrev induced [MonoidalCategoryStruct D] (F : D β₯€ C) [Faithful F] rw [← Category.assoc, Iso.eq_comp_inv] apply F.map_injective simp [fData.associator_eq, fData.whiskerLeft_eq, whisker_exchange_assoc] - comp_whiskerRight := sorry - whiskerRight_id := sorry - whiskerRight_tensor :=sorry - whisker_assoc := sorry - whisker_exchange := sorry + comp_whiskerRight {W X Y} f g Z := F.map_injective <| by cases fData; aesop_cat + whiskerRight_id {X Y} f := by + rw [← Category.assoc, Iso.eq_comp_inv] + apply F.map_injective + simp [fData.rightUnitor_eq, fData.whiskerRight_eq, ← whisker_exchange_assoc] + whiskerRight_tensor {X X'} f Y Z := by + rw [Iso.eq_inv_comp] + apply F.map_injective + simp [fData.associator_eq, fData.whiskerRight_eq, whisker_exchange_assoc] + whisker_assoc X Y Y' f Z := by + rw [← Category.assoc, Iso.eq_comp_inv] + apply F.map_injective + simp [fData.associator_eq, fData.whiskerLeft_eq, fData.whiskerRight_eq, whisker_exchange_assoc] + whisker_exchange f g := F.map_injective <| by + simp [fData.whiskerLeft_eq, fData.whiskerRight_eq, whisker_exchange_assoc] triangle X Y := F.map_injective <| by cases fData; aesop_cat pentagon W X Y Z := F.map_injective <| by simp only [Functor.map_comp, fData.whiskerRight_eq, fData.associator_eq, Iso.trans_assoc, @@ -130,31 +140,6 @@ abbrev induced [MonoidalCategoryStruct D] (F : D β₯€ C) [Faithful F] rw [← MonoidalCategory.whiskerLeft_comp, hom_inv_whiskerRight] rw [whisker_exchange_assoc] simp - -- leftUnitor_naturality {X Y : D} f := F.map_injective <| by - -- have := leftUnitor_naturality (F.map f) - -- simp only [Functor.map_comp, fData.tensorHom_eq, Functor.map_id, fData.leftUnitor_eq, - -- Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, assoc, - -- Iso.hom_inv_id_assoc, id_tensor_comp_tensor_id_assoc, Iso.cancel_iso_inv_left] - -- rw [←this, ←assoc, ←tensor_comp, id_comp, comp_id] - -- rightUnitor_naturality {X Y : D} f := F.map_injective <| by - -- have := rightUnitor_naturality (F.map f) - -- simp only [Functor.map_comp, fData.tensorHom_eq, Functor.map_id, fData.rightUnitor_eq, - -- Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, assoc, - -- Iso.hom_inv_id_assoc, tensor_id_comp_id_tensor_assoc, Iso.cancel_iso_inv_left] - -- rw [←this, ←assoc, ←tensor_comp, id_comp, comp_id] - -- associator_naturality {X₁ Xβ‚‚ X₃ Y₁ Yβ‚‚ Y₃} f₁ fβ‚‚ f₃ := F.map_injective <| by - -- have := associator_naturality (F.map f₁) (F.map fβ‚‚) (F.map f₃) - -- simp [fData.associator_eq, fData.tensorHom_eq] - -- simp_rw [←assoc, ←tensor_comp, assoc, Iso.hom_inv_id, ←assoc] - -- congr 1 - -- conv_rhs => rw [←comp_id (F.map f₁), ←id_comp (F.map f₁)] - -- simp only [tensor_comp] - -- simp only [tensor_id, comp_id, assoc, tensor_hom_inv_id_assoc, id_comp] - -- slice_rhs 2 3 => rw [←this] - -- simp only [← assoc, Iso.inv_hom_id, comp_id] - -- congr 2 - -- simp_rw [←tensor_comp, id_comp] - /-- We can upgrade `F` to a monoidal functor from `D` to `E` with the induced structure. @@ -236,7 +221,7 @@ instance instIsEquivalence_fromTransported (e : C β‰Œ D) : /-- We can upgrade `e.functor` to a monoidal functor from `C` to `D` with the transported structure. -/ -@[simps] +@[simps!] def toTransported (e : C β‰Œ D) : MonoidalFunctor C (Transported e) := monoidalInverse (fromTransported e) #align category_theory.monoidal.to_transported CategoryTheory.Monoidal.toTransported From 329b32e534951e5aae145bcb425bbe351b500658 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 5 Nov 2023 10:35:09 +0900 Subject: [PATCH 38/62] clean up `Transport.lean` --- Mathlib/CategoryTheory/Monoidal/Transport.lean | 7 ------- 1 file changed, 7 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index 1fc2577dc6678..89dc57d6d3e78 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -82,11 +82,6 @@ attribute [nolint docBlame] InducingFunctorData.leftUnitor_eq InducingFunctorData.rightUnitor_eq -theorem InducingFunctorData.leftUnitor_inv_eq [MonoidalCategoryStruct D] {F : D β₯€ C} (fData : InducingFunctorData F) (X : D) : - F.map (Ξ»_ X).inv = (((fData.ΞΌIso (πŸ™_ D) X).symm β‰ͺ≫ (fData.Ξ΅Iso.symm βŠ— Iso.refl (F.obj X))) β‰ͺ≫ Ξ»_ (F.obj X)).inv := by - rw [← Functor.mapIso_inv, Iso.inv_eq_inv] - simp [fData.leftUnitor_eq] - /-- Induce the lawfulness of the monoidal structure along an faithful functor of (plain) categories, where the operations are already defined on the destination type `D`. @@ -101,8 +96,6 @@ abbrev induced [MonoidalCategoryStruct D] (F : D β₯€ C) [Faithful F] tensorHom_def {X₁ Y₁ Xβ‚‚ Yβ‚‚} f g := F.map_injective <| by rw [fData.tensorHom_eq, Functor.map_comp, fData.whiskerRight_eq, fData.whiskerLeft_eq] simp only [tensorHom_def, assoc, Iso.hom_inv_id_assoc] - -- tensor_id X₁ Xβ‚‚ := F.map_injective <| by cases fData; aesop_cat - -- tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚} f₁ fβ‚‚ g₁ gβ‚‚ := F.map_injective <| by cases fData; aesop_cat whiskerLeft_id X Y := F.map_injective <| by simp [fData.whiskerLeft_eq] whiskerLeft_comp W X Y Z f g := F.map_injective <| by cases fData; aesop_cat id_whiskerLeft {X Y} f := by From 76d66b46707ef190a3e71a5dd309dd31303b68de Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 5 Nov 2023 10:49:01 +0900 Subject: [PATCH 39/62] resolve conflict --- .../Category/ModuleCat/Monoidal/Basic.lean | 34 ++++++++++++------- 1 file changed, 22 insertions(+), 12 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean index b328715aa0503..a016c975b101f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean @@ -288,24 +288,32 @@ instance : MonoidalPreadditive (ModuleCat.{u} R) := by Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] - rw [LinearMap.zero_apply, ← id_tensorHom, MonoidalCategory.hom_apply, LinearMap.zero_apply, - TensorProduct.tmul_zero] + rw [LinearMap.zero_apply] + -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 + erw [MonoidalCategory.whiskerLeft_apply] + rw [LinearMap.zero_apply, TensorProduct.tmul_zero] Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] - rw [LinearMap.zero_apply, ← tensorHom_id, MonoidalCategory.hom_apply, LinearMap.zero_apply, - TensorProduct.zero_tmul] + rw [LinearMap.zero_apply] + -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 + erw [MonoidalCategory.whiskerRight_apply] + rw [LinearMap.zero_apply, TensorProduct.zero_tmul] Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] rw [LinearMap.add_apply] - repeat rw [← id_tensorHom, MonoidalCategory.hom_apply] + -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 + erw [MonoidalCategory.whiskerLeft_apply, MonoidalCategory.whiskerLeft_apply] + erw [MonoidalCategory.whiskerLeft_apply] rw [LinearMap.add_apply, TensorProduct.tmul_add] Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] rw [LinearMap.add_apply] - repeat rw [← tensorHom_id, MonoidalCategory.hom_apply] + -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 + erw [MonoidalCategory.whiskerRight_apply, MonoidalCategory.whiskerRight_apply] + erw [MonoidalCategory.whiskerRight_apply] rw [LinearMap.add_apply, TensorProduct.add_tmul] -- Porting note: simp wasn't firing but rw was, annoying @@ -314,14 +322,16 @@ instance : MonoidalLinear R (ModuleCat.{u} R) := by Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] - rw [LinearMap.smul_apply, ← id_tensorHom, MonoidalCategory.hom_apply, - ← id_tensorHom, MonoidalCategory.hom_apply, - LinearMap.smul_apply, TensorProduct.tmul_smul] + rw [LinearMap.smul_apply] + -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 + erw [MonoidalCategory.whiskerLeft_apply, MonoidalCategory.whiskerLeft_apply] + rw [LinearMap.smul_apply, TensorProduct.tmul_smul] Β· dsimp only [autoParam]; intros refine' TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => _) simp only [LinearMap.comprβ‚‚_apply, TensorProduct.mk_apply] - rw [LinearMap.smul_apply, ← tensorHom_id, MonoidalCategory.hom_apply, - ← tensorHom_id, MonoidalCategory.hom_apply, - LinearMap.smul_apply, TensorProduct.smul_tmul, TensorProduct.tmul_smul] + rw [LinearMap.smul_apply] + -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 + erw [MonoidalCategory.whiskerRight_apply, MonoidalCategory.whiskerRight_apply] + rw [LinearMap.smul_apply, TensorProduct.smul_tmul, TensorProduct.tmul_smul] end ModuleCat From 325dc84c9e5099a25d402c7cc9100d8248295325 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 5 Nov 2023 11:52:55 +0900 Subject: [PATCH 40/62] resolve conflict --- Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean | 5 +++++ Mathlib/CategoryTheory/Monoidal/Opposite.lean | 2 -- .../QuadraticForm/QuadraticModuleCat/Monoidal.lean | 10 ++++++++++ 3 files changed, 15 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean index a274031e4c7cf..1dbff1f2d1096 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean @@ -83,6 +83,11 @@ noncomputable instance instMonoidalCategory : MonoidalCategory (AlgebraCat.{u} R simp only [eqToIso_refl, Iso.refl_trans, Iso.refl_symm, Iso.trans_hom, tensorIso_hom, Iso.refl_hom, MonoidalCategory.tensor_id] erw [Category.id_comp, Category.comp_id, MonoidalCategory.tensor_id, Category.id_comp] + leftUnitor_eq := fun X => by + dsimp only [forgetβ‚‚_module_obj, forgetβ‚‚_module_map, Iso.refl_symm, Iso.trans_hom, + Iso.refl_hom, tensorIso_hom] + simp only [MonoidalCategory.leftUnitor_conjugation, Category.id_comp, Iso.hom_inv_id] + rfl rightUnitor_eq := fun X => by dsimp erw [Category.id_comp, MonoidalCategory.tensor_id, Category.id_comp] diff --git a/Mathlib/CategoryTheory/Monoidal/Opposite.lean b/Mathlib/CategoryTheory/Monoidal/Opposite.lean index 8c7667c191838..27f87cb54f518 100644 --- a/Mathlib/CategoryTheory/Monoidal/Opposite.lean +++ b/Mathlib/CategoryTheory/Monoidal/Opposite.lean @@ -178,7 +178,6 @@ instance monoidalCategoryOp : MonoidalCategory Cα΅’α΅– where associator X Y Z := (Ξ±_ (unop X) (unop Y) (unop Z)).symm.op leftUnitor X := (Ξ»_ (unop X)).symm.op rightUnitor X := (ρ_ (unop X)).symm.op - associator X Y Z := (Ξ±_ (unop X) (unop Y) (unop Z)).symm.op whiskerLeft_id X Y := Quiver.Hom.unop_inj (by simp) whiskerLeft_comp W X Y Z f g := Quiver.Hom.unop_inj (by simp) id_whiskerLeft f := Quiver.Hom.unop_inj (by simp) @@ -211,7 +210,6 @@ instance monoidalCategoryMop : MonoidalCategory Cα΄Ήα΅’α΅– where associator X Y Z := (Ξ±_ (unmop Z) (unmop Y) (unmop X)).symm.mop leftUnitor X := (ρ_ (unmop X)).mop rightUnitor X := (Ξ»_ (unmop X)).mop - associator X Y Z := (Ξ±_ (unmop Z) (unmop Y) (unmop X)).symm.mop whiskerLeft_id X Y := unmop_inj (by simp) whiskerLeft_comp W X Y Z f g := unmop_inj (by simp) id_whiskerLeft f := unmop_inj (by simp) diff --git a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat/Monoidal.lean b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat/Monoidal.lean index a064029c7d484..2c4e1a99c04a3 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat/Monoidal.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat/Monoidal.lean @@ -82,6 +82,16 @@ noncomputable instance instMonoidalCategory : MonoidalCategory (QuadraticModuleC (forgetβ‚‚ (QuadraticModuleCat R) (ModuleCat R)) { ΞΌIso := fun X Y => Iso.refl _ Ξ΅Iso := Iso.refl _ + leftUnitor_eq := fun X => by + simp only [forgetβ‚‚_obj, forgetβ‚‚_map, Iso.refl_symm, Iso.trans_assoc, Iso.trans_hom, + Iso.refl_hom, tensorIso_hom, MonoidalCategory.tensorHom_id] + erw [MonoidalCategory.id_whiskerRight, Category.id_comp, Category.id_comp] + rfl + rightUnitor_eq := fun X => by + simp only [forgetβ‚‚_obj, forgetβ‚‚_map, Iso.refl_symm, Iso.trans_assoc, Iso.trans_hom, + Iso.refl_hom, tensorIso_hom, MonoidalCategory.id_tensorHom] + erw [MonoidalCategory.whiskerLeft_id, Category.id_comp, Category.id_comp] + rfl associator_eq := fun X Y Z => by dsimp only [forgetβ‚‚_obj, forgetβ‚‚_map_associator_hom] simp only [eqToIso_refl, Iso.refl_trans, Iso.refl_symm, Iso.trans_hom, tensorIso_hom, From cfe87f9ab2a7b237c9e009fd612883fdf058fa59 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 5 Nov 2023 12:06:46 +0900 Subject: [PATCH 41/62] revert unintended deletion --- Mathlib/CategoryTheory/Monoidal/Transport.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index 89dc57d6d3e78..bdcfe71a75f10 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -152,6 +152,8 @@ def fromInduced [MonoidalCategoryStruct D] (F : D β₯€ C) [Faithful F] left_unitality := by cases fData; aesop_cat right_unitality := by cases fData; aesop_cat } +/-- Transport a monoidal structure along an equivalence of (plain) categories. +-/ @[simps] def transportStruct (e : C β‰Œ D) : MonoidalCategoryStruct.{vβ‚‚} D where tensorObj X Y := e.functor.obj (e.inverse.obj X βŠ— e.inverse.obj Y) From 949e06a0f59c31b85849c6a7a916639d8c20f3b9 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Wed, 8 Nov 2023 16:26:19 +0900 Subject: [PATCH 42/62] Update Mathlib/CategoryTheory/Monoidal/Functorial.lean Co-authored-by: Eric Wieser --- Mathlib/CategoryTheory/Monoidal/Functorial.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Functorial.lean b/Mathlib/CategoryTheory/Monoidal/Functorial.lean index 62dd50047f7de..c13099bb72d50 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functorial.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functorial.lean @@ -82,7 +82,7 @@ class LaxMonoidal (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] extends LaxMonoidalS -- but that isn't the immediate plan. open LaxMonoidalStruct in /-- An unbundled description of lax monoidal functors. -/ -def LaxMonoidal.ofTensorHom (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] [LaxMonoidalStruct F] +abbrev LaxMonoidal.ofTensorHom (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] [LaxMonoidalStruct F] /- naturality -/ (ΞΌ_natural : βˆ€ {X Y X' Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y'), From cac8a2bfebd657583029d1b44a24ece501cc1866 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sat, 6 Jan 2024 15:59:23 +0900 Subject: [PATCH 43/62] Update Action/Monoidal --- .../RepresentationTheory/Action/Monoidal.lean | 39 ++++++++++--------- 1 file changed, 20 insertions(+), 19 deletions(-) diff --git a/Mathlib/RepresentationTheory/Action/Monoidal.lean b/Mathlib/RepresentationTheory/Action/Monoidal.lean index 050d09ba67536..3f1dc0aedbcd1 100644 --- a/Mathlib/RepresentationTheory/Action/Monoidal.lean +++ b/Mathlib/RepresentationTheory/Action/Monoidal.lean @@ -175,7 +175,7 @@ section variable [Preadditive V] [MonoidalPreadditive V] -attribute [local simp] MonoidalPreadditive.tensor_add MonoidalPreadditive.add_tensor +attribute [local simp] MonoidalPreadditive.whiskerLeft_add MonoidalPreadditive.add_whiskerRight instance : MonoidalPreadditive (Action V G) where @@ -387,28 +387,29 @@ variable {W : Type (u + 1)} [LargeCategory W] [MonoidalCategory V] [MonoidalCate /-- A monoidal functor induces a monoidal functor between the categories of `G`-actions within those categories. -/ -@[simps] -def mapAction : MonoidalFunctor (Action V G) (Action W G) where - toFunctor := F.toFunctor.mapAction G - Ξ΅ := +@[simps!] +def mapActionAux : LaxMonoidalFunctor (Action V G) (Action W G) := .ofTensorHom + (F := F.toFunctor.mapAction G) + (Ξ΅ := { hom := F.Ξ΅ comm := fun g => by dsimp [FunctorCategoryEquivalence.inverse, Functor.mapAction] - rw [Category.id_comp, F.map_id, Category.comp_id] } - ΞΌ := fun X Y => + rw [Category.id_comp, F.map_id, Category.comp_id] }) + (ΞΌ := fun X Y => { hom := F.ΞΌ X.V Y.V - comm := fun g => F.toLaxMonoidalFunctor.ΞΌ_natural (X.ρ g) (Y.ρ g) } - -- using `dsimp` before `simp` speeds these up - ΞΌ_natural {X Y X' Y'} f g := by ext; dsimp; simp - associativity X Y Z := by ext; dsimp; simp - left_unitality X := by ext; dsimp; simp - right_unitality X := by - ext - dsimp - simp only [MonoidalCategory.rightUnitor_conjugation, - LaxMonoidalFunctor.right_unitality, Category.id_comp, Category.assoc, - LaxMonoidalFunctor.right_unitality_inv_assoc, Category.comp_id, Iso.hom_inv_id] - rw [← F.map_comp, Iso.inv_hom_id, F.map_id, Category.comp_id] + comm := fun g => F.toLaxMonoidalFunctor.ΞΌ_natural (X.ρ g) (Y.ρ g) }) + (ΞΌ_natural := by intros; ext; simp) + (associativity := by intros; ext; simp) + (left_unitality := by intros; ext; simp) + (right_unitality := by intros; ext; simp) + +/-- A monoidal functor induces a monoidal functor between +the categories of `G`-actions within those categories. -/ +@[simps!] +def mapAction : MonoidalFunctor (Action V G) (Action W G) := + { mapActionAux F G with + Ξ΅_isIso := by dsimp [mapActionAux]; infer_instance + ΞΌ_isIso := by dsimp [mapActionAux]; infer_instance } set_option linter.uppercaseLean3 false in #align category_theory.monoidal_functor.map_Action CategoryTheory.MonoidalFunctor.mapAction From 716fc859c2eb4e9a09ad8ad3b865aa0026389110 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sat, 6 Jan 2024 16:13:21 +0900 Subject: [PATCH 44/62] remove Action.lean --- Mathlib/RepresentationTheory/Action.lean | 1048 ---------------------- 1 file changed, 1048 deletions(-) delete mode 100644 Mathlib/RepresentationTheory/Action.lean diff --git a/Mathlib/RepresentationTheory/Action.lean b/Mathlib/RepresentationTheory/Action.lean deleted file mode 100644 index f957b2ebf25b3..0000000000000 --- a/Mathlib/RepresentationTheory/Action.lean +++ /dev/null @@ -1,1048 +0,0 @@ -/- -Copyright (c) 2020 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison --/ -import Mathlib.Algebra.Category.GroupCat.Basic -import Mathlib.CategoryTheory.SingleObj -import Mathlib.CategoryTheory.Limits.FunctorCategory -import Mathlib.CategoryTheory.Limits.Preserves.Basic -import Mathlib.CategoryTheory.Adjunction.Limits -import Mathlib.CategoryTheory.Monoidal.FunctorCategory -import Mathlib.CategoryTheory.Monoidal.Transport -import Mathlib.CategoryTheory.Monoidal.Rigid.OfEquivalence -import Mathlib.CategoryTheory.Monoidal.Rigid.FunctorCategory -import Mathlib.CategoryTheory.Monoidal.Linear -import Mathlib.CategoryTheory.Monoidal.Braided -import Mathlib.CategoryTheory.Monoidal.Types.Symmetric -import Mathlib.CategoryTheory.Abelian.FunctorCategory -import Mathlib.CategoryTheory.Abelian.Transfer -import Mathlib.CategoryTheory.Conj -import Mathlib.CategoryTheory.Linear.FunctorCategory - -#align_import representation_theory.Action from "leanprover-community/mathlib"@"95a87616d63b3cb49d3fe678d416fbe9c4217bf4" - -/-! -# `Action V G`, the category of actions of a monoid `G` inside some category `V`. - -The prototypical example is `V = ModuleCat R`, -where `Action (ModuleCat R) G` is the category of `R`-linear representations of `G`. - -We check `Action V G β‰Œ (singleObj G β₯€ V)`, -and construct the restriction functors `res {G H : Mon} (f : G ⟢ H) : Action V H β₯€ Action V G`. - -* When `V` has (co)limits so does `Action V G`. -* When `V` is monoidal, braided, or symmetric, so is `Action V G`. -* When `V` is preadditive, linear, or abelian so is `Action V G`. --/ - - -universe u v - -open CategoryTheory Limits - -variable (V : Type (u + 1)) [LargeCategory V] - --- Note: this is _not_ a categorical action of `G` on `V`. -/-- An `Action V G` represents a bundled action of -the monoid `G` on an object of some category `V`. - -As an example, when `V = ModuleCat R`, this is an `R`-linear representation of `G`, -while when `V = Type` this is a `G`-action. --/ -structure Action (G : MonCat.{u}) where - V : V - ρ : G ⟢ MonCat.of (End V) -set_option linter.uppercaseLean3 false in -#align Action Action - -namespace Action - -variable {V} - -@[simp 1100] -theorem ρ_one {G : MonCat.{u}} (A : Action V G) : A.ρ 1 = πŸ™ A.V := by rw [MonoidHom.map_one]; rfl -set_option linter.uppercaseLean3 false in -#align Action.ρ_one Action.ρ_one - -/-- When a group acts, we can lift the action to the group of automorphisms. -/ -@[simps] -def ρAut {G : GroupCat.{u}} (A : Action V (MonCat.of G)) : G ⟢ GroupCat.of (Aut A.V) where - toFun g := - { hom := A.ρ g - inv := A.ρ (g⁻¹ : G) - hom_inv_id := (A.ρ.map_mul (g⁻¹ : G) g).symm.trans (by rw [inv_mul_self, ρ_one]) - inv_hom_id := (A.ρ.map_mul g (g⁻¹ : G)).symm.trans (by rw [mul_inv_self, ρ_one]) } - map_one' := Aut.ext A.ρ.map_one - map_mul' x y := Aut.ext (A.ρ.map_mul x y) -set_option linter.uppercaseLean3 false in -#align Action.ρ_Aut Action.ρAut - --- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing -attribute [nolint simpNF] Action.ρAut_apply_inv Action.ρAut_apply_hom - -variable (G : MonCat.{u}) - -section - -instance inhabited' : Inhabited (Action (Type u) G) := - ⟨⟨PUnit, 1⟩⟩ -set_option linter.uppercaseLean3 false in -#align Action.inhabited' Action.inhabited' - -/-- The trivial representation of a group. -/ -def trivial : Action AddCommGroupCat G where - V := AddCommGroupCat.of PUnit - ρ := 1 -set_option linter.uppercaseLean3 false in -#align Action.trivial Action.trivial - -instance : Inhabited (Action AddCommGroupCat G) := - ⟨trivial G⟩ - -end - -variable {G} - -/-- A homomorphism of `Action V G`s is a morphism between the underlying objects, -commuting with the action of `G`. --/ -@[ext] -structure Hom (M N : Action V G) where - hom : M.V ⟢ N.V - comm : βˆ€ g : G, M.ρ g ≫ hom = hom ≫ N.ρ g := by aesop_cat -set_option linter.uppercaseLean3 false in -#align Action.hom Action.Hom - -namespace Hom - -attribute [reassoc] comm -attribute [local simp] comm comm_assoc - -/-- The identity morphism on an `Action V G`. -/ -@[simps] -def id (M : Action V G) : Action.Hom M M where hom := πŸ™ M.V -set_option linter.uppercaseLean3 false in -#align Action.hom.id Action.Hom.id - -instance (M : Action V G) : Inhabited (Action.Hom M M) := - ⟨id M⟩ - -/-- The composition of two `Action V G` homomorphisms is the composition of the underlying maps. --/ -@[simps] -def comp {M N K : Action V G} (p : Action.Hom M N) (q : Action.Hom N K) : Action.Hom M K where - hom := p.hom ≫ q.hom -set_option linter.uppercaseLean3 false in -#align Action.hom.comp Action.Hom.comp - -end Hom - -instance : Category (Action V G) where - Hom M N := Hom M N - id M := Hom.id M - comp f g := Hom.comp f g - --- porting note: added because `Hom.ext` is not triggered automatically -@[ext] -lemma hom_ext {M N : Action V G} (φ₁ Ο†β‚‚ : M ⟢ N) (h : φ₁.hom = Ο†β‚‚.hom) : φ₁ = Ο†β‚‚ := - Hom.ext _ _ h - -@[simp] -theorem id_hom (M : Action V G) : (πŸ™ M : Hom M M).hom = πŸ™ M.V := - rfl -set_option linter.uppercaseLean3 false in -#align Action.id_hom Action.id_hom - -@[simp] -theorem comp_hom {M N K : Action V G} (f : M ⟢ N) (g : N ⟢ K) : - (f ≫ g : Hom M K).hom = f.hom ≫ g.hom := - rfl -set_option linter.uppercaseLean3 false in -#align Action.comp_hom Action.comp_hom - -/-- Construct an isomorphism of `G` actions/representations -from an isomorphism of the underlying objects, -where the forward direction commutes with the group action. -/ -@[simps] -def mkIso {M N : Action V G} (f : M.V β‰… N.V) - (comm : βˆ€ g : G, M.ρ g ≫ f.hom = f.hom ≫ N.ρ g := by aesop_cat) : M β‰… N where - hom := - { hom := f.hom - comm := comm } - inv := - { hom := f.inv - comm := fun g => by have w := comm g =≫ f.inv; simp at w; simp [w] } -set_option linter.uppercaseLean3 false in -#align Action.mk_iso Action.mkIso - -instance (priority := 100) isIso_of_hom_isIso {M N : Action V G} (f : M ⟢ N) [IsIso f.hom] : - IsIso f := IsIso.of_iso (mkIso (asIso f.hom) f.comm) -set_option linter.uppercaseLean3 false in -#align Action.is_iso_of_hom_is_iso Action.isIso_of_hom_isIso - -instance isIso_hom_mk {M N : Action V G} (f : M.V ⟢ N.V) [IsIso f] (w) : - @IsIso _ _ M N (Hom.mk f w) := - IsIso.of_iso (mkIso (asIso f) w) -set_option linter.uppercaseLean3 false in -#align Action.is_iso_hom_mk Action.isIso_hom_mk - -namespace FunctorCategoryEquivalence - -/-- Auxilliary definition for `functorCategoryEquivalence`. -/ -@[simps] -def functor : Action V G β₯€ SingleObj G β₯€ V where - obj M := - { obj := fun _ => M.V - map := fun g => M.ρ g - map_id := fun _ => M.ρ.map_one - map_comp := fun g h => M.ρ.map_mul h g } - map f := - { app := fun _ => f.hom - naturality := fun _ _ g => f.comm g } -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence.functor Action.FunctorCategoryEquivalence.functor - -/-- Auxilliary definition for `functorCategoryEquivalence`. -/ -@[simps] -def inverse : (SingleObj G β₯€ V) β₯€ Action V G where - obj F := - { V := F.obj PUnit.unit - ρ := - { toFun := fun g => F.map g - map_one' := F.map_id PUnit.unit - map_mul' := fun g h => F.map_comp h g } } - map f := - { hom := f.app PUnit.unit - comm := fun g => f.naturality g } -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence.inverse Action.FunctorCategoryEquivalence.inverse - -/-- Auxilliary definition for `functorCategoryEquivalence`. -/ -@[simps!] -def unitIso : 𝟭 (Action V G) β‰… functor β‹™ inverse := - NatIso.ofComponents fun M => mkIso (Iso.refl _) -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence.unit_iso Action.FunctorCategoryEquivalence.unitIso - -/-- Auxilliary definition for `functorCategoryEquivalence`. -/ -@[simps!] -def counitIso : inverse β‹™ functor β‰… 𝟭 (SingleObj G β₯€ V) := - NatIso.ofComponents fun M => NatIso.ofComponents fun X => Iso.refl _ -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence.counit_iso Action.FunctorCategoryEquivalence.counitIso - -end FunctorCategoryEquivalence - -section - -open FunctorCategoryEquivalence - -variable (V G) - -/-- The category of actions of `G` in the category `V` -is equivalent to the functor category `singleObj G β₯€ V`. --/ -@[simps] -def functorCategoryEquivalence : Action V G β‰Œ SingleObj G β₯€ V where - functor := functor - inverse := inverse - unitIso := unitIso - counitIso := counitIso -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence Action.functorCategoryEquivalence -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence.functor_def Action.functorCategoryEquivalence_functor -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence.inverse_def Action.functorCategoryEquivalence_inverse - -/- -porting note: these two lemmas are redundant with the projections created by the @[simps] -attribute above - -theorem functorCategoryEquivalence.functor_def : - (functorCategoryEquivalence V G).functor = FunctorCategoryEquivalence.functor := - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence.functor_def Action.functorCategoryEquivalence.functor_def - -theorem functorCategoryEquivalence.inverse_def : - (functorCategoryEquivalence V G).inverse = FunctorCategoryEquivalence.inverse := - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence.inverse_def Action.functorCategoryEquivalence.inverse_def-/ - -instance [HasFiniteProducts V] : HasFiniteProducts (Action V G) where - out _ := - Adjunction.hasLimitsOfShape_of_equivalence (Action.functorCategoryEquivalence _ _).functor - -instance [HasFiniteLimits V] : HasFiniteLimits (Action V G) where - out _ _ _ := - Adjunction.hasLimitsOfShape_of_equivalence (Action.functorCategoryEquivalence _ _).functor - -instance [HasLimits V] : HasLimits (Action V G) := - Adjunction.has_limits_of_equivalence (Action.functorCategoryEquivalence _ _).functor - -instance [HasColimits V] : HasColimits (Action V G) := - Adjunction.has_colimits_of_equivalence (Action.functorCategoryEquivalence _ _).functor - -end - -section Forget - -variable (V G) - -/-- (implementation) The forgetful functor from bundled actions to the underlying objects. - -Use the `CategoryTheory.forget` API provided by the `ConcreteCategory` instance below, -rather than using this directly. --/ -@[simps] -def forget : Action V G β₯€ V where - obj M := M.V - map f := f.hom -set_option linter.uppercaseLean3 false in -#align Action.forget Action.forget - -instance : Faithful (forget V G) where map_injective w := Hom.ext _ _ w - -instance [ConcreteCategory V] : ConcreteCategory (Action V G) where - forget := forget V G β‹™ ConcreteCategory.forget - -instance hasForgetToV [ConcreteCategory V] : HasForgetβ‚‚ (Action V G) V where forgetβ‚‚ := forget V G -set_option linter.uppercaseLean3 false in -#align Action.has_forget_to_V Action.hasForgetToV - -/-- The forgetful functor is intertwined by `functorCategoryEquivalence` with -evaluation at `PUnit.star`. -/ -def functorCategoryEquivalenceCompEvaluation : - (functorCategoryEquivalence V G).functor β‹™ (evaluation _ _).obj PUnit.unit β‰… forget V G := - Iso.refl _ -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence_comp_evaluation Action.functorCategoryEquivalenceCompEvaluation - -noncomputable instance instPreservesLimitsForget [HasLimits V] : - Limits.PreservesLimits (forget V G) := - Limits.preservesLimitsOfNatIso (Action.functorCategoryEquivalenceCompEvaluation V G) - -noncomputable instance instPreservesColimitsForget [HasColimits V] : - PreservesColimits (forget V G) := - preservesColimitsOfNatIso (Action.functorCategoryEquivalenceCompEvaluation V G) - --- TODO construct categorical images? -end Forget - -theorem Iso.conj_ρ {M N : Action V G} (f : M β‰… N) (g : G) : - N.ρ g = ((forget V G).mapIso f).conj (M.ρ g) := - by rw [Iso.conj_apply, Iso.eq_inv_comp]; simp [f.hom.comm] -set_option linter.uppercaseLean3 false in -#align Action.iso.conj_ρ Action.Iso.conj_ρ - -section HasZeroMorphisms - -variable [HasZeroMorphisms V] - --- porting note: in order to ease automation, the `Zero` instance is introduced separately, --- and the lemma `zero_hom` was moved just below -instance {X Y : Action V G} : Zero (X ⟢ Y) := ⟨0, by aesop_cat⟩ - -@[simp] -theorem zero_hom {X Y : Action V G} : (0 : X ⟢ Y).hom = 0 := - rfl -set_option linter.uppercaseLean3 false in -#align Action.zero_hom Action.zero_hom - -instance : HasZeroMorphisms (Action V G) where - -instance forget_preservesZeroMorphisms : Functor.PreservesZeroMorphisms (forget V G) where -set_option linter.uppercaseLean3 false in -#align Action.forget_preserves_zero_morphisms Action.forget_preservesZeroMorphisms - -instance forgetβ‚‚_preservesZeroMorphisms [ConcreteCategory V] : - Functor.PreservesZeroMorphisms (forgetβ‚‚ (Action V G) V) where -set_option linter.uppercaseLean3 false in -#align Action.forgetβ‚‚_preserves_zero_morphisms Action.forgetβ‚‚_preservesZeroMorphisms - -instance functorCategoryEquivalence_preservesZeroMorphisms : - Functor.PreservesZeroMorphisms (functorCategoryEquivalence V G).functor where -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence_preserves_zero_morphisms Action.functorCategoryEquivalence_preservesZeroMorphisms - -end HasZeroMorphisms - -section Preadditive - -variable [Preadditive V] - -instance : Preadditive (Action V G) where - homGroup X Y := - { add := fun f g => ⟨f.hom + g.hom, by simp [f.comm, g.comm]⟩ - neg := fun f => ⟨-f.hom, by simp [f.comm]⟩ - zero_add := by intros; ext; exact zero_add _ - add_zero := by intros; ext; exact add_zero _ - add_assoc := by intros; ext; exact add_assoc _ _ _ - add_left_neg := by intros; ext; exact add_left_neg _ - add_comm := by intros; ext; exact add_comm _ _ } - add_comp := by intros; ext; exact Preadditive.add_comp _ _ _ _ _ _ - comp_add := by intros; ext; exact Preadditive.comp_add _ _ _ _ _ _ - -instance forget_additive : Functor.Additive (forget V G) where -set_option linter.uppercaseLean3 false in -#align Action.forget_additive Action.forget_additive - -instance forgetβ‚‚_additive [ConcreteCategory V] : Functor.Additive (forgetβ‚‚ (Action V G) V) where -set_option linter.uppercaseLean3 false in -#align Action.forgetβ‚‚_additive Action.forgetβ‚‚_additive - -instance functorCategoryEquivalence_additive : - Functor.Additive (functorCategoryEquivalence V G).functor where -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence_additive Action.functorCategoryEquivalence_additive - -@[simp] -theorem neg_hom {X Y : Action V G} (f : X ⟢ Y) : (-f).hom = -f.hom := - rfl -set_option linter.uppercaseLean3 false in -#align Action.neg_hom Action.neg_hom - -@[simp] -theorem add_hom {X Y : Action V G} (f g : X ⟢ Y) : (f + g).hom = f.hom + g.hom := - rfl -set_option linter.uppercaseLean3 false in -#align Action.add_hom Action.add_hom - -@[simp] -theorem sum_hom {X Y : Action V G} {ΞΉ : Type*} (f : ΞΉ β†’ (X ⟢ Y)) (s : Finset ΞΉ) : - (s.sum f).hom = s.sum fun i => (f i).hom := - (forget V G).map_sum f s -set_option linter.uppercaseLean3 false in -#align Action.sum_hom Action.sum_hom - -end Preadditive - -section Linear - -variable [Preadditive V] {R : Type*} [Semiring R] [Linear R V] - -instance : Linear R (Action V G) where - homModule X Y := - { smul := fun r f => ⟨r β€’ f.hom, by simp [f.comm]⟩ - one_smul := by intros; ext; exact one_smul _ _ - smul_zero := by intros; ext; exact smul_zero _ - zero_smul := by intros; ext; exact zero_smul _ _ - add_smul := by intros; ext; exact add_smul _ _ _ - smul_add := by intros; ext; exact smul_add _ _ _ - mul_smul := by intros; ext; exact mul_smul _ _ _ } - smul_comp := by intros; ext; exact Linear.smul_comp _ _ _ _ _ _ - comp_smul := by intros; ext; exact Linear.comp_smul _ _ _ _ _ _ - -instance forget_linear : Functor.Linear R (forget V G) where -set_option linter.uppercaseLean3 false in -#align Action.forget_linear Action.forget_linear - -instance forgetβ‚‚_linear [ConcreteCategory V] : Functor.Linear R (forgetβ‚‚ (Action V G) V) where -set_option linter.uppercaseLean3 false in -#align Action.forgetβ‚‚_linear Action.forgetβ‚‚_linear - -instance functorCategoryEquivalence_linear : - Functor.Linear R (functorCategoryEquivalence V G).functor where -set_option linter.uppercaseLean3 false in -#align Action.functor_category_equivalence_linear Action.functorCategoryEquivalence_linear - -@[simp] -theorem smul_hom {X Y : Action V G} (r : R) (f : X ⟢ Y) : (r β€’ f).hom = r β€’ f.hom := - rfl -set_option linter.uppercaseLean3 false in -#align Action.smul_hom Action.smul_hom - -end Linear - -section Abelian - -/-- Auxilliary construction for the `Abelian (Action V G)` instance. -/ -def abelianAux : Action V G β‰Œ ULift.{u} (SingleObj G) β₯€ V := - (functorCategoryEquivalence V G).trans (Equivalence.congrLeft ULift.equivalence) -set_option linter.uppercaseLean3 false in -#align Action.abelian_aux Action.abelianAux - -noncomputable instance [Abelian V] : Abelian (Action V G) := - abelianOfEquivalence abelianAux.functor - -end Abelian - -section Monoidal - -open MonoidalCategory - -variable [MonoidalCategory V] - -instance instMonoidalCategory : MonoidalCategory (Action V G) := - Monoidal.transport (Action.functorCategoryEquivalence _ _).symm - -@[simp] -theorem tensorUnit_v : (πŸ™_ (Action V G)).V = πŸ™_ V := - rfl -set_option linter.uppercaseLean3 false in -#align Action.tensor_unit_V Action.tensorUnit_v - --- porting note: removed @[simp] as the simpNF linter complains -theorem tensorUnit_rho {g : G} : (πŸ™_ (Action V G)).ρ g = πŸ™ (πŸ™_ V) := - rfl -set_option linter.uppercaseLean3 false in -#align Action.tensor_unit_rho Action.tensorUnit_rho - -@[simp] -theorem tensor_v {X Y : Action V G} : (X βŠ— Y).V = X.V βŠ— Y.V := - rfl -set_option linter.uppercaseLean3 false in -#align Action.tensor_V Action.tensor_v - --- porting note: removed @[simp] as the simpNF linter complains -theorem tensor_rho {X Y : Action V G} {g : G} : (X βŠ— Y).ρ g = X.ρ g βŠ— Y.ρ g := - rfl -set_option linter.uppercaseLean3 false in -#align Action.tensor_rho Action.tensor_rho - -@[simp] -theorem tensor_hom {W X Y Z : Action V G} (f : W ⟢ X) (g : Y ⟢ Z) : (f βŠ— g).hom = f.hom βŠ— g.hom := - rfl -set_option linter.uppercaseLean3 false in -#align Action.tensor_hom Action.tensor_hom - -@[simp] -theorem whiskerLeft_hom (X : Action V G) {Y Z : Action V G} (f : Y ⟢ Z) : - (X ◁ f).hom = X.V ◁ f.hom := - rfl - -@[simp] -theorem whiskerRight_hom {X Y : Action V G} (f : X ⟢ Y) (Z : Action V G) : - (f β–· Z).hom = f.hom β–· Z.V := - rfl - -@[simp] -theorem whiskerLeft_v (X : Action V G) {Y Z : Action V G} (f : Y ⟢ Z) : - (X ◁ f).hom = X.V ◁ f.hom := - rfl - -@[simp] -theorem whiskerRight_v {X Y : Action V G} (f : X ⟢ Y) (Z : Action V G) : - (f β–· Z).hom = f.hom β–· Z.V := - rfl - --- porting note: removed @[simp] as the simpNF linter complains -theorem associator_hom_hom {X Y Z : Action V G} : - Hom.hom (Ξ±_ X Y Z).hom = (Ξ±_ X.V Y.V Z.V).hom := by - dsimp - simp -set_option linter.uppercaseLean3 false in -#align Action.associator_hom_hom Action.associator_hom_hom - --- porting note: removed @[simp] as the simpNF linter complains -theorem associator_inv_hom {X Y Z : Action V G} : - Hom.hom (Ξ±_ X Y Z).inv = (Ξ±_ X.V Y.V Z.V).inv := by - dsimp - simp -set_option linter.uppercaseLean3 false in -#align Action.associator_inv_hom Action.associator_inv_hom - --- porting note: removed @[simp] as the simpNF linter complains -theorem leftUnitor_hom_hom {X : Action V G} : Hom.hom (Ξ»_ X).hom = (Ξ»_ X.V).hom := by - dsimp - simp -set_option linter.uppercaseLean3 false in -#align Action.left_unitor_hom_hom Action.leftUnitor_hom_hom - --- porting note: removed @[simp] as the simpNF linter complains -theorem leftUnitor_inv_hom {X : Action V G} : Hom.hom (Ξ»_ X).inv = (Ξ»_ X.V).inv := by - dsimp - simp -set_option linter.uppercaseLean3 false in -#align Action.left_unitor_inv_hom Action.leftUnitor_inv_hom - --- porting note: removed @[simp] as the simpNF linter complains -theorem rightUnitor_hom_hom {X : Action V G} : Hom.hom (ρ_ X).hom = (ρ_ X.V).hom := by - dsimp - simp -set_option linter.uppercaseLean3 false in -#align Action.right_unitor_hom_hom Action.rightUnitor_hom_hom - --- porting note: removed @[simp] as the simpNF linter complains -theorem rightUnitor_inv_hom {X : Action V G} : Hom.hom (ρ_ X).inv = (ρ_ X.V).inv := by - dsimp - simp -set_option linter.uppercaseLean3 false in -#align Action.right_unitor_inv_hom Action.rightUnitor_inv_hom - -/-- Given an object `X` isomorphic to the tensor unit of `V`, `X` equipped with the trivial action -is isomorphic to the tensor unit of `Action V G`. -/ -def tensorUnitIso {X : V} (f : πŸ™_ V β‰… X) : πŸ™_ (Action V G) β‰… Action.mk X 1 := - Action.mkIso f fun _ => by - simp only [MonoidHom.one_apply, End.one_def, Category.id_comp f.hom, tensorUnit_rho, - MonCat.oneHom_apply, MonCat.one_of, Category.comp_id] -set_option linter.uppercaseLean3 false in -#align Action.tensor_unit_iso Action.tensorUnitIso - -variable (V G) - -/-- When `V` is monoidal the forgetful functor `Action V G` to `V` is monoidal. -/ -@[simps] -def forgetMonoidal : MonoidalFunctor (Action V G) V := - { toFunctor := Action.forget _ _ - Ξ΅ := πŸ™ _ - ΞΌ := fun X Y => πŸ™ _ } -set_option linter.uppercaseLean3 false in -#align Action.forget_monoidal Action.forgetMonoidal - -instance forgetMonoidal_faithful : Faithful (forgetMonoidal V G).toFunctor := by - change Faithful (forget V G); infer_instance -set_option linter.uppercaseLean3 false in -#align Action.forget_monoidal_faithful Action.forgetMonoidal_faithful - -section - -variable [BraidedCategory V] - -instance : BraidedCategory (Action V G) := - braidedCategoryOfFaithful (forgetMonoidal V G) (fun X Y => mkIso (Ξ²_ _ _) - (fun g => by simp [FunctorCategoryEquivalence.inverse])) (by aesop_cat) - -/-- When `V` is braided the forgetful functor `Action V G` to `V` is braided. -/ -@[simps!] -def forgetBraided : BraidedFunctor (Action V G) V := - { forgetMonoidal _ _ with } -set_option linter.uppercaseLean3 false in -#align Action.forget_braided Action.forgetBraided - -instance forgetBraided_faithful : Faithful (forgetBraided V G).toFunctor := by - change Faithful (forget V G); infer_instance -set_option linter.uppercaseLean3 false in -#align Action.forget_braided_faithful Action.forgetBraided_faithful - -end - -instance [SymmetricCategory V] : SymmetricCategory (Action V G) := - symmetricCategoryOfFaithful (forgetBraided V G) - -section - -variable [Preadditive V] [MonoidalPreadditive V] - -attribute [local simp] MonoidalPreadditive.whiskerLeft_add MonoidalPreadditive.add_whiskerRight - -instance : MonoidalPreadditive (Action V G) where - -variable {R : Type*} [Semiring R] [Linear R V] [MonoidalLinear R V] - -instance : MonoidalLinear R (Action V G) where - -end - -noncomputable section - -/-- Upgrading the functor `Action V G β₯€ (SingleObj G β₯€ V)` to a monoidal functor. -/ -def functorCategoryMonoidalEquivalence : MonoidalFunctor (Action V G) (SingleObj G β₯€ V) := - Monoidal.fromTransported (Action.functorCategoryEquivalence _ _).symm -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence Action.functorCategoryMonoidalEquivalence - -instance : IsEquivalence (functorCategoryMonoidalEquivalence V G).toFunctor := by - change IsEquivalence (Action.functorCategoryEquivalence _ _).functor; infer_instance - -@[simp] -theorem functorCategoryMonoidalEquivalence.ΞΌ_app (A B : Action V G) : - ((functorCategoryMonoidalEquivalence V G).ΞΌ A B).app PUnit.unit = πŸ™ _ := by - dsimp only [functorCategoryMonoidalEquivalence] - simp only [Monoidal.fromTransported_toLaxMonoidalFunctor_ΞΌ, NatTrans.comp_app] - -- porting note: Lean3 was able to see through some defeq, as the mathlib3 proof was - -- show (πŸ™ A.V βŠ— πŸ™ B.V) ≫ πŸ™ (A.V βŠ— B.V) ≫ (πŸ™ A.V βŠ— πŸ™ B.V) = πŸ™ (A.V βŠ— B.V) - -- simp only [monoidal_category.tensor_id, category.comp_id] - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.ΞΌ_app Action.functorCategoryMonoidalEquivalence.ΞΌ_app - -@[simp] -theorem functorCategoryMonoidalEquivalence.ΞΌIso_inv_app (A B : Action V G) : - ((functorCategoryMonoidalEquivalence V G).ΞΌIso A B).inv.app PUnit.unit = πŸ™ _ := by - rw [← NatIso.app_inv, ← IsIso.Iso.inv_hom] - refine' IsIso.inv_eq_of_hom_inv_id _ - rw [Category.comp_id, NatIso.app_hom, MonoidalFunctor.ΞΌIso_hom, - functorCategoryMonoidalEquivalence.ΞΌ_app] -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.ΞΌ_iso_inv_app Action.functorCategoryMonoidalEquivalence.ΞΌIso_inv_app - -@[simp] -theorem functorCategoryMonoidalEquivalence.Ξ΅_app : - (functorCategoryMonoidalEquivalence V G).Ξ΅.app PUnit.unit = πŸ™ _ := by - dsimp only [functorCategoryMonoidalEquivalence] - simp only [Monoidal.fromTransported_toLaxMonoidalFunctor_Ξ΅] - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.Ξ΅_app Action.functorCategoryMonoidalEquivalence.Ξ΅_app - -@[simp] -theorem functorCategoryMonoidalEquivalence.inv_counit_app_hom (A : Action V G) : - ((functorCategoryMonoidalEquivalence _ _).inv.adjunction.counit.app A).hom = πŸ™ _ := - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.inv_counit_app_hom Action.functorCategoryMonoidalEquivalence.inv_counit_app_hom - -@[simp] -theorem functorCategoryMonoidalEquivalence.counit_app (A : SingleObj G β₯€ V) : - ((functorCategoryMonoidalEquivalence _ _).adjunction.counit.app A).app PUnit.unit = πŸ™ _ := - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.counit_app Action.functorCategoryMonoidalEquivalence.counit_app - -@[simp] -theorem functorCategoryMonoidalEquivalence.inv_unit_app_app (A : SingleObj G β₯€ V) : - ((functorCategoryMonoidalEquivalence _ _).inv.adjunction.unit.app A).app PUnit.unit = πŸ™ _ := - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.inv_unit_app_app Action.functorCategoryMonoidalEquivalence.inv_unit_app_app - -@[simp] -theorem functorCategoryMonoidalEquivalence.unit_app_hom (A : Action V G) : - ((functorCategoryMonoidalEquivalence _ _).adjunction.unit.app A).hom = πŸ™ _ := - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.unit_app_hom Action.functorCategoryMonoidalEquivalence.unit_app_hom - -@[simp] -theorem functorCategoryMonoidalEquivalence.functor_map {A B : Action V G} (f : A ⟢ B) : - (functorCategoryMonoidalEquivalence _ _).map f = FunctorCategoryEquivalence.functor.map f := - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.functor_map Action.functorCategoryMonoidalEquivalence.functor_map - -@[simp] -theorem functorCategoryMonoidalEquivalence.inverse_map {A B : SingleObj G β₯€ V} (f : A ⟢ B) : - (functorCategoryMonoidalEquivalence _ _).inv.map f = FunctorCategoryEquivalence.inverse.map f := - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.inverse_map Action.functorCategoryMonoidalEquivalence.inverse_map - -variable (H : GroupCat.{u}) - -instance [RightRigidCategory V] : RightRigidCategory (SingleObj (H : MonCat.{u}) β₯€ V) := by - change RightRigidCategory (SingleObj H β₯€ V); infer_instance - -/-- If `V` is right rigid, so is `Action V G`. -/ -instance [RightRigidCategory V] : RightRigidCategory (Action V H) := - rightRigidCategoryOfEquivalence (functorCategoryMonoidalEquivalence V _) - -instance [LeftRigidCategory V] : LeftRigidCategory (SingleObj (H : MonCat.{u}) β₯€ V) := by - change LeftRigidCategory (SingleObj H β₯€ V); infer_instance - -/-- If `V` is left rigid, so is `Action V G`. -/ -instance [LeftRigidCategory V] : LeftRigidCategory (Action V H) := - leftRigidCategoryOfEquivalence (functorCategoryMonoidalEquivalence V _) - -instance [RigidCategory V] : RigidCategory (SingleObj (H : MonCat.{u}) β₯€ V) := by - change RigidCategory (SingleObj H β₯€ V); infer_instance - -/-- If `V` is rigid, so is `Action V G`. -/ -instance [RigidCategory V] : RigidCategory (Action V H) := - rigidCategoryOfEquivalence (functorCategoryMonoidalEquivalence V _) - -variable {V H} -variable (X : Action V H) - -@[simp] -theorem rightDual_v [RightRigidCategory V] : Xᘁ.V = X.Vᘁ := - rfl -set_option linter.uppercaseLean3 false in -#align Action.right_dual_V Action.rightDual_v - -@[simp] -theorem leftDual_v [LeftRigidCategory V] : (ᘁX).V = ᘁX.V := - rfl -set_option linter.uppercaseLean3 false in -#align Action.left_dual_V Action.leftDual_v - --- This lemma was always bad, but the linter only noticed after lean4#2644 -@[simp, nolint simpNF] -theorem rightDual_ρ [RightRigidCategory V] (h : H) : Xᘁ.ρ h = (X.ρ (h⁻¹ : H))ᘁ := by - rw [← SingleObj.inv_as_inv]; rfl -set_option linter.uppercaseLean3 false in -#align Action.right_dual_ρ Action.rightDual_ρ - --- This lemma was always bad, but the linter only noticed after lean4#2644 -@[simp, nolint simpNF] -theorem leftDual_ρ [LeftRigidCategory V] (h : H) : (ᘁX).ρ h = ᘁX.ρ (h⁻¹ : H) := by - rw [← SingleObj.inv_as_inv]; rfl -set_option linter.uppercaseLean3 false in -#align Action.left_dual_ρ Action.leftDual_ρ - -end - -end Monoidal - -/-- Actions/representations of the trivial group are just objects in the ambient category. -/ -def actionPunitEquivalence : Action V (MonCat.of PUnit) β‰Œ V where - functor := forget V _ - inverse := - { obj := fun X => ⟨X, 1⟩ - map := fun f => ⟨f, fun ⟨⟩ => by simp⟩ } - unitIso := - NatIso.ofComponents fun X => mkIso (Iso.refl _) fun ⟨⟩ => by - simp only [MonCat.oneHom_apply, MonCat.one_of, End.one_def, id_eq, Functor.comp_obj, - forget_obj, Iso.refl_hom, Category.comp_id] - exact ρ_one X - counitIso := NatIso.ofComponents fun X => Iso.refl _ -set_option linter.uppercaseLean3 false in -#align Action.Action_punit_equivalence Action.actionPunitEquivalence - -variable (V) - -/-- The "restriction" functor along a monoid homomorphism `f : G ⟢ H`, -taking actions of `H` to actions of `G`. - -(This makes sense for any homomorphism, but the name is natural when `f` is a monomorphism.) --/ -@[simps] -def res {G H : MonCat} (f : G ⟢ H) : Action V H β₯€ Action V G where - obj M := - { V := M.V - ρ := f ≫ M.ρ } - map p := - { hom := p.hom - comm := fun g => p.comm (f g) } -set_option linter.uppercaseLean3 false in -#align Action.res Action.res - -/-- The natural isomorphism from restriction along the identity homomorphism to -the identity functor on `Action V G`. --/ -@[simps!] -def resId {G : MonCat} : res V (πŸ™ G) β‰… 𝟭 (Action V G) := - NatIso.ofComponents fun M => mkIso (Iso.refl _) -set_option linter.uppercaseLean3 false in -#align Action.res_id Action.resId - -/-- The natural isomorphism from the composition of restrictions along homomorphisms -to the restriction along the composition of homomorphism. --/ -@[simps!] -def resComp {G H K : MonCat} (f : G ⟢ H) (g : H ⟢ K) : res V g β‹™ res V f β‰… res V (f ≫ g) := - NatIso.ofComponents fun M => mkIso (Iso.refl _) -set_option linter.uppercaseLean3 false in -#align Action.res_comp Action.resComp - --- TODO promote `res` to a pseudofunctor from --- the locally discrete bicategory constructed from `Monα΅’α΅–` to `Cat`, sending `G` to `Action V G`. -variable {H : MonCat.{u}} (f : G ⟢ H) - -instance res_additive [Preadditive V] : (res V f).Additive where -set_option linter.uppercaseLean3 false in -#align Action.res_additive Action.res_additive - -variable {R : Type*} [Semiring R] - -instance res_linear [Preadditive V] [Linear R V] : (res V f).Linear R where -set_option linter.uppercaseLean3 false in -#align Action.res_linear Action.res_linear - -/-- Bundles a type `H` with a multiplicative action of `G` as an `Action`. -/ -def ofMulAction (G H : Type u) [Monoid G] [MulAction G H] : Action (Type u) (MonCat.of G) where - V := H - ρ := @MulAction.toEndHom _ _ _ (by assumption) -set_option linter.uppercaseLean3 false in -#align Action.of_mul_action Action.ofMulAction - -@[simp] -theorem ofMulAction_apply {G H : Type u} [Monoid G] [MulAction G H] (g : G) (x : H) : - (ofMulAction G H).ρ g x = (g β€’ x : H) := - rfl -set_option linter.uppercaseLean3 false in -#align Action.of_mul_action_apply Action.ofMulAction_apply - -/-- Given a family `F` of types with `G`-actions, this is the limit cone demonstrating that the -product of `F` as types is a product in the category of `G`-sets. -/ -def ofMulActionLimitCone {ΞΉ : Type v} (G : Type max v u) [Monoid G] (F : ΞΉ β†’ Type max v u) - [βˆ€ i : ΞΉ, MulAction G (F i)] : - LimitCone (Discrete.functor fun i : ΞΉ => Action.ofMulAction G (F i)) where - cone := - { pt := Action.ofMulAction G (βˆ€ i : ΞΉ, F i) - Ο€ := Discrete.natTrans (fun i => ⟨fun x => x i.as, fun g => rfl⟩) } - isLimit := - { lift := fun s => - { hom := fun x i => (s.Ο€.app ⟨i⟩).hom x - comm := fun g => by - ext x - funext j - exact congr_fun ((s.Ο€.app ⟨j⟩).comm g) x } - fac := fun s j => rfl - uniq := fun s f h => by - ext x - funext j - dsimp at * - rw [← h ⟨j⟩] - rfl } -set_option linter.uppercaseLean3 false in -#align Action.of_mul_action_limit_cone Action.ofMulActionLimitCone - -/-- The `G`-set `G`, acting on itself by left multiplication. -/ -@[simps!] -def leftRegular (G : Type u) [Monoid G] : Action (Type u) (MonCat.of G) := - Action.ofMulAction G G -set_option linter.uppercaseLean3 false in -#align Action.left_regular Action.leftRegular - -/-- The `G`-set `Gⁿ`, acting on itself by left multiplication. -/ -@[simps!] -def diagonal (G : Type u) [Monoid G] (n : β„•) : Action (Type u) (MonCat.of G) := - Action.ofMulAction G (Fin n β†’ G) -set_option linter.uppercaseLean3 false in -#align Action.diagonal Action.diagonal - -/-- We have `fin 1 β†’ G β‰… G` as `G`-sets, with `G` acting by left multiplication. -/ -def diagonalOneIsoLeftRegular (G : Type u) [Monoid G] : diagonal G 1 β‰… leftRegular G := - Action.mkIso (Equiv.funUnique _ _).toIso fun _ => rfl -set_option linter.uppercaseLean3 false in -#align Action.diagonal_one_iso_left_regular Action.diagonalOneIsoLeftRegular - -open MonoidalCategory - -/-- Given `X : Action (Type u) (Mon.of G)` for `G` a group, then `G Γ— X` (with `G` acting as left -multiplication on the first factor and by `X.ρ` on the second) is isomorphic as a `G`-set to -`G Γ— X` (with `G` acting as left multiplication on the first factor and trivially on the second). -The isomorphism is given by `(g, x) ↦ (g, g⁻¹ β€’ x)`. -/ -@[simps] -noncomputable def leftRegularTensorIso (G : Type u) [Group G] (X : Action (Type u) (MonCat.of G)) : - leftRegular G βŠ— X β‰… leftRegular G βŠ— Action.mk X.V 1 where - hom := - { hom := fun g => ⟨g.1, (X.ρ (g.1⁻¹ : G) g.2 : X.V)⟩ - comm := fun (g : G) => by - funext ⟨(x₁ : G), (xβ‚‚ : X.V)⟩ - refine' Prod.ext rfl _ - change (X.ρ ((g * x₁)⁻¹ : G) * X.ρ g) xβ‚‚ = X.ρ _ _ - rw [mul_inv_rev, ← X.ρ.map_mul, inv_mul_cancel_right] } - inv := - { hom := fun g => ⟨g.1, X.ρ g.1 g.2⟩ - comm := fun (g : G) => by - funext ⟨(x₁ : G), (xβ‚‚ : X.V)⟩ - refine' Prod.ext rfl _ - erw [tensor_rho, tensor_rho] - dsimp - -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - erw [leftRegular_ρ_apply] - erw [map_mul] - rfl } - hom_inv_id := by - apply Hom.ext - funext x - refine' Prod.ext rfl _ - change (X.ρ x.1 * X.ρ (x.1⁻¹ : G)) x.2 = x.2 - rw [← X.ρ.map_mul, mul_inv_self, X.ρ.map_one, MonCat.one_of, End.one_def, types_id_apply] - inv_hom_id := by - apply Hom.ext - funext x - refine' Prod.ext rfl _ - change (X.ρ (x.1⁻¹ : G) * X.ρ x.1) x.2 = x.2 - rw [← X.ρ.map_mul, inv_mul_self, X.ρ.map_one, MonCat.one_of, End.one_def, types_id_apply] -set_option linter.uppercaseLean3 false in -#align Action.left_regular_tensor_iso Action.leftRegularTensorIso - -/-- The natural isomorphism of `G`-sets `Gⁿ⁺¹ β‰… G Γ— Gⁿ`, where `G` acts by left multiplication on -each factor. -/ -@[simps!] -noncomputable def diagonalSucc (G : Type u) [Monoid G] (n : β„•) : - diagonal G (n + 1) β‰… leftRegular G βŠ— diagonal G n := - mkIso (Equiv.piFinSuccAboveEquiv _ 0).toIso fun _ => rfl -set_option linter.uppercaseLean3 false in -#align Action.diagonal_succ Action.diagonalSucc - -end Action - -namespace CategoryTheory.Functor - -variable {V} {W : Type (u + 1)} [LargeCategory W] - -/-- A functor between categories induces a functor between -the categories of `G`-actions within those categories. -/ -@[simps] -def mapAction (F : V β₯€ W) (G : MonCat.{u}) : Action V G β₯€ Action W G where - obj M := - { V := F.obj M.V - ρ := - { toFun := fun g => F.map (M.ρ g) - map_one' := by simp only [End.one_def, Action.ρ_one, F.map_id, MonCat.one_of] - map_mul' := fun g h => by - dsimp - rw [map_mul, MonCat.mul_of, End.mul_def, End.mul_def, F.map_comp] } } - map f := - { hom := F.map f.hom - comm := fun g => by dsimp; rw [← F.map_comp, f.comm, F.map_comp] } - map_id M := by ext; simp only [Action.id_hom, F.map_id] - map_comp f g := by ext; simp only [Action.comp_hom, F.map_comp] -set_option linter.uppercaseLean3 false in -#align category_theory.functor.map_Action CategoryTheory.Functor.mapAction - -variable (F : V β₯€ W) (G : MonCat.{u}) [Preadditive V] [Preadditive W] - -instance mapAction_preadditive [F.Additive] : (F.mapAction G).Additive where -set_option linter.uppercaseLean3 false in -#align category_theory.functor.map_Action_preadditive CategoryTheory.Functor.mapAction_preadditive - -variable {R : Type*} [Semiring R] [CategoryTheory.Linear R V] [CategoryTheory.Linear R W] - -instance mapAction_linear [F.Additive] [F.Linear R] : (F.mapAction G).Linear R where -set_option linter.uppercaseLean3 false in -#align category_theory.functor.map_Action_linear CategoryTheory.Functor.mapAction_linear - -end CategoryTheory.Functor - -namespace CategoryTheory.MonoidalFunctor - -open Action - -variable {V} -variable {W : Type (u + 1)} [LargeCategory W] [MonoidalCategory V] [MonoidalCategory W] - (F : MonoidalFunctor V W) (G : MonCat.{u}) - -open MonoidalCategory - -/-- A monoidal functor induces a monoidal functor between -the categories of `G`-actions within those categories. -/ -@[simps!] -def mapActionAux : LaxMonoidalFunctor (Action V G) (Action W G) := .ofTensorHom - (F := F.toFunctor.mapAction G) - (Ξ΅ := - { hom := F.Ξ΅ - comm := fun g => by - dsimp [FunctorCategoryEquivalence.inverse, Functor.mapAction] - rw [Category.id_comp, F.map_id, Category.comp_id] }) - (ΞΌ := fun X Y => - { hom := F.ΞΌ X.V Y.V - comm := fun g => F.toLaxMonoidalFunctor.ΞΌ_natural (X.ρ g) (Y.ρ g) }) - (ΞΌ_natural := by intros; ext; simp) - (associativity := by intros; ext; simp) - (left_unitality := by intros; ext; simp) - (right_unitality := by intros; ext; simp) - -/-- A monoidal functor induces a monoidal functor between -the categories of `G`-actions within those categories. -/ -@[simps!] -def mapAction : MonoidalFunctor (Action V G) (Action W G) := - { mapActionAux F G with - Ξ΅_isIso := by dsimp [mapActionAux]; infer_instance - ΞΌ_isIso := by dsimp [mapActionAux]; infer_instance } -set_option linter.uppercaseLean3 false in -#align category_theory.monoidal_functor.map_Action CategoryTheory.MonoidalFunctor.mapAction - -@[simp] -theorem mapAction_Ξ΅_inv_hom : (inv (F.mapAction G).Ξ΅).hom = inv F.Ξ΅ := by - rw [← cancel_mono F.Ξ΅, IsIso.inv_hom_id, ← F.mapAction_toLaxMonoidalFunctor_Ξ΅_hom G, - ← Action.comp_hom, IsIso.inv_hom_id, Action.id_hom] -set_option linter.uppercaseLean3 false in -#align category_theory.monoidal_functor.map_Action_Ξ΅_inv_hom CategoryTheory.MonoidalFunctor.mapAction_Ξ΅_inv_hom - -@[simp] -theorem mapAction_ΞΌ_inv_hom (X Y : Action V G) : - (inv ((F.mapAction G).ΞΌ X Y)).hom = inv (F.ΞΌ X.V Y.V) := by - rw [← cancel_mono (F.ΞΌ X.V Y.V), IsIso.inv_hom_id, ← F.mapAction_toLaxMonoidalFunctor_ΞΌ_hom G, - ← Action.comp_hom, IsIso.inv_hom_id, Action.id_hom] -set_option linter.uppercaseLean3 false in -#align category_theory.monoidal_functor.map_Action_ΞΌ_inv_hom CategoryTheory.MonoidalFunctor.mapAction_ΞΌ_inv_hom - -end CategoryTheory.MonoidalFunctor From ba6dc88456b76507af2335a2018453fed4e8dc9a Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sat, 6 Jan 2024 16:13:27 +0900 Subject: [PATCH 45/62] lint style --- Mathlib/CategoryTheory/Enriched/Basic.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Functor.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Enriched/Basic.lean b/Mathlib/CategoryTheory/Enriched/Basic.lean index 619fb2f1d6949..7cca58a258848 100644 --- a/Mathlib/CategoryTheory/Enriched/Basic.lean +++ b/Mathlib/CategoryTheory/Enriched/Basic.lean @@ -127,7 +127,7 @@ instance (F : LaxMonoidalFunctor V W) : EnrichedCategory W (TransportEnrichment convert F.map_id _ simp assoc P Q R S := by - simp only [← id_tensorHom, ←tensorHom_id] + simp only [← id_tensorHom, ← tensorHom_id] rw [comp_tensor_id, Category.assoc, ← F.toFunctor.map_id, F.ΞΌ_natural_assoc, F.toFunctor.map_id, ← F.associativity_inv'_assoc, ← F.toFunctor.map_comp, ← F.toFunctor.map_comp, id_tensorHom, tensorHom_id, e_assoc, id_tensor_comp, diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index 1e010dcde7569..a393b3a48a94c 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -422,7 +422,7 @@ def comp : LaxMonoidalFunctor.{v₁, v₃} C E := simp only [comp_whiskerRight, assoc, ΞΌ_natural_left_assoc, MonoidalCategory.whiskerLeft_comp, ΞΌ_natural_right_assoc] slice_rhs 1 3 => rw [← G.associativity] - simp_rw [Category.assoc, ←G.toFunctor.map_comp, F.associativity] } + simp_rw [Category.assoc, ← G.toFunctor.map_comp, F.associativity] } #align category_theory.lax_monoidal_functor.comp CategoryTheory.LaxMonoidalFunctor.comp @[inherit_doc] From a3847440954030118aad53a96acc7bced8083e2f Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sat, 6 Jan 2024 17:45:54 +0900 Subject: [PATCH 46/62] fix style --- Mathlib/CategoryTheory/Monoidal/Limits.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib/CategoryTheory/Monoidal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Limits.lean index 0a26fb47f4f51..e5ccfe46723c3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Limits.lean @@ -60,6 +60,7 @@ instance limitLaxMonoidalStruct : LaxMonoidalStruct fun F : J β₯€ C => limit F w naturality := fun j j' f => by dsimp simp only [Category.id_comp, ← tensor_comp, limit.w] } } + instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorHom (ΞΌ_natural:= fun f g => by ext; dsimp From a57bd7115aa458399a0ac768f3624d2aa10bb2eb Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sat, 6 Jan 2024 18:18:32 +0900 Subject: [PATCH 47/62] revert adding `LaxMonoidalStruct` --- .../Category/ModuleCat/Adjunctions.lean | 17 +++++++-------- .../CategoryTheory/Monoidal/Functorial.lean | 21 ++++++++----------- Mathlib/CategoryTheory/Monoidal/Limits.lean | 12 +++++------ 3 files changed, 21 insertions(+), 29 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index 1f405392a937f..e90fe25c21c54 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -183,23 +183,20 @@ theorem associativity (X Y Z : Type u) : CategoryTheory.associator_hom_apply]; rfl #align Module.free.associativity ModuleCat.Free.associativity -/-- The free R-module functor is lax monoidal. The structure part. -/ -@[simps] -instance : LaxMonoidalStruct.{u} (free R).obj where - -- Send `R` to `PUnit β†’β‚€ R` - Ξ΅ := Ξ΅ R - -- Send `(Ξ± β†’β‚€ R) βŠ— (Ξ² β†’β‚€ R)` to `Ξ± Γ— Ξ² β†’β‚€ R` - ΞΌ X Y := (ΞΌ R X Y).hom - -- In fact, it's strong monoidal, but we don't yet have a typeclass for that. /-- The free R-module functor is lax monoidal. The property part. -/ +@[simps] instance : LaxMonoidal.{u} (free R).obj := .ofTensorHom + -- Send `R` to `PUnit β†’β‚€ R` + (Ξ΅ := Ξ΅ R) + -- Send `(Ξ± β†’β‚€ R) βŠ— (Ξ² β†’β‚€ R)` to `Ξ± Γ— Ξ² β†’β‚€ R` + (ΞΌ := fun X Y => (ΞΌ R X Y).hom) (ΞΌ_natural := fun {_} {_} {_} {_} f g ↦ ΞΌ_natural R f g) (left_unitality := left_unitality R) (right_unitality := right_unitality R) (associativity := associativity R) -instance : IsIso (@LaxMonoidalStruct.Ξ΅ _ _ _ _ _ _ (free R).obj _ _) := by +instance : IsIso (@LaxMonoidal.Ξ΅ _ _ _ _ _ _ (free R).obj _ _) := by refine' ⟨⟨Finsupp.lapply PUnit.unit, ⟨_, _⟩⟩⟩ Β· -- Porting note: broken ext apply LinearMap.ext_ring @@ -228,7 +225,7 @@ variable [CommRing R] def monoidalFree : MonoidalFunctor (Type u) (ModuleCat.{u} R) := { LaxMonoidalFunctor.of (free R).obj with -- Porting note: used to be dsimp - Ξ΅_isIso := inferInstanceAs <| IsIso LaxMonoidalStruct.Ξ΅ + Ξ΅_isIso := inferInstanceAs <| IsIso LaxMonoidal.Ξ΅ ΞΌ_isIso := fun X Y => by dsimp; infer_instance } #align Module.monoidal_free ModuleCat.monoidalFree diff --git a/Mathlib/CategoryTheory/Monoidal/Functorial.lean b/Mathlib/CategoryTheory/Monoidal/Functorial.lean index c13099bb72d50..d22a84eb61cf2 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functorial.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functorial.lean @@ -47,18 +47,14 @@ open MonoidalCategory variable {C : Type u₁} [Category.{v₁} C] [MonoidalCategory.{v₁} C] {D : Type uβ‚‚} [Category.{vβ‚‚} D] [MonoidalCategory.{vβ‚‚} D] -/-- An unbundled description of lax monoidal functors without axioms. See `LaxMonoidal` for -the full description. -/ -class LaxMonoidalStruct (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] where +-- Perhaps in the future we'll redefine `LaxMonoidalFunctor` in terms of this, +-- but that isn't the immediate plan. +/-- An unbundled description of lax monoidal functors. -/ +class LaxMonoidal (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] where /-- unit morphism -/ Ξ΅ : πŸ™_ D ⟢ F (πŸ™_ C) /-- tensorator -/ ΞΌ : βˆ€ X Y : C, F X βŠ— F Y ⟢ F (X βŠ— Y) - --- Perhaps in the future we'll redefine `LaxMonoidalFunctor` in terms of this, --- but that isn't the immediate plan. -/-- An unbundled description of lax monoidal functors. -/ -class LaxMonoidal (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] extends LaxMonoidalStruct F where ΞΌ_natural_left : βˆ€ {X Y : C} (f : X ⟢ Y) (X' : C), (map F f β–· F X') ≫ ΞΌ Y X' = ΞΌ X X' ≫ map F (f β–· X') := by @@ -78,11 +74,12 @@ class LaxMonoidal (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] extends LaxMonoidalS right_unitality : βˆ€ X : C, (ρ_ (F X)).hom = (F X ◁ Ξ΅) ≫ ΞΌ X (πŸ™_ C) ≫ map F (ρ_ X).hom := by aesop_cat --- Perhaps in the future we'll redefine `LaxMonoidalFunctor` in terms of this, --- but that isn't the immediate plan. -open LaxMonoidalStruct in /-- An unbundled description of lax monoidal functors. -/ -abbrev LaxMonoidal.ofTensorHom (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] [LaxMonoidalStruct F] +abbrev LaxMonoidal.ofTensorHom (F : C β†’ D) [Functorial.{v₁, vβ‚‚} F] + /- unit morphism -/ + (Ξ΅ : πŸ™_ D ⟢ F (πŸ™_ C)) + /- tensorator -/ + (ΞΌ : βˆ€ X Y : C, F X βŠ— F Y ⟢ F (X βŠ— Y)) /- naturality -/ (ΞΌ_natural : βˆ€ {X Y X' Y' : C} (f : X ⟢ Y) (g : X' ⟢ Y'), diff --git a/Mathlib/CategoryTheory/Monoidal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Limits.lean index e5ccfe46723c3..680531b275bae 100644 --- a/Mathlib/CategoryTheory/Monoidal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Limits.lean @@ -47,21 +47,19 @@ theorem limitFunctorial_map {F G : J β₯€ C} (Ξ± : F ⟢ G) : variable [MonoidalCategory.{v} C] @[simps] -instance limitLaxMonoidalStruct : LaxMonoidalStruct fun F : J β₯€ C => limit F where - Ξ΅ := +instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorHom + (Ξ΅ := limit.lift _ { pt := _ - Ο€ := { app := fun j => πŸ™ _ } } - ΞΌ F G := + Ο€ := { app := fun j => πŸ™ _ } }) + (ΞΌ := fun F G => limit.lift (F βŠ— G) { pt := limit F βŠ— limit G Ο€ := { app := fun j => limit.Ο€ F j βŠ— limit.Ο€ G j naturality := fun j j' f => by dsimp - simp only [Category.id_comp, ← tensor_comp, limit.w] } } - -instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorHom + simp only [Category.id_comp, ← tensor_comp, limit.w] } }) (ΞΌ_natural:= fun f g => by ext; dsimp simp only [limit.lift_Ο€, Cones.postcompose_obj_Ο€, Monoidal.tensorHom_app, limit.lift_map, From 05038949eda842bff41b1187b2b827ca293c921e Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Wed, 24 Jan 2024 22:01:03 +0900 Subject: [PATCH 48/62] minor fix for the last commit --- Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index e90fe25c21c54..28ab863e9852f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -184,7 +184,7 @@ theorem associativity (X Y Z : Type u) : #align Module.free.associativity ModuleCat.Free.associativity -- In fact, it's strong monoidal, but we don't yet have a typeclass for that. -/-- The free R-module functor is lax monoidal. The property part. -/ +/-- The free R-module functor is lax monoidal. -/ @[simps] instance : LaxMonoidal.{u} (free R).obj := .ofTensorHom -- Send `R` to `PUnit β†’β‚€ R` From f253f83d0a0de258c465b47f93fdd9fcefcfef8f Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 25 Jan 2024 17:26:34 +0900 Subject: [PATCH 49/62] delete empty lines --- Mathlib/CategoryTheory/Monoidal/Functor.lean | 1 - Mathlib/CategoryTheory/Monoidal/Mon_.lean | 1 - 2 files changed, 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index a393b3a48a94c..806088f90207d 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -363,7 +363,6 @@ theorem Ξ΅_hom_inv_id : F.Ξ΅ ≫ F.Ξ΅Iso.inv = πŸ™ _ := noncomputable def commTensorLeft (X : C) : F.toFunctor β‹™ tensorLeft (F.toFunctor.obj X) β‰… tensorLeft X β‹™ F.toFunctor := NatIso.ofComponents (fun Y => F.ΞΌIso X Y) fun f => F.ΞΌ_natural_right X f - #align category_theory.monoidal_functor.comm_tensor_left CategoryTheory.MonoidalFunctor.commTensorLeft /-- Monoidal functors commute with right tensoring up to isomorphism -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index f559eb48a0212..44063f2f33887 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -49,7 +49,6 @@ attribute [reassoc] Mon_.one_mul Mon_.mul_one attribute [simp] Mon_.one_mul Mon_.mul_one - -- We prove a more general `@[simp]` lemma below. attribute [reassoc (attr := simp)] Mon_.mul_assoc From ac50c91b2f9cab6c8f73a096075673594b95c999 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Thu, 25 Jan 2024 22:09:18 +0900 Subject: [PATCH 50/62] add lemmas --- .../Monoidal/OfHasFiniteProducts.lean | 48 +++++++++---------- 1 file changed, 24 insertions(+), 24 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean index 7f35ba5fb3a2a..b5649e03597dd 100644 --- a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean +++ b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean @@ -82,6 +82,14 @@ theorem tensorHom {W X Y Z : C} (f : W ⟢ X) (g : Y ⟢ Z) : f βŠ— g = Limits.p rfl #align category_theory.monoidal_of_has_finite_products.tensor_hom CategoryTheory.monoidalOfHasFiniteProducts.tensorHom +@[simp] +theorem whiskerLeft (X : C) {Y Z : C} (f : Y ⟢ Z) : X ◁ f = Limits.prod.map (πŸ™ X) f := + rfl + +@[simp] +theorem whiskerRight {X Y : C} (f : X ⟢ Y) (Z : C) : f β–· Z = Limits.prod.map f (πŸ™ Z) := + rfl + @[simp] theorem leftUnitor_hom (X : C) : (Ξ»_ X).hom = Limits.prod.snd := rfl @@ -130,18 +138,10 @@ open MonoidalCategory @[simps] def symmetricOfHasFiniteProducts [HasTerminal C] [HasBinaryProducts C] : SymmetricCategory C where braiding X Y := Limits.prod.braiding X Y - braiding_naturality_left f X := by dsimp; simp [← id_tensorHom, ← tensorHom_id] - braiding_naturality_right X _ _ f := by dsimp; simp [← id_tensorHom, ← tensorHom_id] - hexagon_forward X Y Z := by - dsimp [monoidalOfHasFiniteProducts.associator_hom] - simp [← id_tensorHom, ← tensorHom_id] - hexagon_reverse X Y Z := by - dsimp [monoidalOfHasFiniteProducts.associator_inv] - simp [← id_tensorHom, ← tensorHom_id] - -- braiding_naturality_left f X := by simp - -- braiding_naturality_right X _ _ f := by simp - -- hexagon_forward X Y Z := by dsimp [monoidalOfHasFiniteProducts.associator_hom]; simp - -- hexagon_reverse X Y Z := by dsimp [monoidalOfHasFiniteProducts.associator_inv]; simp + braiding_naturality_left f X := by simp + braiding_naturality_right X _ _ f := by simp + hexagon_forward X Y Z := by dsimp [monoidalOfHasFiniteProducts.associator_hom]; simp + hexagon_reverse X Y Z := by dsimp [monoidalOfHasFiniteProducts.associator_inv]; simp symmetry X Y := by dsimp; simp #align category_theory.symmetric_of_has_finite_products CategoryTheory.symmetricOfHasFiniteProducts @@ -187,6 +187,14 @@ theorem tensorHom {W X Y Z : C} (f : W ⟢ X) (g : Y ⟢ Z) : f βŠ— g = Limits.c rfl #align category_theory.monoidal_of_has_finite_coproducts.tensor_hom CategoryTheory.monoidalOfHasFiniteCoproducts.tensorHom +@[simp] +theorem whiskerLeft (X : C) {Y Z : C} (f : Y ⟢ Z) : X ◁ f = Limits.coprod.map (πŸ™ X) f := + rfl + +@[simp] +theorem whiskerRight {X Y : C} (f : X ⟢ Y) (Z : C) : f β–· Z = Limits.coprod.map f (πŸ™ Z) := + rfl + @[simp] theorem leftUnitor_hom (X : C) : (Ξ»_ X).hom = coprod.desc (initial.to X) (πŸ™ _) := rfl @@ -235,18 +243,10 @@ open MonoidalCategory def symmetricOfHasFiniteCoproducts [HasInitial C] [HasBinaryCoproducts C] : SymmetricCategory C where braiding := Limits.coprod.braiding - braiding_naturality_left f g := by dsimp [tensorHom]; simp [← id_tensorHom, ← tensorHom_id] - braiding_naturality_right f g := by dsimp [tensorHom]; simp [← id_tensorHom, ← tensorHom_id] - hexagon_forward X Y Z := by - dsimp [monoidalOfHasFiniteCoproducts.associator_hom] - simp [← id_tensorHom, ← tensorHom_id] - hexagon_reverse X Y Z := by - dsimp [monoidalOfHasFiniteCoproducts.associator_inv] - simp [← id_tensorHom, ← tensorHom_id] - -- braiding_naturality_left f g := by simp - -- braiding_naturality_right f g := by simp - -- hexagon_forward X Y Z := by dsimp [monoidalOfHasFiniteCoproducts.associator_hom]; simp - -- hexagon_reverse X Y Z := by dsimp [monoidalOfHasFiniteCoproducts.associator_inv]; simp + braiding_naturality_left f g := by simp + braiding_naturality_right f g := by simp + hexagon_forward X Y Z := by dsimp [monoidalOfHasFiniteCoproducts.associator_hom]; simp + hexagon_reverse X Y Z := by dsimp [monoidalOfHasFiniteCoproducts.associator_inv]; simp symmetry X Y := by dsimp; simp #align category_theory.symmetric_of_has_finite_coproducts CategoryTheory.symmetricOfHasFiniteCoproducts From 7e0104f23e53065e9714385515950847c3718d2b Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 28 Jan 2024 00:49:12 +0900 Subject: [PATCH 51/62] merge --- Mathlib/CategoryTheory/Monoidal/Category.lean | 16 ++++++++-------- Mathlib/CategoryTheory/Monoidal/Center.lean | 4 ++-- Mathlib/CategoryTheory/Monoidal/Linear.lean | 4 ++-- Mathlib/CategoryTheory/Monoidal/Transport.lean | 2 +- Mathlib/Tactic/CategoryTheory/Coherence.lean | 5 ----- 5 files changed, 13 insertions(+), 18 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 1f216a4de297f..12d3e302dcf28 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -268,42 +268,42 @@ namespace MonoidalCategory @[reassoc (attr := simp)] theorem whiskerLeft_hom_inv (X : C) {Y Z : C} (f : Y β‰… Z) : X ◁ f.hom ≫ X ◁ f.inv = πŸ™ (X βŠ— Y) := by - simp [← id_tensorHom, ← tensor_comp] + rw [← whiskerLeft_comp, hom_inv_id, whiskerLeft_id] @[reassoc (attr := simp)] theorem hom_inv_whiskerRight {X Y : C} (f : X β‰… Y) (Z : C) : f.hom β–· Z ≫ f.inv β–· Z = πŸ™ (X βŠ— Z) := by - simp [← tensorHom_id, ← tensor_comp] + rw [← comp_whiskerRight, hom_inv_id, id_whiskerRight] @[reassoc (attr := simp)] theorem whiskerLeft_inv_hom (X : C) {Y Z : C} (f : Y β‰… Z) : X ◁ f.inv ≫ X ◁ f.hom = πŸ™ (X βŠ— Z) := by - simp [← id_tensorHom, ← tensor_comp] + rw [← whiskerLeft_comp, inv_hom_id, whiskerLeft_id] @[reassoc (attr := simp)] theorem inv_hom_whiskerRight {X Y : C} (f : X β‰… Y) (Z : C) : f.inv β–· Z ≫ f.hom β–· Z = πŸ™ (Y βŠ— Z) := by - simp [← tensorHom_id, ← tensor_comp] + rw [← comp_whiskerRight, inv_hom_id, id_whiskerRight] @[reassoc (attr := simp)] theorem whiskerLeft_hom_inv' (X : C) {Y Z : C} (f : Y ⟢ Z) [IsIso f] : X ◁ f ≫ X ◁ inv f = πŸ™ (X βŠ— Y) := by - simp [← id_tensorHom, ← tensor_comp] + rw [← whiskerLeft_comp, IsIso.hom_inv_id, whiskerLeft_id] @[reassoc (attr := simp)] theorem hom_inv_whiskerRight' {X Y : C} (f : X ⟢ Y) [IsIso f] (Z : C) : f β–· Z ≫ inv f β–· Z = πŸ™ (X βŠ— Z) := by - simp [← tensorHom_id, ← tensor_comp] + rw [← comp_whiskerRight, IsIso.hom_inv_id, id_whiskerRight] @[reassoc (attr := simp)] theorem whiskerLeft_inv_hom' (X : C) {Y Z : C} (f : Y ⟢ Z) [IsIso f] : X ◁ inv f ≫ X ◁ f = πŸ™ (X βŠ— Z) := by - simp [← id_tensorHom, ← tensor_comp] + rw [← whiskerLeft_comp, IsIso.inv_hom_id, whiskerLeft_id] @[reassoc (attr := simp)] theorem inv_hom_whiskerRight' {X Y : C} (f : X ⟢ Y) [IsIso f] (Z : C) : inv f β–· Z ≫ f β–· Z = πŸ™ (Y βŠ— Z) := by - simp [← tensorHom_id, ← tensor_comp] + rw [← comp_whiskerRight, IsIso.inv_hom_id, id_whiskerRight] /-- The left whiskering of an isomorphism is an isomorphism. -/ @[simps] diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index d0064934a79b7..29a64e9dd4fdc 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -273,11 +273,11 @@ theorem tensor_Ξ² (X Y : Center C) (U : C) : @[simp] theorem whiskerLeft_f (X : Center C) {Y₁ Yβ‚‚ : Center C} (f : Y₁ ⟢ Yβ‚‚) : (X ◁ f).f = X.1 ◁ f.f := - id_tensorHom X.1 f.f + rfl @[simp] theorem whiskerRight_f {X₁ Xβ‚‚ : Center C} (f : X₁ ⟢ Xβ‚‚) (Y : Center C) : (f β–· Y).f = f.f β–· Y.1 := - tensorHom_id f.f Y.1 + rfl @[simp] theorem tensor_f {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Center C} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ Yβ‚‚) : (f βŠ— g).f = f.f βŠ— g.f := diff --git a/Mathlib/CategoryTheory/Monoidal/Linear.lean b/Mathlib/CategoryTheory/Monoidal/Linear.lean index 911294f6964a4..c8a2e24043adb 100644 --- a/Mathlib/CategoryTheory/Monoidal/Linear.lean +++ b/Mathlib/CategoryTheory/Monoidal/Linear.lean @@ -32,9 +32,9 @@ variable [MonoidalCategory C] /-- A category is `MonoidalLinear R` if tensoring is `R`-linear in both factors. -/ class MonoidalLinear [MonoidalPreadditive C] : Prop where - whiskerLeft_smul : βˆ€ (X : C) {Y Z : C} (r : R) (f : Y ⟢ Z) , πŸ™ X βŠ— (r β€’ f) = r β€’ (πŸ™ X βŠ— f) := by + whiskerLeft_smul : βˆ€ (X : C) {Y Z : C} (r : R) (f : Y ⟢ Z) , X ◁ (r β€’ f) = r β€’ (X ◁ f) := by aesop_cat - smul_whiskerRight : βˆ€ (r : R) {Y Z : C} (f : Y ⟢ Z) (X : C), (r β€’ f) βŠ— πŸ™ X = r β€’ (f βŠ— πŸ™ X) := by + smul_whiskerRight : βˆ€ (r : R) {Y Z : C} (f : Y ⟢ Z) (X : C), (r β€’ f) β–· X = r β€’ (f β–· X) := by aesop_cat #align category_theory.monoidal_linear CategoryTheory.MonoidalLinear diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index 054b199438dc4..fcf29b82c3b02 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -127,7 +127,7 @@ abbrev induced [MonoidalCategoryStruct D] (F : D β₯€ C) [Faithful F] simp only [Functor.map_comp, fData.whiskerRight_eq, fData.associator_eq, Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom, tensorIso_hom, Iso.refl_hom, tensorHom_id, id_tensorHom, comp_whiskerRight, whisker_assoc, assoc, fData.whiskerLeft_eq, - MonoidalCategory.whiskerLeft_comp, Iso.hom_inv_id_assoc, hom_inv_whiskerLeft_assoc, + MonoidalCategory.whiskerLeft_comp, Iso.hom_inv_id_assoc, whiskerLeft_hom_inv_assoc, hom_inv_whiskerRight_assoc, Iso.inv_hom_id_assoc, Iso.cancel_iso_inv_left] slice_lhs 5 6 => rw [← MonoidalCategory.whiskerLeft_comp, hom_inv_whiskerRight] diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index d9620f18640cf..0abdfb8d9d88a 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -126,11 +126,6 @@ instance whiskerRight (X Y Z : C) [LiftObj X] [LiftObj Y] [LiftObj Z] [MonoidalC MonoidalCoherence (X βŠ— Z) (Y βŠ— Z) := ⟨MonoidalCoherence.hom β–· Z⟩ -@[simps] -instance whiskerRight (X Y Z : C) [LiftObj X] [LiftObj Y] [LiftObj Z] [MonoidalCoherence X Y] : - MonoidalCoherence (X βŠ— Z) (Y βŠ— Z) := - ⟨MonoidalCoherence.hom βŠ— πŸ™ Z⟩ - @[simps] instance tensor_right (X Y : C) [LiftObj X] [LiftObj Y] [MonoidalCoherence (πŸ™_ C) Y] : MonoidalCoherence X (X βŠ— Y) := From 75b4a4199489e778ae2ef685f1d8e78085d2b34c Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sat, 3 Feb 2024 21:05:20 +0900 Subject: [PATCH 52/62] merge --- .../CategoryTheory/Bicategory/SingleObj.lean | 2 -- Mathlib/CategoryTheory/Monoidal/Bimod.lean | 4 +--- Mathlib/CategoryTheory/Monoidal/Braided.lean | 18 ++++++++++-------- Mathlib/CategoryTheory/Monoidal/Category.lean | 5 ++--- Mathlib/CategoryTheory/Monoidal/Center.lean | 2 -- .../Monoidal/CoherenceLemmas.lean | 1 + Mathlib/CategoryTheory/Monoidal/CommMon_.lean | 2 -- Mathlib/CategoryTheory/Monoidal/End.lean | 2 -- .../Monoidal/FunctorCategory.lean | 2 -- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 4 ---- Mathlib/CategoryTheory/Monoidal/Opposite.lean | 2 -- .../Monoidal/Types/Coyoneda.lean | 2 -- Mathlib/Tactic/CategoryTheory/Coherence.lean | 10 +++++----- 13 files changed, 19 insertions(+), 37 deletions(-) diff --git a/Mathlib/CategoryTheory/Bicategory/SingleObj.lean b/Mathlib/CategoryTheory/Bicategory/SingleObj.lean index d19bfd4843cb4..d039080525e0c 100644 --- a/Mathlib/CategoryTheory/Bicategory/SingleObj.lean +++ b/Mathlib/CategoryTheory/Bicategory/SingleObj.lean @@ -68,8 +68,6 @@ protected def star : MonoidalSingleObj C := PUnit.unit #align category_theory.monoidal_single_obj.star CategoryTheory.MonoidalSingleObj.star -attribute [local simp] id_tensorHom tensorHom_id in - /-- The monoidal functor from the endomorphisms of the single object when we promote a monoidal category to a single object bicategory, to the original monoidal category. diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index 4db291ed64e83..79c21c9b55de9 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -180,8 +180,6 @@ set_option linter.uppercaseLean3 false in variable (A) -attribute [local simp] id_tensorHom tensorHom_id - /-- A monoid object as a bimodule over itself. -/ @[simps] def regular : Bimod A A where @@ -868,7 +866,7 @@ theorem id_whiskerLeft_bimod {X Y : Mon_ C} {M N : Bimod X Y} (f : M ⟢ N) : slice_rhs 4 5 => rw [← comp_whiskerRight, Mon_.one_mul] have : (Ξ»_ (X.X βŠ— N.X)).inv ≫ (Ξ±_ (πŸ™_ C) X.X N.X).inv ≫ ((Ξ»_ X.X).hom β–· N.X) = πŸ™ _ := by pure_coherence - slice_rhs 2 4 => rw [← tensorHom_id, this] + slice_rhs 2 4 => rw [this] slice_rhs 1 2 => rw [Category.comp_id] set_option linter.uppercaseLean3 false in #align Bimod.id_whisker_left_Bimod Bimod.id_whiskerLeft_bimod diff --git a/Mathlib/CategoryTheory/Monoidal/Braided.lean b/Mathlib/CategoryTheory/Monoidal/Braided.lean index c5cc16d1385a9..b97b5d4c7e054 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided.lean @@ -138,14 +138,14 @@ def braidedCategoryOfFaithful {C D : Type*} [Category C] [Category D] [MonoidalC intros apply F.map_injective refine (cancel_epi (F.ΞΌ ?_ ?_)).1 ?_ - rw [Functor.map_comp, ← LaxMonoidalFunctor.ΞΌ_natural_left'_assoc, w, Functor.map_comp, - reassoc_of% w, braiding_naturality_left_assoc, LaxMonoidalFunctor.ΞΌ_natural_right'] + rw [Functor.map_comp, ← LaxMonoidalFunctor.ΞΌ_natural_left_assoc, w, Functor.map_comp, + reassoc_of% w, braiding_naturality_left_assoc, LaxMonoidalFunctor.ΞΌ_natural_right] braiding_naturality_right := by intros apply F.map_injective refine (cancel_epi (F.ΞΌ ?_ ?_)).1 ?_ - rw [Functor.map_comp, ← LaxMonoidalFunctor.ΞΌ_natural_right'_assoc, w, Functor.map_comp, - reassoc_of% w, braiding_naturality_right_assoc, LaxMonoidalFunctor.ΞΌ_natural_left'] + rw [Functor.map_comp, ← LaxMonoidalFunctor.ΞΌ_natural_right_assoc, w, Functor.map_comp, + reassoc_of% w, braiding_naturality_right_assoc, LaxMonoidalFunctor.ΞΌ_natural_left] hexagon_forward := by intros apply F.map_injective @@ -225,6 +225,7 @@ theorem braiding_leftUnitor_auxβ‚‚ (X : C) : _ = (ρ_ X).hom β–· πŸ™_ C := by rw [triangle] #align category_theory.braiding_left_unitor_auxβ‚‚ CategoryTheory.braiding_leftUnitor_auxβ‚‚ +@[reassoc] theorem braiding_leftUnitor (X : C) : (Ξ²_ X (πŸ™_ C)).hom ≫ (Ξ»_ X).hom = (ρ_ X).hom := by rw [← whiskerRight_iff, comp_whiskerRight, braiding_leftUnitor_auxβ‚‚] #align category_theory.braiding_left_unitor CategoryTheory.braiding_leftUnitor @@ -256,24 +257,27 @@ theorem braiding_rightUnitor_auxβ‚‚ (X : C) : _ = πŸ™_ C ◁ (Ξ»_ X).hom := by rw [triangle_assoc_comp_right] #align category_theory.braiding_right_unitor_auxβ‚‚ CategoryTheory.braiding_rightUnitor_auxβ‚‚ +@[reassoc] theorem braiding_rightUnitor (X : C) : (Ξ²_ (πŸ™_ C) X).hom ≫ (ρ_ X).hom = (Ξ»_ X).hom := by rw [← whiskerLeft_iff, MonoidalCategory.whiskerLeft_comp, braiding_rightUnitor_auxβ‚‚] #align category_theory.braiding_right_unitor CategoryTheory.braiding_rightUnitor -@[simp] +@[reassoc, simp] theorem braiding_tensorUnit_left (X : C) : (Ξ²_ (πŸ™_ C) X).hom = (Ξ»_ X).hom ≫ (ρ_ X).inv := by simp [← braiding_rightUnitor] +@[reassoc] theorem leftUnitor_inv_braiding (X : C) : (Ξ»_ X).inv ≫ (Ξ²_ (πŸ™_ C) X).hom = (ρ_ X).inv := by simp #align category_theory.left_unitor_inv_braiding CategoryTheory.leftUnitor_inv_braiding +@[reassoc] theorem rightUnitor_inv_braiding (X : C) : (ρ_ X).inv ≫ (Ξ²_ X (πŸ™_ C)).hom = (Ξ»_ X).inv := by apply (cancel_mono (Ξ»_ X).hom).1 simp only [assoc, braiding_leftUnitor, Iso.inv_hom_id] #align category_theory.right_unitor_inv_braiding CategoryTheory.rightUnitor_inv_braiding -@[simp] +@[reassoc, simp] theorem braiding_tensorUnit_right (X : C) : (Ξ²_ X (πŸ™_ C)).hom = (ρ_ X).hom ≫ (Ξ»_ X).inv := by simp [← rightUnitor_inv_braiding] @@ -474,8 +478,6 @@ theorem tensor_ΞΌ_natural {X₁ Xβ‚‚ Y₁ Yβ‚‚ U₁ Uβ‚‚ V₁ Vβ‚‚ : C} (f₁ : simp only [assoc] #align category_theory.tensor_ΞΌ_natural CategoryTheory.tensor_ΞΌ_natural -attribute [local simp] id_tensorHom tensorHom_id - @[reassoc] theorem tensor_ΞΌ_natural_left {X₁ Xβ‚‚ Y₁ Yβ‚‚ : C} (f₁: X₁ ⟢ Y₁) (fβ‚‚ : Xβ‚‚ ⟢ Yβ‚‚) (Z₁ Zβ‚‚ : C) : (f₁ βŠ— fβ‚‚) β–· (Z₁ βŠ— Zβ‚‚) ≫ tensor_ΞΌ C (Y₁, Yβ‚‚) (Z₁, Zβ‚‚) = diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 260b64c40e32d..7321d933f810f 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -443,6 +443,7 @@ theorem leftUnitor_naturality {X Y : C} (f : X ⟢ Y) : theorem leftUnitor_inv_naturality {X Y : C} (f : X ⟢ Y) : f ≫ (Ξ»_ Y).inv = (Ξ»_ X).inv ≫ (_ ◁ f) := by simp +@[reassoc] theorem id_whiskerLeft_symm {X X' : C} (f : X ⟢ X') : f = (Ξ»_ X).inv ≫ πŸ™_ C ◁ f ≫ (Ξ»_ X').hom := by simp @@ -456,7 +457,7 @@ theorem rightUnitor_naturality {X Y : C} (f : X ⟢ Y) : theorem rightUnitor_inv_naturality {X X' : C} (f : X ⟢ X') : f ≫ (ρ_ X').inv = (ρ_ X).inv ≫ (f β–· _) := by simp - +@[reassoc] theorem whiskerRight_id_symm {X Y : C} (f : X ⟢ Y) : f = (ρ_ X).inv ≫ f β–· πŸ™_ C ≫ (ρ_ Y).hom := by simp @@ -951,8 +952,6 @@ def rightUnitorNatIso : tensorUnitRight C β‰… 𝟭 C := section -attribute [local simp] id_tensorHom tensorHom_id whisker_exchange - -- Porting Note: This used to be `variable {C}` but it seems like Lean 4 parses that differently variable {C : Type u} [Category.{v} C] [MonoidalCategory.{v} C] diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index ed2a25e471ab5..c013762e47622 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -217,8 +217,6 @@ def tensorHom {X₁ Y₁ Xβ‚‚ Yβ‚‚ : Center C} (f : X₁ ⟢ Y₁) (g : Xβ‚‚ ⟢ section -attribute [local simp] id_tensorHom tensorHom_id - /-- Auxiliary definition for the `MonoidalCategory` instance on `Center C`. -/ @[simps] def tensorUnit : Center C := diff --git a/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean b/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean index f5cbf0db71747..447e1f050ffc8 100644 --- a/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean +++ b/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean @@ -60,6 +60,7 @@ theorem pentagon_inv_inv_hom (W X Y Z : C) : coherence #align category_theory.monoidal_category.pentagon_inv_inv_hom CategoryTheory.MonoidalCategory.pentagon_inv_inv_hom +@[reassoc] theorem triangle_assoc_comp_right_inv' (X Y : C) : ((ρ_ X).inv βŠ— πŸ™ Y) ≫ (Ξ±_ X (πŸ™_ C) Y).hom = πŸ™ X βŠ— (Ξ»_ Y).inv := by coherence diff --git a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean index 4abc94fdb8686..7e21d4775deb3 100644 --- a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean @@ -171,8 +171,6 @@ def laxBraidedToCommMon : LaxBraidedFunctor (Discrete PUnit.{u + 1}) C β₯€ CommM set_option linter.uppercaseLean3 false in #align CommMon_.equiv_lax_braided_functor_punit.lax_braided_to_CommMon CommMon_.EquivLaxBraidedFunctorPUnit.laxBraidedToCommMon -attribute [local simp] id_tensorHom tensorHom_id - /-- Implementation of `CommMon_.equivLaxBraidedFunctorPunit`. -/ @[simps] def commMonToLaxBraided : CommMon_ C β₯€ LaxBraidedFunctor (Discrete PUnit.{u + 1}) C where diff --git a/Mathlib/CategoryTheory/Monoidal/End.lean b/Mathlib/CategoryTheory/Monoidal/End.lean index ec073400881c0..c66a95eae46b1 100644 --- a/Mathlib/CategoryTheory/Monoidal/End.lean +++ b/Mathlib/CategoryTheory/Monoidal/End.lean @@ -90,8 +90,6 @@ attribute [local instance] endofunctorMonoidalCategory -- porting note: used `dsimp [endofunctorMonoidalCategory]` when necessary instead -- attribute [local reducible] endofunctorMonoidalCategory -attribute [local simp] id_tensorHom tensorHom_id in - /-- Tensoring on the right gives a monoidal functor from `C` into endofunctors of `C`. -/ @[simps!] diff --git a/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean b/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean index b6edc3f851e4f..429b5701b8a6d 100644 --- a/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/FunctorCategory.lean @@ -165,8 +165,6 @@ theorem associator_inv_app {F G H : C β₯€ D} {X} : rfl #align category_theory.monoidal.associator_inv_app CategoryTheory.Monoidal.associator_inv_app -attribute [local simp] id_tensorHom tensorHom_id in - /-- When `C` is any category, and `D` is a monoidal category, the functor category `C β₯€ D` has a natural pointwise monoidal structure, where `(F βŠ— G).obj X = F.obj X βŠ— G.obj X`. diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index 493e2aa048140..44063f2f33887 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -258,8 +258,6 @@ def laxMonoidalToMon : LaxMonoidalFunctor (Discrete PUnit.{u + 1}) C β₯€ Mon_ C map Ξ± := ((mapMonFunctor (Discrete PUnit) C).map Ξ±).app _ #align Mon_.equiv_lax_monoidal_functor_punit.lax_monoidal_to_Mon Mon_.EquivLaxMonoidalFunctorPUnit.laxMonoidalToMon -attribute [local simp] id_tensorHom tensorHom_id - /-- Implementation of `Mon_.equivLaxMonoidalFunctorPUnit`. -/ @[simps] def monToLaxMonoidal : Mon_ C β₯€ LaxMonoidalFunctor (Discrete PUnit.{u + 1}) C where @@ -383,8 +381,6 @@ theorem one_associator {M N P : Mon_ C} : simp #align Mon_.one_associator Mon_.one_associator -attribute [local simp] id_tensorHom tensorHom_id - theorem one_leftUnitor {M : Mon_ C} : ((Ξ»_ (πŸ™_ C)).inv ≫ (πŸ™ (πŸ™_ C) βŠ— M.one)) ≫ (Ξ»_ M.X).hom = M.one := by simp diff --git a/Mathlib/CategoryTheory/Monoidal/Opposite.lean b/Mathlib/CategoryTheory/Monoidal/Opposite.lean index ecf408b8aa7ab..cc3fee59d12ae 100644 --- a/Mathlib/CategoryTheory/Monoidal/Opposite.lean +++ b/Mathlib/CategoryTheory/Monoidal/Opposite.lean @@ -170,8 +170,6 @@ variable [MonoidalCategory.{v₁} C] open Opposite MonoidalCategory -attribute [local simp] id_tensorHom tensorHom_id - instance monoidalCategoryOp : MonoidalCategory Cα΅’α΅– where tensorObj X Y := op (unop X βŠ— unop Y) whiskerLeft X _ _ f := (X.unop ◁ f.unop).op diff --git a/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean b/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean index 40323311fa24f..a6ed99f917dac 100644 --- a/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean +++ b/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean @@ -33,8 +33,6 @@ open MonoidalCategory -- I don't know if that is a problem, might need to change it back in the future, but -- if so it might be better to fix then instead of at the moment of porting. -attribute [local simp] id_tensorHom tensorHom_id - /-- `(πŸ™_ C ⟢ -)` is a lax monoidal functor to `Type`. -/ noncomputable def coyonedaTensorUnit (C : Type u) [Category.{v} C] [MonoidalCategory C] : diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index b1322050f6cea..14db1f057bebe 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -90,11 +90,11 @@ instance LiftHom_comp {X Y Z : C} [LiftObj X] [LiftObj Y] [LiftObj Z] (f : X ⟢ instance liftHom_WhiskerLeft (X : C) [LiftObj X] {Y Z : C} [LiftObj Y] [LiftObj Z] (f : Y ⟢ Z) [LiftHom f] : LiftHom (X ◁ f) where - lift := πŸ™ (LiftObj.lift X) βŠ— LiftHom.lift f + lift := LiftObj.lift X ◁ LiftHom.lift f instance liftHom_WhiskerRight {X Y : C} (f : X ⟢ Y) [LiftObj X] [LiftObj Y] [LiftHom f] {Z : C} [LiftObj Z] : LiftHom (f β–· Z) where - lift := LiftHom.lift f βŠ— πŸ™ (LiftObj.lift Z) + lift := LiftHom.lift f β–· LiftObj.lift Z instance LiftHom_tensor {W X Y Z : C} [LiftObj W] [LiftObj X] [LiftObj Y] [LiftObj Z] (f : W ⟢ X) (g : Y ⟢ Z) [LiftHom f] [LiftHom g] : LiftHom (f βŠ— g) where @@ -129,12 +129,12 @@ instance whiskerRight (X Y Z : C) [LiftObj X] [LiftObj Y] [LiftObj Z] [MonoidalC @[simps] instance tensor_right (X Y : C) [LiftObj X] [LiftObj Y] [MonoidalCoherence (πŸ™_ C) Y] : MonoidalCoherence X (X βŠ— Y) := - ⟨(ρ_ X).inv ≫ (πŸ™ X βŠ— MonoidalCoherence.hom)⟩ + ⟨(ρ_ X).inv ≫ (X ◁ MonoidalCoherence.hom)⟩ @[simps] instance tensor_right' (X Y : C) [LiftObj X] [LiftObj Y] [MonoidalCoherence Y (πŸ™_ C)] : MonoidalCoherence (X βŠ— Y) X := - ⟨(πŸ™ X βŠ— MonoidalCoherence.hom) ≫ (ρ_ X).hom⟩ + ⟨(X ◁ MonoidalCoherence.hom) ≫ (ρ_ X).hom⟩ @[simps] instance left (X Y : C) [LiftObj X] [LiftObj Y] [MonoidalCoherence X Y] : @@ -333,7 +333,7 @@ def insertTrailingIds (g : MVarId) : MetaM MVarId := do -- Porting note: this is an ugly port, using too many `evalTactic`s. -- We can refactor later into either a `macro` (but the flow control is awkward) -- or a `MetaM` tactic. -def coherence_loop (maxSteps := 47) : TacticM Unit := +def coherence_loop (maxSteps := 37) : TacticM Unit := match maxSteps with | 0 => exception' "`coherence` tactic reached iteration limit" | maxSteps' + 1 => do From e6fbaf3b8d032def8eea22e7563877b58acd47b6 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 4 Feb 2024 05:34:47 +0900 Subject: [PATCH 53/62] style consistency --- Mathlib/CategoryTheory/Monoidal/Category.lean | 44 +++++++++---------- 1 file changed, 20 insertions(+), 24 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index 7321d933f810f..5a4c81a2848e3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -248,12 +248,12 @@ theorem tensor_comp {X₁ Y₁ Z₁ Xβ‚‚ Yβ‚‚ Zβ‚‚ : C} @[simp] theorem id_tensorHom (X : C) {Y₁ Yβ‚‚ : C} (f : Y₁ ⟢ Yβ‚‚) : - (πŸ™ X) βŠ— f = X ◁ f := by + πŸ™ X βŠ— f = X ◁ f := by simp [tensorHom_def] @[simp] theorem tensorHom_id {X₁ Xβ‚‚ : C} (f : X₁ ⟢ Xβ‚‚) (Y : C) : - f βŠ— (πŸ™ Y) = f β–· Y := by + f βŠ— πŸ™ Y = f β–· Y := by simp [tensorHom_def] end MonoidalCategory @@ -436,31 +436,27 @@ theorem tensor_whiskerLeft_symm (X Y : C) {Z Z' : C} (f : Z ⟢ Z') : @[reassoc] theorem leftUnitor_naturality {X Y : C} (f : X ⟢ Y) : - πŸ™_ C ◁ f ≫ (Ξ»_ Y).hom = (Ξ»_ X).hom ≫ f := - by simp + πŸ™_ C ◁ f ≫ (Ξ»_ Y).hom = (Ξ»_ X).hom ≫ f := by simp @[reassoc] theorem leftUnitor_inv_naturality {X Y : C} (f : X ⟢ Y) : - f ≫ (Ξ»_ Y).inv = (Ξ»_ X).inv ≫ (_ ◁ f) := by simp + f ≫ (Ξ»_ Y).inv = (Ξ»_ X).inv ≫ _ ◁ f := by simp @[reassoc] theorem id_whiskerLeft_symm {X X' : C} (f : X ⟢ X') : - f = (Ξ»_ X).inv ≫ πŸ™_ C ◁ f ≫ (Ξ»_ X').hom := by - simp + f = (Ξ»_ X).inv ≫ πŸ™_ C ◁ f ≫ (Ξ»_ X').hom := by simp @[reassoc] theorem rightUnitor_naturality {X Y : C} (f : X ⟢ Y) : - f β–· πŸ™_ C ≫ (ρ_ Y).hom = (ρ_ X).hom ≫ f := by - simp + f β–· πŸ™_ C ≫ (ρ_ Y).hom = (ρ_ X).hom ≫ f := by simp @[reassoc] theorem rightUnitor_inv_naturality {X X' : C} (f : X ⟢ X') : - f ≫ (ρ_ X').inv = (ρ_ X).inv ≫ (f β–· _) := by simp + f ≫ (ρ_ X').inv = (ρ_ X).inv ≫ f β–· _ := by simp @[reassoc] theorem whiskerRight_id_symm {X Y : C} (f : X ⟢ Y) : - f = (ρ_ X).inv ≫ f β–· πŸ™_ C ≫ (ρ_ Y).hom := by - simp + f = (ρ_ X).inv ≫ f β–· πŸ™_ C ≫ (ρ_ Y).hom := by simp theorem whiskerLeft_iff {X Y : C} (f g : X ⟢ Y) : πŸ™_ C ◁ f = πŸ™_ C ◁ g ↔ f = g := by simp @@ -500,8 +496,8 @@ theorem pentagon_inv_hom_hom_hom_inv : @[reassoc (attr := simp)] theorem pentagon_hom_inv_inv_inv_inv : W ◁ (Ξ±_ X Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).inv ≫ (Ξ±_ (W βŠ— X) Y Z).inv = - (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z := - by simp [← cancel_epi (W ◁ (Ξ±_ X Y Z).inv)] + (Ξ±_ W (X βŠ— Y) Z).inv ≫ (Ξ±_ W X Y).inv β–· Z := by + simp [← cancel_epi (W ◁ (Ξ±_ X Y Z).inv)] @[reassoc (attr := simp)] theorem pentagon_hom_hom_inv_hom_hom : @@ -525,8 +521,8 @@ theorem pentagon_hom_hom_inv_inv_hom : @[reassoc (attr := simp)] theorem pentagon_inv_hom_hom_hom_hom : (Ξ±_ W X Y).inv β–· Z ≫ (Ξ±_ (W βŠ— X) Y Z).hom ≫ (Ξ±_ W X (Y βŠ— Z)).hom = - (Ξ±_ W (X βŠ— Y) Z).hom ≫ W ◁ (Ξ±_ X Y Z).hom := - by simp [← cancel_epi ((Ξ±_ W X Y).hom β–· Z)] + (Ξ±_ W (X βŠ— Y) Z).hom ≫ W ◁ (Ξ±_ X Y Z).hom := by + simp [← cancel_epi ((Ξ±_ W X Y).hom β–· Z)] @[reassoc (attr := simp)] theorem pentagon_inv_inv_hom_inv_inv : @@ -621,49 +617,49 @@ theorem associator_inv_conjugation {X X' Y Y' Z Z' : C} (f : X ⟢ X') (g : Y @[reassoc (attr := simp)] theorem hom_inv_id_tensor {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : - (f.hom βŠ— g) ≫ (f.inv βŠ— h) = (V ◁ g) ≫ (V ◁ h) := by + (f.hom βŠ— g) ≫ (f.inv βŠ— h) = V ◁ g ≫ V ◁ h := by rw [← tensor_comp, f.hom_inv_id, id_tensorHom, whiskerLeft_comp] #align category_theory.monoidal_category.hom_inv_id_tensor CategoryTheory.MonoidalCategory.hom_inv_id_tensor @[reassoc (attr := simp)] theorem inv_hom_id_tensor {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : - (f.inv βŠ— g) ≫ (f.hom βŠ— h) = (W ◁ g) ≫ (W ◁ h) := by + (f.inv βŠ— g) ≫ (f.hom βŠ— h) = W ◁ g ≫ W ◁ h := by rw [← tensor_comp, f.inv_hom_id, id_tensorHom, whiskerLeft_comp] #align category_theory.monoidal_category.inv_hom_id_tensor CategoryTheory.MonoidalCategory.inv_hom_id_tensor @[reassoc (attr := simp)] theorem tensor_hom_inv_id {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : - (g βŠ— f.hom) ≫ (h βŠ— f.inv) = (g β–· V) ≫ (h β–· V) := by + (g βŠ— f.hom) ≫ (h βŠ— f.inv) = g β–· V ≫ h β–· V := by rw [← tensor_comp, f.hom_inv_id, tensorHom_id, comp_whiskerRight] #align category_theory.monoidal_category.tensor_hom_inv_id CategoryTheory.MonoidalCategory.tensor_hom_inv_id @[reassoc (attr := simp)] theorem tensor_inv_hom_id {V W X Y Z : C} (f : V β‰… W) (g : X ⟢ Y) (h : Y ⟢ Z) : - (g βŠ— f.inv) ≫ (h βŠ— f.hom) = (g β–· W) ≫ (h β–· W) := by + (g βŠ— f.inv) ≫ (h βŠ— f.hom) = g β–· W ≫ h β–· W := by rw [← tensor_comp, f.inv_hom_id, tensorHom_id, comp_whiskerRight] #align category_theory.monoidal_category.tensor_inv_hom_id CategoryTheory.MonoidalCategory.tensor_inv_hom_id @[reassoc (attr := simp)] theorem hom_inv_id_tensor' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : - (f βŠ— g) ≫ (inv f βŠ— h) = (V ◁ g) ≫ (V ◁ h) := by + (f βŠ— g) ≫ (inv f βŠ— h) = V ◁ g ≫ V ◁ h := by rw [← tensor_comp, IsIso.hom_inv_id, id_tensorHom, whiskerLeft_comp] #align category_theory.monoidal_category.hom_inv_id_tensor' CategoryTheory.MonoidalCategory.hom_inv_id_tensor' @[reassoc (attr := simp)] theorem inv_hom_id_tensor' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : - (inv f βŠ— g) ≫ (f βŠ— h) = (W ◁ g) ≫ (W ◁ h) := by + (inv f βŠ— g) ≫ (f βŠ— h) = W ◁ g ≫ W ◁ h := by rw [← tensor_comp, IsIso.inv_hom_id, id_tensorHom, whiskerLeft_comp] #align category_theory.monoidal_category.inv_hom_id_tensor' CategoryTheory.MonoidalCategory.inv_hom_id_tensor' @[reassoc (attr := simp)] theorem tensor_hom_inv_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : - (g βŠ— f) ≫ (h βŠ— inv f) = (g β–· V) ≫ (h β–· V) := by + (g βŠ— f) ≫ (h βŠ— inv f) = g β–· V ≫ h β–· V := by rw [← tensor_comp, IsIso.hom_inv_id, tensorHom_id, comp_whiskerRight] #align category_theory.monoidal_category.tensor_hom_inv_id' CategoryTheory.MonoidalCategory.tensor_hom_inv_id' @[reassoc (attr := simp)] theorem tensor_inv_hom_id' {V W X Y Z : C} (f : V ⟢ W) [IsIso f] (g : X ⟢ Y) (h : Y ⟢ Z) : - (g βŠ— inv f) ≫ (h βŠ— f) = (g β–· W) ≫ (h β–· W) := by + (g βŠ— inv f) ≫ (h βŠ— f) = g β–· W ≫ h β–· W := by rw [← tensor_comp, IsIso.inv_hom_id, tensorHom_id, comp_whiskerRight] #align category_theory.monoidal_category.tensor_inv_hom_id' CategoryTheory.MonoidalCategory.tensor_inv_hom_id' From 3b52996d4a136b9226ea23ed01ee7e4381e4473f Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 4 Mar 2024 21:06:27 +0900 Subject: [PATCH 54/62] fix --- .../CategoryTheory/Monoidal/Free/Basic.lean | 4 --- .../Monoidal/NaturalTransformation.lean | 4 +-- Mathlib/CategoryTheory/Monoidal/Opposite.lean | 30 ++++++++++++------- 3 files changed, 21 insertions(+), 17 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean index 3ad4af7edfdba..db59d13ab79a1 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean @@ -337,11 +337,9 @@ def projectMap (X Y : F C) : (X ⟢ Y) β†’ (projectObj f X ⟢ projectObj f Y) : | l_inv_hom => dsimp only [projectMapAux]; rw [Iso.inv_hom_id] | pentagon => dsimp only [projectMapAux] - simp only [tensorHom_id, id_tensorHom] exact MonoidalCategory.pentagon _ _ _ _ | triangle => dsimp only [projectMapAux] - simp only [tensorHom_id, id_tensorHom] exact MonoidalCategory.triangle _ _ | whisker_exchange => dsimp only [projectMapAux, projectObj]; simp [MonoidalCategory.whisker_exchange] @@ -366,14 +364,12 @@ def project : MonoidalFunctor (F C) D where induction' f using Quotient.recOn Β· dsimp simp - rw [← tensorHom_id, ← tensorHom_id] rfl Β· rfl ΞΌ_natural_right := fun _ f => by induction' f using Quotient.recOn Β· dsimp simp - rw [← id_tensorHom, ← id_tensorHom] rfl Β· rfl #align category_theory.free_monoidal_category.project CategoryTheory.FreeMonoidalCategory.project diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index 61b4f153da4cf..203109c65b96a 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -206,8 +206,8 @@ def monoidalUnit (F : MonoidalFunctor C D) [IsEquivalence F.toFunctor] : Iso.hom_inv_id_app_assoc, tensorHom_def'] slice_rhs 3 4 => erw [Iso.hom_inv_id_app] simp only [id_obj, id_comp, assoc, whisker_exchange_assoc] - simp only [← whiskerLeft_comp_assoc, ← comp_whiskerRight_assoc] - simp [- whiskerLeft_comp, - comp_whiskerRight] } + simp only [← MonoidalCategory.whiskerLeft_comp_assoc, ← comp_whiskerRight_assoc] + simp [- MonoidalCategory.whiskerLeft_comp, - comp_whiskerRight] } #align category_theory.monoidal_unit CategoryTheory.monoidalUnit instance (F : MonoidalFunctor C D) [IsEquivalence F.toFunctor] : IsIso (monoidalUnit F) := diff --git a/Mathlib/CategoryTheory/Monoidal/Opposite.lean b/Mathlib/CategoryTheory/Monoidal/Opposite.lean index cf24459ed65c9..58dbd8fea3e47 100644 --- a/Mathlib/CategoryTheory/Monoidal/Opposite.lean +++ b/Mathlib/CategoryTheory/Monoidal/Opposite.lean @@ -267,17 +267,25 @@ instance monoidalCategoryMop : MonoidalCategory Cα΄Ήα΅’α΅– where associator X Y Z := (Ξ±_ (unmop Z) (unmop Y) (unmop X)).symm.mop leftUnitor X := (ρ_ (unmop X)).mop rightUnitor X := (Ξ»_ (unmop X)).mop - whiskerLeft_id X Y := unmop_inj (by simp) - whiskerLeft_comp W X Y Z f g := unmop_inj (by simp) - id_whiskerLeft f := unmop_inj (by simp) - tensor_whiskerLeft W X Y Z f := unmop_inj (by simp) - id_whiskerRight X Y := unmop_inj (by simp) - comp_whiskerRight f g X := unmop_inj (by simp) - whiskerRight_id f := unmop_inj (by simp) - whisker_assoc W X Y f Z := unmop_inj (by simp) - whisker_exchange f g := unmop_inj (whisker_exchange g.unmop f.unmop).symm - pentagon W X Y Z := unmop_inj (by dsimp; coherence) - triangle X Y := unmop_inj (by dsimp; coherence) + whiskerLeft_id X Y := Quiver.Hom.unmop_inj (by simp) + whiskerLeft_comp W X Y Z f g := Quiver.Hom.unmop_inj (by simp) + id_whiskerLeft f := Quiver.Hom.unmop_inj (by simp) + tensor_whiskerLeft W X Y Z f := Quiver.Hom.unmop_inj (by simp) + id_whiskerRight X Y := Quiver.Hom.unmop_inj (by simp) + comp_whiskerRight f g X := Quiver.Hom.unmop_inj (by simp) + whiskerRight_id f := Quiver.Hom.unmop_inj (by simp) + whisker_assoc W X Y f Z := Quiver.Hom.unmop_inj (by simp) + whisker_exchange f g := Quiver.Hom.unmop_inj (whisker_exchange g.unmop f.unmop).symm + pentagon W X Y Z := Quiver.Hom.unmop_inj (by dsimp; coherence) + triangle X Y := Quiver.Hom.unmop_inj (by dsimp; coherence) +#align category_theory.monoidal_category_mop CategoryTheory.monoidalCategoryMop + +-- it would be nice if we could autogenerate all of these somehow +section MonoidalOppositeLemmas + +@[simp] lemma mop_tensorObj (X Y : C) : mop (X βŠ— Y) = mop Y βŠ— mop X := rfl +@[simp] lemma unmop_tensorObj (X Y : Cα΄Ήα΅’α΅–) : unmop (X βŠ— Y) = unmop Y βŠ— unmop X := rfl + @[simp] lemma mop_tensorUnit : mop (πŸ™_ C) = πŸ™_ Cα΄Ήα΅’α΅– := rfl @[simp] lemma unmop_tensorUnit : unmop (πŸ™_ Cα΄Ήα΅’α΅–) = πŸ™_ C := rfl From b7f60f05f0a665e5903ce5df3448b7ce29ab3563 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 4 Mar 2024 22:32:26 +0900 Subject: [PATCH 55/62] fix --- Mathlib/CategoryTheory/Monoidal/Limits.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Limits.lean index 5cf66283e3449..680531b275bae 100644 --- a/Mathlib/CategoryTheory/Monoidal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Limits.lean @@ -91,7 +91,7 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorH rw [← comp_whiskerRight] erw [limit.lift_Ο€] dsimp - slice_rhs 2 3 => rw [id_tensorHom, leftUnitor_naturality] + slice_rhs 2 3 => rw [leftUnitor_naturality] simp) (right_unitality := fun X => by ext j; dsimp @@ -101,7 +101,7 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorH rw [← MonoidalCategory.whiskerLeft_comp] erw [limit.lift_Ο€] dsimp - slice_rhs 2 3 => rw [tensorHom_id, rightUnitor_naturality] + slice_rhs 2 3 => rw [rightUnitor_naturality] simp) #align category_theory.limits.limit_lax_monoidal CategoryTheory.Limits.limitLaxMonoidal From 39e9a312a4ed03a46fac14de2c9d7be4081cb801 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 4 Mar 2024 22:33:03 +0900 Subject: [PATCH 56/62] remove erw --- .../CategoryTheory/Monoidal/Free/Basic.lean | 26 ++++++++++++ .../Monoidal/Free/Coherence.lean | 41 +++++-------------- 2 files changed, 37 insertions(+), 30 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean index db59d13ab79a1..4a5d91e6bfdbc 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean @@ -219,6 +219,32 @@ instance : MonoidalCategory (F C) where pentagon W X Y Z := Quotient.sound pentagon triangle X Y := Quotient.sound triangle +theorem Hom.inductionOn {motive : {X Y : F C} β†’ (X ⟢ Y) β†’ Prop} {X Y : F C} (t : X ⟢ Y) + (id : (X : F C) β†’ motive (πŸ™ X)) + (Ξ±_hom : (X Y Z : F C) β†’ motive (Ξ±_ X Y Z).hom) + (Ξ±_inv : (X Y Z : F C) β†’ motive (Ξ±_ X Y Z).inv) + (l_hom : (X : F C) β†’ motive (Ξ»_ X).hom) + (l_inv : (X : F C) β†’ motive (Ξ»_ X).inv) + (ρ_hom : (X : F C) β†’ motive (ρ_ X).hom) + (ρ_inv : (X : F C) β†’ motive (ρ_ X).inv) + (comp : {X Y Z : F C} β†’ (f : X ⟢ Y) β†’ (g : Y ⟢ Z) β†’ motive f β†’ motive g β†’ motive (f ≫ g)) + (whiskerLeft : (X : F C) β†’ {Y Z : F C} β†’ (f : Y ⟢ Z) β†’ motive f β†’ motive (X ◁ f)) + (whiskerRight : {X Y : F C} β†’ (f : X ⟢ Y) β†’ (Z : F C) β†’ motive f β†’ motive (f β–· Z)) : + motive t := by + apply Quotient.inductionOn + intro f + induction f with + | id X => exact id X + | Ξ±_hom X Y Z => exact Ξ±_hom X Y Z + | Ξ±_inv X Y Z => exact Ξ±_inv X Y Z + | l_hom X => exact l_hom X + | l_inv X => exact l_inv X + | ρ_hom X => exact ρ_hom X + | ρ_inv X => exact ρ_inv X + | comp f g hf hg => exact comp _ _ (hf ⟦f⟧) (hg ⟦g⟧) + | whiskerLeft X f hf => exact whiskerLeft X _ (hf ⟦f⟧) + | whiskerRight f X hf => exact whiskerRight _ X (hf ⟦f⟧) + @[simp] theorem mk_comp {X Y Z : F C} (f : X ⟢ᡐ Y) (g : Y ⟢ᡐ Z) : ⟦f.comp g⟧ = @CategoryStruct.comp (F C) _ _ _ _ ⟦f⟧ ⟦g⟧ := diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean index 535cb586d842e..e879531460126 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean @@ -285,43 +285,24 @@ theorem normalizeObj_congr (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ Y theorem normalize_naturality (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ Y) : inclusionObj n ◁ f ≫ (normalizeIsoApp' C Y n).hom = (normalizeIsoApp' C X n).hom ≫ - inclusion.map (eqToHom (Discrete.ext _ _ - ((normalizeObj_congr n f)))) := by - induction' f using Quotient.recOn with f'; swap; rfl + inclusion.map (eqToHom (Discrete.ext _ _ (normalizeObj_congr n f))) := by revert n - induction f' with - | id => erw [mk_id]; simp - | Ξ±_hom X Y Z => erw [mk_Ξ±_hom]; simp - | Ξ±_inv X Y Z => erw [mk_Ξ±_inv]; simp - | l_hom X => erw [mk_l_hom]; simp - | l_inv X => erw [mk_l_inv]; simp - | ρ_hom X => erw [mk_ρ_hom]; simp - | ρ_inv X => erw [mk_ρ_inv]; simp - | comp f g ihf ihg => - intro n - erw [mk_comp] - simp only [MonoidalCategory.whiskerLeft_comp] - slice_lhs 2 3 => rw [ihg] - slice_lhs 1 2 => rw [ihf] - simp - | whiskerLeft X f ih => + induction f using Hom.inductionOn + case comp f g ihf ihg => simp [ihg, reassoc_of% (ihf _)] + case whiskerLeft X' X Y f ih => intro n - erw [mk_whiskerLeft] - simp only [tensor_eq_tensor, normalizeObj_tensor, normalizeIsoApp', Iso.trans_hom, - Iso.symm_hom, whiskerRightIso_hom, Function.comp_apply, inclusion_obj, inclusion_map, - Category.assoc] + dsimp only [normalizeObj_tensor, normalizeIsoApp', tensor_eq_tensor, Iso.trans_hom, + Iso.symm_hom, whiskerRightIso_hom, Function.comp_apply, inclusion_obj] rw [associator_inv_naturality_right_assoc, whisker_exchange_assoc, ih] simp - | @whiskerRight X Y h Ξ·' ih => + case whiskerRight X Y h Ξ·' ih => intro n - erw [mk_whiskerRight] - simp only [tensor_eq_tensor, normalizeObj_tensor, normalizeIsoApp', Iso.trans_hom, - Iso.symm_hom, whiskerRightIso_hom, Function.comp_apply, inclusion_obj, inclusion_map, - Category.assoc] + dsimp only [normalizeObj_tensor, normalizeIsoApp', tensor_eq_tensor, Iso.trans_hom, + Iso.symm_hom, whiskerRightIso_hom, Function.comp_apply, inclusion_obj] rw [associator_inv_naturality_middle_assoc, ← comp_whiskerRight_assoc, ih] - have := dcongr_arg (fun x => (normalizeIsoApp' C Ξ·' x).hom) - (normalizeObj_congr n (Quotient.mk (setoidHom X Y) h)) + have := dcongr_arg (fun x => (normalizeIsoApp' C Ξ·' x).hom) (normalizeObj_congr n h) simp [this] + all_goals simp end From ad3031998b48787231e44d102cdc45ace6c050c0 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 4 Mar 2024 22:33:58 +0900 Subject: [PATCH 57/62] style --- .../Monoidal/Free/Coherence.lean | 22 +++++++++---------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean index e879531460126..d7c65ed543398 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean @@ -290,18 +290,18 @@ theorem normalize_naturality (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟢ induction f using Hom.inductionOn case comp f g ihf ihg => simp [ihg, reassoc_of% (ihf _)] case whiskerLeft X' X Y f ih => - intro n - dsimp only [normalizeObj_tensor, normalizeIsoApp', tensor_eq_tensor, Iso.trans_hom, - Iso.symm_hom, whiskerRightIso_hom, Function.comp_apply, inclusion_obj] - rw [associator_inv_naturality_right_assoc, whisker_exchange_assoc, ih] - simp + intro n + dsimp only [normalizeObj_tensor, normalizeIsoApp', tensor_eq_tensor, Iso.trans_hom, + Iso.symm_hom, whiskerRightIso_hom, Function.comp_apply, inclusion_obj] + rw [associator_inv_naturality_right_assoc, whisker_exchange_assoc, ih] + simp case whiskerRight X Y h Ξ·' ih => - intro n - dsimp only [normalizeObj_tensor, normalizeIsoApp', tensor_eq_tensor, Iso.trans_hom, - Iso.symm_hom, whiskerRightIso_hom, Function.comp_apply, inclusion_obj] - rw [associator_inv_naturality_middle_assoc, ← comp_whiskerRight_assoc, ih] - have := dcongr_arg (fun x => (normalizeIsoApp' C Ξ·' x).hom) (normalizeObj_congr n h) - simp [this] + intro n + dsimp only [normalizeObj_tensor, normalizeIsoApp', tensor_eq_tensor, Iso.trans_hom, + Iso.symm_hom, whiskerRightIso_hom, Function.comp_apply, inclusion_obj] + rw [associator_inv_naturality_middle_assoc, ← comp_whiskerRight_assoc, ih] + have := dcongr_arg (fun x => (normalizeIsoApp' C Ξ·' x).hom) (normalizeObj_congr n h) + simp [this] all_goals simp end From 3b70e1f35f75e4de91712b83cc0a0a06c16263f4 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 4 Mar 2024 22:44:05 +0900 Subject: [PATCH 58/62] fix --- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index 3cd885689d61e..624d27a563aa2 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -528,6 +528,7 @@ theorem associator_inv_hom (X Y Z : Mon_ C) : (Ξ±_ X Y Z).inv.hom = (Ξ±_ X.X Y.X instance monMonoidal : MonoidalCategory (Mon_ C) where tensorHom_def := by intros; ext; simp [tensorHom_def] + whisker_exchange := by intros; ext; simp [whisker_exchange] #align Mon_.Mon_monoidal Mon_.monMonoidal end Mon_ From d37570c8f4991375625cfe25747f24209e77edfb Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Mon, 4 Mar 2024 23:27:42 +0900 Subject: [PATCH 59/62] lint --- Mathlib/CategoryTheory/Monoidal/Limits.lean | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Limits.lean index 680531b275bae..58f0f70364d75 100644 --- a/Mathlib/CategoryTheory/Monoidal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Limits.lean @@ -85,7 +85,9 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorH simp) (left_unitality := fun X => by ext j; dsimp - simp + simp only [tensorHom_id, limit.lift_map, Category.assoc, limit.lift_Ο€, Cones.postcompose_obj_pt, + Cones.postcompose_obj_Ο€, NatTrans.comp_app, Functor.const_obj_obj, Monoidal.tensorObj_obj, + Monoidal.tensorUnit_obj, Monoidal.leftUnitor_hom_app] conv_rhs => rw [tensorHom_def _ (limit.Ο€ X j)] slice_rhs 1 2 => rw [← comp_whiskerRight] @@ -95,7 +97,9 @@ instance limitLaxMonoidal : LaxMonoidal fun F : J β₯€ C => limit F := .ofTensorH simp) (right_unitality := fun X => by ext j; dsimp - simp + simp only [id_tensorHom, limit.lift_map, Category.assoc, limit.lift_Ο€, Cones.postcompose_obj_pt, + Cones.postcompose_obj_Ο€, NatTrans.comp_app, Functor.const_obj_obj, Monoidal.tensorObj_obj, + Monoidal.tensorUnit_obj, Monoidal.rightUnitor_hom_app] conv_rhs => rw [tensorHom_def' (limit.Ο€ X j)] slice_rhs 1 2 => rw [← MonoidalCategory.whiskerLeft_comp] From 7a46b940732c0c6d7991f78773909dfd86c51d33 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Fri, 8 Mar 2024 02:45:41 +0900 Subject: [PATCH 60/62] Update Basic.lean --- .../CategoryTheory/Monoidal/Free/Basic.lean | 36 ------------------- 1 file changed, 36 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean index 04aa60d672e44..4c07fa800b853 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean @@ -213,32 +213,6 @@ instance : MonoidalCategory (F C) where pentagon W X Y Z := Quotient.sound pentagon triangle X Y := Quotient.sound triangle -theorem Hom.inductionOn {motive : {X Y : F C} β†’ (X ⟢ Y) β†’ Prop} {X Y : F C} (t : X ⟢ Y) - (id : (X : F C) β†’ motive (πŸ™ X)) - (Ξ±_hom : (X Y Z : F C) β†’ motive (Ξ±_ X Y Z).hom) - (Ξ±_inv : (X Y Z : F C) β†’ motive (Ξ±_ X Y Z).inv) - (l_hom : (X : F C) β†’ motive (Ξ»_ X).hom) - (l_inv : (X : F C) β†’ motive (Ξ»_ X).inv) - (ρ_hom : (X : F C) β†’ motive (ρ_ X).hom) - (ρ_inv : (X : F C) β†’ motive (ρ_ X).inv) - (comp : {X Y Z : F C} β†’ (f : X ⟢ Y) β†’ (g : Y ⟢ Z) β†’ motive f β†’ motive g β†’ motive (f ≫ g)) - (whiskerLeft : (X : F C) β†’ {Y Z : F C} β†’ (f : Y ⟢ Z) β†’ motive f β†’ motive (X ◁ f)) - (whiskerRight : {X Y : F C} β†’ (f : X ⟢ Y) β†’ (Z : F C) β†’ motive f β†’ motive (f β–· Z)) : - motive t := by - apply Quotient.inductionOn - intro f - induction f with - | id X => exact id X - | Ξ±_hom X Y Z => exact Ξ±_hom X Y Z - | Ξ±_inv X Y Z => exact Ξ±_inv X Y Z - | l_hom X => exact l_hom X - | l_inv X => exact l_inv X - | ρ_hom X => exact ρ_hom X - | ρ_inv X => exact ρ_inv X - | comp f g hf hg => exact comp _ _ (hf ⟦f⟧) (hg ⟦g⟧) - | whiskerLeft X f hf => exact whiskerLeft X _ (hf ⟦f⟧) - | whiskerRight f X hf => exact whiskerRight _ X (hf ⟦f⟧) - @[simp] theorem mk_comp {X Y Z : F C} (f : X ⟢ᡐ Y) (g : Y ⟢ᡐ Z) : ⟦f.comp g⟧ = @CategoryStruct.comp (F C) _ _ _ _ ⟦f⟧ ⟦g⟧ := @@ -247,16 +221,6 @@ theorem mk_comp {X Y Z : F C} (f : X ⟢ᡐ Y) (g : Y ⟢ᡐ Z) : #noalign category_theory.free_monoidal_category.mk_tensor -@[simp] -theorem mk_whiskerLeft (X : F C) {Y₁ Yβ‚‚ : F C} (f : Y₁ ⟢ᡐ Yβ‚‚) : - ⟦f.whiskerLeft X⟧ = MonoidalCategory.whiskerLeft (C := F C) (X := X) (f := ⟦f⟧) := - rfl - -@[simp] -theorem mk_whiskerRight {X₁ Xβ‚‚ : F C} (f : X₁ ⟢ᡐ Xβ‚‚) (Y : F C) : - ⟦f.whiskerRight Y⟧ = MonoidalCategory.whiskerRight (C := F C) (f := ⟦f⟧) (Y := Y) := - rfl - @[simp] theorem mk_whiskerLeft (X : F C) {Y₁ Yβ‚‚ : F C} (f : Y₁ ⟢ᡐ Yβ‚‚) : ⟦f.whiskerLeft X⟧ = MonoidalCategory.whiskerLeft (C := F C) (X := X) (f := ⟦f⟧) := From a85b60c9d44d9b6ee42efde8fcb2f6aa1ed4598a Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Fri, 8 Mar 2024 03:17:20 +0900 Subject: [PATCH 61/62] Update NaturalTransformation.lean --- .../Monoidal/NaturalTransformation.lean | 20 ++++++++++--------- 1 file changed, 11 insertions(+), 9 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index 203109c65b96a..bf63ce1a79f34 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -199,15 +199,17 @@ def monoidalUnit (F : MonoidalFunctor C D) [IsEquivalence F.toFunctor] : simp only [Adjunction.homEquiv_unit, Adjunction.homEquiv_naturality_right, id_comp, assoc] simp only [← Functor.map_comp, assoc] - erw [e.counit_app_functor, e.counit_app_functor] - simp only [comp_obj, asEquivalence_functor, asEquivalence_inverse, id_obj, - LaxMonoidalFunctor.ΞΌ_natural_left, LaxMonoidalFunctor.ΞΌ_natural_right_assoc, - IsIso.inv_hom_id_assoc, map_comp, IsEquivalence.inv_fun_map, assoc, - Iso.hom_inv_id_app_assoc, tensorHom_def'] - slice_rhs 3 4 => erw [Iso.hom_inv_id_app] - simp only [id_obj, id_comp, assoc, whisker_exchange_assoc] - simp only [← MonoidalCategory.whiskerLeft_comp_assoc, ← comp_whiskerRight_assoc] - simp [- MonoidalCategory.whiskerLeft_comp, - comp_whiskerRight] } + erw [e.counit_app_functor, e.counit_app_functor, + F.toLaxMonoidalFunctor.ΞΌ_natural, IsIso.inv_hom_id_assoc] + simp only [CategoryTheory.IsEquivalence.inv_fun_map] + slice_rhs 2 3 => erw [Iso.hom_inv_id_app] + dsimp + simp only [CategoryTheory.Category.id_comp] + slice_rhs 1 2 => + rw [← tensor_comp, Iso.hom_inv_id_app, Iso.hom_inv_id_app] + dsimp + rw [tensor_id] + simp } #align category_theory.monoidal_unit CategoryTheory.monoidalUnit instance (F : MonoidalFunctor C D) [IsEquivalence F.toFunctor] : IsIso (monoidalUnit F) := From c7bac228224197f7cee72e2b439304151a5f0319 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Fri, 8 Mar 2024 03:35:28 +0900 Subject: [PATCH 62/62] Update MonoidalComp.lean --- Mathlib/Tactic/CategoryTheory/MonoidalComp.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Tactic/CategoryTheory/MonoidalComp.lean b/Mathlib/Tactic/CategoryTheory/MonoidalComp.lean index 4acfa6699d1f1..998bb02ec7870 100644 --- a/Mathlib/Tactic/CategoryTheory/MonoidalComp.lean +++ b/Mathlib/Tactic/CategoryTheory/MonoidalComp.lean @@ -90,22 +90,22 @@ instance refl (X : C) : MonoidalCoherence X X := βŸ¨πŸ™ _⟩ @[simps] instance whiskerLeft (X Y Z : C) [MonoidalCoherence Y Z] : MonoidalCoherence (X βŠ— Y) (X βŠ— Z) := - βŸ¨πŸ™ X βŠ— βŠ—πŸ™βŸ© + ⟨X ◁ βŠ—πŸ™βŸ© @[simps] instance whiskerRight (X Y Z : C) [MonoidalCoherence X Y] : MonoidalCoherence (X βŠ— Z) (Y βŠ— Z) := - βŸ¨βŠ—πŸ™ βŠ— πŸ™ Z⟩ + βŸ¨βŠ—πŸ™ β–· Z⟩ @[simps] instance tensor_right (X Y : C) [MonoidalCoherence (πŸ™_ C) Y] : MonoidalCoherence X (X βŠ— Y) := - ⟨(ρ_ X).inv ≫ (πŸ™ X βŠ— βŠ—πŸ™)⟩ + ⟨(ρ_ X).inv ≫ X ◁ βŠ—πŸ™βŸ© @[simps] instance tensor_right' (X Y : C) [MonoidalCoherence Y (πŸ™_ C)] : MonoidalCoherence (X βŠ— Y) X := - ⟨(πŸ™ X βŠ— βŠ—πŸ™) ≫ (ρ_ X).hom⟩ + ⟨X ◁ βŠ—πŸ™ ≫ (ρ_ X).hom⟩ @[simps] instance left (X Y : C) [MonoidalCoherence X Y] :