Basic results about purely inseparable extensions #
This file contains basic definitions and results about purely inseparable extensions.
Main definitions #
IsPurelyInseparable
: typeclass for purely inseparable field extensions: an algebraic extensionE / F
is purely inseparable if and only if the minimal polynomial of every element ofE ∖ F
is not separable.
Main results #
IsPurelyInseparable.surjective_algebraMap_of_isSeparable
,IsPurelyInseparable.bijective_algebraMap_of_isSeparable
,IntermediateField.eq_bot_of_isPurelyInseparable_of_isSeparable
: ifE / F
is both purely inseparable and separable, thenalgebraMap F E
is surjective (hence bijective). In particular, if an intermediate field ofE / F
is both purely inseparable and separable, then it is equal toF
.isPurelyInseparable_iff_pow_mem
: a field extensionE / F
of exponential characteristicq
is purely inseparable if and only if for every elementx
ofE
, there exists a natural numbern
such thatx ^ (q ^ n)
is contained inF
.IsPurelyInseparable.trans
: ifE / F
andK / E
are both purely inseparable extensions, thenK / F
is also purely inseparable.isPurelyInseparable_iff_natSepDegree_eq_one
:E / F
is purely inseparable if and only if for every elementx
ofE
, its minimal polynomial has separable degree one.isPurelyInseparable_iff_minpoly_eq_X_pow_sub_C
: a field extensionE / F
of exponential characteristicq
is purely inseparable if and only if for every elementx
ofE
, the minimal polynomial ofx
overF
is of formX ^ (q ^ n) - y
for some natural numbern
and some elementy
ofF
.isPurelyInseparable_iff_minpoly_eq_X_sub_C_pow
: a field extensionE / F
of exponential characteristicq
is purely inseparable if and only if for every elementx
ofE
, the minimal polynomial ofx
overF
is of form(X - x) ^ (q ^ n)
for some natural numbern
.isPurelyInseparable_iff_finSepDegree_eq_one
: an extension is purely inseparable if and only if it has finite separable degree (Field.finSepDegree
) one.IsPurelyInseparable.normal
: a purely inseparable extension is normal.separableClosure.isPurelyInseparable
: ifE / F
is algebraic, thenE
is purely inseparable over the separable closure ofF
inE
.separableClosure_le_iff
: ifE / F
is algebraic, then an intermediate field ofE / F
contains the separable closure ofF
inE
if and only ifE
is purely inseparable over it.eq_separableClosure_iff
: ifE / F
is algebraic, then an intermediate field ofE / F
is equal to the separable closure ofF
inE
if and only if it is separable overF
, andE
is purely inseparable over it.IsPurelyInseparable.injective_comp_algebraMap
: ifE / F
is purely inseparable, then for any reduced ringL
, the map(E →+* L) → (F →+* L)
induced byalgebraMap F E
is injective. In particular, a purely inseparable field extension is an epimorphism in the category of fields.IsPurelyInseparable.of_injective_comp_algebraMap
: ifL
is an algebraically closed field containingE
, such that the map(E →+* L) → (F →+* L)
induced byalgebraMap F E
is injective, thenE / F
is purely inseparable. As a corollary, epimorphisms in the category of fields must be purely inseparable extensions.Field.finSepDegree_eq
: ifE / F
is algebraic, then theField.finSepDegree F E
is equal toField.sepDegree F E
as a natural number. This means that the cardinality ofField.Emb F E
and the degree of(separableClosure F E) / F
are both finite or infinite, and when they are finite, they coincide.Field.finSepDegree_mul_finInsepDegree
: the finite separable degree multiply by the finite inseparable degree is equal to the (finite) field extension degree.
Tags #
separable degree, degree, separable closure, purely inseparable
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.
- isIntegral : Algebra.IsIntegral F E
- inseparable' (x : E) : IsSeparable F x → x ∈ (algebraMap F E).range
Instances
Transfer IsPurelyInseparable
across an AlgEquiv
.
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.
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
.
If K / E / F
is a field extension tower such that K / F
is purely inseparable,
then E / F
is also purely inseparable.
If K / E / F
is a field extension tower such that K / F
is purely inseparable,
then K / E
is also purely inseparable.
If E / F
and K / E
are both purely inseparable extensions, then K / F
is also
purely inseparable.
Stacks Tag 02JJ (See also 00GM)
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
.
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
.
If an extension has finite separable degree one, then it is purely inseparable.
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
.
Equations
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 algebraic extension is purely inseparable if and only if all of its finite dimensional subextensions are purely inseparable.
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.)
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.
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.
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)
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
.