Skip to content

Commit

Permalink
chore(field_theory/finite/basic): move a lemma (#19130)
Browse files Browse the repository at this point in the history
Moving one instance we can erase `field_theory.splitting_field` from the imports of `field_theory.finite.basic)`.
  • Loading branch information
riccardobrasca committed May 31, 2023
1 parent 3283755 commit 12a85fa
Show file tree
Hide file tree
Showing 2 changed files with 19 additions and 21 deletions.
20 changes: 0 additions & 20 deletions src/field_theory/finite/basic.lean
Expand Up @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Chris Hughes, Joey van Langen, Casper Putz
-/
import field_theory.separable
import field_theory.splitting_field
import ring_theory.integral_domain
import tactic.apply_fun

Expand Down Expand Up @@ -219,7 +218,6 @@ begin
... = 0 : by { rw [sum_pow_units K i, if_neg], exact hiq, }
end

section is_splitting_field
open polynomial

section
Expand Down Expand Up @@ -268,24 +266,6 @@ begin
zero_sub] },
end

instance (F : Type*) [field F] [algebra F K] : is_splitting_field F K (X^q - X) :=
{ splits :=
begin
have h : (X^q - X : K[X]).nat_degree = q :=
X_pow_card_sub_X_nat_degree_eq K fintype.one_lt_card,
rw [←splits_id_iff_splits, splits_iff_card_roots, polynomial.map_sub, polynomial.map_pow,
map_X, h, roots_X_pow_card_sub_X K, ←finset.card_def, finset.card_univ],
end,
adjoin_roots :=
begin
classical,
transitivity algebra.adjoin F ((roots (X^q - X : K[X])).to_finset : set K),
{ simp only [polynomial.map_pow, map_X, polynomial.map_sub], },
{ rw [roots_X_pow_card_sub_X, val_to_finset, coe_univ, algebra.adjoin_univ], }
end }

end is_splitting_field

variables {K}

theorem frobenius_pow {p : ℕ} [fact p.prime] [char_p K p] {n : ℕ} (hcard : q = p^n) :
Expand Down
20 changes: 19 additions & 1 deletion src/field_theory/finite/galois_field.lean
Expand Up @@ -8,6 +8,7 @@ import algebra.char_p.algebra
import data.zmod.algebra
import field_theory.finite.basic
import field_theory.galois
import field_theory.splitting_field

/-!
# Galois fields
Expand All @@ -30,9 +31,26 @@ It is a finite field with `p ^ n` elements.

noncomputable theory

open polynomial
open polynomial finset
open_locale polynomial

instance finite_field.has_sub.sub.polynomial.is_splitting_field (K F : Type*) [field K] [fintype K]
[field F] [algebra F K] : is_splitting_field F K (X ^ (fintype.card K) - X) :=
{ splits :=
begin
have h : (X ^ (fintype.card K) - X : K[X]).nat_degree = fintype.card K :=
finite_field.X_pow_card_sub_X_nat_degree_eq K fintype.one_lt_card,
rw [←splits_id_iff_splits, splits_iff_card_roots, polynomial.map_sub, polynomial.map_pow,
map_X, h, finite_field.roots_X_pow_card_sub_X K, ←finset.card_def, finset.card_univ],
end,
adjoin_roots :=
begin
classical,
transitivity algebra.adjoin F ((roots (X ^ (fintype.card K) - X : K[X])).to_finset : set K),
{ simp only [polynomial.map_pow, map_X, polynomial.map_sub], },
{ rw [finite_field.roots_X_pow_card_sub_X, val_to_finset, coe_univ, algebra.adjoin_univ], }
end }

lemma galois_poly_separable {K : Type*} [field K] (p q : ℕ) [char_p K p] (h : p ∣ q) :
separable (X ^ q - X : K[X]) :=
begin
Expand Down

0 comments on commit 12a85fa

Please sign in to comment.