Documentation

Mathlib.CategoryTheory.Monoidal.Action.Basic

Actions from a monoidal category on a category #

Given a monoidal category C, and a category D, we define a left action of C on D as the data of an object c ⊙ₗ d of D for every c : C and d : D, as well as the data required to turn - ⊙ₗ - into a bifunctor, along with structure natural isomorphisms (- ⊗ -) ⊙ₗ - ≅ - ⊙ₗ - ⊙ₗ - and 𝟙_ C ⊙ₗ - ≅ -, subject to coherence conditions.

We also define right actions, for these, the notation for the action of c on d is d ⊙ᵣ c, and the structure isomorphisms are of the form - ⊙ᵣ (- ⊗ -) ≅ (- ⊙ᵣ -) ⊙ᵣ - and - ⊙ₗ 𝟙_ C ≅ -.

References #

TODOs/Projects #

A class that carries the non-Prop data required to define a left action of a monoidal category C on a category D, to set up notations.

  • actionObj : CDD

    The left action on objects. This is denoted c ⊙ₗ d.

  • actionHomLeft {c c' : C} (f : c c') (d : D) : actionObj c d actionObj c' d

    The left action of a map f : c ⟶ c' in C on an object d in D. If we are to consider the action as a functor Α : C ⥤ D ⥤ D, this is (Α.map f).app d. This is denoted f ⊵ₗ d`

  • actionHomRight (c : C) {d d' : D} (f : d d') : actionObj c d actionObj c d'

    The action of an object c : C on a map f : d ⟶ d' in D. If we are to consider the action as a functor Α : C ⥤ D ⥤ D, this is (Α.obj c).map f. This is denoted c ⊴ₗ f`.

  • actionHom {c c' : C} {d d' : D} (f : c c') (g : d d') : actionObj c d actionObj c' d'

    The action of a pair of maps f : c ⟶ c' and d ⟶ d'. By default, this is defined in terms of actionHomLeft and actionHomRight.

  • actionAssocIso (c c' : C) (d : D) : actionObj (tensorObj c c') d actionObj c (actionObj c' d)

    The structural isomorphism (c ⊗ c') ⊙ₗ d ≅ c ⊙ₗ (c' ⊙ₗ d).

  • actionUnitIso (d : D) : actionObj (tensorUnit C) d d

    The structural isomorphism between 𝟙_ C ⊙ₗ d and d.

Instances

    Notation for actionHomLeft, the action of C on morphisms in D.

    Equations
      Instances For

        Notation for actionHomRight, the action of morphism in C on D.

        Equations
          Instances For

            Notation for actionHom, the bifunctorial action of morphisms in C and D on - ⊙ₗ -.

            Equations
              Instances For

                Notation for actionAssocIso, the structural isomorphism - ⊗ - ⊙ₗ - ≅ - ⊙ₗ - ⊙ₗ -.

                Equations
                  Instances For

                    Notation for actionUnitIso, the structural isomorphism 𝟙_ C ⊙ₗ - ≅ -.

                    Equations
                      Instances For

                        Notation for actionUnitIso, the structural isomorphism 𝟙_ C ⊙ₗ - ≅ -, allowing one to specify the acting category.

                        Equations
                          Instances For

                            A MonoidalLeftAction C D is is the data of:

                            • For every object c : C and d : D, an object c ⊙ₗ d of D.
                            • For every morphism f : (c : C) ⟶ c' and every d : D, a morphism f ⊵ₗ d : c ⊙ₗ d ⟶ c' ⊙ₗ d.
                            • For every morphism f : (d : D) ⟶ d' and every c : C, a morphism c ⊴ₗ f : c ⊙ₗ d ⟶ c ⊙ₗ d'.
                            • For every pair of morphisms f : (c : C) ⟶ c' and f : (d : D) ⟶ d', a morphism f ⊙ₗ f' : c ⊙ₗ d ⟶ c' ⊙ₗ d'.
                            • A structure isomorphism αₗ c c' d : c ⊗ c' ⊙ₗ d ≅ c ⊙ₗ c' ⊙ₗ d.
                            • A structure isomorphism λₗ d : (𝟙_ C) ⊙ₗ d ≅ d. Furthermore, we require identities that turn - ⊙ₗ - into a bifunctor, ensure naturality of αₗ and λₗ, and ensure compatibilities with the associator and unitor isomorphisms in C.
                            Instances
                              theorem CategoryTheory.MonoidalCategory.MonoidalLeftAction.actionHom_def_assoc {C : Type u_1} {D : Type u_2} {inst✝ : Category.{u_3, u_1} C} {inst✝¹ : Category.{u_4, u_2} D} {inst✝² : MonoidalCategory C} [self : MonoidalLeftAction C D] {c c' : C} {d d' : D} (f : c c') (g : d d') {Z : D} (h : actionObj c' d' Z) :
                              theorem CategoryTheory.MonoidalCategory.MonoidalLeftAction.id_actionHomLeft_assoc {C : Type u_1} {D : Type u_2} {inst✝ : Category.{u_3, u_1} C} {inst✝¹ : Category.{u_4, u_2} D} {inst✝² : MonoidalCategory C} [self : MonoidalLeftAction C D] (c : C) (d : D) {Z : D} (h : actionObj c d Z) :
                              theorem CategoryTheory.MonoidalCategory.MonoidalLeftAction.actionHom_comp_assoc {C : Type u_1} {D : Type u_2} {inst✝ : Category.{u_3, u_1} C} {inst✝¹ : Category.{u_4, u_2} D} {inst✝² : MonoidalCategory C} [self : MonoidalLeftAction C D] {c c' c'' : C} {d d' d'' : D} (f₁ : c c') (f₂ : c' c'') (g₁ : d d') (g₂ : d' d'') {Z : D} (h : actionObj c'' d'' Z) :
                              theorem CategoryTheory.MonoidalCategory.MonoidalLeftAction.actionAssocIso_hom_naturality_assoc {C : Type u_1} {D : Type u_2} {inst✝ : Category.{u_3, u_1} C} {inst✝¹ : Category.{u_4, u_2} D} {inst✝² : MonoidalCategory C} [self : MonoidalLeftAction C D] {c₁ c₂ c₃ c₄ : C} {d₁ d₂ : D} (f : c₁ c₂) (g : c₃ c₄) (h : d₁ d₂) {Z : D} (h✝ : actionObj c₂ (actionObj c₄ d₂) Z) :
                              @[simp]
                              theorem CategoryTheory.MonoidalCategory.MonoidalLeftAction.associator_actionHom_assoc {C : Type u_1} {D : Type u_2} {inst✝ : Category.{u_3, u_1} C} {inst✝¹ : Category.{u_4, u_2} D} {inst✝² : MonoidalCategory C} [self : MonoidalLeftAction C D] (c₁ c₂ c₃ : C) (d : D) {Z : D} (h : actionObj c₁ (actionObj c₂ (actionObj c₃ d)) Z) :

                              A monoidal category acts on itself on the left through the tensor product.

                              Equations
                                @[simp]
                                theorem CategoryTheory.MonoidalCategory.selfLeftAction_actionHom (C : Type u_1) [Category.{u_3, u_1} C] [MonoidalCategory C] {c✝ c'✝ d✝ d'✝ : C} (f : c✝ c'✝) (g : d✝ d'✝) :
                                @[deprecated CategoryTheory.MonoidalCategory.selfLeftAction (since := "2025-06-13")]

                                Alias of CategoryTheory.MonoidalCategory.selfLeftAction.


                                A monoidal category acts on itself on the left through the tensor product.

                                Equations
                                  Instances For
                                    theorem CategoryTheory.MonoidalCategory.MonoidalLeftAction.actionHom_def' {C : Type u_1} {D : Type u_2} [Category.{u_3, u_1} C] [Category.{u_4, u_2} D] [MonoidalCategory C] [MonoidalLeftAction C D] {x₁ y₁ : C} {x₂ y₂ : D} (f : x₁ y₁) (g : x₂ y₂) :
                                    theorem CategoryTheory.MonoidalCategory.MonoidalLeftAction.actionHom_def'_assoc {C : Type u_1} {D : Type u_2} [Category.{u_3, u_1} C] [Category.{u_4, u_2} D] [MonoidalCategory C] [MonoidalLeftAction C D] {x₁ y₁ : C} {x₂ y₂ : D} (f : x₁ y₁) (g : x₂ y₂) {Z : D} (h : actionObj y₁ y₂ Z) :
                                    theorem CategoryTheory.MonoidalCategory.MonoidalLeftAction.actionAssocIso_inv_naturality {C : Type u_1} {D : Type u_2} [Category.{u_3, u_1} C] [Category.{u_4, u_2} D] [MonoidalCategory C] [MonoidalLeftAction C D] {c₁ c₂ c₃ c₄ : C} {d₁ d₂ : D} (f : c₁ c₂) (g : c₃ c₄) (h : d₁ d₂) :
                                    theorem CategoryTheory.MonoidalCategory.MonoidalLeftAction.actionAssocIso_inv_naturality_assoc {C : Type u_1} {D : Type u_2} [Category.{u_3, u_1} C] [Category.{u_4, u_2} D] [MonoidalCategory C] [MonoidalLeftAction C D] {c₁ c₂ c₃ c₄ : C} {d₁ d₂ : D} (f : c₁ c₂) (g : c₃ c₄) (h : d₁ d₂) {Z : D} (h✝ : actionObj (tensorObj c₂ c₄) d₂ Z) :
                                    @[simp]
                                    theorem CategoryTheory.MonoidalCategory.MonoidalLeftAction.inv_actionHom {C : Type u_1} {D : Type u_2} [Category.{u_3, u_1} C] [Category.{u_4, u_2} D] [MonoidalCategory C] [MonoidalLeftAction C D] {x y : C} {x' y' : D} (f : x y) (g : x' y') [IsIso f] [IsIso g] :
                                    inv (actionHom f g) = actionHom (inv f) (inv g)

                                    Bundle the action of C on D as a functor C ⥤ D ⥤ D.

                                    Equations
                                      Instances For
                                        @[reducible, inline]

                                        Bundle d ↦ c ⊙ₗ d as a functor.

                                        Equations
                                          Instances For
                                            @[reducible, inline]

                                            Bundle c ↦ c ⊙ₗ d as a functor.

                                            Equations
                                              Instances For

                                                Bundle λₗ _ as an isomorphism of functors.

                                                Equations
                                                  Instances For

                                                    A class that carries the non-Prop data required to define a right action of a monoidal category C on a category D, to set up notations.

                                                    • actionObj : DCD

                                                      The right action on objects. This is denoted d ⊙ᵣ c.

                                                    • actionHomRight (d : D) {c c' : C} (f : c c') : actionObj d c actionObj d c'

                                                      The right action of a map f : c ⟶ c' in C on an object d in D. If we are to consider the action as a functor Α : C ⥤ D ⥤ D, this is (Α.map f).app d. This is denoted d ⊴ᵣ f`

                                                    • actionHomLeft {d d' : D} (f : d d') (c : C) : actionObj d c actionObj d' c

                                                      The action of an object c : C on a map f : d ⟶ d' in D. If we are to consider the action as a functor Α : C ⥤ D ⥤ D, this is (Α.obj c).map f. This is denoted f ⊵ᵣ c`.

                                                    • actionHom {c c' : C} {d d' : D} (f : d d') (g : c c') : actionObj d c actionObj d' c'

                                                      The action of a pair of maps f : c ⟶ c' and d ⟶ d'. By default, this is defined in terms of actionHomLeft and actionHomRight.

                                                    • actionAssocIso (d : D) (c c' : C) : actionObj d (tensorObj c c') actionObj (actionObj d c) c'

                                                      The structural isomorphism d ⊙ᵣ (c ⊗ c') ≅ (d ⊙ᵣ c) ⊙ᵣ c'.

                                                    • actionUnitIso (d : D) : actionObj d (tensorUnit C) d

                                                      The structural isomorphism between d ⊙ᵣ 𝟙_ C and d.

                                                    Instances

                                                      Notation for actionHomLeft, the action of D on morphisms in C.

                                                      Equations
                                                        Instances For

                                                          Notation for actionHomRight, the action of morphism in D on C.

                                                          Equations
                                                            Instances For

                                                              Notation for actionHom, the bifunctorial action of morphisms in C and D on - ⊙ -.

                                                              Equations
                                                                Instances For

                                                                  Notation for actionAssocIso, the structural isomorphism - ⊙ᵣ (- ⊗ -) ≅ (- ⊙ᵣ -) ⊙ᵣ -.

                                                                  Equations
                                                                    Instances For

                                                                      Notation for actionUnitIso, the structural isomorphism - ⊙ᵣ 𝟙_ C ≅ -.

                                                                      Equations
                                                                        Instances For

                                                                          Notation for actionUnitIso, the structural isomorphism - ⊙ᵣ 𝟙_ C ≅ -, allowing one to specify the acting category.

                                                                          Equations
                                                                            Instances For

                                                                              A MonoidalRightAction C D is is the data of:

                                                                              • For every object c : C and d : D, an object c ⊙ᵣ d of D.
                                                                              • For every morphism f : (c : C) ⟶ c' and every d : D, a morphism f ⊵ᵣ d : c ⊙ᵣ d ⟶ c' ⊙ᵣ d.
                                                                              • For every morphism f : (d : D) ⟶ d' and every c : C, a morphism c ⊴ᵣ f : c ⊙ᵣ d ⟶ c ⊙ᵣ d'.
                                                                              • For every pair of morphisms f : (c : C) ⟶ c' and f : (d : D) ⟶ d', a morphism f ⊙ᵣₘ f' : c ⊙ᵣ d ⟶ c' ⊙ᵣ d'.
                                                                              • A structure isomorphism αᵣ c c' d : c ⊗ c' ⊙ᵣ d ≅ c ⊙ᵣ c' ⊙ᵣ d.
                                                                              • A structure isomorphism ρᵣ d : (𝟙_ C) ⊙ᵣ d ≅ d. Furthermore, we require identities that turn - ⊙ᵣ - into a bifunctor, ensure naturality of αᵣ and ρᵣ, and ensure compatibilities with the associator and unitor isomorphisms in C.
                                                                              Instances
                                                                                theorem CategoryTheory.MonoidalCategory.MonoidalRightAction.actionHom_def_assoc {C : Type u_1} {D : Type u_2} {inst✝ : Category.{u_3, u_1} C} {inst✝¹ : Category.{u_4, u_2} D} {inst✝² : MonoidalCategory C} [self : MonoidalRightAction C D] {c c' : C} {d d' : D} (f : d d') (g : c c') {Z : D} (h : actionObj d' c' Z) :
                                                                                theorem CategoryTheory.MonoidalCategory.MonoidalRightAction.actionHom_comp_assoc {C : Type u_1} {D : Type u_2} {inst✝ : Category.{u_3, u_1} C} {inst✝¹ : Category.{u_4, u_2} D} {inst✝² : MonoidalCategory C} [self : MonoidalRightAction C D] {c c' c'' : C} {d d' d'' : D} (f₁ : d d') (f₂ : d' d'') (g₁ : c c') (g₂ : c' c'') {Z : D} (h : actionObj d'' c'' Z) :
                                                                                theorem CategoryTheory.MonoidalCategory.MonoidalRightAction.actionAssocIso_hom_naturality_assoc {C : Type u_1} {D : Type u_2} {inst✝ : Category.{u_3, u_1} C} {inst✝¹ : Category.{u_4, u_2} D} {inst✝² : MonoidalCategory C} [self : MonoidalRightAction C D] {d₁ d₂ : D} {c₁ c₂ c₃ c₄ : C} (f : d₁ d₂) (g : c₁ c₂) (h : c₃ c₄) {Z : D} (h✝ : actionObj (actionObj d₂ c₂) c₄ Z) :
                                                                                @[simp]
                                                                                theorem CategoryTheory.MonoidalCategory.MonoidalRightAction.actionHom_associator_assoc {C : Type u_1} {D : Type u_2} {inst✝ : Category.{u_3, u_1} C} {inst✝¹ : Category.{u_4, u_2} D} {inst✝² : MonoidalCategory C} [self : MonoidalRightAction C D] (c₁ c₂ c₃ : C) (d : D) {Z : D} (h : actionObj (actionObj (actionObj d c₁) c₂) c₃ Z) :

                                                                                A monoidal category acts on itself through the tensor product.

                                                                                Equations
                                                                                  @[simp]
                                                                                  theorem CategoryTheory.MonoidalCategory.selRightfAction_actionHom (C : Type u_1) [Category.{u_3, u_1} C] [MonoidalCategory C] {c✝ c'✝ d✝ d'✝ : C} (f : d✝ d'✝) (g : c✝ c'✝) :
                                                                                  theorem CategoryTheory.MonoidalCategory.MonoidalRightAction.actionHom_def' {C : Type u_1} {D : Type u_2} [Category.{u_4, u_1} C] [Category.{u_3, u_2} D] [MonoidalCategory C] [MonoidalRightAction C D] {x₁ y₁ : D} {x₂ y₂ : C} (f : x₁ y₁) (g : x₂ y₂) :
                                                                                  theorem CategoryTheory.MonoidalCategory.MonoidalRightAction.actionHom_def'_assoc {C : Type u_1} {D : Type u_2} [Category.{u_4, u_1} C] [Category.{u_3, u_2} D] [MonoidalCategory C] [MonoidalRightAction C D] {x₁ y₁ : D} {x₂ y₂ : C} (f : x₁ y₁) (g : x₂ y₂) {Z : D} (h : actionObj y₁ y₂ Z) :
                                                                                  theorem CategoryTheory.MonoidalCategory.MonoidalRightAction.actionAssocIso_inv_naturality {C : Type u_1} {D : Type u_2} [Category.{u_4, u_1} C] [Category.{u_3, u_2} D] [MonoidalCategory C] [MonoidalRightAction C D] {d₁ d₂ : D} {c₁ c₂ c₃ c₄ : C} (f : d₁ d₂) (g : c₁ c₂) (h : c₃ c₄) :
                                                                                  theorem CategoryTheory.MonoidalCategory.MonoidalRightAction.actionAssocIso_inv_naturality_assoc {C : Type u_1} {D : Type u_2} [Category.{u_4, u_1} C] [Category.{u_3, u_2} D] [MonoidalCategory C] [MonoidalRightAction C D] {d₁ d₂ : D} {c₁ c₂ c₃ c₄ : C} (f : d₁ d₂) (g : c₁ c₂) (h : c₃ c₄) {Z : D} (h✝ : actionObj d₂ (tensorObj c₂ c₄) Z) :
                                                                                  @[simp]
                                                                                  theorem CategoryTheory.MonoidalCategory.MonoidalRightAction.inv_actionHom {C : Type u_1} {D : Type u_2} [Category.{u_4, u_1} C] [Category.{u_3, u_2} D] [MonoidalCategory C] [MonoidalRightAction C D] {x y : D} {x' y' : C} (f : x y) (g : x' y') [IsIso f] [IsIso g] :
                                                                                  inv (actionHom f g) = actionHom (inv f) (inv g)

                                                                                  Bundle the action of C on D as a functor C ⥤ D ⥤ D.

                                                                                  Equations
                                                                                    Instances For
                                                                                      @[reducible, inline]

                                                                                      Bundle d ↦ d ⊙ᵣ c as a functor.

                                                                                      Equations
                                                                                        Instances For
                                                                                          @[reducible, inline]

                                                                                          Bundle c ↦ d ⊙ᵣ c as a functor.

                                                                                          Equations
                                                                                            Instances For

                                                                                              Bundle ρᵣ _ as an isomorphism of functors.

                                                                                              Equations
                                                                                                Instances For