Documentation

Mathlib.Algebra.Polynomial.Degree.Operations

Lemmas for calculating the degree of univariate polynomials #

Main results #

theorem Polynomial.supDegree_eq_degree {R : Type u} [Semiring R] (p : Polynomial R) :
AddMonoidAlgebra.supDegree WithBot.some p.toFinsupp = p.degree
theorem Polynomial.degree_lt_wf {R : Type u} [Semiring R] :
WellFounded fun (p q : Polynomial R) => p.degree < q.degree
Equations
  • Polynomial.instWellFoundedRelation = { rel := fun (p q : Polynomial R) => p.degree < q.degree, wf := }
theorem Polynomial.natDegree_of_subsingleton {R : Type u} [Semiring R] {p : Polynomial R} [Subsingleton R] :
p.natDegree = 0
theorem Polynomial.le_natDegree_of_ne_zero {R : Type u} {n : } [Semiring R] {p : Polynomial R} (h : p.coeff n 0) :
n p.natDegree
theorem Polynomial.degree_eq_of_le_of_coeff_ne_zero {R : Type u} {n : } [Semiring R] {p : Polynomial R} (pn : p.degree n) (p1 : p.coeff n 0) :
p.degree = n
theorem Polynomial.natDegree_eq_of_le_of_coeff_ne_zero {R : Type u} {n : } [Semiring R] {p : Polynomial R} (pn : p.natDegree n) (p1 : p.coeff n 0) :
p.natDegree = n
theorem Polynomial.natDegree_lt_natDegree {R : Type u} [Semiring R] {p q : Polynomial R} (hp : p 0) (hpq : p.degree < q.degree) :
p.natDegree < q.natDegree
theorem Polynomial.coeff_eq_zero_of_degree_lt {R : Type u} {n : } [Semiring R] {p : Polynomial R} (h : p.degree < n) :
p.coeff n = 0
theorem Polynomial.coeff_eq_zero_of_natDegree_lt {R : Type u} [Semiring R] {p : Polynomial R} {n : } (h : p.natDegree < n) :
p.coeff n = 0
theorem Polynomial.ext_iff_natDegree_le {R : Type u} [Semiring R] {p q : Polynomial R} {n : } (hp : p.natDegree n) (hq : q.natDegree n) :
p = q in, p.coeff i = q.coeff i
theorem Polynomial.ext_iff_degree_le {R : Type u} [Semiring R] {p q : Polynomial R} {n : } (hp : p.degree n) (hq : q.degree n) :
p = q in, p.coeff i = q.coeff i
@[simp]
theorem Polynomial.coeff_natDegree_succ_eq_zero {R : Type u} [Semiring R] {p : Polynomial R} :
p.coeff (p.natDegree + 1) = 0
theorem Polynomial.ite_le_natDegree_coeff {R : Type u} [Semiring R] (p : Polynomial R) (n : ) (I : Decidable (n < 1 + p.natDegree)) :
(if n < 1 + p.natDegree then p.coeff n else 0) = p.coeff n
theorem Polynomial.coeff_mul_X_sub_C {R : Type u} [Ring R] {p : Polynomial R} {r : R} {a : } :
(p * (Polynomial.X - Polynomial.C r)).coeff (a + 1) = p.coeff a - p.coeff (a + 1) * r
theorem Polynomial.coeff_natDegree_eq_zero_of_degree_lt {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.degree < q.degree) :
p.coeff q.natDegree = 0
theorem Polynomial.ne_zero_of_degree_gt {R : Type u} [Semiring R] {p : Polynomial R} {n : WithBot } (h : n < p.degree) :
p 0
theorem Polynomial.ne_zero_of_degree_ge_degree {R : Type u} [Semiring R] {p q : Polynomial R} (hpq : p.degree q.degree) (hp : p 0) :
q 0
theorem Polynomial.ne_zero_of_natDegree_gt {R : Type u} [Semiring R] {p : Polynomial R} {n : } (h : n < p.natDegree) :
p 0
theorem Polynomial.degree_lt_degree {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.natDegree < q.natDegree) :
p.degree < q.degree
theorem Polynomial.natDegree_lt_natDegree_iff {R : Type u} [Semiring R] {p q : Polynomial R} (hp : p 0) :
p.natDegree < q.natDegree p.degree < q.degree
theorem Polynomial.eq_C_of_degree_le_zero {R : Type u} [Semiring R] {p : Polynomial R} (h : p.degree 0) :
p = Polynomial.C (p.coeff 0)
theorem Polynomial.eq_C_of_degree_eq_zero {R : Type u} [Semiring R] {p : Polynomial R} (h : p.degree = 0) :
p = Polynomial.C (p.coeff 0)
theorem Polynomial.degree_le_zero_iff {R : Type u} [Semiring R] {p : Polynomial R} :
p.degree 0 p = Polynomial.C (p.coeff 0)
theorem Polynomial.degree_add_eq_left_of_degree_lt {R : Type u} [Semiring R] {p q : Polynomial R} (h : q.degree < p.degree) :
(p + q).degree = p.degree
theorem Polynomial.degree_add_eq_right_of_degree_lt {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.degree < q.degree) :
(p + q).degree = q.degree
theorem Polynomial.natDegree_add_eq_left_of_degree_lt {R : Type u} [Semiring R] {p q : Polynomial R} (h : q.degree < p.degree) :
(p + q).natDegree = p.natDegree
theorem Polynomial.natDegree_add_eq_left_of_natDegree_lt {R : Type u} [Semiring R] {p q : Polynomial R} (h : q.natDegree < p.natDegree) :
(p + q).natDegree = p.natDegree
theorem Polynomial.natDegree_add_eq_right_of_degree_lt {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.degree < q.degree) :
(p + q).natDegree = q.natDegree
theorem Polynomial.natDegree_add_eq_right_of_natDegree_lt {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.natDegree < q.natDegree) :
(p + q).natDegree = q.natDegree
theorem Polynomial.degree_add_C {R : Type u} {a : R} [Semiring R] {p : Polynomial R} (hp : 0 < p.degree) :
(p + Polynomial.C a).degree = p.degree
@[simp]
theorem Polynomial.natDegree_add_C {R : Type u} [Semiring R] {p : Polynomial R} {a : R} :
(p + Polynomial.C a).natDegree = p.natDegree
@[simp]
theorem Polynomial.natDegree_C_add {R : Type u} [Semiring R] {p : Polynomial R} {a : R} :
(Polynomial.C a + p).natDegree = p.natDegree
theorem Polynomial.degree_add_eq_of_leadingCoeff_add_ne_zero {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.leadingCoeff + q.leadingCoeff 0) :
(p + q).degree = p.degree q.degree
theorem Polynomial.natDegree_eq_of_natDegree_add_lt_left {R : Type u} [Semiring R] (p q : Polynomial R) (H : (p + q).natDegree < p.natDegree) :
p.natDegree = q.natDegree
theorem Polynomial.natDegree_eq_of_natDegree_add_lt_right {R : Type u} [Semiring R] (p q : Polynomial R) (H : (p + q).natDegree < q.natDegree) :
p.natDegree = q.natDegree
theorem Polynomial.natDegree_eq_of_natDegree_add_eq_zero {R : Type u} [Semiring R] (p q : Polynomial R) (H : (p + q).natDegree = 0) :
p.natDegree = q.natDegree
theorem Polynomial.monic_of_natDegree_le_of_coeff_eq_one {R : Type u} [Semiring R] {p : Polynomial R} (n : ) (pn : p.natDegree n) (p1 : p.coeff n = 1) :
p.Monic
theorem Polynomial.monic_of_degree_le_of_coeff_eq_one {R : Type u} [Semiring R] {p : Polynomial R} (n : ) (pn : p.degree n) (p1 : p.coeff n = 1) :
p.Monic
theorem Polynomial.leadingCoeff_add_of_degree_lt {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.degree < q.degree) :
(p + q).leadingCoeff = q.leadingCoeff
theorem Polynomial.leadingCoeff_add_of_degree_lt' {R : Type u} [Semiring R] {p q : Polynomial R} (h : q.degree < p.degree) :
(p + q).leadingCoeff = p.leadingCoeff
theorem Polynomial.leadingCoeff_add_of_degree_eq {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.degree = q.degree) (hlc : p.leadingCoeff + q.leadingCoeff 0) :
(p + q).leadingCoeff = p.leadingCoeff + q.leadingCoeff
@[simp]
theorem Polynomial.coeff_mul_degree_add_degree {R : Type u} [Semiring R] (p q : Polynomial R) :
(p * q).coeff (p.natDegree + q.natDegree) = p.leadingCoeff * q.leadingCoeff
theorem Polynomial.degree_mul' {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.leadingCoeff * q.leadingCoeff 0) :
(p * q).degree = p.degree + q.degree
theorem Polynomial.Monic.degree_mul {R : Type u} [Semiring R] {p q : Polynomial R} (hq : q.Monic) :
(p * q).degree = p.degree + q.degree
theorem Polynomial.natDegree_mul' {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.leadingCoeff * q.leadingCoeff 0) :
(p * q).natDegree = p.natDegree + q.natDegree
theorem Polynomial.leadingCoeff_mul' {R : Type u} [Semiring R] {p q : Polynomial R} (h : p.leadingCoeff * q.leadingCoeff 0) :
(p * q).leadingCoeff = p.leadingCoeff * q.leadingCoeff
theorem Polynomial.leadingCoeff_pow' {R : Type u} {n : } [Semiring R] {p : Polynomial R} :
p.leadingCoeff ^ n 0(p ^ n).leadingCoeff = p.leadingCoeff ^ n
theorem Polynomial.degree_pow' {R : Type u} [Semiring R] {p : Polynomial R} {n : } :
p.leadingCoeff ^ n 0(p ^ n).degree = n p.degree
theorem Polynomial.natDegree_pow' {R : Type u} [Semiring R] {p : Polynomial R} {n : } (h : p.leadingCoeff ^ n 0) :
(p ^ n).natDegree = n * p.natDegree
theorem Polynomial.leadingCoeff_monic_mul {R : Type u} [Semiring R] {p q : Polynomial R} (hp : p.Monic) :
(p * q).leadingCoeff = q.leadingCoeff
theorem Polynomial.leadingCoeff_mul_monic {R : Type u} [Semiring R] {p q : Polynomial R} (hq : q.Monic) :
(p * q).leadingCoeff = p.leadingCoeff
@[simp]
theorem Polynomial.leadingCoeff_mul_X_pow {R : Type u} [Semiring R] {p : Polynomial R} {n : } :
(p * Polynomial.X ^ n).leadingCoeff = p.leadingCoeff
@[simp]
theorem Polynomial.leadingCoeff_mul_X {R : Type u} [Semiring R] {p : Polynomial R} :
(p * Polynomial.X).leadingCoeff = p.leadingCoeff
@[simp]
theorem Polynomial.coeff_pow_mul_natDegree {R : Type u} [Semiring R] (p : Polynomial R) (n : ) :
(p ^ n).coeff (n * p.natDegree) = p.leadingCoeff ^ n
theorem Polynomial.coeff_mul_add_eq_of_natDegree_le {R : Type u} [Semiring R] {df dg : } {f g : Polynomial R} (hdf : f.natDegree df) (hdg : g.natDegree dg) :
(f * g).coeff (df + dg) = f.coeff df * g.coeff dg
theorem Polynomial.degree_smul_le {R : Type u} [Semiring R] (a : R) (p : Polynomial R) :
(a p).degree p.degree
theorem Polynomial.natDegree_smul_le {R : Type u} [Semiring R] (a : R) (p : Polynomial R) :
(a p).natDegree p.natDegree
theorem Polynomial.degree_lt_degree_mul_X {R : Type u} [Semiring R] {p : Polynomial R} (hp : p 0) :
p.degree < (p * Polynomial.X).degree
theorem Polynomial.eq_C_of_natDegree_le_zero {R : Type u} [Semiring R] {p : Polynomial R} (h : p.natDegree 0) :
p = Polynomial.C (p.coeff 0)
theorem Polynomial.eq_C_of_natDegree_eq_zero {R : Type u} [Semiring R] {p : Polynomial R} (h : p.natDegree = 0) :
p = Polynomial.C (p.coeff 0)
theorem Polynomial.natDegree_eq_zero {R : Type u} [Semiring R] {p : Polynomial R} :
p.natDegree = 0 ∃ (x : R), Polynomial.C x = p
theorem Polynomial.eq_C_coeff_zero_iff_natDegree_eq_zero {R : Type u} [Semiring R] {p : Polynomial R} :
p = Polynomial.C (p.coeff 0) p.natDegree = 0
theorem Polynomial.eq_one_of_monic_natDegree_zero {R : Type u} [Semiring R] {p : Polynomial R} (hf : p.Monic) (hfd : p.natDegree = 0) :
p = 1
theorem Polynomial.Monic.natDegree_eq_zero {R : Type u} [Semiring R] {p : Polynomial R} (hf : p.Monic) :
p.natDegree = 0 p = 1
theorem Polynomial.degree_sum_fin_lt {R : Type u} [Semiring R] {n : } (f : Fin nR) :
(∑ i : Fin n, Polynomial.C (f i) * Polynomial.X ^ i).degree < n
theorem Polynomial.degree_C_lt_degree_C_mul_X {R : Type u} {a b : R} [Semiring R] (ha : a 0) :
(Polynomial.C b).degree < (Polynomial.C a * Polynomial.X).degree
@[simp]
theorem Polynomial.natDegree_mul_X {R : Type u} [Semiring R] [Nontrivial R] {p : Polynomial R} (hp : p 0) :
(p * Polynomial.X).natDegree = p.natDegree + 1
@[simp]
theorem Polynomial.natDegree_X_mul {R : Type u} [Semiring R] [Nontrivial R] {p : Polynomial R} (hp : p 0) :
(Polynomial.X * p).natDegree = p.natDegree + 1
@[simp]
theorem Polynomial.natDegree_mul_X_pow {R : Type u} [Semiring R] [Nontrivial R] {p : Polynomial R} (n : ) (hp : p 0) :
(p * Polynomial.X ^ n).natDegree = p.natDegree + n
@[simp]
theorem Polynomial.natDegree_X_pow_mul {R : Type u} [Semiring R] [Nontrivial R] {p : Polynomial R} (n : ) (hp : p 0) :
(Polynomial.X ^ n * p).natDegree = p.natDegree + n
theorem Polynomial.natDegree_X_pow_le {R : Type u_1} [Semiring R] (n : ) :
(Polynomial.X ^ n).natDegree n
theorem Polynomial.not_isUnit_X {R : Type u} [Semiring R] [Nontrivial R] :
¬IsUnit Polynomial.X
@[simp]
theorem Polynomial.degree_mul_X {R : Type u} [Semiring R] [Nontrivial R] {p : Polynomial R} :
(p * Polynomial.X).degree = p.degree + 1
@[simp]
theorem Polynomial.degree_mul_X_pow {R : Type u} [Semiring R] [Nontrivial R] {p : Polynomial R} (n : ) :
(p * Polynomial.X ^ n).degree = p.degree + n
theorem Polynomial.degree_sub_C {R : Type u} {a : R} [Ring R] {p : Polynomial R} (hp : 0 < p.degree) :
(p - Polynomial.C a).degree = p.degree
@[simp]
theorem Polynomial.natDegree_sub_C {R : Type u} [Ring R] {p : Polynomial R} {a : R} :
(p - Polynomial.C a).natDegree = p.natDegree
theorem Polynomial.leadingCoeff_sub_of_degree_lt {R : Type u} [Ring R] {p q : Polynomial R} (h : q.degree < p.degree) :
(p - q).leadingCoeff = p.leadingCoeff
theorem Polynomial.leadingCoeff_sub_of_degree_lt' {R : Type u} [Ring R] {p q : Polynomial R} (h : p.degree < q.degree) :
(p - q).leadingCoeff = -q.leadingCoeff
theorem Polynomial.leadingCoeff_sub_of_degree_eq {R : Type u} [Ring R] {p q : Polynomial R} (h : p.degree = q.degree) (hlc : p.leadingCoeff q.leadingCoeff) :
(p - q).leadingCoeff = p.leadingCoeff - q.leadingCoeff
theorem Polynomial.degree_sub_eq_left_of_degree_lt {R : Type u} [Ring R] {p q : Polynomial R} (h : q.degree < p.degree) :
(p - q).degree = p.degree
theorem Polynomial.degree_sub_eq_right_of_degree_lt {R : Type u} [Ring R] {p q : Polynomial R} (h : p.degree < q.degree) :
(p - q).degree = q.degree
theorem Polynomial.natDegree_sub_eq_left_of_natDegree_lt {R : Type u} [Ring R] {p q : Polynomial R} (h : q.natDegree < p.natDegree) :
(p - q).natDegree = p.natDegree
theorem Polynomial.natDegree_sub_eq_right_of_natDegree_lt {R : Type u} [Ring R] {p q : Polynomial R} (h : p.natDegree < q.natDegree) :
(p - q).natDegree = q.natDegree
@[simp]
theorem Polynomial.degree_X_add_C {R : Type u} [Nontrivial R] [Semiring R] (a : R) :
(Polynomial.X + Polynomial.C a).degree = 1
theorem Polynomial.natDegree_X_add_C {R : Type u} [Nontrivial R] [Semiring R] (x : R) :
(Polynomial.X + Polynomial.C x).natDegree = 1
@[simp]
theorem Polynomial.nextCoeff_X_add_C {S : Type v} [Semiring S] (c : S) :
(Polynomial.X + Polynomial.C c).nextCoeff = c
theorem Polynomial.degree_X_pow_add_C {R : Type u} [Nontrivial R] [Semiring R] {n : } (hn : 0 < n) (a : R) :
(Polynomial.X ^ n + Polynomial.C a).degree = n
theorem Polynomial.X_pow_add_C_ne_zero {R : Type u} [Nontrivial R] [Semiring R] {n : } (hn : 0 < n) (a : R) :
Polynomial.X ^ n + Polynomial.C a 0
theorem Polynomial.X_add_C_ne_zero {R : Type u} [Nontrivial R] [Semiring R] (r : R) :
Polynomial.X + Polynomial.C r 0
theorem Polynomial.zero_nmem_multiset_map_X_add_C {R : Type u} [Nontrivial R] [Semiring R] {α : Type u_1} (m : Multiset α) (f : αR) :
0Multiset.map (fun (a : α) => Polynomial.X + Polynomial.C (f a)) m
theorem Polynomial.natDegree_X_pow_add_C {R : Type u} [Nontrivial R] [Semiring R] {n : } {r : R} :
(Polynomial.X ^ n + Polynomial.C r).natDegree = n
theorem Polynomial.X_pow_add_C_ne_one {R : Type u} [Nontrivial R] [Semiring R] {n : } (hn : 0 < n) (a : R) :
Polynomial.X ^ n + Polynomial.C a 1
theorem Polynomial.X_add_C_ne_one {R : Type u} [Nontrivial R] [Semiring R] (r : R) :
Polynomial.X + Polynomial.C r 1
@[simp]
theorem Polynomial.leadingCoeff_X_pow_add_C {R : Type u} [Semiring R] {n : } (hn : 0 < n) {r : R} :
(Polynomial.X ^ n + Polynomial.C r).leadingCoeff = 1
@[simp]
theorem Polynomial.leadingCoeff_X_add_C {S : Type v} [Semiring S] (r : S) :
(Polynomial.X + Polynomial.C r).leadingCoeff = 1
@[simp]
theorem Polynomial.leadingCoeff_X_pow_add_one {R : Type u} [Semiring R] {n : } (hn : 0 < n) :
(Polynomial.X ^ n + 1).leadingCoeff = 1
@[simp]
theorem Polynomial.leadingCoeff_pow_X_add_C {R : Type u} [Semiring R] (r : R) (i : ) :
((Polynomial.X + Polynomial.C r) ^ i).leadingCoeff = 1
@[simp]
theorem Polynomial.degree_mul {R : Type u} [Semiring R] [NoZeroDivisors R] {p q : Polynomial R} :
(p * q).degree = p.degree + q.degree

