mathlib documentation

data.​polynomial.​eval

data.​polynomial.​eval

Theory of univariate polynomials

The main defs here are eval₂, eval, and map. We give several lemmas about their interaction with each other and with module operations.

def polynomial.​eval₂ {R : Type u} {S : Type v} [semiring R] [semiring S] :
(R →+* S)S → polynomial R → S

Evaluate a polynomial p given a ring hom f from the scalar ring to the target and a value x for the variable in the target

Equations
theorem polynomial.​eval₂_eq_sum {R : Type u} {S : Type v} [semiring R] {p : polynomial R} [semiring S] {f : R →+* S} {x : S} :
polynomial.eval₂ f x p = finsupp.sum p (λ (e : ) (a : R), f a * x ^ e)

@[simp]
theorem polynomial.​eval₂_zero {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) (x : S) :

@[simp]
theorem polynomial.​eval₂_C {R : Type u} {S : Type v} {a : R} [semiring R] [semiring S] (f : R →+* S) (x : S) :

@[simp]
theorem polynomial.​eval₂_X {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) (x : S) :

@[simp]
theorem polynomial.​eval₂_monomial {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) (x : S) {n : } {r : R} :

@[simp]
theorem polynomial.​eval₂_X_pow {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) (x : S) {n : } :

@[simp]
theorem polynomial.​eval₂_add {R : Type u} {S : Type v} [semiring R] {p q : polynomial R} [semiring S] (f : R →+* S) (x : S) :

@[simp]
theorem polynomial.​eval₂_one {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) (x : S) :

@[simp]
theorem polynomial.​eval₂_bit0 {R : Type u} {S : Type v} [semiring R] {p : polynomial R} [semiring S] (f : R →+* S) (x : S) :

@[simp]
theorem polynomial.​eval₂_bit1 {R : Type u} {S : Type v} [semiring R] {p : polynomial R} [semiring S] (f : R →+* S) (x : S) :

@[simp]
theorem polynomial.​eval₂_smul {R : Type u} {S : Type v} [semiring R] [semiring S] (g : R →+* S) (p : polynomial R) (x : S) {s : R} :

@[instance]
def polynomial.​eval₂.​is_add_monoid_hom {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) (x : S) :

Equations
  • _ = _
@[simp]
theorem polynomial.​eval₂_nat_cast {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) (x : S) (n : ) :

theorem polynomial.​eval₂_sum {R : Type u} {S : Type v} {T : Type w} [semiring R] [semiring S] (f : R →+* S) [semiring T] (p : polynomial T) (g : T → polynomial R) (x : S) :
polynomial.eval₂ f x (finsupp.sum p g) = finsupp.sum p (λ (n : ) (a : T), polynomial.eval₂ f x (g n a))

theorem polynomial.​eval₂_finset_sum {R : Type u} {S : Type v} {ι : Type y} [semiring R] [semiring S] (f : R →+* S) (s : finset ι) (g : ι → polynomial R) (x : S) :
polynomial.eval₂ f x (s.sum (λ (i : ι), g i)) = s.sum (λ (i : ι), polynomial.eval₂ f x (g i))

theorem polynomial.​eval₂_mul_noncomm {R : Type u} {S : Type v} [semiring R] {p q : polynomial R} [semiring S] (f : R →+* S) (x : S) :
(∀ (b : R) (a : S), a * f b = f b * a)polynomial.eval₂ f x (p * q) = polynomial.eval₂ f x p * polynomial.eval₂ f x q

theorem polynomial.​eval₂_list_prod_noncomm {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) (x : S) (ps : list (polynomial R)) :
(∀ (b : R) (a : S), a * f b = f b * a)polynomial.eval₂ f x ps.prod = (list.map (polynomial.eval₂ f x) ps).prod

def polynomial.​eval₂_ring_hom_noncomm {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) :
(∀ (b : R) (a : S), a * f b = f b * a)S → polynomial R →+* S

eval₂ as a ring_hom for noncommutative rings

Equations

We next prove that eval₂ is multiplicative as long as target ring is commutative (even if the source ring is not).

@[simp]
theorem polynomial.​eval₂_mul {R : Type u} {S : Type v} [semiring R] {p q : polynomial R} [comm_semiring S] (f : R →+* S) (x : S) :

