Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat(ring_theory/trace): Tr(x) is the sum of embeddings σ x into an algebraically closed field #8762

Closed
wants to merge 10 commits into from
5 changes: 4 additions & 1 deletion src/field_theory/adjoin.lean
Original file line number Diff line number Diff line change
Expand Up @@ -530,7 +530,10 @@ lemma card_alg_hom_adjoin_integral (h : is_integral F α) (h_sep : (minpoly F α
(h_splits : (minpoly F α).splits (algebra_map F K)) :
@fintype.card (F⟮α⟯ →ₐ[F] K) (fintype_of_alg_hom_adjoin_integral F h) =
(minpoly F α).nat_degree :=
by { rw alg_hom.card_of_power_basis, all_goals { rwa [adjoin.power_basis_gen, minpoly_gen h] } }
begin
rw alg_hom.card_of_power_basis;
simp only [adjoin.power_basis_dim, adjoin.power_basis_gen, minpoly_gen h, h_sep, h_splits],
end

end adjoin_integral_element

Expand Down
13 changes: 11 additions & 2 deletions src/field_theory/intermediate_field.lean
Original file line number Diff line number Diff line change
Expand Up @@ -208,12 +208,21 @@ end
instance algebra : algebra K S :=
S.to_subalgebra.algebra

instance to_algebra : algebra S L :=
instance to_algebra {R : Type*} [semiring R] [algebra L R] : algebra S R :=
S.to_subalgebra.to_algebra

instance : is_scalar_tower K S L :=
instance is_scalar_tower_bot {R : Type*} [semiring R] [algebra L R] :
is_scalar_tower S L R :=
is_scalar_tower.subalgebra _ _ _ S.to_subalgebra

instance is_scalar_tower_mid {R : Type*} [semiring R] [algebra L R] [algebra K R]
[is_scalar_tower K L R] : is_scalar_tower K S R :=
is_scalar_tower.subalgebra' _ _ _ S.to_subalgebra

/-- Specialize `is_scalar_tower_mid` to the common case where the top field is `L` -/
instance is_scalar_tower_mid' : is_scalar_tower K S L :=
S.is_scalar_tower_mid

variables {L' : Type*} [field L'] [algebra K L']

/-- If `f : L →+* L'` fixes `K`, `S.map f` is the intermediate field between `L'` and `K`
Expand Down
8 changes: 8 additions & 0 deletions src/field_theory/polynomial_galois_group.lean
Original file line number Diff line number Diff line change
Expand Up @@ -97,6 +97,14 @@ instance [h : fact (p.splits (algebra_map F E))] : is_scalar_tower F p.splitting
is_scalar_tower.of_algebra_map_eq
(λ x, ((is_splitting_field.lift p.splitting_field p h.1).commutes x).symm)

-- The `algebra p.splitting_field E` instance above behaves badly when
-- `E := p.splitting_field`, since it may result in a unification problem
-- `is_splitting_field.lift.to_ring_hom.to_algebra =?= algebra.id`,
-- which takes an extremely long time to resolve, causing timeouts.
-- Since we don't really care about this definition, marking it as irreducible
-- causes that unification to error out early.
attribute [irreducible] gal.algebra

/-- Restrict from a superfield automorphism into a member of `gal p`. -/
def restrict [fact (p.splits (algebra_map F E))] : (E ≃ₐ[F] E) →* p.gal :=
alg_equiv.restrict_normal_hom p.splitting_field
Expand Down
14 changes: 14 additions & 0 deletions src/field_theory/primitive_element.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,6 +5,7 @@ Authors: Thomas Browning, Patrick Lutz
-/

import field_theory.adjoin
import field_theory.algebraic_closure
import field_theory.separable

/-!
Expand Down Expand Up @@ -194,4 +195,17 @@ let α := (exists_primitive_element F E).some,
have e : F⟮α⟯ = ⊤ := (exists_primitive_element F E).some_spec,
pb.map ((intermediate_field.equiv_of_eq e).trans intermediate_field.top_equiv)

/-- If `E / F` is a finite separable extension, then there are finitely many
embeddings from `E` into `K` that fix `F`, corresponding to the number of
conjugate roots of the primitive element generating `F`. -/
instance {K : Type*} [field K] [algebra F K] : fintype (E →ₐ[F] K) :=
power_basis.alg_hom.fintype (power_basis_of_finite_of_separable F E)

end field

@[simp] lemma alg_hom.card (F E K : Type*) [field F] [field E] [field K] [is_alg_closed K]
[algebra F E] [finite_dimensional F E] [is_separable F E] [algebra F K] :
fintype.card (E →ₐ[F] K) = finrank F E :=
(alg_hom.card_of_power_basis (field.power_basis_of_finite_of_separable F E)
(is_separable.separable _ _) (is_alg_closed.splits_codomain _)).trans
(power_basis.finrank _).symm
4 changes: 2 additions & 2 deletions src/field_theory/separable.lean
Original file line number Diff line number Diff line change
Expand Up @@ -657,12 +657,12 @@ variables [algebra K S] [algebra K L]

lemma alg_hom.card_of_power_basis (pb : power_basis K S) (h_sep : (minpoly K pb.gen).separable)
(h_splits : (minpoly K pb.gen).splits (algebra_map K L)) :
@fintype.card (S →ₐ[K] L) (power_basis.alg_hom.fintype pb) = (minpoly K pb.gen).nat_degree :=
@fintype.card (S →ₐ[K] L) (power_basis.alg_hom.fintype pb) = pb.dim :=
begin
let s := ((minpoly K pb.gen).map (algebra_map K L)).roots.to_finset,
have H := λ x, multiset.mem_to_finset,
rw [fintype.card_congr pb.lift_equiv', fintype.card_of_subtype s H,
nat_degree_eq_card_roots h_splits, multiset.to_finset_card_of_nodup],
← pb.nat_degree_minpoly, nat_degree_eq_card_roots h_splits, multiset.to_finset_card_of_nodup],
exact nodup_roots ((separable_map (algebra_map K L)).mpr h_sep)
end

Expand Down
4 changes: 1 addition & 3 deletions src/linear_algebra/free_module_pid.lean
Original file line number Diff line number Diff line change
Expand Up @@ -437,9 +437,7 @@ begin
(show (set.range (λ ψ : M →ₗ[R] R, ψ.submodule_image N)).nonempty,
from ⟨_, set.mem_range.mpr ⟨0, rfl⟩⟩),
obtain ⟨ϕ, rfl⟩ := set.mem_range.mp P_eq,
use ϕ,
intros ψ hψ,
exact P_max _ ⟨_, rfl⟩ hψ },
exact ⟨ϕ, λ ψ hψ, P_max _ ⟨_, rfl⟩ hψ⟩ },
Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Not only is this slightly shorter, it runs noticeably faster which was useful when messing around with some instances. (See also #8761.)

let ϕ := this.some,
have ϕ_max := this.some_spec,
-- Since `ϕ(N)` is a `R`-submodule of the PID `R`,
Expand Down
1 change: 1 addition & 0 deletions src/ring_theory/power_basis.lean
Original file line number Diff line number Diff line change
Expand Up @@ -329,6 +329,7 @@ noncomputable def lift_equiv (pb : power_basis A S) :

/-- `pb.lift_equiv'` states that elements of the root set of the minimal
polynomial of `pb.gen` correspond to maps sending `pb.gen` to that root. -/
@[simps {fully_applied := ff}]
noncomputable def lift_equiv' (pb : power_basis A S) :
(S →ₐ[A] B) ≃ {y : B // y ∈ ((minpoly A pb.gen).map (algebra_map A B)).roots} :=
pb.lift_equiv.trans ((equiv.refl _).subtype_equiv (λ x,
Expand Down
88 changes: 79 additions & 9 deletions src/ring_theory/trace.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,9 +6,10 @@ Authors: Anne Baanen

import linear_algebra.bilinear_form
import linear_algebra.char_poly.coeff
import linear_algebra.determinant
import linear_algebra.trace
import field_theory.adjoin
import field_theory.algebraic_closure
import field_theory.primitive_element
import ring_theory.power_basis

/-!
Expand Down Expand Up @@ -232,20 +233,29 @@ end

end intermediate_field.adjoin_simple

lemma trace_eq_sum_roots [finite_dimensional K L]
Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

trace_eq_sum_roots still exists with the same statement, it's just that the first part of the proof was split off into trace_eq_trace_adjoin and the diff is having trouble showing that clearly.

{x : L} (hF : (minpoly K x).splits (algebra_map K F)) :
algebra_map K F (algebra.trace K L x) =
finrank K⟮x⟯ L • ((minpoly K x).map (algebra_map K _)).roots.sum :=
open intermediate_field

variables (K)

lemma trace_eq_trace_adjoin [finite_dimensional K L] (x : L) :
algebra.trace K L x = finrank K⟮x⟯ L • trace K K⟮x⟯ (adjoin_simple.gen K x) :=
begin
haveI : finite_dimensional K⟮x⟯ L := finite_dimensional.right K _ L,
rw ← @trace_trace _ _ K K⟮x⟯ _ _ _ _ _ _ _ _ x,
conv in x { rw ← intermediate_field.adjoin_simple.algebra_map_gen K x },
rw [trace_algebra_map, ← is_scalar_tower.algebra_map_smul K, (algebra.trace K K⟮x⟯).map_smul,
smul_eq_mul, ring_hom.map_mul, ← is_scalar_tower.algebra_map_apply ℕ K _, ← smul_def,
intermediate_field.adjoin_simple.trace_gen_eq_sum_roots _ hF],
all_goals { apply_instance }
smul_eq_mul, algebra.smul_def],
apply_instance
end

variables {K}

lemma trace_eq_sum_roots [finite_dimensional K L]
{x : L} (hF : (minpoly K x).splits (algebra_map K F)) :
algebra_map K F (algebra.trace K L x) =
finrank K⟮x⟯ L • ((minpoly K x).map (algebra_map K _)).roots.sum :=
by rw [trace_eq_trace_adjoin K x, algebra.smul_def, ring_hom.map_mul, ← algebra.smul_def,
intermediate_field.adjoin_simple.trace_gen_eq_sum_roots _ hF, is_scalar_tower.algebra_map_smul]

end eq_sum_roots

variables {F : Type*} [field F]
Expand All @@ -268,3 +278,63 @@ begin
{ apply is_alg_closed.splits_codomain },
{ apply_instance }
end

section eq_sum_embeddings

variables [algebra K F] [is_scalar_tower K L F]

open algebra intermediate_field

variables (F) (E : Type*) [field E] [algebra K E]

lemma trace_eq_sum_embeddings_gen
(pb : power_basis K L)
(hE : (minpoly K pb.gen).splits (algebra_map K E)) (hfx : (minpoly K pb.gen).separable) :
algebra_map K E (algebra.trace K L pb.gen) =
(@@finset.univ (power_basis.alg_hom.fintype pb)).sum (λ σ, σ pb.gen) :=
begin
letI := classical.dec_eq E,
rw [pb.trace_gen_eq_sum_roots hE, fintype.sum_equiv pb.lift_equiv', finset.sum_mem_multiset,
finset.sum_eq_multiset_sum, multiset.to_finset_val,
multiset.erase_dup_eq_self.mpr (nodup_roots ((separable_map _).mpr hfx)), multiset.map_id],
{ intro x, refl },
{ intro σ, rw [power_basis.lift_equiv'_apply_coe, id.def] }
end

variables [is_alg_closed E]

lemma sum_embeddings_eq_finrank_mul [finite_dimensional K F] [is_separable K F]
(pb : power_basis K L) :
∑ σ : F →ₐ[K] E, σ (algebra_map L F pb.gen) =
finrank L F • (@@finset.univ (power_basis.alg_hom.fintype pb)).sum
(λ σ : L →ₐ[K] E, σ pb.gen) :=
begin
haveI : finite_dimensional L F := finite_dimensional.right K L F,
haveI : is_separable L F := is_separable_tower_top_of_is_separable K L F,
letI : fintype (L →ₐ[K] E) := power_basis.alg_hom.fintype pb,
letI : ∀ (f : L →ₐ[K] E), fintype (@@alg_hom L F E _ _ _ _ f.to_ring_hom.to_algebra) :=
_, -- will be solved by unification
rw [fintype.sum_equiv alg_hom_equiv_sigma (λ (σ : F →ₐ[K] E), _) (λ σ, σ.1 pb.gen),
← finset.univ_sigma_univ, finset.sum_sigma, ← finset.sum_nsmul],
refine finset.sum_congr rfl (λ σ _, _),
{ letI : algebra L E := σ.to_ring_hom.to_algebra,
simp only [finset.sum_const, finset.card_univ],
rw alg_hom.card L F E },
{ intros σ,
simp only [alg_hom_equiv_sigma, equiv.coe_fn_mk, alg_hom.restrict_domain, alg_hom.comp_apply,
is_scalar_tower.coe_to_alg_hom'] }
end

lemma trace_eq_sum_embeddings [finite_dimensional K L] [is_separable K L]
{x : L} (hx : is_integral K x) :
algebra_map K E (algebra.trace K L x) = ∑ σ : L →ₐ[K] E, σ x :=
begin
rw [trace_eq_trace_adjoin K x, algebra.smul_def, ring_hom.map_mul, ← adjoin.power_basis_gen hx,
trace_eq_sum_embeddings_gen E (adjoin.power_basis hx) (is_alg_closed.splits_codomain _),
← algebra.smul_def, algebra_map_smul],
{ exact (sum_embeddings_eq_finrank_mul L E (adjoin.power_basis hx)).symm },
{ haveI := is_separable_tower_bot_of_is_separable K K⟮x⟯ L,
exact is_separable.separable K _ }
end

end eq_sum_embeddings