Documentation

Mathlib.CategoryTheory.Monoidal.Bimod

The category of bimodule objects over a pair of monoid objects. #

theorem id_tensor_π_preserves_coequalizer_inv_colimMap_desc {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] [CategoryTheory.MonoidalCategory C] [CategoryTheory.Limits.HasCoequalizers C] [∀ (X : C), CategoryTheory.Limits.PreservesColimitsOfSize.{0, 0, v₁, v₁, u₁, u₁} (CategoryTheory.MonoidalCategory.tensorLeft X)] {X Y Z X' Y' Z' : C} (f g : X Y) (f' g' : X' Y') (p : CategoryTheory.MonoidalCategoryStruct.tensorObj Z X X') (q : CategoryTheory.MonoidalCategoryStruct.tensorObj Z Y Y') (wf : CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategoryStruct.whiskerLeft Z f) q = CategoryTheory.CategoryStruct.comp p f') (wg : CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategoryStruct.whiskerLeft Z g) q = CategoryTheory.CategoryStruct.comp p g') (h : Y' Z') (wh : CategoryTheory.CategoryStruct.comp f' h = CategoryTheory.CategoryStruct.comp g' h) :
theorem π_tensor_id_preserves_coequalizer_inv_colimMap_desc {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] [CategoryTheory.MonoidalCategory C] [CategoryTheory.Limits.HasCoequalizers C] [∀ (X : C), CategoryTheory.Limits.PreservesColimitsOfSize.{0, 0, v₁, v₁, u₁, u₁} (CategoryTheory.MonoidalCategory.tensorRight X)] {X Y Z X' Y' Z' : C} (f g : X Y) (f' g' : X' Y') (p : CategoryTheory.MonoidalCategoryStruct.tensorObj X Z X') (q : CategoryTheory.MonoidalCategoryStruct.tensorObj Y Z Y') (wf : CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategoryStruct.whiskerRight f Z) q = CategoryTheory.CategoryStruct.comp p f') (wg : CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategoryStruct.whiskerRight g Z) q = CategoryTheory.CategoryStruct.comp p g') (h : Y' Z') (wh : CategoryTheory.CategoryStruct.comp f' h = CategoryTheory.CategoryStruct.comp g' h) :
structure Bimod {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] [CategoryTheory.MonoidalCategory C] (A B : Mon_ C) :
Type (max u₁ v₁)

A bimodule object for a pair of monoid objects, all internal to some monoidal category.

structure Bimod.Hom {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] [CategoryTheory.MonoidalCategory C] {A B : Mon_ C} (M N : Bimod A B) :
Type v₁

A morphism of bimodule objects.

theorem Bimod.Hom.ext_iff {C : Type u₁} {inst✝ : CategoryTheory.Category.{v₁, u₁} C} {inst✝¹ : CategoryTheory.MonoidalCategory C} {A B : Mon_ C} {M N : Bimod A B} {x y : M.Hom N} :
x = y x.hom = y.hom
theorem Bimod.Hom.ext {C : Type u₁} {inst✝ : CategoryTheory.Category.{v₁, u₁} C} {inst✝¹ : CategoryTheory.MonoidalCategory C} {A B : Mon_ C} {M N : Bimod A B} {x y : M.Hom N} (hom : x.hom = y.hom) :
x = y

The identity morphism on a bimodule object.

Equations
Equations
def Bimod.comp {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] [CategoryTheory.MonoidalCategory C] {A B : Mon_ C} {M N O : Bimod A B} (f : M.Hom N) (g : N.Hom O) :
M.Hom O

Composition of bimodule object morphisms.

Equations
@[simp]
Equations
  • One or more equations did not get rendered due to their size.
theorem Bimod.hom_ext {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] [CategoryTheory.MonoidalCategory C] {A B : Mon_ C} {M N : Bimod A B} (f g : M N) (h : f.hom = g.hom) :
f = g
theorem Bimod.hom_ext_iff {C : Type u₁} [CategoryTheory.Category.{v₁, u₁} C] [CategoryTheory.MonoidalCategory C] {A B : Mon_ C} {M N : Bimod A B} {f g : M N} :
f = g f.hom = g.hom

Construct an isomorphism of bimodules by giving an isomorphism between the underlying objects and checking compatibility with left and right actions only in the forward direction.

Equations
  • One or more equations did not get rendered due to their size.

A monoid object as a bimodule over itself.

Equations
  • Bimod.regular A = { X := A.X, actLeft := A.mul, one_actLeft := , left_assoc := , actRight := A.mul, actRight_one := , right_assoc := , middle_assoc := }

The forgetful functor from bimodule objects to the ambient category.

Equations

The underlying object of the tensor product of two bimodules.

Equations
  • One or more equations did not get rendered due to their size.

Left action for the tensor product of two bimodules.

Equations
  • One or more equations did not get rendered due to their size.

Right action for the tensor product of two bimodules.

Equations
  • One or more equations did not get rendered due to their size.

Tensor product of two bimodule objects as a bimodule object.

Equations
  • One or more equations did not get rendered due to their size.

Left whiskering for morphisms of bimodule objects.

Equations
  • One or more equations did not get rendered due to their size.

Right whiskering for morphisms of bimodule objects.

Equations
  • One or more equations did not get rendered due to their size.

An auxiliary morphism for the definition of the underlying morphism of the forward component of the associator isomorphism.

Equations
  • One or more equations did not get rendered due to their size.

An auxiliary morphism for the definition of the underlying morphism of the inverse component of the associator isomorphism.

Equations
  • One or more equations did not get rendered due to their size.

The underlying morphism of the forward component of the left unitor isomorphism.

Equations

The underlying morphism of the inverse component of the left unitor isomorphism.

Equations
  • One or more equations did not get rendered due to their size.

The underlying morphism of the forward component of the right unitor isomorphism.

Equations

The underlying morphism of the inverse component of the right unitor isomorphism.

Equations
  • One or more equations did not get rendered due to their size.