Documentation

Mathlib.RingTheory.Frobenius

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.

Let G be a finite group acting on a ring S, and R is the fixed subring of S.

def AlgHom.IsArithFrobAt {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] (Ο† : S →ₐ[R] S) (Q : Ideal S) :

Ο† : S →ₐ[R] S is an (arithmetic) Frobenius at Q if Ο† x ≑ x ^ #(R/p) (mod Q) for all x : S (AlgHom.IsArithFrobAt).

Equations
Instances For
    theorem AlgHom.IsArithFrobAt.mk_apply {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) (x : S) :
    theorem AlgHom.IsArithFrobAt.finite_quotient {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) :
    theorem AlgHom.IsArithFrobAt.card_pos {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) :
    theorem AlgHom.IsArithFrobAt.le_comap {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) :
    def AlgHom.IsArithFrobAt.restrict {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) :

    A Frobenius element at Q restricts to the Frobenius map on S β§Έ Q.

    Equations
    Instances For
      theorem AlgHom.IsArithFrobAt.restrict_apply {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) (x : S β§Έ Q) :
      theorem AlgHom.IsArithFrobAt.restrict_mk {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) (x : S) :
      theorem AlgHom.IsArithFrobAt.restrict_injective {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) [Q.IsPrime] :
      theorem AlgHom.IsArithFrobAt.comap_eq {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) [Q.IsPrime] :
      Ideal.comap Ο† Q = Q
      theorem AlgHom.IsArithFrobAt.apply_of_pow_eq_one {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) [IsDomain S] {ΞΆ : S} {m : β„•} (hΞΆ : ΞΆ ^ m = 1) (hk' : ↑m βˆ‰ Q) :
      Ο† ΞΆ = ΞΆ ^ Nat.card (R β§Έ Ideal.under R Q)

      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.

      noncomputable def AlgHom.IsArithFrobAt.localize {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) [Q.IsPrime] :

      A Frobenius element at Q restricts to an automorphism of S_Q.

      Equations
      Instances For
        @[simp]
        theorem AlgHom.IsArithFrobAt.localize_algebraMap {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) [Q.IsPrime] (x : S) :
        theorem AlgHom.IsArithFrobAt.eq_of_isUnramifiedAt {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {Ο† ψ : S →ₐ[R] S} {Q : Ideal S} (H : Ο†.IsArithFrobAt Q) (H' : ψ.IsArithFrobAt Q) [Q.IsPrime] (hQ : Q.primeCompl ≀ nonZeroDivisors S) [Algebra.IsUnramifiedAt R Q] [IsNoetherianRing S] :
        Ο† = ψ

        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.

        @[reducible, inline]
        abbrev IsArithFrobAt (R : Type u_1) {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {M : Type u_3} [Monoid M] [MulSemiringAction M S] [SMulCommClass M R S] (Οƒ : M) (Q : Ideal S) :

        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
        Instances For
          theorem IsArithFrobAt.mul_inv_mem_inertia {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {G : Type u_3} [Group G] [MulSemiringAction G S] [SMulCommClass G R S] {Q : Ideal S} {Οƒ Οƒ' : G} (H : IsArithFrobAt R Οƒ Q) (H' : IsArithFrobAt R Οƒ' Q) :
          theorem IsArithFrobAt.conj {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] {G : Type u_3} [Group G] [MulSemiringAction G S] [SMulCommClass G R S] {Q : Ideal S} {Οƒ : G} (H : IsArithFrobAt R Οƒ Q) (Ο„ : G) :
          IsArithFrobAt R (Ο„ * Οƒ * τ⁻¹) (Ο„ β€’ Q)
          theorem IsArithFrobAt.exists_of_isInvariant (R : Type u_1) {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] (G : Type u_3) [Group G] [MulSemiringAction G S] [SMulCommClass G R S] (Q : Ideal S) [Finite G] [Algebra.IsInvariant R S G] [Q.IsPrime] [Finite (S β§Έ Q)] :
          βˆƒ (Οƒ : G), IsArithFrobAt R Οƒ Q

          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.

          theorem IsArithFrobAt.exists_primesOver_isConj {R : Type u_1} (S : Type u_2) [CommRing R] [CommRing S] [Algebra R S] (G : Type u_3) [Group G] [MulSemiringAction G S] [SMulCommClass G R S] [Finite G] [Algebra.IsInvariant R S G] (P : Ideal R) (hP : βˆƒ (Q : ↑(P.primesOver S)), Finite (S β§Έ ↑Q)) :
          βˆƒ (Οƒ : ↑(P.primesOver S) β†’ G), (βˆ€ (Q : ↑(P.primesOver S)), IsArithFrobAt R (Οƒ Q) ↑Q) ∧ βˆ€ (Q₁ Qβ‚‚ : ↑(P.primesOver S)), IsConj (Οƒ Q₁) (Οƒ Qβ‚‚)
          noncomputable def arithFrobAt (R : Type u_1) {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] (G : Type u_3) [Group G] [MulSemiringAction G S] [SMulCommClass G R S] (Q : Ideal S) [Finite G] [Algebra.IsInvariant R S G] [Q.IsPrime] [Finite (S β§Έ Q)] :
          G

          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
          Instances For
            theorem IsArithFrobAt.arithFrobAt (R : Type u_1) {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] (G : Type u_3) [Group G] [MulSemiringAction G S] [SMulCommClass G R S] (Q : Ideal S) [Finite G] [Algebra.IsInvariant R S G] [Q.IsPrime] [Finite (S β§Έ Q)] :
            theorem isConj_arithFrobAt (R : Type u_1) {S : Type u_2} [CommRing R] [CommRing S] [Algebra R S] (G : Type u_3) [Group G] [MulSemiringAction G S] [SMulCommClass G R S] (Q : Ideal S) [Finite G] [Algebra.IsInvariant R S G] [Q.IsPrime] [Finite (S β§Έ Q)] (Q' : Ideal S) [Q'.IsPrime] [Finite (S β§Έ Q')] (H : Ideal.under R Q = Ideal.under R Q') :
            IsConj (arithFrobAt R G Q) (arithFrobAt R G Q')