Documentation

Mathlib.RingTheory.Polynomial.Basic

Ring-theoretic supplement of Algebra.Polynomial. #

Main results #

instance Polynomial.instCharP {R : Type u} [Semiring R] (p : ) [h : CharP R p] :
instance Polynomial.instExpChar {R : Type u} [Semiring R] (p : ) [h : ExpChar R p] :

The R-submodule of R[X] consisting of polynomials of degree ≤ n.

Equations

The R-submodule of R[X] consisting of polynomials of degree < n.

Equations
theorem Polynomial.degreeLE_mono {R : Type u} [Semiring R] {m n : WithBot } (H : m n) :
theorem Polynomial.degreeLE_eq_span_X_pow {R : Type u} [Semiring R] [DecidableEq R] {n : } :
degreeLE R n = Submodule.span R (Finset.image (fun (n : ) => X ^ n) (Finset.range (n + 1)))
theorem Polynomial.mem_degreeLT {R : Type u} [Semiring R] {n : } {f : Polynomial R} :
f degreeLT R n f.degree < n
theorem Polynomial.degreeLT_mono {R : Type u} [Semiring R] {m n : } (H : m n) :
def Polynomial.degreeLTEquiv (R : Type u_2) [Semiring R] (n : ) :
(degreeLT R n) ≃ₗ[R] Fin nR

The first n coefficients on degreeLT n form a linear equivalence with Fin n → R.

Equations
  • One or more equations did not get rendered due to their size.
theorem Polynomial.degreeLTEquiv_eq_zero_iff_eq_zero {R : Type u} [Semiring R] {n : } {p : Polynomial R} (hp : p degreeLT R n) :
(degreeLTEquiv R n) p, hp = 0 p = 0
theorem Polynomial.eval_eq_sum_degreeLTEquiv {R : Type u} [Semiring R] {n : } {p : Polynomial R} (hp : p degreeLT R n) (x : R) :
eval x p = i : Fin n, (degreeLTEquiv R n) p, hp i * x ^ i

The equivalence between monic polynomials of degree n and polynomials of degree less than n, formed by adding a term X ^ n.

Equations
  • One or more equations did not get rendered due to their size.
theorem Polynomial.exists_degree_le_of_mem_span {R : Type u} [Semiring R] {s : Set (Polynomial R)} {p : Polynomial R} (hs : s.Nonempty) (hp : p Submodule.span R s) :
p's, p.degree p'.degree

For every polynomial p in the span of a set s : Set R[X], there exists a polynomial of p' ∈ s with higher degree. See also Polynomial.exists_degree_le_of_mem_span_of_finite.

theorem Polynomial.exists_degree_le_of_mem_span_of_finite {R : Type u} [Semiring R] {s : Set (Polynomial R)} (s_fin : s.Finite) (hs : s.Nonempty) :
p's, pSubmodule.span R s, p.degree p'.degree

A stronger version of Polynomial.exists_degree_le_of_mem_span under the assumption that the set s : R[X] is finite. There exists a polynomial p' ∈ s whose degree dominates the degree of every element of p ∈ span R s.

theorem Polynomial.span_le_degreeLE_of_finite {R : Type u} [Semiring R] {s : Set (Polynomial R)} (s_fin : s.Finite) :
∃ (n : ), Submodule.span R s degreeLE R n

The span of every finite set of polynomials is contained in a degreeLE n for some n.

theorem Polynomial.span_of_finite_le_degreeLT {R : Type u} [Semiring R] {s : Set (Polynomial R)} (s_fin : s.Finite) :
∃ (n : ), Submodule.span R s degreeLT R n

The span of every finite set of polynomials is contained in a degreeLT n for some n.

If R is a nontrivial ring, the polynomials R[X] are not finite as an R-module. When R is a field, this is equivalent to R[X] being an infinite-dimensional vector space over R.

theorem Polynomial.geom_sum_X_comp_X_add_one_eq_sum {R : Type u} [Semiring R] (n : ) :
(∑ iFinset.range n, X ^ i).comp (X + 1) = iFinset.range n, (n.choose (i + 1)) * X ^ i
theorem Polynomial.Monic.geom_sum {R : Type u} [Semiring R] {P : Polynomial R} (hP : P.Monic) (hdeg : 0 < P.natDegree) {n : } (hn : n 0) :
(∑ iFinset.range n, P ^ i).Monic
theorem Polynomial.Monic.geom_sum' {R : Type u} [Semiring R] {P : Polynomial R} (hP : P.Monic) (hdeg : 0 < P.degree) {n : } (hn : n 0) :
(∑ iFinset.range n, P ^ i).Monic
theorem Polynomial.monic_geom_sum_X {R : Type u} [Semiring R] {n : } (hn : n 0) :
(∑ iFinset.range n, X ^ i).Monic

