Documentation

Mathlib.RingTheory.Polynomial.Vieta

Vieta's Formula #

The main result is Multiset.prod_X_add_C_eq_sum_esymm, which shows that the product of linear terms X + λ with λ in a Multiset s is equal to a linear combination of the symmetric functions esymm s.

From this, we deduce MvPolynomial.prod_X_add_C_eq_sum_esymm which is the equivalent formula for the product of linear terms X + X i with i in a Fintype σ as a linear combination of the symmetric polynomials esymm σ R j.

For R be an integral domain (so that p.roots is defined for any p : R[X] as a multiset), we derive Polynomial.coeff_eq_esymm_roots_of_card, the relationship between the coefficients and the roots of p for a polynomial p that splits (i.e. having as many roots as its degree).

theorem Multiset.prod_X_add_C_eq_sum_esymm {R : Type u_1} [CommSemiring R] (s : Multiset R) :
(Multiset.map (fun (r : R) => Polynomial.X + Polynomial.C r) s).prod = (Finset.range (Multiset.card s + 1)).sum fun (j : ) => Polynomial.C (s.esymm j) * Polynomial.X ^ (Multiset.card s - j)

A sum version of Vieta's formula for Multiset: the product of the linear terms X + λ where λ runs through a multiset s is equal to a linear combination of the symmetric functions esymm s of the λ's .

theorem Multiset.prod_X_add_C_coeff {R : Type u_1} [CommSemiring R] (s : Multiset R) {k : } (h : k Multiset.card s) :
(Multiset.map (fun (r : R) => Polynomial.X + Polynomial.C r) s).prod.coeff k = s.esymm (Multiset.card s - k)

Vieta's formula for the coefficients of the product of linear terms X + λ where λ runs through a multiset s : the kth coefficient is the symmetric function esymm (card s - k) s.

theorem Multiset.prod_X_add_C_coeff' {R : Type u_1} [CommSemiring R] {σ : Type u_2} (s : Multiset σ) (r : σR) {k : } (h : k Multiset.card s) :
(Multiset.map (fun (i : σ) => Polynomial.X + Polynomial.C (r i)) s).prod.coeff k = (Multiset.map r s).esymm (Multiset.card s - k)
theorem Finset.prod_X_add_C_coeff {R : Type u_1} [CommSemiring R] {σ : Type u_2} (s : Finset σ) (r : σR) {k : } (h : k s.card) :
(s.prod fun (i : σ) => Polynomial.X + Polynomial.C (r i)).coeff k = (Finset.powersetCard (s.card - k) s).sum fun (t : Finset σ) => t.prod fun (i : σ) => r i
theorem Multiset.esymm_neg {R : Type u_1} [CommRing R] (s : Multiset R) (k : ) :
(Multiset.map Neg.neg s).esymm k = (-1) ^ k * s.esymm k
theorem Multiset.prod_X_sub_X_eq_sum_esymm {R : Type u_1} [CommRing R] (s : Multiset R) :
(Multiset.map (fun (t : R) => Polynomial.X - Polynomial.C t) s).prod = (Finset.range (Multiset.card s + 1)).sum fun (j : ) => (-1) ^ j * (Polynomial.C (s.esymm j) * Polynomial.X ^ (Multiset.card s - j))
theorem Multiset.prod_X_sub_C_coeff {R : Type u_1} [CommRing R] (s : Multiset R) {k : } (h : k Multiset.card s) :
(Multiset.map (fun (t : R) => Polynomial.X - Polynomial.C t) s).prod.coeff k = (-1) ^ (Multiset.card s - k) * s.esymm (Multiset.card s - k)
theorem Polynomial.coeff_eq_esymm_roots_of_card {R : Type u_1} [CommRing R] [IsDomain R] {p : Polynomial R} (hroots : Multiset.card p.roots = p.natDegree) {k : } (h : k p.natDegree) :
p.coeff k = p.leadingCoeff * (-1) ^ (p.natDegree - k) * p.roots.esymm (p.natDegree - k)

Vieta's formula for the coefficients and the roots of a polynomial over an integral domain with as many roots as its degree.

theorem Polynomial.coeff_eq_esymm_roots_of_splits {F : Type u_2} [Field F] {p : Polynomial F} (hsplit : Polynomial.Splits (RingHom.id F) p) {k : } (h : k p.natDegree) :
p.coeff k = p.leadingCoeff * (-1) ^ (p.natDegree - k) * p.roots.esymm (p.natDegree - k)

Vieta's formula for split polynomials over a field.

theorem MvPolynomial.prod_C_add_X_eq_sum_esymm (R : Type u_1) (σ : Type u_2) [CommSemiring R] [Fintype σ] :
(Finset.univ.prod fun (i : σ) => Polynomial.X + Polynomial.C (MvPolynomial.X i)) = (Finset.range (Fintype.card σ + 1)).sum fun (j : ) => Polynomial.C (MvPolynomial.esymm σ R j) * Polynomial.X ^ (Fintype.card σ - j)

A sum version of Vieta's formula for MvPolynomial: viewing X i as variables, the product of linear terms λ + X i is equal to a linear combination of the symmetric polynomials esymm σ R j.

theorem MvPolynomial.prod_X_add_C_coeff (R : Type u_1) (σ : Type u_2) [CommSemiring R] [Fintype σ] (k : ) (h : k Fintype.card σ) :
(Finset.univ.prod fun (i : σ) => Polynomial.X + Polynomial.C (MvPolynomial.X i)).coeff k = MvPolynomial.esymm σ R (Fintype.card σ - k)