Documentation

Mathlib.RingTheory.MvPolynomial.Symmetric.Defs

Symmetric Polynomials and Elementary Symmetric Polynomials #

This file defines symmetric MvPolynomials and the bases of elementary, complete homogeneous, power sum, and monomial symmetric MvPolynomials. We also prove some basic facts about them.

Main declarations #

Notation #

As in other polynomial files, we typically use the notation:

def Multiset.esymm {R : Type u_1} [CommSemiring R] (s : Multiset R) (n : ) :
R

The nth elementary symmetric function evaluated at the elements of s

Equations
Instances For
    theorem Finset.esymm_map_val {R : Type u_1} [CommSemiring R] {σ : Type u_2} (f : σR) (s : Finset σ) (n : ) :
    (Multiset.map f s.val).esymm n = tFinset.powersetCard n s, t.prod f
    def MvPolynomial.IsSymmetric {σ : Type u_1} {R : Type u_3} [CommSemiring R] (φ : MvPolynomial σ R) :

    A MvPolynomial φ is symmetric if it is invariant under permutations of its variables by the rename operation

    Equations
    Instances For

      The subalgebra of symmetric MvPolynomials.

      Equations
      Instances For
        @[simp]
        theorem MvPolynomial.mem_symmetricSubalgebra {σ : Type u_1} {R : Type u_3} [CommSemiring R] (p : MvPolynomial σ R) :
        @[simp]
        theorem MvPolynomial.IsSymmetric.C {σ : Type u_1} {R : Type u_3} [CommSemiring R] (r : R) :
        (MvPolynomial.C r).IsSymmetric
        theorem MvPolynomial.IsSymmetric.add {σ : Type u_1} {R : Type u_3} [CommSemiring R] {φ : MvPolynomial σ R} {ψ : MvPolynomial σ R} (hφ : φ.IsSymmetric) (hψ : ψ.IsSymmetric) :
        (φ + ψ).IsSymmetric
        theorem MvPolynomial.IsSymmetric.mul {σ : Type u_1} {R : Type u_3} [CommSemiring R] {φ : MvPolynomial σ R} {ψ : MvPolynomial σ R} (hφ : φ.IsSymmetric) (hψ : ψ.IsSymmetric) :
        (φ * ψ).IsSymmetric
        theorem MvPolynomial.IsSymmetric.smul {σ : Type u_1} {R : Type u_3} [CommSemiring R] {φ : MvPolynomial σ R} (r : R) (hφ : φ.IsSymmetric) :
        (r φ).IsSymmetric
        @[simp]
        theorem MvPolynomial.IsSymmetric.map {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] {φ : MvPolynomial σ R} (hφ : φ.IsSymmetric) (f : R →+* S) :
        ((MvPolynomial.map f) φ).IsSymmetric
        theorem MvPolynomial.IsSymmetric.rename {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] {φ : MvPolynomial σ R} (hφ : φ.IsSymmetric) (e : σ τ) :
        ((MvPolynomial.rename e) φ).IsSymmetric
        @[simp]
        theorem MvPolynomial.isSymmetric_rename {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] {φ : MvPolynomial σ R} {e : σ τ} :
        ((MvPolynomial.rename e) φ).IsSymmetric φ.IsSymmetric
        theorem MvPolynomial.IsSymmetric.neg {σ : Type u_1} {R : Type u_3} [CommRing R] {φ : MvPolynomial σ R} (hφ : φ.IsSymmetric) :
        (-φ).IsSymmetric
        theorem MvPolynomial.IsSymmetric.sub {σ : Type u_1} {R : Type u_3} [CommRing R] {φ : MvPolynomial σ R} {ψ : MvPolynomial σ R} (hφ : φ.IsSymmetric) (hψ : ψ.IsSymmetric) :
        (φ - ψ).IsSymmetric
        @[simp]
        theorem MvPolynomial.renameSymmetricSubalgebra_symm_apply_coe {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (e : σ τ) (a : { x : MvPolynomial τ R // x MvPolynomial.symmetricSubalgebra τ R }) :
        @[simp]
        theorem MvPolynomial.renameSymmetricSubalgebra_apply_coe {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (e : σ τ) (a : { x : MvPolynomial σ R // x MvPolynomial.symmetricSubalgebra σ R }) :
        def MvPolynomial.renameSymmetricSubalgebra {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (e : σ τ) :

        MvPolynomial.rename induces an isomorphism between the symmetric subalgebras.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          def MvPolynomial.esymm (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] (n : ) :

          The nth elementary symmetric MvPolynomial σ R. It is the sum over all the degree n squarefree monomials in MvPolynomial σ R.

          Equations
          Instances For
            def MvPolynomial.esymmPart (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] {n : } (μ : n.Partition) :

            esymmPart is the product of the symmetric polynomials esymm μᵢ, where μ = (μ₁, μ₂, ...) is a partition.

            Equations
            Instances For
              theorem MvPolynomial.esymm_eq_multiset_esymm (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] :
              MvPolynomial.esymm σ R = (Multiset.map MvPolynomial.X Finset.univ.val).esymm

              The nth elementary symmetric MvPolynomial σ R is obtained by evaluating the nth elementary symmetric at the Multiset of the monomials

              theorem MvPolynomial.aeval_esymm_eq_multiset_esymm {S : Type u_4} (σ : Type u_5) (R : Type u_6) [CommSemiring R] [CommSemiring S] [Fintype σ] [Algebra R S] (n : ) (f : σS) :
              (MvPolynomial.aeval f) (MvPolynomial.esymm σ R n) = (Multiset.map f Finset.univ.val).esymm n
              theorem MvPolynomial.esymm_eq_sum_subtype (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] (n : ) :
              MvPolynomial.esymm σ R n = t : { s : Finset σ // s.card = n }, it, MvPolynomial.X i

              We can define esymm σ R n by summing over a subtype instead of over powerset_len.

              theorem MvPolynomial.esymm_eq_sum_monomial (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] (n : ) :
              MvPolynomial.esymm σ R n = tFinset.powersetCard n Finset.univ, (MvPolynomial.monomial (∑ it, Finsupp.single i 1)) 1

              We can define esymm σ R n as a sum over explicit monomials

              @[simp]
              theorem MvPolynomial.esymm_zero (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] :
              @[simp]
              theorem MvPolynomial.esymm_one (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] :
              MvPolynomial.esymm σ R 1 = i : σ, MvPolynomial.X i
              theorem MvPolynomial.map_esymm {S : Type u_4} (σ : Type u_5) (R : Type u_6) [CommSemiring R] [CommSemiring S] [Fintype σ] (n : ) (f : R →+* S) :
              theorem MvPolynomial.rename_esymm {τ : Type u_2} (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [Fintype τ] (n : ) (e : σ τ) :
              theorem MvPolynomial.esymm_isSymmetric (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] (n : ) :
              (MvPolynomial.esymm σ R n).IsSymmetric
              theorem MvPolynomial.support_esymm'' (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [DecidableEq σ] [Nontrivial R] (n : ) :
              (MvPolynomial.esymm σ R n).support = (Finset.powersetCard n Finset.univ).biUnion fun (t : Finset σ) => (Finsupp.single (∑ it, Finsupp.single i 1) 1).support
              theorem MvPolynomial.support_esymm' (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [DecidableEq σ] [Nontrivial R] (n : ) :
              (MvPolynomial.esymm σ R n).support = (Finset.powersetCard n Finset.univ).biUnion fun (t : Finset σ) => {it, Finsupp.single i 1}
              theorem MvPolynomial.support_esymm (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [DecidableEq σ] [Nontrivial R] (n : ) :
              (MvPolynomial.esymm σ R n).support = Finset.image (fun (t : Finset σ) => it, Finsupp.single i 1) (Finset.powersetCard n Finset.univ)
              theorem MvPolynomial.degrees_esymm (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [Nontrivial R] {n : } (hpos : 0 < n) (hn : n Fintype.card σ) :
              (MvPolynomial.esymm σ R n).degrees = Finset.univ.val
              def MvPolynomial.hsymm (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [DecidableEq σ] (n : ) :

              The nth complete homogeneous symmetric MvPolynomial σ R. It is the sum over all the degree n monomials in MvPolynomial σ R.

              Equations
              Instances For
                def MvPolynomial.hsymmPart (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [DecidableEq σ] {n : } (μ : n.Partition) :

                hsymmPart is the product of the symmetric polynomials hsymm μᵢ, where μ = (μ₁, μ₂, ...) is a partition.

                Equations
                Instances For
                  @[simp]
                  theorem MvPolynomial.hsymm_zero (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [DecidableEq σ] :
                  @[simp]
                  theorem MvPolynomial.hsymm_one (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [DecidableEq σ] :
                  MvPolynomial.hsymm σ R 1 = i : σ, MvPolynomial.X i
                  theorem MvPolynomial.map_hsymm {S : Type u_4} (σ : Type u_5) (R : Type u_6) [CommSemiring R] [CommSemiring S] [Fintype σ] [DecidableEq σ] (n : ) (f : R →+* S) :
                  theorem MvPolynomial.rename_hsymm {τ : Type u_2} (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [Fintype τ] [DecidableEq σ] [DecidableEq τ] (n : ) (e : σ τ) :
                  theorem MvPolynomial.hsymm_isSymmetric (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [DecidableEq σ] (n : ) :
                  (MvPolynomial.hsymm σ R n).IsSymmetric
                  def MvPolynomial.psum (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] (n : ) :

                  The degree-n power sum symmetric MvPolynomial σ R. It is the sum over all the n-th powers of the variables.

                  Equations
                  Instances For
                    def MvPolynomial.psumPart (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] {n : } (μ : n.Partition) :

                    psumPart is the product of the symmetric polynomials psum μᵢ, where μ = (μ₁, μ₂, ...) is a partition.

                    Equations
                    Instances For
                      @[simp]
                      theorem MvPolynomial.psum_zero (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] :
                      @[simp]
                      theorem MvPolynomial.psum_one (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] :
                      MvPolynomial.psum σ R 1 = i : σ, MvPolynomial.X i
                      @[simp]
                      @[simp]
                      theorem MvPolynomial.rename_psum {τ : Type u_2} (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [Fintype τ] (n : ) (e : σ τ) :
                      theorem MvPolynomial.psum_isSymmetric (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] (n : ) :
                      (MvPolynomial.psum σ R n).IsSymmetric
                      def MvPolynomial.msymm (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [DecidableEq σ] {n : } (μ : n.Partition) :

                      The monomial symmetric MvPolynomial σ R with exponent set μ. It is the sum over all the monomials in MvPolynomial σ R such that the multiset of exponents is equal to the multiset of parts of μ.

                      Equations
                      Instances For
                        @[simp]
                        @[simp]
                        theorem MvPolynomial.rename_msymm {τ : Type u_2} (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [Fintype τ] [DecidableEq σ] [DecidableEq τ] {n : } (μ : n.Partition) (e : σ τ) :
                        theorem MvPolynomial.msymm_isSymmetric (σ : Type u_5) (R : Type u_6) [CommSemiring R] [Fintype σ] [DecidableEq σ] {n : } (μ : n.Partition) :
                        (MvPolynomial.msymm σ R μ).IsSymmetric