Given a polynomial, return the polynomial whose coefficients are in the ring closure of the original coefficients.

Equations
@[simp]
theorem Polynomial.coeff_restriction {R : Type u} [Ring R] {p : Polynomial R} {n : } :
(p.restriction.coeff n) = p.coeff n
theorem Polynomial.coeff_restriction' {R : Type u} [Ring R] {p : Polynomial R} {n : } :
(p.restriction.coeff n) = p.coeff n
@[simp]
@[simp]
theorem Polynomial.eval₂_restriction {R : Type u} {S : Type u_1} [Ring R] [Semiring S] {f : R →+* S} {x : S} {p : Polynomial R} :
def Polynomial.toSubring {R : Type u} [Ring R] (p : Polynomial R) (T : Subring R) (hp : p.coeffs T) :

Given a polynomial p and a subring T that contains the coefficients of p, return the corresponding polynomial whose coefficients are in T.

Equations
@[simp]
theorem Polynomial.coeff_toSubring {R : Type u} [Ring R] (p : Polynomial R) (T : Subring R) (hp : p.coeffs T) {n : } :
((p.toSubring T hp).coeff n) = p.coeff n
theorem Polynomial.coeff_toSubring' {R : Type u} [Ring R] (p : Polynomial R) (T : Subring R) (hp : p.coeffs T) {n : } :
((p.toSubring T hp).coeff n) = p.coeff n
@[simp]
theorem Polynomial.support_toSubring {R : Type u} [Ring R] (p : Polynomial R) (T : Subring R) (hp : p.coeffs T) :
@[simp]
theorem Polynomial.degree_toSubring {R : Type u} [Ring R] (p : Polynomial R) (T : Subring R) (hp : p.coeffs T) :
@[simp]
theorem Polynomial.natDegree_toSubring {R : Type u} [Ring R] (p : Polynomial R) (T : Subring R) (hp : p.coeffs T) :
@[simp]
theorem Polynomial.monic_toSubring {R : Type u} [Ring R] (p : Polynomial R) (T : Subring R) (hp : p.coeffs T) :
@[simp]
theorem Polynomial.toSubring_zero {R : Type u} [Ring R] (T : Subring R) :
toSubring 0 T = 0
@[simp]
theorem Polynomial.toSubring_one {R : Type u} [Ring R] (T : Subring R) :
toSubring 1 T = 1
@[simp]
theorem Polynomial.map_toSubring {R : Type u} [Ring R] (p : Polynomial R) (T : Subring R) (hp : p.coeffs T) :
map T.subtype (p.toSubring T hp) = p
def Polynomial.ofSubring {R : Type u} [Ring R] (T : Subring R) (p : Polynomial T) :

Given a polynomial whose coefficients are in some subring, return the corresponding polynomial whose coefficients are in the ambient ring.

Equations
theorem Polynomial.coeff_ofSubring {R : Type u} [Ring R] (T : Subring R) (p : Polynomial T) (n : ) :
(ofSubring T p).coeff n = (p.coeff n)
@[simp]
theorem Polynomial.coeffs_ofSubring {R : Type u} [Ring R] (T : Subring R) {p : Polynomial T} :
(ofSubring T p).coeffs T

Transport an ideal of R[X] to an R-submodule of R[X].

Equations
def Ideal.degreeLE {R : Type u} [Semiring R] (I : Ideal (Polynomial R)) (n : WithBot ) :

Given an ideal I of R[X], make the R-submodule of I consisting of polynomials of degree ≤ n.

Equations
def Ideal.leadingCoeffNth {R : Type u} [Semiring R] (I : Ideal (Polynomial R)) (n : ) :

Given an ideal I of R[X], make the ideal in R of leading coefficients of polynomials in I with degree ≤ n.

Equations
def Ideal.leadingCoeff {R : Type u} [Semiring R] (I : Ideal (Polynomial R)) :

Given an ideal I in R[X], make the ideal in R of the leading coefficients in I.

Equations
theorem Ideal.polynomial_mem_ideal_of_coeff_mem_ideal {R : Type u} [CommSemiring R] (I : Ideal (Polynomial R)) (p : Polynomial R) (hp : ∀ (n : ), p.coeff n comap Polynomial.C I) :
p I

