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 #
MvPolynomial.expand: expand a polynomial by a factor of p, so∑ aₙ xⁿbecomes∑ aₙ xⁿᵖ.
Expand the polynomial by a factor of p, so ∑ aₙ xⁿ becomes ∑ aₙ xⁿᵖ.
See also Polynomial.expand.
Equations
- MvPolynomial.expand p = MvPolynomial.bind₁ fun (i : σ) => MvPolynomial.X i ^ p
Instances For
@[simp]
theorem
MvPolynomial.expand_zero_apply
{σ : Type u_1}
{R : Type u_3}
[CommSemiring R]
(p : MvPolynomial σ R)
:
@[simp]
theorem
MvPolynomial.expand_one_apply
{σ : Type u_1}
{R : Type u_3}
[CommSemiring R]
(f : MvPolynomial σ R)
:
theorem
MvPolynomial.expand_comp_bind₁
{σ : Type u_1}
{τ : Type u_2}
{R : Type u_3}
[CommSemiring R]
(p : ℕ)
(f : σ → MvPolynomial τ R)
:
theorem
MvPolynomial.expand_bind₁
{σ : Type u_1}
{τ : Type u_2}
{R : Type u_3}
[CommSemiring R]
(p : ℕ)
(f : σ → MvPolynomial τ R)
(φ : MvPolynomial σ R)
:
@[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)
:
@[simp]
theorem
MvPolynomial.rename_expand
{σ : Type u_1}
{τ : Type u_2}
{R : Type u_3}
[CommSemiring R]
(f : σ → τ)
(p : ℕ)
(φ : MvPolynomial σ R)
:
theorem
MvPolynomial.eval₂Hom_comp_expand
{σ : Type u_1}
{R : Type u_3}
{S : Type u_4}
[CommSemiring R]
[CommSemiring S]
(f : R →+* S)
(g : σ → S)
(p : ℕ)
:
@[simp]
theorem
MvPolynomial.eval₂_expand
{σ : Type u_1}
{R : Type u_3}
{S : Type u_4}
[CommSemiring R]
[CommSemiring S]
(f : R →+* S)
(g : σ → S)
(φ : MvPolynomial σ R)
(p : ℕ)
:
@[simp]
theorem
MvPolynomial.aeval_comp_expand
{σ : Type u_1}
{R : Type u_3}
[CommSemiring R]
{A : Type u_5}
[CommSemiring A]
[Algebra R A]
(f : σ → A)
(p : ℕ)
:
@[simp]
theorem
MvPolynomial.aeval_expand
{σ : Type u_1}
{R : Type u_3}
[CommSemiring R]
{A : Type u_5}
[CommSemiring A]
[Algebra R A]
(f : σ → A)
(φ : MvPolynomial σ R)
(p : ℕ)
:
@[simp]
theorem
MvPolynomial.eval_expand
{σ : Type u_1}
{R : Type u_3}
[CommSemiring R]
(f : σ → R)
(φ : MvPolynomial σ R)
(p : ℕ)
:
theorem
MvPolynomial.expand_mul
{σ : Type u_1}
{R : Type u_3}
[CommSemiring R]
(p q : ℕ)
(φ : MvPolynomial σ R)
:
@[simp]
theorem
MvPolynomial.coeff_expand_smul
{σ : Type u_1}
{R : Type u_3}
[CommSemiring R]
(φ : MvPolynomial σ R)
{p : ℕ}
(hp : p ≠ 0)
(m : σ →₀ ℕ)
:
theorem
MvPolynomial.support_expand_subset
{σ : Type u_1}
{R : Type u_3}
[CommSemiring R]
[DecidableEq σ]
(φ : MvPolynomial σ R)
(p : ℕ)
:
theorem
MvPolynomial.coeff_expand_of_not_dvd
{σ : Type u_1}
{R : Type u_3}
[CommSemiring R]
(φ : MvPolynomial σ R)
{p : ℕ}
{m : σ →₀ ℕ}
{i : σ}
(h : ¬p ∣ m i)
:
theorem
MvPolynomial.support_expand
{σ : Type u_1}
{R : Type u_3}
[CommSemiring R]
[DecidableEq σ]
(φ : MvPolynomial σ R)
{p : ℕ}
(hp : p ≠ 0)
: