Documentation

Mathlib.Algebra.MvPolynomial.Basic

Multivariate polynomials #

This file defines polynomial rings over a base ring (or even semiring), with variables from a general type σ (which could be infinite).

Important definitions #

Let R be a commutative ring (or a semiring) and let σ be an arbitrary type. This file creates the type MvPolynomial σ R, which mathematicians might denote $R[X_i : i \in σ]$. It is the type of multivariate (a.k.a. multivariable) polynomials, with variables corresponding to the terms in σ, and coefficients in R.

Notation #

In the definitions below, we use the following notation:

Definitions #

Implementation notes #

Recall that if Y has a zero, then X →₀ Y is the type of functions from X to Y with finite support, i.e. such that only finitely many elements of X get sent to non-zero terms in Y. The definition of MvPolynomial σ R is (σ →₀ ℕ) →₀ R; here σ →₀ ℕ denotes the space of all monomials in the variables, and the function to R sends a monomial to its coefficient in the polynomial being represented.

Tags #

polynomial, multivariate polynomial, multivariable polynomial

def MvPolynomial (σ : Type u_1) (R : Type u_2) [CommSemiring R] :
Type (max u_1 u_2)

Multivariate polynomial, where σ is the index set of the variables and R is the coefficient ring

Equations
Instances For
    instance MvPolynomial.inhabited {R : Type u} {σ : Type u_1} [CommSemiring R] :
    Equations
    instance MvPolynomial.faithfulSMul {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring S₁] [SMulZeroClass R S₁] [FaithfulSMul R S₁] :
    instance MvPolynomial.module {R : Type u} {S₁ : Type v} {σ : Type u_1} [Semiring R] [CommSemiring S₁] [Module R S₁] :
    Module R (MvPolynomial σ S₁)
    Equations
    instance MvPolynomial.isScalarTower {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring S₂] [SMul R S₁] [SMulZeroClass R S₂] [SMulZeroClass S₁ S₂] [IsScalarTower R S₁ S₂] :
    IsScalarTower R S₁ (MvPolynomial σ S₂)
    instance MvPolynomial.smulCommClass {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring S₂] [SMulZeroClass R S₂] [SMulZeroClass S₁ S₂] [SMulCommClass R S₁ S₂] :
    SMulCommClass R S₁ (MvPolynomial σ S₂)
    instance MvPolynomial.isCentralScalar {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring S₁] [SMulZeroClass R S₁] [SMulZeroClass Rᵐᵒᵖ S₁] [IsCentralScalar R S₁] :
    instance MvPolynomial.algebra {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] :
    Algebra R (MvPolynomial σ S₁)
    Equations
    instance MvPolynomial.isScalarTower_right {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring S₁] [DistribSMul R S₁] [IsScalarTower R S₁ S₁] :
    instance MvPolynomial.smulCommClass_right {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring S₁] [DistribSMul R S₁] [SMulCommClass R S₁ S₁] :
    instance MvPolynomial.unique {R : Type u} {σ : Type u_1} [CommSemiring R] [Subsingleton R] :

    If R is a subsingleton, then MvPolynomial σ R has a unique element

    Equations
    def MvPolynomial.monomial {R : Type u} {σ : Type u_1} [CommSemiring R] (s : σ →₀ ) :

    monomial s a is the monomial with coefficient a and exponents given by s

    Equations
    Instances For
      theorem MvPolynomial.mul_def {R : Type u} {σ : Type u_1} [CommSemiring R] {p q : MvPolynomial σ R} :
      p * q = Finsupp.sum p fun (m : σ →₀ ) (a : R) => Finsupp.sum q fun (n : σ →₀ ) (b : R) => (MvPolynomial.monomial (m + n)) (a * b)
      def MvPolynomial.C {R : Type u} {σ : Type u_1} [CommSemiring R] :

      C a is the constant polynomial with value a

      Equations
      Instances For
        def MvPolynomial.X {R : Type u} {σ : Type u_1} [CommSemiring R] (n : σ) :

        X n is the degree 1 monomial $X_n$.

        Equations
        Instances For
          theorem MvPolynomial.monomial_left_injective {R : Type u} {σ : Type u_1} [CommSemiring R] {r : R} (hr : r 0) :
          @[simp]
          theorem MvPolynomial.monomial_left_inj {R : Type u} {σ : Type u_1} [CommSemiring R] {s t : σ →₀ } {r : R} (hr : r 0) :
          theorem MvPolynomial.C_apply {R : Type u} {σ : Type u_1} {a : R} [CommSemiring R] :
          MvPolynomial.C a = (MvPolynomial.monomial 0) a
          @[simp]
          theorem MvPolynomial.C_0 {R : Type u} {σ : Type u_1} [CommSemiring R] :
          MvPolynomial.C 0 = 0
          @[simp]
          theorem MvPolynomial.C_1 {R : Type u} {σ : Type u_1} [CommSemiring R] :
          MvPolynomial.C 1 = 1
          theorem MvPolynomial.C_mul_monomial {R : Type u} {σ : Type u_1} {a a' : R} {s : σ →₀ } [CommSemiring R] :
          MvPolynomial.C a * (MvPolynomial.monomial s) a' = (MvPolynomial.monomial s) (a * a')
          @[simp]
          theorem MvPolynomial.C_add {R : Type u} {σ : Type u_1} {a a' : R} [CommSemiring R] :
          MvPolynomial.C (a + a') = MvPolynomial.C a + MvPolynomial.C a'
          @[simp]
          theorem MvPolynomial.C_mul {R : Type u} {σ : Type u_1} {a a' : R} [CommSemiring R] :
          MvPolynomial.C (a * a') = MvPolynomial.C a * MvPolynomial.C a'
          @[simp]
          theorem MvPolynomial.C_pow {R : Type u} {σ : Type u_1} [CommSemiring R] (a : R) (n : ) :
          MvPolynomial.C (a ^ n) = MvPolynomial.C a ^ n
          @[simp]
          theorem MvPolynomial.C_inj {σ : Type u_2} (R : Type u_3) [CommSemiring R] (r s : R) :
          MvPolynomial.C r = MvPolynomial.C s r = s
          @[simp]
          theorem MvPolynomial.C_eq_zero {R : Type u} {σ : Type u_1} {a : R} [CommSemiring R] :
          MvPolynomial.C a = 0 a = 0
          theorem MvPolynomial.C_ne_zero {R : Type u} {σ : Type u_1} {a : R} [CommSemiring R] :
          MvPolynomial.C a 0 a 0
          theorem MvPolynomial.C_eq_coe_nat {R : Type u} {σ : Type u_1} [CommSemiring R] (n : ) :
          MvPolynomial.C n = n
          theorem MvPolynomial.C_mul' {R : Type u} {σ : Type u_1} {a : R} [CommSemiring R] {p : MvPolynomial σ R} :
          MvPolynomial.C a * p = a p
          theorem MvPolynomial.smul_eq_C_mul {R : Type u} {σ : Type u_1} [CommSemiring R] (p : MvPolynomial σ R) (a : R) :
          a p = MvPolynomial.C a * p
          theorem MvPolynomial.C_eq_smul_one {R : Type u} {σ : Type u_1} {a : R} [CommSemiring R] :
          MvPolynomial.C a = a 1
          theorem MvPolynomial.smul_monomial {R : Type u} {σ : Type u_1} {a : R} {s : σ →₀ } [CommSemiring R] {S₁ : Type u_2} [SMulZeroClass S₁ R] (r : S₁) :
          @[simp]
          theorem MvPolynomial.X_inj {R : Type u} {σ : Type u_1} [CommSemiring R] [Nontrivial R] (m n : σ) :
          theorem MvPolynomial.monomial_pow {R : Type u} {σ : Type u_1} {a : R} {e : } {s : σ →₀ } [CommSemiring R] :
          @[simp]
          theorem MvPolynomial.monomial_mul {R : Type u} {σ : Type u_1} [CommSemiring R] {s s' : σ →₀ } {a b : R} :

          fun s ↦ monomial s 1 as a homomorphism.

          Equations
          Instances For
            theorem MvPolynomial.monomial_add_single {R : Type u} {σ : Type u_1} {a : R} {e : } {n : σ} {s : σ →₀ } [CommSemiring R] :
            theorem MvPolynomial.monomial_single_add {R : Type u} {σ : Type u_1} {a : R} {e : } {n : σ} {s : σ →₀ } [CommSemiring R] :
            theorem MvPolynomial.C_mul_X_pow_eq_monomial {R : Type u} {σ : Type u_1} [CommSemiring R] {s : σ} {a : R} {n : } :
            MvPolynomial.C a * MvPolynomial.X s ^ n = (MvPolynomial.monomial (Finsupp.single s n)) a
            theorem MvPolynomial.C_mul_X_eq_monomial {R : Type u} {σ : Type u_1} [CommSemiring R] {s : σ} {a : R} :
            @[simp]
            theorem MvPolynomial.monomial_zero {R : Type u} {σ : Type u_1} [CommSemiring R] {s : σ →₀ } :
            @[simp]
            theorem MvPolynomial.monomial_eq_zero {R : Type u} {σ : Type u_1} [CommSemiring R] {s : σ →₀ } {b : R} :
            @[simp]
            theorem MvPolynomial.sum_monomial_eq {R : Type u} {σ : Type u_1} [CommSemiring R] {A : Type u_2} [AddCommMonoid A] {u : σ →₀ } {r : R} {b : (σ →₀ )RA} (w : b u 0 = 0) :
            @[simp]
            theorem MvPolynomial.sum_C {R : Type u} {σ : Type u_1} {a : R} [CommSemiring R] {A : Type u_2} [AddCommMonoid A] {b : (σ →₀ )RA} (w : b 0 0 = 0) :
            Finsupp.sum (MvPolynomial.C a) b = b 0 a
            theorem MvPolynomial.monomial_sum_one {R : Type u} {σ : Type u_1} [CommSemiring R] {α : Type u_2} (s : Finset α) (f : ασ →₀ ) :
            (MvPolynomial.monomial (∑ is, f i)) 1 = is, (MvPolynomial.monomial (f i)) 1
            theorem MvPolynomial.monomial_sum_index {R : Type u} {σ : Type u_1} [CommSemiring R] {α : Type u_2} (s : Finset α) (f : ασ →₀ ) (a : R) :
            (MvPolynomial.monomial (∑ is, f i)) a = MvPolynomial.C a * is, (MvPolynomial.monomial (f i)) 1
            theorem MvPolynomial.monomial_finsupp_sum_index {R : Type u} {σ : Type u_1} [CommSemiring R] {α : Type u_2} {β : Type u_3} [Zero β] (f : α →₀ β) (g : αβσ →₀ ) (a : R) :
            (MvPolynomial.monomial (f.sum g)) a = MvPolynomial.C a * f.prod fun (a : α) (b : β) => (MvPolynomial.monomial (g a b)) 1
            theorem MvPolynomial.monomial_eq_monomial_iff {R : Type u} [CommSemiring R] {α : Type u_2} (a₁ a₂ : α →₀ ) (b₁ b₂ : R) :
            (MvPolynomial.monomial a₁) b₁ = (MvPolynomial.monomial a₂) b₂ a₁ = a₂ b₁ = b₂ b₁ = 0 b₂ = 0
            theorem MvPolynomial.monomial_eq {R : Type u} {σ : Type u_1} {a : R} {s : σ →₀ } [CommSemiring R] :
            (MvPolynomial.monomial s) a = MvPolynomial.C a * s.prod fun (n : σ) (e : ) => MvPolynomial.X n ^ e
            @[simp]
            theorem MvPolynomial.prod_X_pow_eq_monomial {R : Type u} {σ : Type u_1} {s : σ →₀ } [CommSemiring R] :
            xs.support, MvPolynomial.X x ^ s x = (MvPolynomial.monomial s) 1
            theorem MvPolynomial.induction_on_monomial {R : Type u} {σ : Type u_1} [CommSemiring R] {M : MvPolynomial σ RProp} (h_C : ∀ (a : R), M (MvPolynomial.C a)) (h_X : ∀ (p : MvPolynomial σ R) (n : σ), M pM (p * MvPolynomial.X n)) (s : σ →₀ ) (a : R) :
            theorem MvPolynomial.induction_on' {R : Type u} {σ : Type u_1} [CommSemiring R] {P : MvPolynomial σ RProp} (p : MvPolynomial σ R) (h1 : ∀ (u : σ →₀ ) (a : R), P ((MvPolynomial.monomial u) a)) (h2 : ∀ (p q : MvPolynomial σ R), P pP qP (p + q)) :
            P p

            Analog of Polynomial.induction_on'. To prove something about mv_polynomials, it suffices to show the condition is closed under taking sums, and it holds for monomials.

            theorem MvPolynomial.induction_on''' {R : Type u} {σ : Type u_1} [CommSemiring R] {M : MvPolynomial σ RProp} (p : MvPolynomial σ R) (h_C : ∀ (a : R), M (MvPolynomial.C a)) (h_add_weak : ∀ (a : σ →₀ ) (b : R) (f : (σ →₀ ) →₀ R), af.supportb 0M fM ((let_fun this := (MvPolynomial.monomial a) b; this) + f)) :
            M p

            Similar to MvPolynomial.induction_on but only a weak form of h_add is required.

            theorem MvPolynomial.induction_on'' {R : Type u} {σ : Type u_1} [CommSemiring R] {M : MvPolynomial σ RProp} (p : MvPolynomial σ R) (h_C : ∀ (a : R), M (MvPolynomial.C a)) (h_add_weak : ∀ (a : σ →₀ ) (b : R) (f : (σ →₀ ) →₀ R), af.supportb 0M fM ((MvPolynomial.monomial a) b)M ((let_fun this := (MvPolynomial.monomial a) b; this) + f)) (h_X : ∀ (p : MvPolynomial σ R) (n : σ), M pM (p * MvPolynomial.X n)) :
            M p

            Similar to MvPolynomial.induction_on but only a yet weaker form of h_add is required.

            theorem MvPolynomial.induction_on {R : Type u} {σ : Type u_1} [CommSemiring R] {M : MvPolynomial σ RProp} (p : MvPolynomial σ R) (h_C : ∀ (a : R), M (MvPolynomial.C a)) (h_add : ∀ (p q : MvPolynomial σ R), M pM qM (p + q)) (h_X : ∀ (p : MvPolynomial σ R) (n : σ), M pM (p * MvPolynomial.X n)) :
            M p

            Analog of Polynomial.induction_on.

            theorem MvPolynomial.ringHom_ext {R : Type u} {σ : Type u_1} [CommSemiring R] {A : Type u_2} [Semiring A] {f g : MvPolynomial σ R →+* A} (hC : ∀ (r : R), f (MvPolynomial.C r) = g (MvPolynomial.C r)) (hX : ∀ (i : σ), f (MvPolynomial.X i) = g (MvPolynomial.X i)) :
            f = g
            theorem MvPolynomial.ringHom_ext' {R : Type u} {σ : Type u_1} [CommSemiring R] {A : Type u_2} [Semiring A] {f g : MvPolynomial σ R →+* A} (hC : f.comp MvPolynomial.C = g.comp MvPolynomial.C) (hX : ∀ (i : σ), f (MvPolynomial.X i) = g (MvPolynomial.X i)) :
            f = g

            See note [partially-applied ext lemmas].

            theorem MvPolynomial.hom_eq_hom {R : Type u} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [Semiring S₂] (f g : MvPolynomial σ R →+* S₂) (hC : f.comp MvPolynomial.C = g.comp MvPolynomial.C) (hX : ∀ (n : σ), f (MvPolynomial.X n) = g (MvPolynomial.X n)) (p : MvPolynomial σ R) :
            f p = g p
            theorem MvPolynomial.is_id {R : Type u} {σ : Type u_1} [CommSemiring R] (f : MvPolynomial σ R →+* MvPolynomial σ R) (hC : f.comp MvPolynomial.C = MvPolynomial.C) (hX : ∀ (n : σ), f (MvPolynomial.X n) = MvPolynomial.X n) (p : MvPolynomial σ R) :
            f p = p
            theorem MvPolynomial.algHom_ext' {R : Type u} {σ : Type u_1} [CommSemiring R] {A : Type u_2} {B : Type u_3} [CommSemiring A] [CommSemiring B] [Algebra R A] [Algebra R B] {f g : MvPolynomial σ A →ₐ[R] B} (h₁ : f.comp (IsScalarTower.toAlgHom R A (MvPolynomial σ A)) = g.comp (IsScalarTower.toAlgHom R A (MvPolynomial σ A))) (h₂ : ∀ (i : σ), f (MvPolynomial.X i) = g (MvPolynomial.X i)) :
            f = g
            theorem MvPolynomial.algHom_ext {R : Type u} {σ : Type u_1} [CommSemiring R] {A : Type u_2} [Semiring A] [Algebra R A] {f g : MvPolynomial σ R →ₐ[R] A} (hf : ∀ (i : σ), f (MvPolynomial.X i) = g (MvPolynomial.X i)) :
            f = g
            @[simp]
            theorem MvPolynomial.algHom_C {R : Type u} {σ : Type u_1} [CommSemiring R] {A : Type u_2} [Semiring A] [Algebra R A] (f : MvPolynomial σ R →ₐ[R] A) (r : R) :
            f (MvPolynomial.C r) = (algebraMap R A) r
            theorem MvPolynomial.linearMap_ext {R : Type u} {σ : Type u_1} [CommSemiring R] {M : Type u_2} [AddCommMonoid M] [Module R M] {f g : MvPolynomial σ R →ₗ[R] M} (h : ∀ (s : σ →₀ ), f ∘ₗ MvPolynomial.monomial s = g ∘ₗ MvPolynomial.monomial s) :
            f = g
            def MvPolynomial.support {R : Type u} {σ : Type u_1} [CommSemiring R] (p : MvPolynomial σ R) :

            The finite set of all m : σ →₀ ℕ such that X^m has a non-zero coefficient.

            Equations
            • p.support = p.support
            Instances For
              theorem MvPolynomial.finsupp_support_eq_support {R : Type u} {σ : Type u_1} [CommSemiring R] (p : MvPolynomial σ R) :
              p.support = p.support
              theorem MvPolynomial.support_monomial {R : Type u} {σ : Type u_1} {a : R} {s : σ →₀ } [CommSemiring R] [h : Decidable (a = 0)] :
              ((MvPolynomial.monomial s) a).support = if a = 0 then else {s}
              theorem MvPolynomial.support_monomial_subset {R : Type u} {σ : Type u_1} {a : R} {s : σ →₀ } [CommSemiring R] :
              ((MvPolynomial.monomial s) a).support {s}
              theorem MvPolynomial.support_add {R : Type u} {σ : Type u_1} [CommSemiring R] {p q : MvPolynomial σ R} [DecidableEq σ] :
              (p + q).support p.support q.support
              theorem MvPolynomial.support_X {R : Type u} {σ : Type u_1} {n : σ} [CommSemiring R] [Nontrivial R] :
              (MvPolynomial.X n).support = {Finsupp.single n 1}
              theorem MvPolynomial.support_X_pow {R : Type u} {σ : Type u_1} [CommSemiring R] [Nontrivial R] (s : σ) (n : ) :
              (MvPolynomial.X s ^ n).support = {Finsupp.single s n}
              theorem MvPolynomial.support_smul {R : Type u} {σ : Type u_1} [CommSemiring R] {S₁ : Type u_2} [SMulZeroClass S₁ R] {a : S₁} {f : MvPolynomial σ R} :
              (a f).support f.support
              theorem MvPolynomial.support_sum {R : Type u} {σ : Type u_1} [CommSemiring R] {α : Type u_2} [DecidableEq σ] {s : Finset α} {f : αMvPolynomial σ R} :
              (∑ xs, f x).support s.biUnion fun (x : α) => (f x).support
              def MvPolynomial.coeff {R : Type u} {σ : Type u_1} [CommSemiring R] (m : σ →₀ ) (p : MvPolynomial σ R) :
              R

              The coefficient of the monomial m in the multi-variable polynomial p.

              Equations
              Instances For
                @[simp]
                theorem MvPolynomial.mem_support_iff {R : Type u} {σ : Type u_1} [CommSemiring R] {p : MvPolynomial σ R} {m : σ →₀ } :
                m p.support MvPolynomial.coeff m p 0
                theorem MvPolynomial.not_mem_support_iff {R : Type u} {σ : Type u_1} [CommSemiring R] {p : MvPolynomial σ R} {m : σ →₀ } :
                mp.support MvPolynomial.coeff m p = 0
                theorem MvPolynomial.sum_def {R : Type u} {σ : Type u_1} [CommSemiring R] {A : Type u_2} [AddCommMonoid A] {p : MvPolynomial σ R} {b : (σ →₀ )RA} :
                Finsupp.sum p b = mp.support, b m (MvPolynomial.coeff m p)
                theorem MvPolynomial.support_mul {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (p q : MvPolynomial σ R) :
                (p * q).support p.support + q.support
                theorem MvPolynomial.ext {R : Type u} {σ : Type u_1} [CommSemiring R] (p q : MvPolynomial σ R) :
                (∀ (m : σ →₀ ), MvPolynomial.coeff m p = MvPolynomial.coeff m q)p = q
                @[simp]
                theorem MvPolynomial.coeff_add {R : Type u} {σ : Type u_1} [CommSemiring R] (m : σ →₀ ) (p q : MvPolynomial σ R) :
                @[simp]
                theorem MvPolynomial.coeff_smul {R : Type u} {σ : Type u_1} [CommSemiring R] {S₁ : Type u_2} [SMulZeroClass S₁ R] (m : σ →₀ ) (C : S₁) (p : MvPolynomial σ R) :
                @[simp]
                theorem MvPolynomial.coeff_zero {R : Type u} {σ : Type u_1} [CommSemiring R] (m : σ →₀ ) :
                @[simp]
                theorem MvPolynomial.coeff_zero_X {R : Type u} {σ : Type u_1} [CommSemiring R] (i : σ) :
                def MvPolynomial.coeffAddMonoidHom {R : Type u} {σ : Type u_1} [CommSemiring R] (m : σ →₀ ) :

                MvPolynomial.coeff m but promoted to an AddMonoidHom.

                Equations
                Instances For
                  def MvPolynomial.lcoeff (R : Type u) {σ : Type u_1} [CommSemiring R] (m : σ →₀ ) :

                  MvPolynomial.coeff m but promoted to a LinearMap.

                  Equations
                  Instances For
                    @[simp]
                    theorem MvPolynomial.lcoeff_apply (R : Type u) {σ : Type u_1} [CommSemiring R] (m : σ →₀ ) (p : MvPolynomial σ R) :
                    theorem MvPolynomial.coeff_sum {R : Type u} {σ : Type u_1} [CommSemiring R] {X : Type u_2} (s : Finset X) (f : XMvPolynomial σ R) (m : σ →₀ ) :
                    MvPolynomial.coeff m (∑ xs, f x) = xs, MvPolynomial.coeff m (f x)
                    theorem MvPolynomial.monic_monomial_eq {R : Type u} {σ : Type u_1} [CommSemiring R] (m : σ →₀ ) :
                    (MvPolynomial.monomial m) 1 = m.prod fun (n : σ) (e : ) => MvPolynomial.X n ^ e
                    @[simp]
                    theorem MvPolynomial.coeff_monomial {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (m n : σ →₀ ) (a : R) :
                    MvPolynomial.coeff m ((MvPolynomial.monomial n) a) = if n = m then a else 0
                    @[simp]
                    theorem MvPolynomial.coeff_C {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (m : σ →₀ ) (a : R) :
                    MvPolynomial.coeff m (MvPolynomial.C a) = if 0 = m then a else 0
                    theorem MvPolynomial.eq_C_of_isEmpty {R : Type u} {σ : Type u_1} [CommSemiring R] [IsEmpty σ] (p : MvPolynomial σ R) :
                    p = MvPolynomial.C (MvPolynomial.coeff 0 p)
                    theorem MvPolynomial.coeff_one {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (m : σ →₀ ) :
                    MvPolynomial.coeff m 1 = if 0 = m then 1 else 0
                    @[simp]
                    theorem MvPolynomial.coeff_zero_C {R : Type u} {σ : Type u_1} [CommSemiring R] (a : R) :
                    MvPolynomial.coeff 0 (MvPolynomial.C a) = a
                    @[simp]
                    theorem MvPolynomial.coeff_X_pow {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (i : σ) (m : σ →₀ ) (k : ) :
                    MvPolynomial.coeff m (MvPolynomial.X i ^ k) = if Finsupp.single i k = m then 1 else 0
                    theorem MvPolynomial.coeff_X' {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (i : σ) (m : σ →₀ ) :
                    MvPolynomial.coeff m (MvPolynomial.X i) = if Finsupp.single i 1 = m then 1 else 0
                    @[simp]
                    theorem MvPolynomial.coeff_X {R : Type u} {σ : Type u_1} [CommSemiring R] (i : σ) :
                    @[simp]
                    theorem MvPolynomial.coeff_C_mul {R : Type u} {σ : Type u_1} [CommSemiring R] (m : σ →₀ ) (a : R) (p : MvPolynomial σ R) :
                    MvPolynomial.coeff m (MvPolynomial.C a * p) = a * MvPolynomial.coeff m p
                    theorem MvPolynomial.coeff_mul {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (p q : MvPolynomial σ R) (n : σ →₀ ) :
                    @[simp]
                    theorem MvPolynomial.coeff_mul_monomial {R : Type u} {σ : Type u_1} [CommSemiring R] (m s : σ →₀ ) (r : R) (p : MvPolynomial σ R) :
                    @[simp]
                    theorem MvPolynomial.coeff_monomial_mul {R : Type u} {σ : Type u_1} [CommSemiring R] (m s : σ →₀ ) (r : R) (p : MvPolynomial σ R) :
                    @[simp]
                    theorem MvPolynomial.coeff_mul_X {R : Type u} {σ : Type u_1} [CommSemiring R] (m : σ →₀ ) (s : σ) (p : MvPolynomial σ R) :
                    @[simp]
                    theorem MvPolynomial.coeff_X_mul {R : Type u} {σ : Type u_1} [CommSemiring R] (m : σ →₀ ) (s : σ) (p : MvPolynomial σ R) :
                    theorem MvPolynomial.coeff_single_X_pow {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (s s' : σ) (n n' : ) :
                    MvPolynomial.coeff (Finsupp.single s' n') (MvPolynomial.X s ^ n) = if s = s' n = n' n = 0 n' = 0 then 1 else 0
                    @[simp]
                    theorem MvPolynomial.coeff_single_X {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (s s' : σ) (n : ) :
                    MvPolynomial.coeff (Finsupp.single s' n) (MvPolynomial.X s) = if n = 1 s = s' then 1 else 0
                    @[simp]
                    theorem MvPolynomial.support_mul_X {R : Type u} {σ : Type u_1} [CommSemiring R] (s : σ) (p : MvPolynomial σ R) :
                    @[simp]
                    theorem MvPolynomial.support_X_mul {R : Type u} {σ : Type u_1} [CommSemiring R] (s : σ) (p : MvPolynomial σ R) :
                    @[simp]
                    theorem MvPolynomial.support_smul_eq {R : Type u} {σ : Type u_1} [CommSemiring R] {S₁ : Type u_2} [Semiring S₁] [Module S₁ R] [NoZeroSMulDivisors S₁ R] {a : S₁} (h : a 0) (p : MvPolynomial σ R) :
                    (a p).support = p.support
                    theorem MvPolynomial.support_sdiff_support_subset_support_add {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (p q : MvPolynomial σ R) :
                    p.support \ q.support (p + q).support
                    theorem MvPolynomial.support_symmDiff_support_subset_support_add {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (p q : MvPolynomial σ R) :
                    symmDiff p.support q.support (p + q).support
                    theorem MvPolynomial.coeff_mul_monomial' {R : Type u} {σ : Type u_1} [CommSemiring R] (m s : σ →₀ ) (r : R) (p : MvPolynomial σ R) :
                    MvPolynomial.coeff m (p * (MvPolynomial.monomial s) r) = if s m then MvPolynomial.coeff (m - s) p * r else 0
                    theorem MvPolynomial.coeff_monomial_mul' {R : Type u} {σ : Type u_1} [CommSemiring R] (m s : σ →₀ ) (r : R) (p : MvPolynomial σ R) :
                    MvPolynomial.coeff m ((MvPolynomial.monomial s) r * p) = if s m then r * MvPolynomial.coeff (m - s) p else 0
                    theorem MvPolynomial.coeff_mul_X' {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (m : σ →₀ ) (s : σ) (p : MvPolynomial σ R) :
                    MvPolynomial.coeff m (p * MvPolynomial.X s) = if s m.support then MvPolynomial.coeff (m - Finsupp.single s 1) p else 0
                    theorem MvPolynomial.coeff_X_mul' {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (m : σ →₀ ) (s : σ) (p : MvPolynomial σ R) :
                    MvPolynomial.coeff m (MvPolynomial.X s * p) = if s m.support then MvPolynomial.coeff (m - Finsupp.single s 1) p else 0
                    theorem MvPolynomial.eq_zero_iff {R : Type u} {σ : Type u_1} [CommSemiring R] {p : MvPolynomial σ R} :
                    p = 0 ∀ (d : σ →₀ ), MvPolynomial.coeff d p = 0
                    theorem MvPolynomial.ne_zero_iff {R : Type u} {σ : Type u_1} [CommSemiring R] {p : MvPolynomial σ R} :
                    p 0 ∃ (d : σ →₀ ), MvPolynomial.coeff d p 0
                    @[simp]
                    theorem MvPolynomial.X_ne_zero {R : Type u} {σ : Type u_1} [CommSemiring R] [Nontrivial R] (s : σ) :
                    @[simp]
                    theorem MvPolynomial.support_eq_empty {R : Type u} {σ : Type u_1} [CommSemiring R] {p : MvPolynomial σ R} :
                    p.support = p = 0
                    @[simp]
                    theorem MvPolynomial.support_nonempty {R : Type u} {σ : Type u_1} [CommSemiring R] {p : MvPolynomial σ R} :
                    p.support.Nonempty p 0
                    theorem MvPolynomial.exists_coeff_ne_zero {R : Type u} {σ : Type u_1} [CommSemiring R] {p : MvPolynomial σ R} (h : p 0) :
                    ∃ (d : σ →₀ ), MvPolynomial.coeff d p 0
                    theorem MvPolynomial.C_dvd_iff_dvd_coeff {R : Type u} {σ : Type u_1} [CommSemiring R] (r : R) (φ : MvPolynomial σ R) :
                    MvPolynomial.C r φ ∀ (i : σ →₀ ), r MvPolynomial.coeff i φ
                    @[simp]
                    theorem MvPolynomial.isRegular_X {R : Type u} {σ : Type u_1} {n : σ} [CommSemiring R] :
                    @[simp]
                    theorem MvPolynomial.isRegular_X_pow {R : Type u} {σ : Type u_1} {n : σ} [CommSemiring R] (k : ) :
                    @[simp]
                    theorem MvPolynomial.isRegular_prod_X {R : Type u} {σ : Type u_1} [CommSemiring R] (s : Finset σ) :
                    IsRegular (∏ ns, MvPolynomial.X n)
                    def MvPolynomial.coeffs {R : Type u} {σ : Type u_1} [CommSemiring R] (p : MvPolynomial σ R) :

                    The finset of nonzero coefficients of a multivariate polynomial.

                    Equations
                    Instances For
                      theorem MvPolynomial.mem_coeffs_iff {R : Type u} {σ : Type u_1} [CommSemiring R] {p : MvPolynomial σ R} {c : R} :
                      c p.coeffs np.support, c = MvPolynomial.coeff n p
                      theorem MvPolynomial.coeff_mem_coeffs {R : Type u} {σ : Type u_1} [CommSemiring R] {p : MvPolynomial σ R} (m : σ →₀ ) (h : MvPolynomial.coeff m p 0) :
                      MvPolynomial.coeff m p p.coeffs
                      theorem MvPolynomial.zero_not_mem_coeffs {R : Type u} {σ : Type u_1} [CommSemiring R] (p : MvPolynomial σ R) :
                      0p.coeffs

                      constantCoeff p returns the constant term of the polynomial p, defined as coeff 0 p. This is a ring homomorphism.

                      Equations
                      Instances For
                        @[simp]
                        theorem MvPolynomial.constantCoeff_C {R : Type u} (σ : Type u_1) [CommSemiring R] (r : R) :
                        MvPolynomial.constantCoeff (MvPolynomial.C r) = r
                        @[simp]
                        theorem MvPolynomial.constantCoeff_X (R : Type u) {σ : Type u_1} [CommSemiring R] (i : σ) :
                        MvPolynomial.constantCoeff (MvPolynomial.X i) = 0
                        @[simp]
                        theorem MvPolynomial.constantCoeff_smul {S₁ : Type v} {σ : Type u_1} [CommSemiring S₁] {R : Type u_2} [SMulZeroClass R S₁] (a : R) (f : MvPolynomial σ S₁) :
                        MvPolynomial.constantCoeff (a f) = a MvPolynomial.constantCoeff f
                        theorem MvPolynomial.constantCoeff_monomial {R : Type u} {σ : Type u_1} [CommSemiring R] [DecidableEq σ] (d : σ →₀ ) (r : R) :
                        MvPolynomial.constantCoeff ((MvPolynomial.monomial d) r) = if d = 0 then r else 0
                        @[simp]
                        theorem MvPolynomial.support_sum_monomial_coeff {R : Type u} {σ : Type u_1} [CommSemiring R] (p : MvPolynomial σ R) :
                        vp.support, (MvPolynomial.monomial v) (MvPolynomial.coeff v p) = p
                        theorem MvPolynomial.as_sum {R : Type u} {σ : Type u_1} [CommSemiring R] (p : MvPolynomial σ R) :
                        p = vp.support, (MvPolynomial.monomial v) (MvPolynomial.coeff v p)