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

Commit 791852e

Browse files
committed
feat(polynomial/field_division): Add root_set_prod and clean up lemma statements (#16035)
This PR adds `root_set_prod` (based on `roots_prod`) and cleans up lemma statements (the file already has `{R S : Type*} [field R]`).
1 parent 0398787 commit 791852e

File tree

1 file changed

+14
-5
lines changed

1 file changed

+14
-5
lines changed

src/data/polynomial/field_division.lean

Lines changed: 14 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -339,19 +339,19 @@ theorem is_coprime_map [field k] (f : R →+* k) :
339339
is_coprime (p.map f) (q.map f) ↔ is_coprime p q :=
340340
by rw [← euclidean_domain.gcd_is_unit_iff, ← euclidean_domain.gcd_is_unit_iff, gcd_map, is_unit_map]
341341

342-
lemma mem_roots_map [field k] {f : R →+* k} {x : k} (hp : p ≠ 0) :
342+
lemma mem_roots_map [comm_ring k] [is_domain k] {f : R →+* k} {x : k} (hp : p ≠ 0) :
343343
x ∈ (p.map f).roots ↔ p.eval₂ f x = 0 :=
344344
begin
345345
rw mem_roots (show p.map f ≠ 0, by exact map_ne_zero hp),
346346
dsimp only [is_root],
347347
rw polynomial.eval_map,
348348
end
349349

350-
lemma mem_root_set [field k] [algebra R k] {x : k} (hp : p ≠ 0) :
350+
lemma mem_root_set [comm_ring k] [is_domain k] [algebra R k] {x : k} (hp : p ≠ 0) :
351351
x ∈ p.root_set k ↔ aeval x p = 0 :=
352352
iff.trans multiset.mem_to_finset (mem_roots_map hp)
353353

354-
lemma root_set_C_mul_X_pow {R S : Type*} [field R] [field S] [algebra R S]
354+
lemma root_set_C_mul_X_pow [comm_ring S] [is_domain S] [algebra R S]
355355
{n : ℕ} (hn : n ≠ 0) {a : R} (ha : a ≠ 0) : (C a * X ^ n).root_set S = {0} :=
356356
begin
357357
ext x,
@@ -361,14 +361,23 @@ begin
361361
{ exact mul_ne_zero (mt C_eq_zero.mp ha) (pow_ne_zero n X_ne_zero) },
362362
end
363363

364-
lemma root_set_monomial {R S : Type*} [field R] [field S] [algebra R S]
364+
lemma root_set_monomial [comm_ring S] [is_domain S] [algebra R S]
365365
{n : ℕ} (hn : n ≠ 0) {a : R} (ha : a ≠ 0) : (monomial n a).root_set S = {0} :=
366366
by rw [←C_mul_X_pow_eq_monomial, root_set_C_mul_X_pow hn ha]
367367

368-
lemma root_set_X_pow {R S : Type*} [field R] [field S] [algebra R S]
368+
lemma root_set_X_pow [comm_ring S] [is_domain S] [algebra R S]
369369
{n : ℕ} (hn : n ≠ 0) : (X ^ n : R[X]).root_set S = {0} :=
370370
by { rw [←one_mul (X ^ n : R[X]), ←C_1, root_set_C_mul_X_pow hn], exact one_ne_zero }
371371

372+
lemma root_set_prod [comm_ring S] [is_domain S] [algebra R S]
373+
{ι : Type*} (f : ι → R[X]) (s : finset ι) (h : s.prod f ≠ 0) :
374+
(s.prod f).root_set S = ⋃ (i ∈ s), (f i).root_set S :=
375+
begin
376+
simp only [root_set, ←finset.mem_coe],
377+
rw [polynomial.map_prod, roots_prod, finset.bind_to_finset, s.val_to_finset, finset.coe_bUnion],
378+
rwa [←polynomial.map_prod, ne, map_eq_zero],
379+
end
380+
372381
lemma exists_root_of_degree_eq_one (h : degree p = 1) : ∃ x, is_root p x :=
373382
⟨-(p.coeff 0 / p.coeff 1),
374383
have p.coeff 10,

0 commit comments

Comments
 (0)