Skip to content
This repository has been archived by the owner on Jul 24, 2024. It is now read-only.

[Merged by Bors] - feat(ring_theory/polynomial/vieta): add version of prod_X_add_C_eq_sum_esymm for multiset #15008

Closed
wants to merge 101 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
101 commits
Select commit Hold shift + click to select a range
c6c2e3c
Complete version of multiset.prod_X_add_C_eq_sum_esymm
xroblot Jun 21, 2022
141b56d
Fix corrupted proof
xroblot Jun 27, 2022
a5f521d
Proof clean finished
xroblot Jun 27, 2022
9a4f586
Finished notation clean up
xroblot Jun 27, 2022
594f8ce
should not have touch this file...
xroblot Jun 27, 2022
3c13fa8
Lint
xroblot Jun 27, 2022
76ca33b
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jun 28, 2022
cb1c0ce
Put things into place
xroblot Jun 28, 2022
b3bc3fc
Trying to find a place for everything
xroblot Jun 28, 2022
df38653
Merge branch 'master' of https://github.com/leanprover-community/math…
xroblot Jun 28, 2022
4640587
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jun 29, 2022
08d49e6
Some more cleaning and golfing
xroblot Jun 29, 2022
ce8cc05
Putting things into place
xroblot Jun 29, 2022
415b2c1
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jun 29, 2022
f5a8162
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jun 30, 2022
20367c5
Prepare for PR
xroblot Jun 30, 2022
aae2136
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jun 30, 2022
0b0adb1
typo
xroblot Jun 30, 2022
de852f5
Add [simp] to a lemma
xroblot Jun 30, 2022
4be60cd
clean up
xroblot Jun 30, 2022
1130170
add prod_X_add_C_coeff
xroblot Jul 1, 2022
5074cc2
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jul 2, 2022
45ff359
review
xroblot Jul 2, 2022
bfe9a11
Review
xroblot Jul 2, 2022
14e56bc
New proofs
xroblot Jul 3, 2022
3a0f810
clean up
xroblot Jul 3, 2022
4500444
Clean up
xroblot Jul 3, 2022
247b803
Revert
xroblot Jul 3, 2022
07bb04b
Putting things into places
xroblot Jul 3, 2022
a611091
Remove unneeded(?) functions
xroblot Jul 3, 2022
1592a91
add multiset.prod_X_add_C_coeff
xroblot Jul 3, 2022
3b922a4
Small golf
xroblot Jul 3, 2022
9772ee5
Lint
xroblot Jul 4, 2022
d9345ff
Add link between esymm for multiset and mv_polynomial
xroblot Jul 4, 2022
939eb38
Proves mv_polynomial versions from multiset versions
xroblot Jul 4, 2022
93cd67f
Change text
xroblot Jul 5, 2022
64b53cb
Typo
xroblot Jul 5, 2022
e6ab642
Add proofs for X_sub_C versions
xroblot Jul 5, 2022
c53bc8f
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jul 5, 2022
0aae0b1
WIP
xroblot Jul 5, 2022
4682f33
Add add_le_of_le_of_disjoint
xroblot Jul 6, 2022
3d9d99f
Settings
xroblot Jul 6, 2022
33ffa49
Golf + typos
xroblot Jul 6, 2022
1b3097c
Don't touch this file
xroblot Jul 6, 2022
1564c70
Fix name
xroblot Jul 6, 2022
7600155
small golf
xroblot Jul 6, 2022
76ebcc5
Lint
xroblot Jul 6, 2022
2d689eb
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jul 8, 2022
c5dc54c
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jul 14, 2022
fe51fe6
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jul 19, 2022
a2af8a3
Change statement (and proof) of finset_sum_le_of_le_of_disjoint
xroblot Jul 19, 2022
731ffc2
Fix name
xroblot Jul 19, 2022
e004329
Remove unneeded argument
xroblot Jul 19, 2022
ceb66c3
Adapt to finset_sum_eq_sup_of_disjoint
xroblot Jul 19, 2022
2939e67
Lint
xroblot Jul 20, 2022
c725ec7
Update src/data/multiset/antidiagonal.lean
xroblot Jul 20, 2022
65a1525
Initial commit
xroblot Jul 20, 2022
fdf64bd
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jul 20, 2022
9fa48e0
Merge with #15556
xroblot Jul 20, 2022
3039cad
Fix unnecessary decidable_eq assumption
xroblot Jul 20, 2022
2be00db
Remove unnecessary `classical`
xroblot Jul 20, 2022
8c49030
Make small lemmas iff
xroblot Jul 20, 2022
83924da
Complete proof of disjoint_finset_iff_sum_eq_sup
xroblot Jul 21, 2022
10e5c0d
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Jul 27, 2022
9ec2002
Merge branch 'xfr_bigop_disjoint' into xfr_vieta
xroblot Jul 27, 2022
002ffb9
Add cons_sub_of_le
xroblot Jul 27, 2022
e8d1222
Clean up after merge
xroblot Jul 27, 2022
558ea8f
Merge with #15556
xroblot Jul 27, 2022
372169f
Flip iffs
xroblot Jul 27, 2022
c3ebeb1
Use simpa for proof of symmetric statements
xroblot Jul 27, 2022
87ccc35
Update src/algebra/big_operators/basic.lean
xroblot Jul 27, 2022
87b87a4
Add add_eq_union_left_of_le and use it to simplify the proof of finse…
xroblot Jul 28, 2022
b104f37
Small golf
xroblot Jul 29, 2022
4e333ae
Update src/algebra/big_operators/basic.lean
xroblot Jul 29, 2022
e4dbc51
Typo
xroblot Jul 29, 2022
a24932a
Merge branch 'xfr_bigop_disjoint', remote-tracking branch 'origin' in…
xroblot Jul 29, 2022
81f1197
Small golf
xroblot Jul 30, 2022
6425582
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Aug 1, 2022
ef4c7e1
Lint
xroblot Aug 1, 2022
e911b1e
First commit
xroblot Aug 2, 2022
69e6f32
Add sup_powerset_len
xroblot Aug 3, 2022
b24295a
Merge branch 'xfr_powerset', remote-tracking branch 'origin' into xfr…
xroblot Aug 3, 2022
12bf528
Fix typo
xroblot Aug 3, 2022
3b1ec0e
Fix indentation
xroblot Aug 3, 2022
11c5b54
More indents...
xroblot Aug 3, 2022
f7efdd4
add disjoint_powerset_len
xroblot Aug 3, 2022
1f64bfa
feat(data/list/perm): binding all the sublists of a given length give…
eric-wieser Aug 3, 2022
f0da1b9
fix
eric-wieser Aug 3, 2022
9f60747
one more lemma
eric-wieser Aug 3, 2022
a75f021
Merge branch 'eric-wieser/sublists_len_join' into xfr_powerset
xroblot Aug 4, 2022
ce5c9d5
Prove bind_powerset_len using #15834
xroblot Aug 4, 2022
c83e0b5
Re-add sup_powerset_len
xroblot Aug 5, 2022
18ff11b
Merge branch 'master' into xfr_powerset
eric-wieser Aug 9, 2022
6716514
Update src/data/multiset/powerset.lean
eric-wieser Aug 9, 2022
010e179
Update src/algebra/big_operators/basic.lean
xroblot Aug 9, 2022
7347aa8
Update src/algebra/big_operators/basic.lean
xroblot Aug 10, 2022
bbe1df9
Merge remote-tracking branch 'refs/remotes/origin/master' into xfr_vieta
xroblot Aug 10, 2022
7b59aa9
Merge remote-tracking branch 'origin/xfr_powerset' into xfr_vieta
xroblot Aug 10, 2022
150caa6
Cleanup after merge
xroblot Aug 10, 2022
efb6c73
Minor changes
xroblot Aug 10, 2022
c3ddf28
Merge remote-tracking branch 'origin/master' into xfr_vieta
eric-wieser Aug 15, 2022
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
20 changes: 20 additions & 0 deletions src/ring_theory/polynomial/symmetric.lean
Original file line number Diff line number Diff line change
Expand Up @@ -44,6 +44,15 @@ open equiv (perm)
open_locale big_operators
noncomputable theory

