Symmetric tensor power of a semimodule over a commutative semiring #
We define the ι
-indexed symmetric tensor power of M
as the PiTensorProduct
quotiented by
the relation that the tprod
of ι
elements is equal to the tprod
of the same elements permuted
by a permutation of ι
. We denote this space by Sym[R] ι M
, and the canonical multilinear map
from ι → M
to Sym[R] ι M
by ⨂ₛ[R] i, f i
. We also reserve the notation Sym[R]^n M
for the
n
th symmetric tensor power of M
, which is the symmetric tensor power indexed by Fin n
.
Main definitions: #
SymmetricPower.module
: the symmetric tensor power is a module overR
.
TODO: #
- Grading: show that there is a map
Sym[R]^i M × Sym[R]^j M → Sym[R]^(i + j) M
that is associative and commutative, and thatn ↦ Sym[R]^n M
is a graded (semi)ring and algebra. - Universal property: linear maps from
Sym[R]^n M
toN
correspond to symmetric multilinear mapsM ^ n
toN
. - Relate to homogneous (multivariate) polynomials of degree
n
.
The relation on the ι
-indexed tensor power of M
where two tensors are equal
if they are related by a permutation of ι
.
- perm {R ι : Type u} [CommSemiring R] {M : Type v} [AddCommMonoid M] [Module R M] (e : Equiv.Perm ι) (f : ι → M) : Rel R ι M ((PiTensorProduct.tprod R) fun (i : ι) => f i) ((PiTensorProduct.tprod R) fun (i : ι) => f (e i))
Instances For
The ι
-indexed symmetric tensor power of a semimodule M
over a commutative semiring R
is the quotient of the ι
-indexed tensor power of M
by the relation that two tensors are equal
if they are related by a permutation of ι
.
Equations
- SymmetricPower R ι M = (addConGen (SymmetricPower.Rel R ι M)).Quotient
Instances For
The ι
-indexed symmetric tensor power of a semimodule M
over a commutative semiring R
is the quotient of the ι
-indexed tensor power of M
by the relation that two tensors are equal
if they are related by a permutation of ι
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The n
th symmetric tensor power of a semimodule M
over a commutative semiring R
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- SymmetricPower.instAddCommMonoid R ι M = (addConGen (SymmetricPower.Rel R ι M)).addCommMonoid
Equations
- SymmetricPower.instAddCommGroup ι R M = (addConGen (SymmetricPower.Rel R ι M)).addCommGroup
Scalar multiplication by r : R
. Use •
instead.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
The canonical map from the ι
-indexed tensor power to the symmetric tensor power.
Equations
- SymmetricPower.mk R ι M = { toFun := (↑(addConGen (SymmetricPower.Rel R ι M)).mk').toFun, map_add' := ⋯, map_smul' := ⋯ }
Instances For
The multilinear map that takes ι
-indexed elements of M
and
returns their symmetric tensor power. Denoted ⨂ₛ[R] i, f i
.
Equations
Instances For
The multilinear map that takes ι
-indexed elements of M
and
returns their symmetric tensor power. Denoted ⨂ₛ[R] i, f i
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The pure tensors (i.e. the elements of the image of SymmetricPower.tprod
) span the symmetric
tensor power.