Documentation

Mathlib.RingTheory.Coalgebra.TensorProduct

Tensor products of coalgebras #

Suppose S is an R-algebra. Given an S-coalgebra A and R-coalgebra B, we can define a natural comultiplication map Δ : A ⊗[R] B → (A ⊗[R] B) ⊗[S] (A ⊗[R] B) and counit map ε : A ⊗[R] B → S induced by the comultiplication and counit maps of A and B.

In this file we show that Δ, ε satisfy the axioms of a coalgebra, and also define other data in the monoidal structure on R-coalgebras, like the tensor product of two coalgebra morphisms as a coalgebra morphism.

In particular, when R = S we get tensor products of coalgebras, and when A = S we get the base change S ⊗[R] B as an S-coalgebra.

noncomputable instance TensorProduct.instCoalgebraStruct {R : Type u_1} {S : Type u_2} {A : Type u_3} {B : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid A] [AddCommMonoid B] [Algebra R S] [Module R A] [Module S A] [Module R B] [Coalgebra R B] [Coalgebra S A] [IsScalarTower R S A] :
Equations
  • One or more equations did not get rendered due to their size.
@[deprecated TensorProduct.comul_def (since := "2025-04-09")]

Alias of TensorProduct.comul_def.

@[deprecated TensorProduct.counit_def (since := "2025-04-09")]

Alias of TensorProduct.counit_def.

@[simp]
theorem TensorProduct.comul_tmul {R : Type u_1} {S : Type u_2} {A : Type u_3} {B : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid A] [AddCommMonoid B] [Algebra R S] [Module R A] [Module S A] [Module R B] [Coalgebra R B] [Coalgebra S A] [IsScalarTower R S A] (x : A) (y : B) :
@[simp]
theorem TensorProduct.counit_tmul {R : Type u_1} {S : Type u_2} {A : Type u_3} {B : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid A] [AddCommMonoid B] [Algebra R S] [Module R A] [Module S A] [Module R B] [Coalgebra R B] [Coalgebra S A] [IsScalarTower R S A] (x : A) (y : B) :