namespace multiset

variables {R : Type*} [comm_semiring R]

/-- The `n`th elementary symmetric function evaluated at the elements of `s` -/
def esymm (s : multiset R) (n : ℕ) : R := ((s.powerset_len n).map multiset.prod).sum

end multiset
xroblot marked this conversation as resolved.
Show resolved Hide resolved

namespace mv_polynomial

variables {σ : Type*} {R : Type*}
Expand Down Expand Up @@ -121,6 +130,17 @@ variables (σ R) [comm_semiring R] [comm_semiring S] [fintype σ] [fintype τ]
def esymm (n : ℕ) : mv_polynomial σ R :=
∑ t in powerset_len n univ, ∏ i in t, X i

/-- The `n`th elementary symmetric `mv_polynomial σ R` is obtained by evaluating the
`n`th elementary symmetric at the `multiset` of the monomials -/
lemma esymm_eq_multiset.esymm (n : ℕ) :
esymm σ R n =
multiset.esymm (multiset.map (λ i : σ, (X i : mv_polynomial σ R)) finset.univ.val) n :=
begin
rw [esymm, multiset.esymm, finset.sum_eq_multiset_sum],
conv_lhs { congr, congr, funext, rw finset.prod_eq_multiset_prod },
rw [multiset.powerset_len_map, ←map_val_val_powerset_len, multiset.map_map, multiset.map_map],
end