If every coefficient of a polynomial is in an ideal I, then so is the polynomial itself

theorem Ideal.mem_map_C_iff {R : Type u} [CommSemiring R] {I : Ideal R} {f : Polynomial R} :
f map Polynomial.C I ∀ (n : ), f.coeff n I

The push-forward of an ideal I of R to R[X] via inclusion is exactly the set of polynomials whose coefficients are in I

theorem Ideal.mem_leadingCoeffNth {R : Type u} [CommSemiring R] (I : Ideal (Polynomial R)) (n : ) (x : R) :
x I.leadingCoeffNth n pI, p.degree n p.leadingCoeff = x
theorem Ideal.mem_leadingCoeff {R : Type u} [CommSemiring R] (I : Ideal (Polynomial R)) (x : R) :
x I.leadingCoeff pI, p.leadingCoeff = x
theorem Polynomial.coeff_prod_mem_ideal_pow_tsub {R : Type u} [CommSemiring R] {ι : Type u_2} (s : Finset ι) (f : ιPolynomial R) (I : Ideal R) (n : ι) (h : is, ∀ (k : ), (f i).coeff k I ^ (n i - k)) (k : ) :
(s.prod f).coeff k I ^ (s.sum n - k)

If I is an ideal, and pᵢ is a finite family of polynomials each satisfying ∀ k, (pᵢ)ₖ ∈ Iⁿⁱ⁻ᵏ for some nᵢ, then p = ∏ pᵢ also satisfies ∀ k, pₖ ∈ Iⁿ⁻ᵏ with n = ∑ nᵢ.

R[X] is never a field for any ring R.

theorem Ideal.eq_zero_of_constant_mem_of_maximal {R : Type u} [Ring R] (hR : IsField R) (I : Ideal (Polynomial R)) [hI : I.IsMaximal] (x : R) (hx : Polynomial.C x I) :
x = 0

The only constant in a maximal ideal over a field is 0.

If P is a prime ideal of R, then P.R[x] is a prime ideal of R[x].

If P is a prime ideal of R, then P.R[x] is a prime ideal of R[x].

theorem Ideal.is_fg_degreeLE {R : Type u} [CommRing R] [IsNoetherianRing R] (I : Ideal (Polynomial R)) (n : ) :
(I.degreeLE n).FG
theorem span_le_of_C_coeff_mem {R : Type u} [Semiring R] {f : Polynomial R} {I : Ideal (Polynomial R)} (cf : ∀ (i : ), Polynomial.C (f.coeff i) I) :
Ideal.span {g : Polynomial R | ∃ (i : ), g = Polynomial.C (f.coeff i)} I

If the coefficients of a polynomial belong to an ideal, then that ideal contains the ideal spanned by the coefficients of the polynomial.

theorem mem_span_C_coeff {R : Type u} [Semiring R] {f : Polynomial R} :
f Ideal.span {g : Polynomial R | ∃ (i : ), g = Polynomial.C (f.coeff i)}
theorem exists_C_coeff_not_mem {R : Type u} [Semiring R] {f : Polynomial R} {I : Ideal (Polynomial R)} :
fI∃ (i : ), Polynomial.C (f.coeff i)I
theorem Polynomial.prime_C_iff {R : Type u} [CommRing R] {r : R} :
theorem MvPolynomial.prime_C_iff {R : Type u} (σ : Type v) [CommRing R] {r : R} :
theorem MvPolynomial.prime_rename_iff {R : Type u} {σ : Type v} [CommRing R] (s : Set σ) {p : MvPolynomial (↑s) R} :

Hilbert basis theorem: a polynomial ring over a Noetherian ring is a Noetherian ring.

theorem Polynomial.linearIndependent_powers_iff_aeval {R : Type u} {M : Type w} [CommRing R] [AddCommGroup M] [Module R M] (f : M →ₗ[R] M) (v : M) :
(LinearIndependent R fun (n : ) => (f ^ n) v) ∀ (p : Polynomial R), ((aeval f) p) v = 0p = 0
theorem Polynomial.disjoint_ker_aeval_of_isCoprime {R : Type u} {M : Type w} [CommRing R] [AddCommGroup M] [Module R M] (f : M →ₗ[R] M) {p q : Polynomial R} (hpq : IsCoprime p q) :
@[deprecated Polynomial.disjoint_ker_aeval_of_isCoprime (since := "2025-01-23")]
theorem Polynomial.disjoint_ker_aeval_of_coprime {R : Type u} {M : Type w} [CommRing R] [AddCommGroup M] [Module R M] (f : M →ₗ[R] M) {p q : Polynomial R} (hpq : IsCoprime p q) :

