Documentation

Mathlib.Algebra.MvPolynomial.Eval

Multivariate polynomials #

This file defines functions for evaluating multivariate polynomials. These include generically evaluating a polynomial given a valuation of all its variables, and more advanced evaluations that allow one to map the coefficients to different rings.

Notation #

In the definitions below, we use the following notation:

Definitions #

def MvPolynomial.eval₂ {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (p : MvPolynomial σ R) :
S₁

Evaluate a polynomial p given a valuation g of all the variables and a ring hom f from the scalar ring to the target

Equations
Instances For
    theorem MvPolynomial.eval₂_eq {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (g : R →+* S₁) (X : σS₁) (f : MvPolynomial σ R) :
    MvPolynomial.eval₂ g X f = df.support, g (MvPolynomial.coeff d f) * id.support, X i ^ d i
    theorem MvPolynomial.eval₂_eq' {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Fintype σ] (g : R →+* S₁) (X : σS₁) (f : MvPolynomial σ R) :
    MvPolynomial.eval₂ g X f = df.support, g (MvPolynomial.coeff d f) * i : σ, X i ^ d i
    @[simp]
    theorem MvPolynomial.eval₂_zero {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) :
    @[simp]
    theorem MvPolynomial.eval₂_add {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] {p q : MvPolynomial σ R} (f : R →+* S₁) (g : σS₁) :
    @[simp]
    theorem MvPolynomial.eval₂_monomial {R : Type u} {S₁ : Type v} {σ : Type u_1} {a : R} {s : σ →₀ } [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) :
    MvPolynomial.eval₂ f g ((MvPolynomial.monomial s) a) = f a * s.prod fun (n : σ) (e : ) => g n ^ e
    @[simp]
    theorem MvPolynomial.eval₂_C {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (a : R) :
    MvPolynomial.eval₂ f g (MvPolynomial.C a) = f a
    @[simp]
    theorem MvPolynomial.eval₂_one {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) :
    @[simp]
    theorem MvPolynomial.eval₂_natCast {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (n : ) :
    MvPolynomial.eval₂ f g n = n
    @[simp]
    theorem MvPolynomial.eval₂_ofNat {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (n : ) [n.AtLeastTwo] :
    @[simp]
    theorem MvPolynomial.eval₂_X {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (n : σ) :
    theorem MvPolynomial.eval₂_mul_monomial {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] {p : MvPolynomial σ R} (f : R →+* S₁) (g : σS₁) {s : σ →₀ } {a : R} :
    MvPolynomial.eval₂ f g (p * (MvPolynomial.monomial s) a) = MvPolynomial.eval₂ f g p * f a * s.prod fun (n : σ) (e : ) => g n ^ e
    theorem MvPolynomial.eval₂_mul_C {R : Type u} {S₁ : Type v} {σ : Type u_1} {a : R} [CommSemiring R] [CommSemiring S₁] {p : MvPolynomial σ R} (f : R →+* S₁) (g : σS₁) :
    MvPolynomial.eval₂ f g (p * MvPolynomial.C a) = MvPolynomial.eval₂ f g p * f a
    @[simp]
    theorem MvPolynomial.eval₂_mul {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] {q : MvPolynomial σ R} (f : R →+* S₁) (g : σS₁) {p : MvPolynomial σ R} :
    @[simp]
    theorem MvPolynomial.eval₂_pow {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) {p : MvPolynomial σ R} {n : } :
    def MvPolynomial.eval₂Hom {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) :

    MvPolynomial.eval₂ as a RingHom.

    Equations
    Instances For
      @[simp]
      theorem MvPolynomial.coe_eval₂Hom {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) :
      theorem MvPolynomial.eval₂Hom_congr {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] {f₁ f₂ : R →+* S₁} {g₁ g₂ : σS₁} {p₁ p₂ : MvPolynomial σ R} :
      f₁ = f₂g₁ = g₂p₁ = p₂(MvPolynomial.eval₂Hom f₁ g₁) p₁ = (MvPolynomial.eval₂Hom f₂ g₂) p₂
      @[simp]
      theorem MvPolynomial.eval₂Hom_C {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (r : R) :
      (MvPolynomial.eval₂Hom f g) (MvPolynomial.C r) = f r
      @[simp]
      theorem MvPolynomial.eval₂Hom_X' {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (i : σ) :
      @[simp]
      theorem MvPolynomial.comp_eval₂Hom {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [CommSemiring S₂] (f : R →+* S₁) (g : σS₁) (φ : S₁ →+* S₂) :
      φ.comp (MvPolynomial.eval₂Hom f g) = MvPolynomial.eval₂Hom (φ.comp f) fun (i : σ) => φ (g i)
      theorem MvPolynomial.map_eval₂Hom {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [CommSemiring S₂] (f : R →+* S₁) (g : σS₁) (φ : S₁ →+* S₂) (p : MvPolynomial σ R) :
      φ ((MvPolynomial.eval₂Hom f g) p) = (MvPolynomial.eval₂Hom (φ.comp f) fun (i : σ) => φ (g i)) p
      theorem MvPolynomial.eval₂Hom_monomial {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (d : σ →₀ ) (r : R) :
      (MvPolynomial.eval₂Hom f g) ((MvPolynomial.monomial d) r) = f r * d.prod fun (i : σ) (k : ) => g i ^ k
      theorem MvPolynomial.eval₂_comp_left {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] {S₂ : Type u_2} [CommSemiring S₂] (k : S₁ →+* S₂) (f : R →+* S₁) (g : σS₁) (p : MvPolynomial σ R) :
      k (MvPolynomial.eval₂ f g p) = MvPolynomial.eval₂ (k.comp f) (k g) p
      theorem MvPolynomial.eval₂_congr {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] {p : MvPolynomial σ R} (f : R →+* S₁) (g₁ g₂ : σS₁) (h : ∀ {i : σ} {c : σ →₀ }, i c.supportMvPolynomial.coeff c p 0g₁ i = g₂ i) :
      @[simp]
      theorem MvPolynomial.eval₂_sum {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (s : Finset S₂) (p : S₂MvPolynomial σ R) :
      MvPolynomial.eval₂ f g (∑ xs, p x) = xs, MvPolynomial.eval₂ f g (p x)
      @[simp]
      theorem MvPolynomial.eval₂_prod {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (s : Finset S₂) (p : S₂MvPolynomial σ R) :
      MvPolynomial.eval₂ f g (∏ xs, p x) = xs, MvPolynomial.eval₂ f g (p x)
      theorem MvPolynomial.eval₂_assoc {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (q : S₂MvPolynomial σ R) (p : MvPolynomial S₂ R) :
      def MvPolynomial.eval {R : Type u} {σ : Type u_1} [CommSemiring R] (f : σR) :

      Evaluate a polynomial p given a valuation f of all the variables

      Equations
      Instances For
        theorem MvPolynomial.eval_eq {R : Type u} {σ : Type u_1} [CommSemiring R] (X : σR) (f : MvPolynomial σ R) :
        (MvPolynomial.eval X) f = df.support, MvPolynomial.coeff d f * id.support, X i ^ d i
        theorem MvPolynomial.eval_eq' {R : Type u} {σ : Type u_1} [CommSemiring R] [Fintype σ] (X : σR) (f : MvPolynomial σ R) :
        (MvPolynomial.eval X) f = df.support, MvPolynomial.coeff d f * i : σ, X i ^ d i
        theorem MvPolynomial.eval_monomial {R : Type u} {σ : Type u_1} {a : R} {s : σ →₀ } [CommSemiring R] {f : σR} :
        (MvPolynomial.eval f) ((MvPolynomial.monomial s) a) = a * s.prod fun (n : σ) (e : ) => f n ^ e
        @[simp]
        theorem MvPolynomial.eval_C {R : Type u} {σ : Type u_1} [CommSemiring R] {f : σR} (a : R) :
        (MvPolynomial.eval f) (MvPolynomial.C a) = a
        @[simp]
        theorem MvPolynomial.eval_X {R : Type u} {σ : Type u_1} [CommSemiring R] {f : σR} (n : σ) :
        @[simp]
        theorem MvPolynomial.eval_ofNat {R : Type u} {σ : Type u_1} [CommSemiring R] {f : σR} (n : ) [n.AtLeastTwo] :
        @[simp]
        theorem MvPolynomial.smul_eval {R : Type u} {σ : Type u_1} [CommSemiring R] (x : σR) (p : MvPolynomial σ R) (s : R) :
        theorem MvPolynomial.eval_add {R : Type u} {σ : Type u_1} [CommSemiring R] {p q : MvPolynomial σ R} {f : σR} :
        theorem MvPolynomial.eval_mul {R : Type u} {σ : Type u_1} [CommSemiring R] {p q : MvPolynomial σ R} {f : σR} :
        theorem MvPolynomial.eval_pow {R : Type u} {σ : Type u_1} [CommSemiring R] {p : MvPolynomial σ R} {f : σR} (n : ) :
        theorem MvPolynomial.eval_sum {R : Type u} {σ : Type u_1} [CommSemiring R] {ι : Type u_2} (s : Finset ι) (f : ιMvPolynomial σ R) (g : σR) :
        (MvPolynomial.eval g) (∑ is, f i) = is, (MvPolynomial.eval g) (f i)
        theorem MvPolynomial.eval_prod {R : Type u} {σ : Type u_1} [CommSemiring R] {ι : Type u_2} (s : Finset ι) (f : ιMvPolynomial σ R) (g : σR) :
        (MvPolynomial.eval g) (∏ is, f i) = is, (MvPolynomial.eval g) (f i)
        theorem MvPolynomial.eval_assoc {R : Type u} {σ : Type u_1} [CommSemiring R] {τ : Type u_2} (f : σMvPolynomial τ R) (g : τR) (p : MvPolynomial σ R) :
        @[simp]
        theorem MvPolynomial.eval₂_id {R : Type u} {σ : Type u_1} [CommSemiring R] {g : σR} (p : MvPolynomial σ R) :
        theorem MvPolynomial.eval_eval₂ {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {τ : Type u_3} {x : τS} [CommSemiring S] (f : R →+* MvPolynomial τ S) (g : σMvPolynomial τ S) (p : MvPolynomial σ R) :
        def MvPolynomial.map {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) :

        map f p maps a polynomial p across a ring hom f

        Equations
        Instances For
          @[simp]
          theorem MvPolynomial.map_monomial {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (s : σ →₀ ) (a : R) :
          @[simp]
          theorem MvPolynomial.map_C {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (a : R) :
          (MvPolynomial.map f) (MvPolynomial.C a) = MvPolynomial.C (f a)
          @[simp]
          theorem MvPolynomial.map_ofNat {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (n : ) [n.AtLeastTwo] :
          @[simp]
          theorem MvPolynomial.map_X {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (n : σ) :
          theorem MvPolynomial.map_id {R : Type u} {σ : Type u_1} [CommSemiring R] (p : MvPolynomial σ R) :
          theorem MvPolynomial.map_map {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) [CommSemiring S₂] (g : S₁ →+* S₂) (p : MvPolynomial σ R) :
          theorem MvPolynomial.eval₂_eq_eval_map {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (p : MvPolynomial σ R) :
          theorem MvPolynomial.eval₂_comp_right {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] {S₂ : Type u_2} [CommSemiring S₂] (k : S₁ →+* S₂) (f : R →+* S₁) (g : σS₁) (p : MvPolynomial σ R) :
          theorem MvPolynomial.map_eval₂ {R : Type u} {S₁ : Type v} {S₂ : Type w} {S₃ : Type x} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : S₂MvPolynomial S₃ R) (p : MvPolynomial S₂ R) :
          theorem MvPolynomial.coeff_map {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (p : MvPolynomial σ R) (m : σ →₀ ) :
          theorem MvPolynomial.map_injective {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (hf : Function.Injective f) :
          theorem MvPolynomial.map_surjective {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (hf : Function.Surjective f) :
          theorem MvPolynomial.map_leftInverse {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] {f : R →+* S₁} {g : S₁ →+* R} (hf : Function.LeftInverse f g) :

          If f is a left-inverse of g then map f is a left-inverse of map g.

          theorem MvPolynomial.map_rightInverse {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] {f : R →+* S₁} {g : S₁ →+* R} (hf : Function.RightInverse f g) :

          If f is a right-inverse of g then map f is a right-inverse of map g.

          @[simp]
          theorem MvPolynomial.eval_map {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σS₁) (p : MvPolynomial σ R) :
          theorem MvPolynomial.eval₂_comp {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : σR) (p : MvPolynomial σ R) :
          @[simp]
          theorem MvPolynomial.eval₂_map {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [CommSemiring S₂] (f : R →+* S₁) (g : σS₂) (φ : S₁ →+* S₂) (p : MvPolynomial σ R) :
          @[simp]
          theorem MvPolynomial.eval₂Hom_map_hom {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [CommSemiring S₂] (f : R →+* S₁) (g : σS₂) (φ : S₁ →+* S₂) (p : MvPolynomial σ R) :
          @[simp]
          theorem MvPolynomial.constantCoeff_map {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (φ : MvPolynomial σ R) :
          MvPolynomial.constantCoeff ((MvPolynomial.map f) φ) = f (MvPolynomial.constantCoeff φ)
          theorem MvPolynomial.support_map_subset {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (p : MvPolynomial σ R) :
          ((MvPolynomial.map f) p).support p.support
          theorem MvPolynomial.support_map_of_injective {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (p : MvPolynomial σ R) {f : R →+* S₁} (hf : Function.Injective f) :
          ((MvPolynomial.map f) p).support = p.support
          theorem MvPolynomial.C_dvd_iff_map_hom_eq_zero {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (q : R →+* S₁) (r : R) (hr : ∀ (r' : R), q r' = 0 r r') (φ : MvPolynomial σ R) :
          MvPolynomial.C r φ (MvPolynomial.map q) φ = 0
          theorem MvPolynomial.map_mapRange_eq_iff {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] (f : R →+* S₁) (g : S₁R) (hg : g 0 = 0) (φ : MvPolynomial σ S₁) :
          (MvPolynomial.map f) (Finsupp.mapRange g hg φ) = φ ∀ (d : σ →₀ ), f (g (MvPolynomial.coeff d φ)) = MvPolynomial.coeff d φ
          def MvPolynomial.mapAlgHom {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [CommSemiring S₂] [Algebra R S₁] [Algebra R S₂] (f : S₁ →ₐ[R] S₂) :

          If f : S₁ →ₐ[R] S₂ is a morphism of R-algebras, then so is MvPolynomial.map f.

          Equations
          Instances For
            @[simp]
            theorem MvPolynomial.mapAlgHom_apply {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [CommSemiring S₂] [Algebra R S₁] [Algebra R S₂] (f : S₁ →ₐ[R] S₂) (p : MvPolynomial σ S₁) :
            @[simp]
            theorem MvPolynomial.mapAlgHom_id {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] :
            @[simp]
            theorem MvPolynomial.mapAlgHom_coe_ringHom {R : Type u} {S₁ : Type v} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [CommSemiring S₂] [Algebra R S₁] [Algebra R S₂] (f : S₁ →ₐ[R] S₂) :

            The algebra of multivariate polynomials #

            @[simp]
            theorem MvPolynomial.algebraMap_apply {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (r : R) :
            (algebraMap R (MvPolynomial σ S₁)) r = MvPolynomial.C ((algebraMap R S₁) r)
            def MvPolynomial.aeval {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) :

            A map σ → S₁ where S₁ is an algebra over R generates an R-algebra homomorphism from multivariate polynomials over σ to S₁.

            Equations
            Instances For
              theorem MvPolynomial.aeval_def {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) (p : MvPolynomial σ R) :
              theorem MvPolynomial.aeval_eq_eval₂Hom {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) (p : MvPolynomial σ R) :
              @[simp]
              theorem MvPolynomial.coe_aeval_eq_eval {S₁ : Type v} {σ : Type u_1} [CommSemiring S₁] (f : σS₁) :
              @[simp]
              theorem MvPolynomial.aeval_X {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) (s : σ) :
              theorem MvPolynomial.aeval_C {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) (r : R) :
              (MvPolynomial.aeval f) (MvPolynomial.C r) = (algebraMap R S₁) r
              @[simp]
              theorem MvPolynomial.aeval_ofNat {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) (n : ) [n.AtLeastTwo] :
              theorem MvPolynomial.aeval_unique {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (φ : MvPolynomial σ R →ₐ[R] S₁) :
              theorem MvPolynomial.comp_aeval {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) {B : Type u_2} [CommSemiring B] [Algebra R B] (φ : S₁ →ₐ[R] B) :
              φ.comp (MvPolynomial.aeval f) = MvPolynomial.aeval fun (i : σ) => φ (f i)
              theorem MvPolynomial.comp_aeval_apply {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) {B : Type u_2} [CommSemiring B] [Algebra R B] (φ : S₁ →ₐ[R] B) (p : MvPolynomial σ R) :
              φ ((MvPolynomial.aeval f) p) = (MvPolynomial.aeval fun (i : σ) => φ (f i)) p
              @[simp]
              theorem MvPolynomial.map_aeval {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] {B : Type u_2} [CommSemiring B] (g : σS₁) (φ : S₁ →+* B) (p : MvPolynomial σ R) :
              φ ((MvPolynomial.aeval g) p) = (MvPolynomial.eval₂Hom (φ.comp (algebraMap R S₁)) fun (i : σ) => φ (g i)) p
              @[simp]
              theorem MvPolynomial.eval₂Hom_zero {R : Type u} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₂] (f : R →+* S₂) :
              @[simp]
              theorem MvPolynomial.eval₂Hom_zero' {R : Type u} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₂] (f : R →+* S₂) :
              theorem MvPolynomial.eval₂Hom_zero_apply {R : Type u} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₂] (f : R →+* S₂) (p : MvPolynomial σ R) :
              (MvPolynomial.eval₂Hom f 0) p = f (MvPolynomial.constantCoeff p)
              theorem MvPolynomial.eval₂Hom_zero'_apply {R : Type u} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₂] (f : R →+* S₂) (p : MvPolynomial σ R) :
              (MvPolynomial.eval₂Hom f fun (x : σ) => 0) p = f (MvPolynomial.constantCoeff p)
              @[simp]
              theorem MvPolynomial.eval₂_zero_apply {R : Type u} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₂] (f : R →+* S₂) (p : MvPolynomial σ R) :
              MvPolynomial.eval₂ f 0 p = f (MvPolynomial.constantCoeff p)
              @[simp]
              theorem MvPolynomial.eval₂_zero'_apply {R : Type u} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₂] (f : R →+* S₂) (p : MvPolynomial σ R) :
              MvPolynomial.eval₂ f (fun (x : σ) => 0) p = f (MvPolynomial.constantCoeff p)
              @[simp]
              theorem MvPolynomial.aeval_zero {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (p : MvPolynomial σ R) :
              (MvPolynomial.aeval 0) p = (algebraMap R S₁) (MvPolynomial.constantCoeff p)
              @[simp]
              theorem MvPolynomial.aeval_zero' {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (p : MvPolynomial σ R) :
              (MvPolynomial.aeval fun (x : σ) => 0) p = (algebraMap R S₁) (MvPolynomial.constantCoeff p)
              @[simp]
              theorem MvPolynomial.eval_zero' {R : Type u} {σ : Type u_1} [CommSemiring R] :
              theorem MvPolynomial.aeval_monomial {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (g : σS₁) (d : σ →₀ ) (r : R) :
              (MvPolynomial.aeval g) ((MvPolynomial.monomial d) r) = (algebraMap R S₁) r * d.prod fun (i : σ) (k : ) => g i ^ k
              theorem MvPolynomial.eval₂Hom_eq_zero {R : Type u} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₂] (f : R →+* S₂) (g : σS₂) (φ : MvPolynomial σ R) (h : ∀ (d : σ →₀ ), MvPolynomial.coeff d φ 0id.support, g i = 0) :
              theorem MvPolynomial.aeval_eq_zero {R : Type u} {S₂ : Type w} {σ : Type u_1} [CommSemiring R] [CommSemiring S₂] [Algebra R S₂] (f : σS₂) (φ : MvPolynomial σ R) (h : ∀ (d : σ →₀ ), MvPolynomial.coeff d φ 0id.support, f i = 0) :
              theorem MvPolynomial.aeval_sum {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) {ι : Type u_2} (s : Finset ι) (φ : ιMvPolynomial σ R) :
              (MvPolynomial.aeval f) (∑ is, φ i) = is, (MvPolynomial.aeval f) (φ i)
              theorem MvPolynomial.aeval_prod {R : Type u} {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) {ι : Type u_2} (s : Finset ι) (φ : ιMvPolynomial σ R) :
              (MvPolynomial.aeval f) (∏ is, φ i) = is, (MvPolynomial.aeval f) (φ i)
              theorem Algebra.adjoin_range_eq_range_aeval (R : Type u) {S₁ : Type v} {σ : Type u_1} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (f : σS₁) :
              theorem Algebra.adjoin_eq_range (R : Type u) {S₁ : Type v} [CommSemiring R] [CommSemiring S₁] [Algebra R S₁] (s : Set S₁) :
              def MvPolynomial.aevalTower {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {A : Type u_3} [CommSemiring S] [CommSemiring A] [Algebra S R] [Algebra S A] (f : R →ₐ[S] A) (X : σA) :

              Version of aeval for defining algebra homs out of MvPolynomial σ R over a smaller base ring than R.

              Equations
              Instances For
                @[simp]
                theorem MvPolynomial.aevalTower_X {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {A : Type u_3} [CommSemiring S] [CommSemiring A] [Algebra S R] [Algebra S A] (g : R →ₐ[S] A) (y : σA) (i : σ) :
                @[simp]
                theorem MvPolynomial.aevalTower_C {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {A : Type u_3} [CommSemiring S] [CommSemiring A] [Algebra S R] [Algebra S A] (g : R →ₐ[S] A) (y : σA) (x : R) :
                (MvPolynomial.aevalTower g y) (MvPolynomial.C x) = g x
                @[simp]
                theorem MvPolynomial.aevalTower_ofNat {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {A : Type u_3} [CommSemiring S] [CommSemiring A] [Algebra S R] [Algebra S A] (g : R →ₐ[S] A) (y : σA) (n : ) [n.AtLeastTwo] :
                @[simp]
                theorem MvPolynomial.aevalTower_comp_C {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {A : Type u_3} [CommSemiring S] [CommSemiring A] [Algebra S R] [Algebra S A] (g : R →ₐ[S] A) (y : σA) :
                theorem MvPolynomial.aevalTower_algebraMap {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {A : Type u_3} [CommSemiring S] [CommSemiring A] [Algebra S R] [Algebra S A] (g : R →ₐ[S] A) (y : σA) (x : R) :
                theorem MvPolynomial.aevalTower_comp_algebraMap {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {A : Type u_3} [CommSemiring S] [CommSemiring A] [Algebra S R] [Algebra S A] (g : R →ₐ[S] A) (y : σA) :
                (↑(MvPolynomial.aevalTower g y)).comp (algebraMap R (MvPolynomial σ R)) = g
                theorem MvPolynomial.aevalTower_toAlgHom {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {A : Type u_3} [CommSemiring S] [CommSemiring A] [Algebra S R] [Algebra S A] (g : R →ₐ[S] A) (y : σA) (x : R) :
                @[simp]
                theorem MvPolynomial.aevalTower_comp_toAlgHom {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {A : Type u_3} [CommSemiring S] [CommSemiring A] [Algebra S R] [Algebra S A] (g : R →ₐ[S] A) (y : σA) :
                theorem MvPolynomial.eval₂_mem {R : Type u} {σ : Type u_1} [CommSemiring R] {S : Type u_2} {subS : Type u_3} [CommSemiring S] [SetLike subS S] [SubsemiringClass subS S] {f : R →+* S} {p : MvPolynomial σ R} {s : subS} (hs : ip.support, f (MvPolynomial.coeff i p) s) {v : σS} (hv : ∀ (i : σ), v i s) :
                theorem MvPolynomial.eval_mem {σ : Type u_1} {S : Type u_2} {subS : Type u_3} [CommSemiring S] [SetLike subS S] [SubsemiringClass subS S] {p : MvPolynomial σ S} {s : subS} (hs : ip.support, MvPolynomial.coeff i p s) {v : σS} (hv : ∀ (i : σ), v i s) :
                theorem MvPolynomial.aeval_sum_elim {R : Type u} [CommSemiring R] {S : Type u_2} {T : Type u_3} [CommSemiring S] [Algebra R S] [CommSemiring T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] {σ : Type u_4} {τ : Type u_5} (p : MvPolynomial (σ τ) R) (f : τS) (g : σT) :