hopf_tensor_induction x with x₁ x₂ attempts to replace x by x₁ ⊗ₜ x₂ via linearity. This is an implementation detail that is used to set up tensor products of coalgebras, bialgebras, and hopf algebras, and shouldn't be relied on downstream.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    noncomputable instance TensorProduct.instCoalgebra {R : Type u_1} {S : Type u_2} {A : Type u_3} {B : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid A] [AddCommMonoid B] [Algebra R S] [Module R A] [Module S A] [Module R B] [Coalgebra R B] [Coalgebra S A] [IsScalarTower R S A] :
    Equations
    instance TensorProduct.instIsCocomm {R : Type u_1} {S : Type u_2} {A : Type u_3} {B : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid A] [AddCommMonoid B] [Algebra R S] [Module R A] [Module S A] [Module R B] [Coalgebra R B] [Coalgebra S A] [IsScalarTower R S A] [Coalgebra.IsCocomm S A] [Coalgebra.IsCocomm R B] :
    noncomputable def Coalgebra.TensorProduct.map {R : Type u_5} {S : Type u_6} {M : Type u_7} {N : Type u_8} {P : Type u_9} {Q : Type u_10} [CommSemiring R] [CommSemiring S] [Algebra R S] [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] [AddCommMonoid Q] [Module R M] [Module R N] [Module R P] [Module R Q] [Module S M] [IsScalarTower R S M] [Coalgebra S M] [Module S N] [IsScalarTower R S N] [Coalgebra S N] [Coalgebra R P] [Coalgebra R Q] (f : M →ₗc[S] N) (g : P →ₗc[R] Q) :

    The tensor product of two coalgebra morphisms as a coalgebra morphism.

    Equations
    Instances For
      @[simp]
      theorem Coalgebra.TensorProduct.map_tmul {R : Type u_5} {S : Type u_6} {M : Type u_7} {N : Type u_8} {P : Type u_9} {Q : Type u_10} [CommSemiring R] [CommSemiring S] [Algebra R S] [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] [AddCommMonoid Q] [Module R M] [Module R N] [Module R P] [Module R Q] [Module S M] [IsScalarTower R S M] [Coalgebra S M] [Module S N] [IsScalarTower R S N] [Coalgebra S N] [Coalgebra R P] [Coalgebra R Q] (f : M →ₗc[S] N) (g : P →ₗc[R] Q) (x : M) (y : P) :
      (map f g) (x ⊗ₜ[R] y) = f x ⊗ₜ[R] g y
      @[simp]
      theorem Coalgebra.TensorProduct.map_toLinearMap {R : Type u_5} {S : Type u_6} {M : Type u_7} {N : Type u_8} {P : Type u_9} {Q : Type u_10} [CommSemiring R] [CommSemiring S] [Algebra R S] [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] [AddCommMonoid Q] [Module R M] [Module R N] [Module R P] [Module R Q] [Module S M] [IsScalarTower R S M] [Coalgebra S M] [Module S N] [IsScalarTower R S N] [Coalgebra S N] [Coalgebra R P] [Coalgebra R Q] (f : M →ₗc[S] N) (g : P →ₗc[R] Q) :
      noncomputable def Coalgebra.TensorProduct.assoc (R : Type u_5) (S : Type u_6) (M : Type u_7) (N : Type u_8) (P : Type u_9) [CommSemiring R] [CommSemiring S] [Algebra R S] [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] [Module R M] [Module R N] [Module R P] [Module S M] [IsScalarTower R S M] [Coalgebra S M] [Module S N] [IsScalarTower R S N] [Coalgebra S N] [Coalgebra R P] :

      The associator for tensor products of R-coalgebras, as a coalgebra equivalence.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        theorem Coalgebra.TensorProduct.assoc_tmul {R : Type u_5} {S : Type u_6} {M : Type u_7} {N : Type u_8} {P : Type u_9} [CommSemiring R] [CommSemiring S] [Algebra R S] [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] [Module R M] [Module R N] [Module R P] [Module S M] [IsScalarTower R S M] [Coalgebra S M] [Module S N] [IsScalarTower R S N] [Coalgebra S N] [Coalgebra R P] (x : M) (y : N) (z : P) :
        (TensorProduct.assoc R S M N P) ((x ⊗ₜ[S] y) ⊗ₜ[R] z) = x ⊗ₜ[S] y ⊗ₜ[R] z
        @[simp]
        theorem Coalgebra.TensorProduct.assoc_symm_tmul {R : Type u_5} {S : Type u_6} {M : Type u_7} {N : Type u_8} {P : Type u_9} [CommSemiring R] [CommSemiring S] [Algebra R S] [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] [Module R M] [Module R N] [Module R P] [Module S M] [IsScalarTower R S M] [Coalgebra S M] [Module S N] [IsScalarTower R S N] [Coalgebra S N] [Coalgebra R P] (x : M) (y : N) (z : P) :
        (TensorProduct.assoc R S M N P).symm (x ⊗ₜ[S] y ⊗ₜ[R] z) = (x ⊗ₜ[S] y) ⊗ₜ[R] z
        @[simp]
        theorem Coalgebra.TensorProduct.assoc_toLinearEquiv {R : Type u_5} {S : Type u_6} {M : Type u_7} {N : Type u_8} {P : Type u_9} [CommSemiring R] [CommSemiring S] [Algebra R S] [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] [Module R M] [Module R N] [Module R P] [Module S M] [IsScalarTower R S M] [Coalgebra S M] [Module S N] [IsScalarTower R S N] [Coalgebra S N] [Coalgebra R P] :
        noncomputable def Coalgebra.TensorProduct.lid (R : Type u_5) (P : Type u_9) [CommSemiring R] [AddCommMonoid P] [Module R P] [Coalgebra R P] :

        The base ring is a left identity for the tensor product of coalgebras, up to coalgebra equivalence.

        Equations
        Instances For
          @[simp]
          theorem Coalgebra.TensorProduct.lid_tmul {R : Type u_5} {P : Type u_9} [CommSemiring R] [AddCommMonoid P] [Module R P] [Coalgebra R P] (r : R) (a : P) :
          (TensorProduct.lid R P) (r ⊗ₜ[R] a) = r a
          @[simp]
          theorem Coalgebra.TensorProduct.lid_symm_apply {R : Type u_5} {P : Type u_9} [CommSemiring R] [AddCommMonoid P] [Module R P] [Coalgebra R P] (a : P) :
          noncomputable def Coalgebra.TensorProduct.rid (R : Type u_5) (S : Type u_6) (M : Type u_7) [CommSemiring R] [CommSemiring S] [Algebra R S] [AddCommMonoid M] [Module R M] [Module S M] [IsScalarTower R S M] [Coalgebra S M] :

          The base ring is a right identity for the tensor product of coalgebras, up to coalgebra equivalence.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            @[simp]
            theorem Coalgebra.TensorProduct.rid_tmul {R : Type u_5} {S : Type u_6} {M : Type u_7} [CommSemiring R] [CommSemiring S] [Algebra R S] [AddCommMonoid M] [Module R M] [Module S M] [IsScalarTower R S M] [Coalgebra S M] (r : R) (a : M) :
            (TensorProduct.rid R S M) (a ⊗ₜ[R] r) = r a
            @[simp]
            theorem Coalgebra.TensorProduct.rid_symm_apply {R : Type u_5} {S : Type u_6} {M : Type u_7} [CommSemiring R] [CommSemiring S] [Algebra R S] [AddCommMonoid M] [Module R M] [Module S M] [IsScalarTower R S M] [Coalgebra S M] (a : M) :
            @[reducible, inline]
            noncomputable abbrev CoalgHom.lTensor {R : Type u_5} (M : Type u_6) {N : Type u_7} {P : Type u_8} [CommRing R] [AddCommGroup M] [AddCommGroup N] [AddCommGroup P] [Module R M] [Module R N] [Module R P] [Coalgebra R M] [Coalgebra R N] [Coalgebra R P] (f : N →ₗc[R] P) :

            lTensor M f : M ⊗ N →ₗc M ⊗ P is the natural coalgebra morphism induced by f : N →ₗc P.

            Equations
            Instances For
              @[reducible, inline]
              noncomputable abbrev CoalgHom.rTensor {R : Type u_5} (M : Type u_6) {N : Type u_7} {P : Type u_8} [CommRing R] [AddCommGroup M] [AddCommGroup N] [AddCommGroup P] [Module R M] [Module R N] [Module R P] [Coalgebra R M] [Coalgebra R N] [Coalgebra R P] (f : N →ₗc[R] P) :

              rTensor M f : N ⊗ M →ₗc P ⊗ M is the natural coalgebra morphism induced by f : N →ₗc P.

              Equations
              Instances For