@[instance]
def polynomial.​eval₂.​is_semiring_hom {R : Type u} {S : Type v} [semiring R] [comm_semiring S] (f : R →+* S) (x : S) :

Equations
  • _ = _
def polynomial.​eval₂_ring_hom {R : Type u} {S : Type v} [semiring R] [comm_semiring S] :
(R →+* S)S → polynomial R →+* S

eval₂ as a ring_hom

Equations
@[simp]
theorem polynomial.​coe_eval₂_ring_hom {R : Type u} {S : Type v} [semiring R] [comm_semiring S] (f : R →+* S) (x : S) :

theorem polynomial.​eval₂_pow {R : Type u} {S : Type v} [semiring R] {p : polynomial R} [comm_semiring S] (f : R →+* S) (x : S) (n : ) :

def polynomial.​eval {R : Type u} [semiring R] :
R → polynomial R → R

eval x p is the evaluation of the polynomial p at x

Equations
theorem polynomial.​eval_eq_sum {R : Type u} [semiring R] {p : polynomial R} {x : R} :
polynomial.eval x p = finsupp.sum p (λ (e : ) (a : R), a * x ^ e)

@[simp]
theorem polynomial.​eval_C {R : Type u} {a : R} [semiring R] {x : R} :

@[simp]
theorem polynomial.​eval_nat_cast {R : Type u} [semiring R] {x : R} {n : } :

@[simp]
theorem polynomial.​eval_X {R : Type u} [semiring R] {x : R} :

@[simp]
theorem polynomial.​eval_monomial {R : Type u} [semiring R] {x : R} {n : } {a : R} :

@[simp]
theorem polynomial.​eval_zero {R : Type u} [semiring R] {x : R} :

@[simp]
theorem polynomial.​eval_add {R : Type u} [semiring R] {p q : polynomial R} {x : R} :

@[simp]
theorem polynomial.​eval_one {R : Type u} [semiring R] {x : R} :

@[simp]
theorem polynomial.​eval_bit0 {R : Type u} [semiring R] {p : polynomial R} {x : R} :

@[simp]
theorem polynomial.​eval_bit1 {R : Type u} [semiring R] {p : polynomial R} {x : R} :

@[simp]
theorem polynomial.​eval_smul {R : Type u} [semiring R] (p : polynomial R) (x : R) {s : R} :

theorem polynomial.​eval_sum {R : Type u} [semiring R] (p : polynomial R) (f : R → polynomial R) (x : R) :
polynomial.eval x (finsupp.sum p f) = finsupp.sum p (λ (n : ) (a : R), polynomial.eval x (f n a))

def polynomial.​is_root {R : Type u} [semiring R] :
polynomial RR → Prop

is_root p x implies x is a root of p. The evaluation of p at x is zero

Equations
@[instance]
def polynomial.​decidable {R : Type u} {a : R} [semiring R] {p : polynomial R} [decidable_eq R] :

Equations
@[simp]
theorem polynomial.​is_root.​def {R : Type u} {a : R} [semiring R] {p : polynomial R} :

theorem polynomial.​zero_is_root_of_coeff_zero_eq_zero {R : Type u} [semiring R] {p : polynomial R} :
p.coeff 0 = 0p.is_root 0

def polynomial.​comp {R : Type u} [semiring R] :

The composition of polynomials as a polynomial.

Equations
theorem polynomial.​comp_eq_sum_left {R : Type u} [semiring R] {p q : polynomial R} :
p.comp q = finsupp.sum p (λ (e : ) (a : R), polynomial.C a * q ^ e)

@[simp]
theorem polynomial.​comp_X {R : Type u} [semiring R] {p : polynomial R} :

@[simp]
theorem polynomial.​X_comp {R : Type u} [semiring R] {p : polynomial R} :

@[simp]
theorem polynomial.​comp_C {R : Type u} {a : R} [semiring R] {p : polynomial R} :

@[simp]
theorem polynomial.​C_comp {R : Type u} {a : R} [semiring R] {p : polynomial R} :

@[simp]
theorem polynomial.​comp_zero {R : Type u} [semiring R] {p : polynomial R} :

@[simp]
theorem polynomial.​zero_comp {R : Type u} [semiring R] {p : polynomial R} :
0.comp p = 0