/-- We can define `esymm σ R n` by summing over a subtype instead of over `powerset_len`. -/
lemma esymm_eq_sum_subtype (n : ℕ) : esymm σ R n =
∑ t : {s : finset σ // s.card = n}, ∏ i in (t : finset σ), X i :=
Expand Down
182 changes: 112 additions & 70 deletions src/ring_theory/polynomial/vieta.lean
Original file line number Diff line number Diff line change
Expand Up @@ -9,95 +9,137 @@ import ring_theory.polynomial.symmetric
/-!
# Vieta's Formula

The main result is `vieta.prod_X_add_C_eq_sum_esymm`, which shows that the product of linear terms
`λ + X i` is equal to a linear combination of the symmetric polynomials `esymm σ R j`.
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`.

## Implementation Notes:

We first take the viewpoint where the "roots" `X i` are variables. This means we work over
`polynomial (mv_polynomial σ R)`, which enables us to talk about linear combinations of
`esymm σ R j`. We then derive Vieta's formula in `polynomial R` by giving a
valuation from each `X i` to `r i`.
From this, we deduce `mv_polynomial.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`.

-/

open_locale big_operators polynomial

open finset polynomial fintype
namespace multiset

namespace mv_polynomial
section semiring

variables {R : Type*} [comm_semiring R]
variables (σ : Type*) [fintype σ]

/-- A sum version of Vieta's formula. 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`. -/
lemma prod_X_add_C_eq_sum_esymm :
(∏ i : σ, (polynomial.C (X i) + polynomial.X) : polynomial (mv_polynomial σ R) )=
∑ j in range (card σ + 1),
(polynomial.C (esymm σ R j) * polynomial.X ^ (card σ - 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 .-/
lemma prod_X_add_C_eq_sum_esymm (s : multiset R) :
(s.map (λ r, polynomial.X + polynomial.C r)).prod =
∑ j in finset.range (s.card + 1), (polynomial.C (s.esymm j) * polynomial.X^(s.card - j)) :=
begin
classical,
rw [prod_add, sum_powerset],
refine sum_congr begin congr end (λ j hj, _),
rw [esymm, map_sum, sum_mul],
refine sum_congr rfl (λ t ht, _),
have h : (univ \ t).card = card σ - j :=
by { rw card_sdiff (mem_powerset_len.mp ht).1, congr, exact (mem_powerset_len.mp ht).2 },
rw [map_prod, prod_const, ← h],
rw [prod_map_add, antidiagonal_eq_map_powerset, map_map, ←bind_powerset_len, function.comp,
map_bind, sum_bind, finset.sum_eq_multiset_sum, finset.range_coe, map_congr (eq.refl _)],
intros _ _,
rw [esymm, ←sum_hom', ←sum_map_mul_right, map_congr (eq.refl _)],
intros _ ht,
rw mem_powerset_len at ht,
simp [ht, map_const, prod_repeat, prod_hom', map_id', card_sub],
end

/-- A fully expanded sum version of Vieta's formula, evaluated at the roots.
The product of linear terms `X + r i` is equal to `∑ j in range (n + 1), e_j * X ^ (n - j)`,
where `e_j` is the `j`th symmetric polynomial of the constant terms `r i`. -/
lemma prod_X_add_C_eval (r : σ → R) : ∏ i : σ, (polynomial.C (r i) + polynomial.X) =
∑ i in range (card σ + 1), (∑ t in powerset_len i (univ : finset σ),
∏ i in t, polynomial.C (r i)) * polynomial.X ^ (card σ - i) :=
/-- Vieta's formula for the coefficients of the product of linear terms `X + λ` where `λ` runs
through a multiset `s` : the `k`th coefficient is the symmetric function `esymm (card s - k) s`. -/
lemma prod_X_add_C_coeff (s : multiset R) {k : ℕ} (h : k ≤ s.card):
polynomial.coeff (s.map (λ r, polynomial.X + polynomial.C r)).prod k =
s.esymm (s.card - k) :=
begin
classical,
have h := @prod_X_add_C_eq_sum_esymm _ _ σ _,
apply_fun (polynomial.map (eval r)) at h,
rw [polynomial.map_prod, polynomial.map_sum] at h,
convert h,
simp only [eval_X, polynomial.map_add, polynomial.map_C, polynomial.map_X, eq_self_iff_true],
funext,
simp only [function.funext_iff, esymm, polynomial.map_C, polynomial.map_sum, map_sum,
polynomial.map_C, polynomial.map_pow, polynomial.map_X, polynomial.map_mul],
congr,
funext,
simp only [eval_prod, eval_X, map_prod],
convert polynomial.ext_iff.mp (prod_X_add_C_eq_sum_esymm s) k,
simp_rw [polynomial.finset_sum_coeff, polynomial.coeff_C_mul_X_pow],
rw finset.sum_eq_single_of_mem (s.card - k) _,
{ rw if_pos (nat.sub_sub_self h).symm, },
{ intros j hj1 hj2,
suffices : k ≠ card s - j,
{ rw if_neg this, },
{ intro hn,
rw [hn, nat.sub_sub_self (nat.lt_succ_iff.mp (finset.mem_range.mp hj1))] at hj2,
exact ne.irrefl hj2, }},
{ rw finset.mem_range,
exact nat.sub_lt_succ s.card k }
end

end semiring

section ring

variables {R : Type*} [comm_ring R]

lemma esymm_neg (s : multiset R) (k : ℕ) :
(map has_neg.neg s).esymm k = (-1)^k * esymm s k :=
begin
rw [esymm, esymm, ←multiset.sum_map_mul_left, multiset.powerset_len_map, multiset.map_map,
map_congr (eq.refl _)],
intros x hx,
rw [(by { exact (mem_powerset_len.mp hx).right.symm }), ←prod_repeat, ←multiset.map_const],
nth_rewrite 2 ←map_id' x,
rw [←prod_map_mul, map_congr (eq.refl _)],
exact λ z _, neg_one_mul z,
end

lemma prod_X_sub_C_eq_sum_esymm (s : multiset R) :
(s.map (λ t, polynomial.X - polynomial.C t)).prod =
∑ j in finset.range (s.card + 1),
(-1)^j* (polynomial.C (s.esymm j) * polynomial.X ^ (s.card - j)) :=
begin
conv_lhs { congr, congr, funext, rw sub_eq_add_neg, rw ←map_neg polynomial.C _, },
convert prod_X_add_C_eq_sum_esymm (map (λ t, -t) s) using 1,
{ rwa map_map, },
{ simp only [esymm_neg, card_map, mul_assoc, map_mul, map_pow, map_neg, map_one], },
end

lemma esymm_to_sum (r : σ → R) (j : ℕ) : polynomial.C (eval r (esymm σ R j)) =
∑ t in powerset_len j (univ : finset σ), ∏ i in t, polynomial.C (r i) :=
by simp only [esymm, eval_sum, eval_prod, eval_X, map_sum, map_prod]
lemma prod_X_sub_C_coeff (s : multiset R) {k : ℕ} (h : k ≤ s.card):
polynomial.coeff (s.map (λ t, polynomial.X - polynomial.C t)).prod k =
(-1)^(s.card - k) * s.esymm (s.card - k) :=
begin
conv_lhs { congr, congr, congr, funext, rw sub_eq_add_neg, rw ←map_neg polynomial.C _, },
convert prod_X_add_C_coeff (map (λ t, -t) s) _ using 1,
{ rwa map_map, },
{ rwa [esymm_neg, card_map] },
{ rwa card_map },
end

end ring

end multiset

namespace mv_polynomial

open finset polynomial fintype

variables (R σ : Type*) [comm_semiring R] [fintype σ]

/-- A sum version of Vieta's formula for `mv_polynomial`: 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`. -/
lemma prod_C_add_X_eq_sum_esymm :
(∏ i : σ, (polynomial.X + polynomial.C (X i)) : polynomial (mv_polynomial σ R) )=
∑ j in range (card σ + 1),
(polynomial.C (esymm σ R j) * polynomial.X ^ (card σ - j)) :=
begin
let s := (multiset.map (λ i : σ, (X i : mv_polynomial σ R)) finset.univ.val),
rw (_ : card σ = s.card),
{ simp_rw [esymm_eq_multiset.esymm σ R _, finset.prod_eq_multiset_prod],
convert multiset.prod_X_add_C_eq_sum_esymm s,
rwa multiset.map_map, },
{ rw multiset.card_map, exact rfl, }
end

/-- Vieta's formula for the coefficients of the product of linear terms `X + r i`,
The `k`th coefficient is `∑ t in powerset_len (card σ - k) (univ : finset σ), ∏ i in t, r i`,
i.e. the symmetric polynomial `esymm σ R (card σ - k)` of the constant terms `r i`. -/
lemma prod_X_add_C_coeff (r : σ → R) (k : ℕ) (h : k ≤ card σ):
polynomial.coeff (∏ i : σ, (polynomial.C (r i) + polynomial.X)) k =
∑ t in powerset_len (card σ - k) (univ : finset σ), ∏ i in t, r i :=
lemma prod_X_add_C_coeff (k : ℕ) (h : k ≤ card σ):
(∏ i : σ, (polynomial.X + polynomial.C (X i)) : polynomial (mv_polynomial σ R) ).coeff k =
esymm σ R (card σ - k) :=
begin
have hk : filter (λ (x : ℕ), k = card σ - x) (range (card σ + 1)) = {card σ - k} :=
begin
refine finset.ext (λ a, ⟨λ ha, _, λ ha, _ ⟩),
rw mem_singleton,
have hσ := (tsub_eq_iff_eq_add_of_le (mem_range_succ_iff.mp
(mem_filter.mp ha).1)).mp ((mem_filter.mp ha).2).symm,
symmetry,
rwa [(tsub_eq_iff_eq_add_of_le h), add_comm],
rw mem_filter,
have haσ : a ∈ range (card σ + 1) :=
by { rw mem_singleton.mp ha, exact mem_range_succ_iff.mpr (@tsub_le_self _ _ _ _ _ k) },
refine ⟨haσ, eq.symm _⟩,
rw tsub_eq_iff_eq_add_of_le (mem_range_succ_iff.mp haσ),
have hσ := (tsub_eq_iff_eq_add_of_le h).mp (mem_singleton.mp ha).symm,
rwa add_comm,
end,
simp only [prod_X_add_C_eval, ← esymm_to_sum, finset_sum_coeff, coeff_C_mul_X_pow, sum_ite, hk,
sum_singleton, esymm, eval_sum, eval_prod, eval_X, add_zero, sum_const_zero],
let s := (multiset.map (λ i : σ, (X i : mv_polynomial σ R)) finset.univ.val),
rw (_ : card σ = s.card) at ⊢ h,
{ rw [esymm_eq_multiset.esymm σ R (s.card - k), finset.prod_eq_multiset_prod],
convert multiset.prod_X_add_C_coeff s h,
rwa multiset.map_map },
repeat { rw multiset.card_map, exact rfl, },
end

end mv_polynomial