Alias of Polynomial.disjoint_ker_aeval_of_isCoprime.

theorem Polynomial.sup_aeval_range_eq_top_of_isCoprime {R : Type u} {M : Type w} [CommRing R] [AddCommGroup M] [Module R M] (f : M →ₗ[R] M) {p q : Polynomial R} (hpq : IsCoprime p q) :
@[deprecated Polynomial.sup_aeval_range_eq_top_of_isCoprime (since := "2025-01-23")]
theorem Polynomial.sup_aeval_range_eq_top_of_coprime {R : Type u} {M : Type w} [CommRing R] [AddCommGroup M] [Module R M] (f : M →ₗ[R] M) {p q : Polynomial R} (hpq : IsCoprime p q) :

Alias of Polynomial.sup_aeval_range_eq_top_of_isCoprime.

theorem Polynomial.sup_ker_aeval_le_ker_aeval_mul {R : Type u} {M : Type w} [CommRing R] [AddCommGroup M] [Module R M] {f : M →ₗ[R] M} {p q : Polynomial R} :
LinearMap.ker ((aeval f) p)LinearMap.ker ((aeval f) q) LinearMap.ker ((aeval f) (p * q))
theorem Polynomial.sup_ker_aeval_eq_ker_aeval_mul_of_coprime {R : Type u} {M : Type w} [CommRing R] [AddCommGroup M] [Module R M] (f : M →ₗ[R] M) {p q : Polynomial R} (hpq : IsCoprime p q) :
LinearMap.ker ((aeval f) p)LinearMap.ker ((aeval f) q) = LinearMap.ker ((aeval f) (p * q))
theorem MvPolynomial.aeval_natDegree_le {σ : Type v} {R : Type u_2} [CommSemiring R] {m n : } (F : MvPolynomial σ R) (hF : F.totalDegree m) (f : σPolynomial R) (hf : ∀ (i : σ), (f i).natDegree n) :
((aeval f) F).natDegree m * n

The multivariate polynomial ring in finitely many variables over a noetherian ring is itself a noetherian ring.

Auxiliary lemma: Multivariate polynomials over an integral domain with variables indexed by Fin n form an integral domain. This fact is proven inductively, and then used to prove the general case without any finiteness hypotheses. See MvPolynomial.noZeroDivisors for the general case.

Auxiliary definition: Multivariate polynomials in finitely many variables over an integral domain form an integral domain. This fact is proven by transport of structure from the MvPolynomial.noZeroDivisors_fin, and then used to prove the general case without finiteness hypotheses. See MvPolynomial.noZeroDivisors for the general case.

instance MvPolynomial.isDomain {R : Type u} {σ : Type v} [CommRing R] [IsDomain R] :

The multivariate polynomial ring over an integral domain is an integral domain.

theorem MvPolynomial.map_mvPolynomial_eq_eval₂ {R : Type u} {σ : Type v} [CommRing R] {S : Type u_2} [CommSemiring S] [Finite σ] (ϕ : MvPolynomial σ R →+* S) (p : MvPolynomial σ R) :
ϕ p = eval₂ (ϕ.comp C) (fun (s : σ) => ϕ (X s)) p
theorem MvPolynomial.mem_ideal_of_coeff_mem_ideal {R : Type u} {σ : Type v} [CommRing R] (I : Ideal (MvPolynomial σ R)) (p : MvPolynomial σ R) (hcoe : ∀ (m : σ →₀ ), coeff m p Ideal.comap C I) :
p I

If every coefficient of a polynomial is in an ideal I, then so is the polynomial itself, multivariate version.

theorem MvPolynomial.mem_map_C_iff {R : Type u} {σ : Type v} [CommRing R] {I : Ideal R} {f : MvPolynomial σ R} :
f Ideal.map C I ∀ (m : σ →₀ ), coeff m f I

The push-forward of an ideal I of R to MvPolynomial σ R via inclusion is exactly the set of polynomials whose coefficients are in I

theorem MvPolynomial.ker_map {R : Type u} {S : Type u_1} {σ : Type v} [CommRing R] [CommRing S] (f : R →+* S) :
theorem MvPolynomial.ker_mapAlgHom {R : Type u} [CommRing R] {S₁ : Type u_2} {S₂ : Type u_3} {σ : Type u_4} [CommRing S₁] [CommRing S₂] [Algebra R S₁] [Algebra R S₂] (f : S₁ →ₐ[R] S₂) :