The exponent of purely inseparable extensions #
This file defines the exponent of a purely inseparable extension (if one exists) and some related results.
Most results are stated using ringExpChar K
rather than using [ExpChar K p]
parameter because
it gives cleaner API. To use the results in a context with [ExpChar K p]
, consider using
ringExpChar.eq K p
for substitution.
Main definitions #
IsPurelyInseparable.HasExponent
: typeclass to assert a purely inseparable field extensionL / K
has an exponent, that is a smallest natural numbere
such thata ^ ringExpChar K ^ e ∈ K
for alla ∈ L
.IsPurelyInseparable.exponent
: the exponent of a purely inseparable field extension.IsPurelyInseparable.elemExponent
: the exponent of an element of a purely inseparable field extension, that is the smallest natural numbere
such thata ^ ringExpChar K ^ e ∈ K
.IsPurelyInseparable.iterateFrobenius
: the iterated Frobenius map (ring homomorphism)L →+* K
for purely inseparable field extensionL / K
with exponent; forn ≥ exponent K L
, it acts likex ↦ x ^ p ^ n
but the codomain is the base fieldK
.IsPurelyInseparable.iterateFrobeniusₛₗ
: version ofiterateFrobenius
as a semilinear map over a subfieldF
ofK
, wrt the iterated Frobenius homomorphism onF
.
Tags #
purely inseparable
A predicate class on a ring extension saying that there is a natural number e
such that a ^ ringExpChar K ^ e ∈ K
for all a ∈ L
.
- has_exponent : ∃ (e : ℕ), ∀ (a : L), a ^ ringExpChar K ^ e ∈ (algebraMap K L).range
Instances
Version of hasExponent_iff
using ExpChar
.
The exponent of a purely inseparable extension is the smallest
natural number e
such that a ^ ringExpChar K ^ e ∈ K
for all a ∈ L
.
Equations
Instances For
Version of exponent_def
using ExpChar
.
Version of exponent_min
using ExpChar
.
The exponent of an element a ∈ L
of a purely inseparable field extension L / K
is the smallest natural number e
such that a ^ ringExpChar K ^ e ∈ K
.
Equations
Instances For
The element y
of the base field K
such that
a ^ ringExpChar K ^ elemExponent K a = algebraMap K L y
.
See IsPurelyInseparable.algebraMap_elemReduct_eq
.
Equations
Instances For
Version of minpoly_eq
using ExpChar
.
The degree of the minimal polynomial of an element a ∈ L
equals
ringExpChar K ^ elemExponent K a
.
Version of minpoly_natDegree_eq
using ExpChar
.
Version of algebraMap_elemReduct_eq
using ExpChar
.
Version of elemExponent_def
using ExpChar
.
Version of elemExponent_le_of_pow_mem
using ExpChar
.
Version of elemExponent_min
using ExpChar
.
An exponent of an element is less or equal than exponent of the extension.
Iterated Frobenius map (ring homomorphism) for purely inseparable field extension with exponent.
If n ≥ exponent K L
, it acts like x ↦ x ^ p ^ n
but the codomain is the base field K
.
Equations
- IsPurelyInseparable.iterateFrobenius K L p hn = { toFun := IsPurelyInseparable.iterateFrobeniusAux✝ K L p n, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Action of iterateFrobenius
on the top field.
Action of iterateFrobenius
on the bottom field.
Version of iterateFrobenius
as a semilinear map over a subfield F
of K
, wrt the
iterated Frobenius homomorphism on F
.
Equations
- IsPurelyInseparable.iterateFrobeniusₛₗ F K L p hn = { toFun := (↑↑(IsPurelyInseparable.iterateFrobenius K L p hn)).toFun, map_add' := ⋯, map_smul' := ⋯ }
Instances For
Action of iterateFrobeniusₛₗ
on the top field.
Action of iterateFrobeniusₛₗ
on the bottom field.
Action of iterateFrobeniusₛₗ
on the base field.