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 #
Algebra.TensorProduct.semiring: the ring structure onA ⊗[R] Bfor twoR-algebrasA,B.Algebra.TensorProduct.leftAlgebra: theS-algebra structure onA ⊗[R] B, for whenAis additionally anSalgebra.
References #
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
- LinearMap.liftBaseChangeEquiv A = (LinearMap.ringLmapEquivSelf A A (M →ₗ[R] N)).symm.trans (TensorProduct.AlgebraTensorModule.lift.equiv R A A A M N)
Instances For
If N is an A module, we may lift a linear map M →ₗ[R] N to A ⊗[R] M →ₗ[A] N
Equations
Instances For
The R-algebra structure on A ⊗[R] B #
Equations
- One or more equations did not get rendered due to their size.
(Implementation detail)
The multiplication map on A ⊗[R] B,
as an R-bilinear map.
Equations
Instances For
Equations
- Algebra.TensorProduct.instMul = { mul := fun (a b : TensorProduct R A B) => (Algebra.TensorProduct.mul a) b }
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
- Algebra.TensorProduct.instNonUnitalSemiring = { toNonUnitalNonAssocSemiring := Algebra.TensorProduct.instNonUnitalNonAssocSemiring, mul_assoc := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
The ring morphism A →+* A ⊗[R] B sending a to a ⊗ₜ 1.
Equations
Instances For
Equations
- Algebra.TensorProduct.leftAlgebra = { toSMul := TensorProduct.leftHasSMul, algebraMap := Algebra.TensorProduct.includeLeftRingHom.comp (algebraMap S A), commutes' := ⋯, smul_def' := ⋯ }
The tensor product of two R-algebras is an R-algebra.
The R-algebra morphism A →ₐ[R] A ⊗[R] B sending a to a ⊗ₜ 1.
Equations
- Algebra.TensorProduct.includeLeft = { toRingHom := Algebra.TensorProduct.includeLeftRingHom, commutes' := ⋯ }
Instances For
The algebra morphism B →ₐ[R] A ⊗[R] B sending b to 1 ⊗ₜ b.
Equations
Instances For
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].
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.
Equations
- One or more equations did not get rendered due to their size.
Equations
- Algebra.TensorProduct.instCommSemiring = { toSemiring := inferInstance, mul_comm := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- Algebra.TensorProduct.instCommRing = { toRing := inferInstance, mul_comm := ⋯ }
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.
Instances For
If M is a B-module that is also an A-module, the canonical map
M →ₗ[A] B ⊗[A] M is injective.
An auxiliary definition, used for constructing the Module (A ⊗[R] B) M in
TensorProduct.Algebra.module below.
Equations
- TensorProduct.Algebra.moduleAux = TensorProduct.lift { toFun := fun (a : A) => a • (Algebra.lsmul R R M).toLinearMap, map_add' := ⋯, map_smul' := ⋯ }
Instances For
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
Equations
- TensorProduct.instStarMul = { toInvolutiveStar := TensorProduct.instInvolutiveStar, star_mul := ⋯ }
Equations
- TensorProduct.instStarRing = { toStarMul := TensorProduct.instStarMul, star_add := ⋯ }