@[simp]
theorem polynomial.​comp_one {R : Type u} [semiring R] {p : polynomial R} :

@[simp]
theorem polynomial.​one_comp {R : Type u} [semiring R] {p : polynomial R} :
1.comp p = 1

@[simp]
theorem polynomial.​add_comp {R : Type u} [semiring R] {p q r : polynomial R} :
(p + q).comp r = p.comp r + q.comp r

def polynomial.​map {R : Type u} {S : Type v} [semiring R] [semiring S] :
(R →+* S)polynomial Rpolynomial S

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

Equations
@[instance]
def polynomial.​is_semiring_hom_C_f {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) :

Equations
  • _ = _
@[simp]
theorem polynomial.​map_C {R : Type u} {S : Type v} {a : R} [semiring R] [semiring S] (f : R →+* S) :

@[simp]
theorem polynomial.​map_X {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) :

@[simp]
theorem polynomial.​map_monomial {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) {n : } {a : R} :

@[simp]
theorem polynomial.​map_zero {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) :

@[simp]
theorem polynomial.​map_add {R : Type u} {S : Type v} [semiring R] {p q : polynomial R} [semiring S] (f : R →+* S) :

@[simp]
theorem polynomial.​map_one {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) :

@[simp]
theorem polynomial.​map_nat_cast {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) (n : ) :

@[simp]
theorem polynomial.​coeff_map {R : Type u} {S : Type v} [semiring R] {p : polynomial R} [semiring S] (f : R →+* S) (n : ) :
(polynomial.map f p).coeff n = f (p.coeff n)

theorem polynomial.​map_map {R : Type u} {S : Type v} {T : Type w} [semiring R] [semiring S] (f : R →+* S) [semiring T] (g : S →+* T) (p : polynomial R) :

@[simp]
theorem polynomial.​map_id {R : Type u} [semiring R] {p : polynomial R} :

theorem polynomial.​eval₂_eq_eval_map {R : Type u} {S : Type v} [semiring R] {p : polynomial R} [semiring S] (f : R →+* S) {x : S} :

theorem polynomial.​map_injective {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) :

theorem polynomial.​map_monic_eq_zero_iff {R : Type u} {S : Type v} [semiring R] {p : polynomial R} [semiring S] {f : R →+* S} :
p.monic(polynomial.map f p = 0 ∀ (x : R), f x = 0)

theorem polynomial.​map_monic_ne_zero {R : Type u} {S : Type v} [semiring R] {p : polynomial R} [semiring S] {f : R →+* S} (hp : p.monic) [nontrivial S] :

@[simp]
theorem polynomial.​map_mul {R : Type u} {S : Type v} [semiring R] {p q : polynomial R} [semiring S] (f : R →+* S) :

@[instance]
def polynomial.​map.​is_semiring_hom {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) :

Equations
  • _ = _
theorem polynomial.​map_list_prod {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) (L : list (polynomial R)) :

@[simp]
theorem polynomial.​map_pow {R : Type u} {S : Type v} [semiring R] {p : polynomial R} [semiring S] (f : R →+* S) (n : ) :

theorem polynomial.​mem_map_range {R : Type u} {S : Type v} [semiring R] [semiring S] (f : R →+* S) {p : polynomial S} :

theorem polynomial.​eval₂_map {R : Type u} {S : Type v} {T : Type w} [semiring R] {p : polynomial R} [semiring S] (f : R →+* S) [semiring T] (g : S →+* T) (x : T) :

theorem polynomial.​eval_map {R : Type u} {S : Type v} [semiring R] {p : polynomial R} [semiring S] (f : R →+* S) (x : S) :

After having set up the basic theory of eval₂, eval, comp, and map, we make eval₂ irreducible.

Perhaps we can make the others irreducible too?

theorem polynomial.​hom_eval₂ {R : Type u} {S : Type v} {T : Type w} [semiring R] (p : polynomial R) [comm_semiring S] [comm_semiring T] (f : R →+* S) (g : S →+* T) (x : S) :

@[simp]
theorem polynomial.​eval_mul {R : Type u} [comm_semiring R] {p q : polynomial R} {x : R} :

@[simp]
theorem polynomial.​eval_pow {R : Type u} [comm_semiring R] {p : polynomial R} {x : R} (n : ) :

