Skip to content

Commit

Permalink
feat: port FieldTheory.Normal (#4856)
Browse files Browse the repository at this point in the history
Co-authored-by: Vierkantor <vierkantor@vierkantor.com>
Co-authored-by: Scott Morrison <scott.morrison@anu.edu.au>
Co-authored-by: Thomas Browning <tb65536@uw.edu>
Co-authored-by: Scott Morrison <scott@tqft.net>
  • Loading branch information
5 people committed Jun 9, 2023
1 parent f8e9600 commit 62b6b1a
Show file tree
Hide file tree
Showing 4 changed files with 572 additions and 2 deletions.
1 change: 1 addition & 0 deletions Mathlib.lean
Expand Up @@ -1673,6 +1673,7 @@ import Mathlib.FieldTheory.Laurent
import Mathlib.FieldTheory.Minpoly.Basic
import Mathlib.FieldTheory.Minpoly.Field
import Mathlib.FieldTheory.MvPolynomial
import Mathlib.FieldTheory.Normal
import Mathlib.FieldTheory.PerfectClosure
import Mathlib.FieldTheory.RatFunc
import Mathlib.FieldTheory.Separable
Expand Down
5 changes: 5 additions & 0 deletions Mathlib/Data/Polynomial/RingDivision.lean
Expand Up @@ -589,6 +589,11 @@ theorem isRoot_of_mem_roots (h : a ∈ p.roots) : IsRoot p a :=
(mem_roots'.1 h).2
#align polynomial.is_root_of_mem_roots Polynomial.isRoot_of_mem_roots

-- Porting note: added during port.
lemma mem_roots_iff_aeval_eq_zero (w : p ≠ 0) : x ∈ roots p ↔ aeval x p = 0 := by
rw [mem_roots w, IsRoot.def, aeval_def, eval₂_eq_eval_map]
simp

theorem card_le_degree_of_subset_roots {p : R[X]} {Z : Finset R} (h : Z.val ⊆ p.roots) :
Z.card ≤ p.natDegree :=
(Multiset.card_le_of_le (Finset.val_le_iff_val_subset.2 h)).trans (Polynomial.card_roots' p)
Expand Down

0 comments on commit 62b6b1a

Please sign in to comment.