Documentation

Mathlib.LinearAlgebra.TensorPower.Symmetric

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 nth symmetric tensor power of M, which is the symmetric tensor power indexed by Fin n.

Main definitions: #

TODO: #

inductive SymmetricPower.Rel (R ι : Type u) [CommSemiring R] (M : Type v) [AddCommMonoid M] [Module R M] :
(PiTensorProduct R fun (x : ι) => M)(PiTensorProduct R fun (x : ι) => M)Prop

The relation on the ι-indexed tensor power of M where two tensors are equal if they are related by a permutation of ι.

Instances For
    noncomputable def SymmetricPower (R ι : Type u) [CommSemiring R] (M : Type v) [AddCommMonoid M] [Module R M] :
    Type (max u v)

    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
    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 nth 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
          theorem SymmetricPower.smul {R ι : Type u} [CommSemiring R] {M : Type v} [AddCommMonoid M] [Module R M] (r : R) (x y : PiTensorProduct R fun (x : ι) => M) (h : (addConGen (Rel R ι M)) x y) :
          (addConGen (Rel R ι M)) (r x) (r y)
          noncomputable def SymmetricPower.smul' {R : Type u} (ι : Type u) [CommSemiring R] (M : Type v) [AddCommMonoid M] [Module R M] (r : R) :

          Scalar multiplication by r : R. Use instead.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            noncomputable instance SymmetricPower.module (R ι : Type u) [CommSemiring R] (M : Type v) [AddCommMonoid M] [Module R M] :
            Equations
            • One or more equations did not get rendered due to their size.
            noncomputable def SymmetricPower.mk (R ι : Type u) [CommSemiring R] (M : Type v) [AddCommMonoid M] [Module R M] :
            (PiTensorProduct R fun (x : ι) => M) →ₗ[R] SymmetricPower R ι M

            The canonical map from the ι-indexed tensor power to the symmetric tensor power.

            Equations
            Instances For
              noncomputable def SymmetricPower.tprod (R : Type u) {ι : Type u} [CommSemiring R] {M : Type v} [AddCommMonoid M] [Module R M] :
              MultilinearMap R (fun (x : ι) => M) (SymmetricPower R ι M)

              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
                  @[simp]
                  theorem SymmetricPower.tprod_equiv {R ι : Type u} [CommSemiring R] {M : Type v} [AddCommMonoid M] [Module R M] (e : Equiv.Perm ι) (f : ιM) :
                  (⨂ₛ[R] (i : ι), f (e i)) = ⨂ₛ[R] (i : ι), f i
                  theorem SymmetricPower.range_mk (R ι : Type u) [CommSemiring R] (M : Type v) [AddCommMonoid M] [Module R M] :

                  The pure tensors (i.e. the elements of the image of SymmetricPower.tprod) span the symmetric tensor power.