Documentation

Mathlib.FieldTheory.PurelyInseparable.Basic

Basic results about purely inseparable extensions #

This file contains basic definitions and results about purely inseparable extensions.

Main definitions #

Main results #

Tags #

separable degree, degree, separable closure, purely inseparable

class IsPurelyInseparable (F : Type u_1) (E : Type u_2) [CommRing F] [Ring E] [Algebra F E] :

Typeclass for purely inseparable field extensions: an algebraic extension E / F is purely inseparable if and only if the minimal polynomial of every element of E ∖ F is not separable.

We define this for general (commutative) rings and only assume F and E are fields if this is needed for a proof.

Instances
    theorem IsPurelyInseparable.isIntegral' (F : Type u_1) {E : Type u_2} [CommRing F] [Ring E] [Algebra F E] [IsPurelyInseparable F E] (x : E) :
    theorem IsPurelyInseparable.inseparable (F : Type u_1) {E : Type u_2} [CommRing F] [Ring E] [Algebra F E] [IsPurelyInseparable F E] (x : E) :
    theorem isPurelyInseparable_iff {F : Type u_1} {E : Type u_2} [CommRing F] [Ring E] [Algebra F E] :
    IsPurelyInseparable F E ∀ (x : E), IsIntegral F x (IsSeparable F xx (algebraMap F E).range)
    theorem AlgEquiv.isPurelyInseparable {F : Type u_1} {E : Type u_2} [CommRing F] [Ring E] [Algebra F E] {K : Type u_3} [Ring K] [Algebra F K] (e : K ≃ₐ[F] E) [IsPurelyInseparable F K] :

    Transfer IsPurelyInseparable across an AlgEquiv.

    theorem AlgEquiv.isPurelyInseparable_iff {F : Type u_1} {E : Type u_2} [CommRing F] [Ring E] [Algebra F E] {K : Type u_3} [Ring K] [Algebra F K] (e : K ≃ₐ[F] E) :

    If E / F is an algebraic extension, F is separably closed, then E / F is purely inseparable.

    If E / F is both purely inseparable and separable, then algebraMap F E is surjective.

    If E / F is both purely inseparable and separable, then algebraMap F E is bijective.

    If a subalgebra of E / F is both purely inseparable and separable, then it is equal to F.

    If an intermediate field of E / F is both purely inseparable and separable, then it is equal to F.

    If E / F is purely inseparable, then the separable closure of F in E is equal to F.

    If E / F is an algebraic extension, then the separable closure of F in E is equal to F if and only if E / F is purely inseparable.

    theorem isPurelyInseparable_iff_pow_mem (F : Type u) {E : Type v} [Field F] [Ring E] [IsDomain E] [Algebra F E] (q : ) [ExpChar F q] :
    IsPurelyInseparable F E ∀ (x : E), ∃ (n : ), x ^ q ^ n (algebraMap F E).range

    A field extension E / F of exponential characteristic q is purely inseparable if and only if for every element x of E, there exists a natural number n such that x ^ (q ^ n) is contained in F.

    Stacks Tag 09HE

    theorem IsPurelyInseparable.pow_mem (F : Type u) {E : Type v} [Field F] [Ring E] [IsDomain E] [Algebra F E] (q : ) [ExpChar F q] (x : E) [IsPurelyInseparable F E] :
    ∃ (n : ), x ^ q ^ n (algebraMap F E).range
    theorem IsPurelyInseparable.tower_bot (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] (K : Type w) [Field K] [Algebra F K] [Algebra E K] [IsScalarTower F E K] [IsPurelyInseparable F K] :

    If K / E / F is a field extension tower such that K / F is purely inseparable, then E / F is also purely inseparable.

    theorem IsPurelyInseparable.tower_top (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] (K : Type w) [Field K] [Algebra F K] [Algebra E K] [IsScalarTower F E K] [h : IsPurelyInseparable F K] :

    If K / E / F is a field extension tower such that K / F is purely inseparable, then K / E is also purely inseparable.

    theorem IsPurelyInseparable.trans (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] (K : Type w) [Field K] [Algebra F K] [Algebra E K] [IsScalarTower F E K] [h1 : IsPurelyInseparable F E] [h2 : IsPurelyInseparable E K] :

    If E / F and K / E are both purely inseparable extensions, then K / F is also purely inseparable.

    Stacks Tag 02JJ (See also 00GM)

    theorem isPurelyInseparable_iff_natSepDegree_eq_one (F : Type u) {E : Type v} [Field F] [Field E] [Algebra F E] :
    IsPurelyInseparable F E ∀ (x : E), (minpoly F x).natSepDegree = 1

    A field extension E / F is purely inseparable if and only if for every element x of E, its minimal polynomial has separable degree one.

    theorem isPurelyInseparable_iff_minpoly_eq_X_pow_sub_C (F : Type u) {E : Type v} [Field F] [Field E] [Algebra F E] (q : ) [hF : ExpChar F q] :
    IsPurelyInseparable F E ∀ (x : E), ∃ (n : ) (y : F), minpoly F x = Polynomial.X ^ q ^ n - Polynomial.C y

    A field extension E / F of exponential characteristic q is purely inseparable if and only if for every element x of E, the minimal polynomial of x over F is of form X ^ (q ^ n) - y for some natural number n and some element y of F.

    theorem IsPurelyInseparable.minpoly_eq_X_pow_sub_C (F : Type u) {E : Type v} [Field F] [Field E] [Algebra F E] (q : ) [ExpChar F q] [IsPurelyInseparable F E] (x : E) :
    ∃ (n : ) (y : F), minpoly F x = Polynomial.X ^ q ^ n - Polynomial.C y
    theorem isPurelyInseparable_iff_minpoly_eq_X_sub_C_pow (F : Type u) {E : Type v} [Field F] [Field E] [Algebra F E] (q : ) [hF : ExpChar F q] :
    IsPurelyInseparable F E ∀ (x : E), ∃ (n : ), Polynomial.map (algebraMap F E) (minpoly F x) = (Polynomial.X - Polynomial.C x) ^ q ^ n

    A field extension E / F of exponential characteristic q is purely inseparable if and only if for every element x of E, the minimal polynomial of x over F is of form (X - x) ^ (q ^ n) for some natural number n.

    theorem IsPurelyInseparable.minpoly_eq_X_sub_C_pow (F : Type u) {E : Type v} [Field F] [Field E] [Algebra F E] (q : ) [ExpChar F q] [IsPurelyInseparable F E] (x : E) :

    If an extension has finite separable degree one, then it is purely inseparable.

    theorem IsPurelyInseparable.injective_comp_algebraMap (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] [IsPurelyInseparable F E] (L : Type u_2) [CommRing L] [IsReduced L] :
    Function.Injective fun (f : E →+* L) => f.comp (algebraMap F E)

    If E / F is purely inseparable, then for any reduced ring L, the map (E →+* L) → (F →+* L) induced by algebraMap F E is injective. In particular, a purely inseparable field extension is an epimorphism in the category of fields.

    If E / F is purely inseparable, then for any reduced F-algebra L, there exists at most one F-algebra homomorphism from E to L.

    If E / F is purely inseparable, then Field.Emb F E has exactly one element.

    Equations

    A purely inseparable extension has finite separable degree one.

    A purely inseparable extension has separable degree one.

    A purely inseparable extension has inseparable degree equal to degree.

    A purely inseparable extension has finite inseparable degree equal to degree.

    An extension is purely inseparable if and only if it has finite separable degree one.

    An algebraic extension is purely inseparable if and only if all of its finite dimensional subextensions are purely inseparable.

    instance IsPurelyInseparable.normal (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] [IsPurelyInseparable F E] :
    Normal F E

    A purely inseparable extension is normal.

    If E / F is algebraic, then E is purely inseparable over the separable closure of F in E.

    Stacks Tag 030K ($E/E_{sep}$ is purely inseparable.)

    theorem separableClosure_le (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] (L : IntermediateField F E) [h : IsPurelyInseparable (↥L) E] :

    An intermediate field of E / F contains the separable closure of F in E if E is purely inseparable over it.

    If E / F is algebraic, then an intermediate field of E / F contains the separable closure of F in E if and only if E is purely inseparable over it.

    theorem eq_separableClosure (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] (L : IntermediateField F E) [Algebra.IsSeparable F L] [IsPurelyInseparable (↥L) E] :

    If an intermediate field of E / F is separable over F, and E is purely inseparable over it, then it is equal to the separable closure of F in E.

    If E / F is algebraic, then an intermediate field of E / F is equal to the separable closure of F in E if and only if it is separable over F, and E is purely inseparable over it.

    theorem IsPurelyInseparable.of_injective_comp_algebraMap (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] (L : Type w) [Field L] [IsAlgClosed L] [Nonempty (E →+* L)] (h : Function.Injective fun (f : E →+* L) => f.comp (algebraMap F E)) :

    If L is an algebraically closed field containing E, such that the map (E →+* L) → (F →+* L) induced by algebraMap F E is injective, then E / F is purely inseparable. As a corollary, epimorphisms in the category of fields must be purely inseparable extensions.

    If E is an algebraic closure of F, then F is separably closed if and only if E / F is purely inseparable.

    If E / F is an algebraic extension, F is separably closed, then E is also separably closed.

    If E / F is algebraic, then the Field.finSepDegree F E is equal to Field.sepDegree F E as a natural number. This means that the cardinality of Field.Emb F E and the degree of (separableClosure F E) / F are both finite or infinite, and when they are finite, they coincide.

    Stacks Tag 09HJ (sepDegree is defined as the right hand side of 09HJ)

    The finite separable degree multiply by the finite inseparable degree is equal to the (finite) field extension degree.

    If K / E / F is a field extension tower, such that E / F is algebraic and K / E is separable, then E adjoin separableClosure F K is equal to K. It is a special case of separableClosure.adjoin_eq_of_isAlgebraic, and is an intermediate result used to prove it.

    If K / E / F is a field extension tower, such that E / F is algebraic, then E adjoin separableClosure F K is equal to separableClosure E K.