Documentation

Mathlib.RingTheory.TensorProduct.Basic

The tensor product of R-algebras #

This file provides results about the multiplicative structure on A ⊗[R] B when R is a commutative (semi)ring and A and B are both R-algebras. On these tensor products, multiplication is characterized by (a₁ ⊗ₜ b₁) * (a₂ ⊗ₜ b₂) = (a₁ * a₂) ⊗ₜ (b₁ * b₂).

Main declarations #

References #

def LinearMap.liftBaseChangeEquiv {R : Type u_1} {M : Type u_2} {N : Type u_3} (A : Type u_4) [CommSemiring R] [CommSemiring A] [Algebra R A] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Module A N] [IsScalarTower R A N] :

If M is an R-module and N is an A-module, then A-linear maps A ⊗[R] M →ₗ[A] N correspond to R linear maps M →ₗ[R] N by composing with M → A ⊗ M, x ↦ 1 ⊗ x.

Equations
Instances For
    @[reducible, inline]
    abbrev LinearMap.liftBaseChange {R : Type u_1} {M : Type u_2} {N : Type u_3} (A : Type u_4) [CommSemiring R] [CommSemiring A] [Algebra R A] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Module A N] [IsScalarTower R A N] (l : M →ₗ[R] N) :

    If N is an A module, we may lift a linear map M →ₗ[R] N to A ⊗[R] M →ₗ[A] N

    Equations
    Instances For
      @[simp]
      theorem LinearMap.liftBaseChange_tmul {R : Type u_1} {M : Type u_2} {N : Type u_3} (A : Type u_4) [CommSemiring R] [CommSemiring A] [Algebra R A] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Module A N] [IsScalarTower R A N] (l : M →ₗ[R] N) (x : A) (y : M) :
      (liftBaseChange A l) (x ⊗ₜ[R] y) = x l y
      theorem LinearMap.liftBaseChange_one_tmul {R : Type u_1} {M : Type u_2} {N : Type u_3} (A : Type u_4) [CommSemiring R] [CommSemiring A] [Algebra R A] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Module A N] [IsScalarTower R A N] (l : M →ₗ[R] N) (y : M) :
      (liftBaseChange A l) (1 ⊗ₜ[R] y) = l y
      @[simp]
      theorem LinearMap.liftBaseChangeEquiv_symm_apply {R : Type u_4} {M : Type u_2} {N : Type u_3} (A : Type u_1) [CommSemiring R] [CommSemiring A] [Algebra R A] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Module A N] [IsScalarTower R A N] (l : TensorProduct R A M →ₗ[A] N) (x : M) :
      ((liftBaseChangeEquiv A).symm l) x = l (1 ⊗ₜ[R] x)
      theorem LinearMap.liftBaseChange_comp {R : Type u_3} {M : Type u_4} {N : Type u_5} (A : Type u_2) [CommSemiring R] [CommSemiring A] [Algebra R A] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Module A N] [IsScalarTower R A N] {P : Type u_1} [AddCommMonoid P] [Module A P] [Module R P] [IsScalarTower R A P] (l : M →ₗ[R] N) (l' : N →ₗ[A] P) :
      @[simp]
      theorem LinearMap.range_liftBaseChange {R : Type u_1} {M : Type u_2} {N : Type u_3} (A : Type u_4) [CommSemiring R] [CommSemiring A] [Algebra R A] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Module A N] [IsScalarTower R A N] (l : M →ₗ[R] N) :

      The R-algebra structure on A ⊗[R] B #

      Equations
      • One or more equations did not get rendered due to their size.
      theorem Algebra.TensorProduct.natCast_def {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [AddCommMonoidWithOne A] [Module R A] [AddCommMonoidWithOne B] [Module R B] (n : ) :
      n = n ⊗ₜ[R] 1
      theorem Algebra.TensorProduct.natCast_def' {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [AddCommMonoidWithOne A] [Module R A] [AddCommMonoidWithOne B] [Module R B] (n : ) :
      n = 1 ⊗ₜ[R] n
      @[irreducible]

      (Implementation detail) The multiplication map on A ⊗[R] B, as an R-bilinear map.

      Equations
      Instances For
        @[simp]
        theorem Algebra.TensorProduct.mul_apply {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] [NonUnitalNonAssocSemiring B] [Module R B] [SMulCommClass R B B] [IsScalarTower R B B] (a₁ a₂ : A) (b₁ b₂ : B) :
        (mul (a₁ ⊗ₜ[R] b₁)) (a₂ ⊗ₜ[R] b₂) = (a₁ * a₂) ⊗ₜ[R] (b₁ * b₂)
        @[simp]
        theorem Algebra.TensorProduct.tmul_mul_tmul {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] [NonUnitalNonAssocSemiring B] [Module R B] [SMulCommClass R B B] [IsScalarTower R B B] (a₁ a₂ : A) (b₁ b₂ : B) :
        a₁ ⊗ₜ[R] b₁ * a₂ ⊗ₜ[R] b₂ = (a₁ * a₂) ⊗ₜ[R] (b₁ * b₂)
        theorem SemiconjBy.tmul {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] [NonUnitalNonAssocSemiring B] [Module R B] [SMulCommClass R B B] [IsScalarTower R B B] {a₁ a₂ a₃ : A} {b₁ b₂ b₃ : B} (ha : SemiconjBy a₁ a₂ a₃) (hb : SemiconjBy b₁ b₂ b₃) :
        SemiconjBy (a₁ ⊗ₜ[R] b₁) (a₂ ⊗ₜ[R] b₂) (a₃ ⊗ₜ[R] b₃)
        theorem Commute.tmul {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] [NonUnitalNonAssocSemiring B] [Module R B] [SMulCommClass R B B] [IsScalarTower R B B] {a₁ a₂ : A} {b₁ b₂ : B} (ha : Commute a₁ a₂) (hb : Commute b₁ b₂) :
        Commute (a₁ ⊗ₜ[R] b₁) (a₂ ⊗ₜ[R] b₂)
        Equations
        • One or more equations did not get rendered due to their size.
        theorem Algebra.TensorProduct.one_mul {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [NonAssocSemiring A] [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] [NonAssocSemiring B] [Module R B] [SMulCommClass R B B] [IsScalarTower R B B] (x : TensorProduct R A B) :
        (mul (1 ⊗ₜ[R] 1)) x = x
        theorem Algebra.TensorProduct.mul_one {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [NonAssocSemiring A] [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] [NonAssocSemiring B] [Module R B] [SMulCommClass R B B] [IsScalarTower R B B] (x : TensorProduct R A B) :
        (mul x) (1 ⊗ₜ[R] 1) = x
        Equations
        • One or more equations did not get rendered due to their size.
        theorem Algebra.TensorProduct.mul_assoc {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [NonUnitalSemiring A] [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] [NonUnitalSemiring B] [Module R B] [SMulCommClass R B B] [IsScalarTower R B B] (x y z : TensorProduct R A B) :
        (mul ((mul x) y)) z = (mul x) ((mul y) z)
        instance Algebra.TensorProduct.instSemiring {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] :
        Equations
        • One or more equations did not get rendered due to their size.
        @[simp]
        theorem Algebra.TensorProduct.tmul_pow {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (a : A) (b : B) (k : ) :
        a ⊗ₜ[R] b ^ k = (a ^ k) ⊗ₜ[R] (b ^ k)

        The ring morphism A →+* A ⊗[R] B sending a to a ⊗ₜ 1.

        Equations
        Instances For
          @[simp]
          theorem Algebra.TensorProduct.includeLeftRingHom_apply {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (a : A) :
          instance Algebra.TensorProduct.leftAlgebra {R : Type uR} {S : Type uS} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CommSemiring S] [Algebra S A] [SMulCommClass R S A] :
          Equations
          instance Algebra.TensorProduct.instAlgebra {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] :

          The tensor product of two R-algebras is an R-algebra.

          Equations
          @[simp]
          theorem Algebra.TensorProduct.algebraMap_apply {R : Type uR} {S : Type uS} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CommSemiring S] [Algebra S A] [SMulCommClass R S A] (r : S) :
          (algebraMap S (TensorProduct R A B)) r = (algebraMap S A) r ⊗ₜ[R] 1
          theorem Algebra.TensorProduct.algebraMap_apply' {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (r : R) :
          (algebraMap R (TensorProduct R A B)) r = 1 ⊗ₜ[R] (algebraMap R B) r
          def Algebra.TensorProduct.includeLeft {R : Type uR} {S : Type uS} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CommSemiring S] [Algebra S A] [SMulCommClass R S A] :

          The R-algebra morphism A →ₐ[R] A ⊗[R] B sending a to a ⊗ₜ 1.

          Equations
          Instances For
            @[simp]
            theorem Algebra.TensorProduct.includeLeft_apply {R : Type uR} {S : Type uS} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CommSemiring S] [Algebra S A] [SMulCommClass R S A] (a : A) :
            def Algebra.TensorProduct.includeRight {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] :

            The algebra morphism B →ₐ[R] A ⊗[R] B sending b to 1 ⊗ₜ b.

            Equations
            Instances For
              @[simp]
              theorem Algebra.TensorProduct.includeRight_apply {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (b : B) :
              theorem Algebra.TensorProduct.ext {R : Type uR} {S : Type uS} {A : Type uA} {B : Type uB} {C : Type uC} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Semiring C] [Algebra R C] [CommSemiring S] [Algebra S A] [Algebra R S] [Algebra S C] [IsScalarTower R S A] [IsScalarTower R S C] f g : TensorProduct R A B →ₐ[S] C (ha : f.comp includeLeft = g.comp includeLeft) (hb : (AlgHom.restrictScalars R f).comp includeRight = (AlgHom.restrictScalars R g).comp includeRight) :
              f = g

              A version of TensorProduct.ext for AlgHom.

              Using this as the @[ext] lemma instead of Algebra.TensorProduct.ext' allows ext to apply lemmas specific to A →ₐ[S] _ and B →ₐ[R] _; notably this allows recursion into nested tensor products of algebras.

              See note [partially-applied ext lemmas].

              theorem Algebra.TensorProduct.ext' {R : Type uR} {S : Type uS} {A : Type uA} {B : Type uB} {C : Type uC} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Semiring C] [Algebra R C] [CommSemiring S] [Algebra S A] [Algebra R S] [Algebra S C] [IsScalarTower R S A] [IsScalarTower R S C] {g h : TensorProduct R A B →ₐ[S] C} (H : ∀ (a : A) (b : B), g (a ⊗ₜ[R] b) = h (a ⊗ₜ[R] b)) :
              g = h
              Equations
              • One or more equations did not get rendered due to their size.
              theorem Algebra.TensorProduct.intCast_def {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [AddCommGroupWithOne A] [Module R A] [AddCommMonoidWithOne B] [Module R B] (z : ) :
              z = z ⊗ₜ[R] 1
              Equations
              • One or more equations did not get rendered due to their size.
              Equations
              • One or more equations did not get rendered due to their size.
              Equations
              • One or more equations did not get rendered due to their size.
              instance Algebra.TensorProduct.instRing {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Ring A] [Algebra R A] [Semiring B] [Algebra R B] :
              Equations
              • One or more equations did not get rendered due to their size.
              theorem Algebra.TensorProduct.intCast_def' {R : Type uR} {A : Type uA} [CommSemiring R] [Ring A] [Algebra R A] {B : Type u_3} [Ring B] [Algebra R B] (z : ) :
              z = 1 ⊗ₜ[R] z
              instance Algebra.TensorProduct.instCommRing {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [CommRing A] [Algebra R A] [CommSemiring B] [Algebra R B] :
              Equations
              @[reducible, inline]
              abbrev Algebra.TensorProduct.rightAlgebra {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [CommSemiring B] [Algebra R B] :

              S ⊗[R] T has a T-algebra structure. This is not a global instance or else the action of S on S ⊗[R] S would be ambiguous.

              Equations
              Instances For
                theorem Algebra.TensorProduct.right_algebraMap_apply {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Algebra R A] [CommSemiring B] [Algebra R B] (b : B) :
                (algebraMap B (TensorProduct R A B)) b = 1 ⊗ₜ[R] b

                If M is a B-module that is also an A-module, the canonical map M →ₗ[A] B ⊗[A] M is injective.

                theorem Algebra.baseChange_lmul {R : Type u_1} {B : Type u_2} [CommSemiring R] [Semiring B] [Algebra R B] {A : Type u_3} [CommSemiring A] [Algebra R A] (f : B) :
                LinearMap.baseChange A ((lmul R B) f) = (lmul A (TensorProduct R A B)) (1 ⊗ₜ[R] f)
                def TensorProduct.Algebra.moduleAux {R : Type u_1} {A : Type u_2} {B : Type u_3} {M : Type u_4} [CommSemiring R] [AddCommMonoid M] [Module R M] [Semiring A] [Semiring B] [Module A M] [Module B M] [Algebra R A] [Algebra R B] [IsScalarTower R A M] [IsScalarTower R B M] :

                An auxiliary definition, used for constructing the Module (A ⊗[R] B) M in TensorProduct.Algebra.module below.

                Equations
                Instances For
                  theorem TensorProduct.Algebra.moduleAux_apply {R : Type u_1} {A : Type u_2} {B : Type u_3} {M : Type u_4} [CommSemiring R] [AddCommMonoid M] [Module R M] [Semiring A] [Semiring B] [Module A M] [Module B M] [Algebra R A] [Algebra R B] [IsScalarTower R A M] [IsScalarTower R B M] (a : A) (b : B) (m : M) :
                  (moduleAux (a ⊗ₜ[R] b)) m = a b m
                  def TensorProduct.Algebra.module {R : Type u_1} {A : Type u_2} {B : Type u_3} {M : Type u_4} [CommSemiring R] [AddCommMonoid M] [Module R M] [Semiring A] [Semiring B] [Module A M] [Module B M] [Algebra R A] [Algebra R B] [IsScalarTower R A M] [IsScalarTower R B M] [SMulCommClass A B M] :

                  If M is a representation of two different R-algebras A and B whose actions commute, then it is a representation the R-algebra A ⊗[R] B.

                  An important example arises from a semiring S; allowing S to act on itself via left and right multiplication, the roles of R, A, B, M are played by , S, Sᵐᵒᵖ, S. This example is important because a submodule of S as a Module over S ⊗[ℕ] Sᵐᵒᵖ is a two-sided ideal.

                  NB: This is not an instance because in the case B = A and M = A ⊗[R] A we would have a diamond of smul actions. Furthermore, this would not be a mere definitional diamond but a true mathematical diamond in which A ⊗[R] A had two distinct scalar actions on itself: one from its multiplication, and one from this would-be instance. Arguably we could live with this but in any case the real fix is to address the ambiguity in notation, probably along the lines outlined here: https://leanprover.zulipchat.com/#narrow/stream/144837-PR-reviews/topic/.234773.20base.20change/near/240929258

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    theorem TensorProduct.Algebra.smul_def {R : Type u_1} {A : Type u_2} {B : Type u_3} {M : Type u_4} [CommSemiring R] [AddCommMonoid M] [Module R M] [Semiring A] [Semiring B] [Module A M] [Module B M] [Algebra R A] [Algebra R B] [IsScalarTower R A M] [IsScalarTower R B M] [SMulCommClass A B M] (a : A) (b : B) (m : M) :
                    a ⊗ₜ[R] b m = a b m
                    theorem TensorProduct.mk_surjective (R : Type u_1) (M : Type u_2) (S : Type u_3) [CommSemiring R] [AddCommMonoid M] [Module R M] [Semiring S] [Algebra R S] (h : Function.Surjective (algebraMap R S)) :
                    Function.Surjective ((mk R S M) 1)
                    theorem TensorProduct.flip_mk_surjective {R : Type u_1} (S : Type u_3) {T : Type u_4} [CommSemiring R] [Semiring S] [Algebra R S] [Ring T] [Algebra R T] (h : Function.Surjective (algebraMap R T)) :
                    Function.Surjective ((mk R S T).flip 1)
                    theorem LinearMap.mulLeft_tmul {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [NonUnitalNonAssocSemiring A] [NonUnitalNonAssocSemiring B] [Module R A] [Module R B] [SMulCommClass R A A] [SMulCommClass R B B] [IsScalarTower R A A] [IsScalarTower R B B] (a : A) (b : B) :
                    theorem LinearMap.mulRight_tmul {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [NonUnitalNonAssocSemiring A] [NonUnitalNonAssocSemiring B] [Module R A] [Module R B] [SMulCommClass R A A] [SMulCommClass R B B] [IsScalarTower R A A] [IsScalarTower R B B] (a : A) (b : B) :
                    noncomputable instance TensorProduct.instStarMul {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [NonUnitalNonAssocSemiring A] [NonUnitalNonAssocSemiring B] [Module R A] [Module R B] [SMulCommClass R A A] [SMulCommClass R B B] [IsScalarTower R A A] [IsScalarTower R B B] [StarRing R] [StarRing A] [StarRing B] [StarModule R A] [StarModule R B] :
                    Equations
                    noncomputable instance TensorProduct.instStarRing {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [NonUnitalNonAssocSemiring A] [NonUnitalNonAssocSemiring B] [Module R A] [Module R B] [SMulCommClass R A A] [SMulCommClass R B B] [IsScalarTower R A A] [IsScalarTower R B B] [StarRing R] [StarRing A] [StarRing B] [StarModule R A] [StarModule R B] :
                    Equations