degree as a monoid homomorphism between R[X] and Multiplicative (WithBot ℕ). This is useful to prove results about multiplication and degree.

Equations
  • Polynomial.degreeMonoidHom = { toFun := Polynomial.degree, map_one' := , map_mul' := }
Instances For
    @[simp]
    theorem Polynomial.degree_pow {R : Type u} [Semiring R] [NoZeroDivisors R] [Nontrivial R] (p : Polynomial R) (n : ) :
    (p ^ n).degree = n p.degree
    @[simp]
    theorem Polynomial.leadingCoeff_mul {R : Type u} [Semiring R] [NoZeroDivisors R] (p q : Polynomial R) :
    (p * q).leadingCoeff = p.leadingCoeff * q.leadingCoeff

    Polynomial.leadingCoeff bundled as a MonoidHom when R has NoZeroDivisors, and thus leadingCoeff is multiplicative

    Equations
    • Polynomial.leadingCoeffHom = { toFun := Polynomial.leadingCoeff, map_one' := , map_mul' := }
    Instances For
      @[simp]
      theorem Polynomial.leadingCoeffHom_apply {R : Type u} [Semiring R] [NoZeroDivisors R] (p : Polynomial R) :
      Polynomial.leadingCoeffHom p = p.leadingCoeff
      @[simp]
      theorem Polynomial.leadingCoeff_pow {R : Type u} [Semiring R] [NoZeroDivisors R] (p : Polynomial R) (n : ) :
      (p ^ n).leadingCoeff = p.leadingCoeff ^ n
      theorem Polynomial.leadingCoeff_dvd_leadingCoeff {R : Type u} [Semiring R] [NoZeroDivisors R] {a p : Polynomial R} (hap : a p) :
      a.leadingCoeff p.leadingCoeff
      theorem Polynomial.degree_le_mul_left {R : Type u} [Semiring R] [NoZeroDivisors R] {q : Polynomial R} (p : Polynomial R) (hq : q 0) :
      p.degree (p * q).degree
      theorem Polynomial.Monic.natDegree_pos {R : Type u} [CommSemiring R] {p : Polynomial R} (hp : p.Monic) :
      0 < p.natDegree p 1
      theorem Polynomial.Monic.degree_pos {R : Type u} [CommSemiring R] {p : Polynomial R} (hp : p.Monic) :
      0 < p.degree p 1
      @[simp]
      theorem Polynomial.leadingCoeff_X_pow_sub_C {R : Type u} [Ring R] {n : } (hn : 0 < n) {r : R} :
      (Polynomial.X ^ n - Polynomial.C r).leadingCoeff = 1
      @[simp]
      theorem Polynomial.leadingCoeff_X_pow_sub_one {R : Type u} [Ring R] {n : } (hn : 0 < n) :
      (Polynomial.X ^ n - 1).leadingCoeff = 1
      @[simp]
      theorem Polynomial.degree_X_sub_C {R : Type u} [Ring R] [Nontrivial R] (a : R) :
      (Polynomial.X - Polynomial.C a).degree = 1
      theorem Polynomial.natDegree_X_sub_C {R : Type u} [Ring R] [Nontrivial R] (x : R) :
      (Polynomial.X - Polynomial.C x).natDegree = 1
      @[simp]
      theorem Polynomial.nextCoeff_X_sub_C {S : Type v} [Ring S] (c : S) :
      (Polynomial.X - Polynomial.C c).nextCoeff = -c
      theorem Polynomial.degree_X_pow_sub_C {R : Type u} [Ring R] [Nontrivial R] {n : } (hn : 0 < n) (a : R) :
      (Polynomial.X ^ n - Polynomial.C a).degree = n
      theorem Polynomial.X_pow_sub_C_ne_zero {R : Type u} [Ring R] [Nontrivial R] {n : } (hn : 0 < n) (a : R) :
      Polynomial.X ^ n - Polynomial.C a 0
      theorem Polynomial.X_sub_C_ne_zero {R : Type u} [Ring R] [Nontrivial R] (r : R) :
      Polynomial.X - Polynomial.C r 0
      theorem Polynomial.zero_nmem_multiset_map_X_sub_C {R : Type u} [Ring R] [Nontrivial R] {α : Type u_1} (m : Multiset α) (f : αR) :
      0Multiset.map (fun (a : α) => Polynomial.X - Polynomial.C (f a)) m
      theorem Polynomial.natDegree_X_pow_sub_C {R : Type u} [Ring R] [Nontrivial R] {n : } {r : R} :
      (Polynomial.X ^ n - Polynomial.C r).natDegree = n
      @[simp]
      theorem Polynomial.leadingCoeff_X_sub_C {S : Type v} [Ring S] (r : S) :
      (Polynomial.X - Polynomial.C r).leadingCoeff = 1