Documentation

Mathlib.Algebra.MvPolynomial.Expand

Expand multivariate polynomials #

Given a multivariate polynomial φ, one may replace every occurrence of X i by X i ^ n, for some natural number n. This operation is called MvPolynomial.expand and it is an algebra homomorphism.

Main declaration #

noncomputable def MvPolynomial.expand {σ : Type u_1} {R : Type u_3} [CommSemiring R] (p : ) :

Expand the polynomial by a factor of p, so ∑ aₙ xⁿ becomes ∑ aₙ xⁿᵖ.

See also Polynomial.expand.

Equations
Instances For
    theorem MvPolynomial.expand_C {σ : Type u_1} {R : Type u_3} [CommSemiring R] (p : ) (r : R) :
    (expand p) (C r) = C r
    @[simp]
    theorem MvPolynomial.expand_X {σ : Type u_1} {R : Type u_3} [CommSemiring R] (p : ) (i : σ) :
    (expand p) (X i) = X i ^ p
    @[simp]
    theorem MvPolynomial.expand_monomial {σ : Type u_1} {R : Type u_3} [CommSemiring R] (p : ) (d : σ →₀ ) (r : R) :
    (expand p) ((monomial d) r) = C r * id.support, (X i ^ p) ^ d i
    theorem MvPolynomial.expand_one_apply {σ : Type u_1} {R : Type u_3} [CommSemiring R] (f : MvPolynomial σ R) :
    (expand 1) f = f
    @[simp]
    theorem MvPolynomial.expand_one {σ : Type u_1} {R : Type u_3} [CommSemiring R] :
    theorem MvPolynomial.expand_comp_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (p : ) (f : σMvPolynomial τ R) :
    (expand p).comp (bind₁ f) = bind₁ fun (i : σ) => (expand p) (f i)
    theorem MvPolynomial.expand_bind₁ {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (p : ) (f : σMvPolynomial τ R) (φ : MvPolynomial σ R) :
    (expand p) ((bind₁ f) φ) = (bind₁ fun (i : σ) => (expand p) (f i)) φ
    @[simp]
    theorem MvPolynomial.map_expand {σ : Type u_1} {R : Type u_3} {S : Type u_4} [CommSemiring R] [CommSemiring S] (f : R →+* S) (p : ) (φ : MvPolynomial σ R) :
    (map f) ((expand p) φ) = (expand p) ((map f) φ)
    @[simp]
    theorem MvPolynomial.rename_expand {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : στ) (p : ) (φ : MvPolynomial σ R) :
    (rename f) ((expand p) φ) = (expand p) ((rename f) φ)
    @[simp]
    theorem MvPolynomial.rename_comp_expand {σ : Type u_1} {τ : Type u_2} {R : Type u_3} [CommSemiring R] (f : στ) (p : ) :
    (rename f).comp (expand p) = (expand p).comp (rename f)