Frobenius elements #
In algebraic number theory, if L/K
is a finite Galois extension of number fields, with rings of
integers πL/πK
, and if q
is prime ideal of πL
lying over a prime ideal p
of πK
, then
there exists a Frobenius element Frob p
in Gal(L/K)
with the property that
Frob p x β‘ x ^ #(πK/p) (mod q)
for all x β πL
.
Following RingTheory/Invariant.lean
, we develop the theory in the setting that
there is a finite group G
acting on a ring S
, and R
is the fixed subring of S
.
Main results #
Let S/R
be an extension of rings, Q
be a prime of S
,
and P := R β© Q
with finite residue field of cardinality q
.
AlgHom.IsArithFrobAt
: We say that aΟ : S ββ[R] S
is an (arithmetic) Frobenius atQ
ifΟ x β‘ x ^ q (mod Q)
for allx : S
.AlgHom.IsArithFrobAt.apply_of_pow_eq_one
: SupposeS
is a domain andΟ
is a Frobenius atQ
, thenΟ ΞΆ = ΞΆ ^ q
for anym
-th root of unityΞΆ
withq β€ m
.AlgHom.IsArithFrobAt.eq_of_isUnramifiedAt
: SupposeS
is noetherian,Q
contains all zero-divisors, and the extension is unramified atQ
. Then the Frobenius is unique (if exists).
Let G
be a finite group acting on a ring S
, and R
is the fixed subring of S
.
IsArithFrobAt
: We say that aΟ : G
is an (arithmetic) Frobenius atQ
ifΟ β’ x β‘ x ^ q (mod Q)
for allx : S
.IsArithFrobAt.mul_inv_mem_inertia
: Two Frobenius elements atQ
differ by an element in the inertia subgroup ofQ
.IsArithFrobAt.conj
: IfΟ
is a Frobenius atQ
, thenΟΟΟβ»ΒΉ
is a Frobenius atΟ β’ Q
.IsArithFrobAt.exists_of_isInvariant
: Frobenius element exists.
Ο : S ββ[R] S
is an (arithmetic) Frobenius at Q
if
Ο x β‘ x ^ #(R/p) (mod Q)
for all x : S
(AlgHom.IsArithFrobAt
).
Equations
- Ο.IsArithFrobAt Q = β (x : S), Ο x - x ^ Nat.card (R β§Έ Ideal.under R Q) β Q
Instances For
A Frobenius element at Q
restricts to the Frobenius map on S β§Έ Q
.
Equations
- H.restrict = { toRingHom := Ideal.quotientMap Q βΟ β―, commutes' := β― }
Instances For
Suppose S
is a domain, and Ο : S ββ[R] S
is a Frobenius at Q : Ideal S
.
Let ΞΆ
be a m
-th root of unity with Q β€ m
, then Ο
sends ΞΆ
to ΞΆ ^ q
.
A Frobenius element at Q
restricts to an automorphism of S_Q
.
Equations
- H.localize = { toRingHom := Localization.localRingHom Q Q βΟ β―, commutes' := β― }
Instances For
Suppose S
is noetherian and Q
is a prime of S
containing all zero divisors.
If S/R
is unramified at Q
, then the Frobenius Ο : S ββ[R] S
over Q
is unique.
Suppose S
is an R
algebra, M
is a monoid acting on S
whose action is trivial on R
Ο : M
is an (arithmetic) Frobenius at an ideal Q
of S
if Ο β’ x β‘ x ^ q (mod Q)
for all x
.
Equations
- IsArithFrobAt R Ο Q = (MulSemiringAction.toAlgHom R S Ο).IsArithFrobAt Q
Instances For
Let G
be a finite group acting on S
, and R
be the fixed subring.
If Q
is a prime of S
with finite residue field,
then there exists a Frobenius element Ο : G
at Q
.
Let G
be a finite group acting on S
, R
be the fixed subring, and Q
be a prime of S
with finite residue field. This is an arbitrary choice of a Frobenius over Q
. It is chosen so that
the Frobenius elements of Qβ
and Qβ
are conjugate if they lie over the same prime.
Equations
- arithFrobAt R G Q = β―.choose β¨Q, β―β©