theorem polynomial.​eval₂_hom {R : Type u} {S : Type v} [comm_semiring R] {p : polynomial R} [comm_semiring S] (f : R →+* S) (x : R) :

theorem polynomial.​root_mul_left_of_is_root {R : Type u} {a : R} [comm_semiring R] (p : polynomial R) {q : polynomial R} :
q.is_root a(p * q).is_root a

theorem polynomial.​root_mul_right_of_is_root {R : Type u} {a : R} [comm_semiring R] {p : polynomial R} (q : polynomial R) :
p.is_root a(p * q).is_root a

theorem polynomial.​eval_prod {R : Type u} [comm_semiring R] {ι : Type u_1} (s : finset ι) (p : ι → polynomial R) (x : R) :
polynomial.eval x (s.prod (λ (j : ι), p j)) = s.prod (λ (j : ι), polynomial.eval x (p j))

Polynomial evaluation commutes with finset.prod

theorem polynomial.​map_multiset_prod {R : Type u} {S : Type v} [comm_semiring R] [comm_semiring S] (f : R →+* S) (m : multiset (polynomial R)) :

theorem polynomial.​map_prod {R : Type u} {S : Type v} [comm_semiring R] [comm_semiring S] (f : R →+* S) {ι : Type u_1} (g : ι → polynomial R) (s : finset ι) :
polynomial.map f (s.prod (λ (i : ι), g i)) = s.prod (λ (i : ι), polynomial.map f (g i))

theorem polynomial.​map_sum {R : Type u} {S : Type v} [comm_semiring R] [comm_semiring S] (f : R →+* S) {ι : Type u_1} (g : ι → polynomial R) (s : finset ι) :
polynomial.map f (s.sum (λ (i : ι), g i)) = s.sum (λ (i : ι), polynomial.map f (g i))

theorem polynomial.​support_map_subset {R : Type u} {S : Type v} [comm_semiring R] [comm_semiring S] (f : R →+* S) (p : polynomial R) :

theorem polynomial.​map_comp {R : Type u} {S : Type v} [comm_semiring R] [comm_semiring S] (f : R →+* S) (p q : polynomial R) :

theorem polynomial.​C_neg {R : Type u} {a : R} [ring R] :

theorem polynomial.​C_sub {R : Type u} {a b : R} [ring R] :

@[instance]
def polynomial.​map.​is_ring_hom {R : Type u} [ring R] {S : Type u_1} [ring S] (f : R →+* S) :

Equations
  • _ = _
@[simp]
theorem polynomial.​map_sub {R : Type u} [ring R] {p q : polynomial R} {S : Type u_1} [comm_ring S] (f : R →+* S) :

@[simp]
theorem polynomial.​map_neg {R : Type u} [ring R] {p : polynomial R} {S : Type u_1} [comm_ring S] (f : R →+* S) :

@[simp]
theorem polynomial.​eval_int_cast {R : Type u} [ring R] {n : } {x : R} :

@[simp]
theorem polynomial.​eval₂_neg {R : Type u} [ring R] {p : polynomial R} {S : Type u_1} [ring S] (f : R →+* S) {x : S} :

@[simp]
theorem polynomial.​eval₂_sub {R : Type u} [ring R] {p q : polynomial R} {S : Type u_1} [ring S] (f : R →+* S) {x : S} :

@[simp]
theorem polynomial.​eval_neg {R : Type u} [ring R] (p : polynomial R) (x : R) :

@[simp]
theorem polynomial.​eval_sub {R : Type u} [ring R] (p q : polynomial R) (x : R) :

theorem polynomial.​root_X_sub_C {R : Type u} {a b : R} [ring R] :

@[instance]
def polynomial.​eval₂.​is_ring_hom {R : Type u} [comm_ring R] {S : Type u_1} [comm_ring S] (f : R →+* S) {x : S} :

Equations
  • _ = _
@[instance]

Equations
theorem polynomial.​eval₂_endomorphism_algebra_map {R : Type u} [comm_ring R] {M : Type w} [add_comm_group M] [module R M] (f : M →ₗ[R] M) (v : M) (p : polynomial R) :
(polynomial.eval₂ (algebra_map R (M →ₗ[R] M)) f p) v = finsupp.sum p (λ (n : ) (b : R), b (f ^ n) v)