Symmetric Algebras #
Given a commutative semiring R
, and an R
-module M
, we construct the symmetric algebra of M
.
This is the free commutative R
-algebra generated (R
-linearly) by the module M
.
Notation #
SymmetricAlgebra R M
: a concrete construction of the symmetric algebra defined as a quotient of the tensor algebra. It is endowed with an R-algebra structure and a commutative ring structure.SymmetricAlgebra.ι R
: the canonical R-linear mapM →ₗ[R] SymmetricAlgebra R M
.
Note #
See SymAlg R
instead if you are looking for the symmetrized algebra, which gives a commutative
multiplication on R
by $a \circ b = \frac{1}{2}(ab + ba)$.
Relation on the tensor algebra which will yield the symmetric algebra when quotiented out by.
- mul_comm {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] (x y : M) : SymRel R M ((ι R) x * (ι R) y) ((ι R) y * (ι R) x)
Instances For
Concrete construction of the symmetric algebra of M
by quotienting out
the tensor algebra by the commutativity relation.
Equations
- SymmetricAlgebra R M = RingQuot (TensorAlgebra.SymRel R M)
Instances For
Algebra homomorphism from the tensor algebra over M
to the symmetric algebra over M
.
Equations
Instances For
Canonical inclusion of M
into the symmetric algebra SymmetricAlgebra R M
.
Equations
- SymmetricAlgebra.ι R M = ↑(SymmetricAlgebra.algHom R M) ∘ₗ TensorAlgebra.ι R
Instances For
Equations
- SymmetricAlgebra.instCommSemiring R M = { toSemiring := RingQuot.instSemiring (TensorAlgebra.SymRel R M), mul_comm := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
For any linear map f : M →ₗ[R] A
, SymmetricAlgebra.lift f
lifts the linear map to an
R-algebra homomorphism from SymmetricAlgebra R M
to A
.
Equations
Instances For
The left-inverse of algebraMap
.
Equations
Instances For
A SymmetricAlgebra
over a nontrivial semiring is nontrivial.