Adjoining Elements to Fields #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
In this file we introduce the notion of adjoining elements to fields.
This isn't quite the same as adjoining elements to rings.
For example, algebra.adjoin K {x} might not include x⁻¹.
Main results #
- adjoin_adjoin_left: adjoining S and then T is the same as adjoining- S ∪ T.
- bot_eq_top_of_rank_adjoin_eq_one: if- F⟮x⟯has dimension- 1over- Ffor every- xin- Ethen- F = E
Notation #
- F⟮α⟯: adjoin a single element- αto- F.
adjoin F S extends a field F by adjoining a set S ⊆ E.
Equations
- intermediate_field.adjoin F S = {to_subalgebra := {carrier := (subfield.closure (set.range ⇑(algebra_map F E) ∪ S)).carrier, mul_mem' := _, one_mem' := _, add_mem' := _, zero_mem' := _, algebra_map_mem' := _}, neg_mem' := _, inv_mem' := _}
Instances for ↥intermediate_field.adjoin
        
    Galois insertion between adjoin and coe.
Equations
- intermediate_field.gi = {choice := λ (s : set E) (hs : ↑(intermediate_field.adjoin F s) ≤ s), (intermediate_field.adjoin F s).copy s _, gc := _, le_l_u := _, choice_eq := _}
Equations
Construct an algebra isomorphism from an equality of intermediate fields
The top intermediate_field is isomorphic to the field.
This is the intermediate field version of subalgebra.top_equiv.
Equations
- intermediate_field.adjoin.field_coe F S = {coe := λ (x : F), ⟨⇑(algebra_map F E) x, _⟩}
If K is a field with F ⊆ K and S ⊆ K then adjoin F S ≤ K.
Variation on set.insert to enable good notation for adjoining elements to fields.
Used to preferentially use singleton rather than insert when adjoining one element.
Instances of this typeclass
Instances of other typeclasses for intermediate_field.insert
        
        - intermediate_field.insert.has_sizeof_inst
Equations
- intermediate_field.insert_empty = {insert := λ (x : α), {x}}
Equations
- intermediate_field.insert_nonempty s = {insert := λ (x : α), has_insert.insert x s}
A compositum of splitting fields is a splitting field
algebra isomorphism between adjoin_root and F⟮α⟯
Equations
The elements 1, x, ..., x ^ (d - 1) form a basis for K⟮x⟯,
where d is the degree of the minimal polynomial of x.
The power basis 1, x, ..., x ^ (d - 1) for K⟮x⟯,
where d is the degree of the minimal polynomial of x.
Equations
- intermediate_field.adjoin.power_basis hx = {gen := intermediate_field.adjoin_simple.gen K x, dim := (minpoly K x).nat_degree, basis := intermediate_field.power_basis_aux hx, basis_eq_pow := _}
Algebra homomorphism F⟮α⟯ →ₐ[F] K are in bijection with the set of roots
of minpoly α in K.
Equations
Fintype of algebra homomorphism F⟮α⟯ →ₐ[F] K
An intermediate field S is finitely generated if there exists t : finset E such that
intermediate_field.adjoin F t = S.
Equations
- intermediate_field.lifts.partial_order = {le := λ (x y : intermediate_field.lifts F E K), x.fst ≤ y.fst ∧ ∀ (s : ↥(x.fst)) (t : ↥(y.fst)), ↑s = ↑t → ⇑(x.snd) s = ⇑(y.snd) t, lt := preorder.lt._default (λ (x y : intermediate_field.lifts F E K), x.fst ≤ y.fst ∧ ∀ (s : ↥(x.fst)) (t : ↥(y.fst)), ↑s = ↑t → ⇑(x.snd) s = ⇑(y.snd) t), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _}
Equations
- intermediate_field.lifts.order_bot = {bot := ⟨⊥, (algebra.of_id F K).comp (intermediate_field.bot_equiv F E).to_alg_hom⟩, bot_le := _}
An upper bound on a chain of lifts
Equations
- intermediate_field.lifts.upper_bound_intermediate_field hc = {to_subalgebra := {carrier := λ (s : E), ∃ (x : intermediate_field.lifts F E K), x ∈ has_insert.insert ⊥ c ∧ s ∈ x.fst, mul_mem' := _, one_mem' := _, add_mem' := _, zero_mem' := _, algebra_map_mem' := _}, neg_mem' := _, inv_mem' := _}
An upper bound on a chain of lifts
Extend a lift x : lifts F E K to an element s : E whose conjugates are all in K
Equations
- x.lift_of_splits h1 h2 = let h3 : is_integral ↥(x.fst) s := _, key : polynomial.splits (alg_hom.to_ring_hom x.snd) (minpoly ↥(x.fst) s) := _ in ⟨intermediate_field.restrict_scalars F (↥(x.fst))⟮s⟯, alg_hom_equiv_sigma.inv_fun ⟨x.snd, (intermediate_field.alg_hom_adjoin_integral_equiv ↥(x.fst) h3).inv_fun ⟨polynomial.root_of_splits (alg_hom.to_ring_hom x.snd) key _, _⟩⟩⟩
A compositum of algebraic extensions is algebraic
pb.equiv_adjoin_simple is the equivalence between K⟮pb.gen⟯ and L itself.