Documentation

Mathlib.NumberTheory.LSeries.PrimesInAP

Dirichlet's Theorem on primes in arithmetic progression #

The goal of this file is to prove Dirichlet's Theorem: If q is a positive natural number and a : ZMod q is invertible, then there are infinitely many prime numbers p such that (p : ZMod q) = a.

The main steps of the proof are as follows.

  1. Define ArithmeticFunction.vonMangoldt.residueClass a for a : ZMod q, which is a function ℕ → ℝ taking the value zero when (n : ℤMod q) ≠ a and Λ n else (where Λ is the von Mangoldt function ArithmeticFunction.vonMangoldt; we have Λ (p^k) = log p for prime powers and Λ n = 0 otherwise.)
  2. Show that this function can be written as a linear combination of functions of the form χ * Λ (pointwise product) with Dirichlet characters χ mod q. See ArithmeticFunction.vonMangoldt.residueClass_eq.
  3. This implies that the L-series of ArithmeticFunction.vonMangoldt.residueClass a agrees (on re s > 1) with the corresponding linear combination of negative logarithmic derivatives of Dirichlet L-functions. See ArithmeticFunction.vonMangoldt.LSeries_residueClass_eq.
  4. Define an auxiliary function ArithmeticFunction.vonMangoldt.LFunctionResidueClassAux a that is this linear combination of negative logarithmic derivatives of L-functions minus (q.totient)⁻¹/(s-1), which cancels the pole at s = 1. See ArithmeticFunction.vonMangoldt.eqOn_LFunctionResidueClassAux for the statement that the auxiliary function agrees with the L-series of ArithmeticFunction.vonMangoldt.residueClass up to the term (q.totient)⁻¹/(s-1).
  5. Show that the auxiliary function is continuous on re s ≥ 1; see ArithmeticFunction.vonMangoldt.continuousOn_LFunctionResidueClassAux. This relies heavily on the non-vanishing of Dirichlet L-functions on the closed half-plane re s ≥ 1 (DirichletCharacter.LFunction_ne_zero_of_one_le_re), which in turn can only be stated since we know that the L-series of a Dirichlet character extends to an entire function (unless the character is trivial; then there is a simple pole at s = 1); see DirichletCharacter.LFunction_eq_LSeries (contributed by David Loeffler).
  6. Show that the sum of Λ n / n over any residue class, but excluding the primes, converges. See ArithmeticFunction.vonMangoldt.summable_residueClass_non_primes_div.
  7. Combining these ingredients, we can deduce that the sum of Λ n / n over the primes in a residue class must diverge. See ArithmeticFunction.vonMangoldt.not_summable_residueClass_prime_div.
  8. This finally easily implies that there must be infinitely many primes in the residue class.

Definitions #

Main Result #

We give two versions of Dirichlet's Theorem:

Tags #

prime number, arithmetic progression, residue class, Dirichlet's Theorem

Auxiliary statements #

An infinite product or sum over a function supported in prime powers can be written as an iterated product or sum over primes and natural numbers.

theorem tprod_eq_tprod_primes_of_mulSupport_subset_prime_powers {α : Type u_1} [CommGroup α] [UniformSpace α] [UniformGroup α] [CompleteSpace α] [T0Space α] {f : α} (hfm : Multipliable f) (hf : Function.mulSupport f {n : | IsPrimePow n}) :
∏' (n : ), f n = ∏' (p : Nat.Primes) (k : ), f (p ^ (k + 1))
theorem tsum_eq_tsum_primes_of_support_subset_prime_powers {α : Type u_1} [AddCommGroup α] [UniformSpace α] [UniformAddGroup α] [CompleteSpace α] [T0Space α] {f : α} (hfm : Summable f) (hf : Function.support f {n : | IsPrimePow n}) :
∑' (n : ), f n = ∑' (p : Nat.Primes) (k : ), f (p ^ (k + 1))
theorem tprod_eq_tprod_primes_mul_tprod_primes_of_mulSupport_subset_prime_powers {α : Type u_1} [CommGroup α] [UniformSpace α] [UniformGroup α] [CompleteSpace α] [T0Space α] {f : α} (hfm : Multipliable f) (hf : Function.mulSupport f {n : | IsPrimePow n}) :
∏' (n : ), f n = (∏' (p : Nat.Primes), f p) * ∏' (p : Nat.Primes) (k : ), f (p ^ (k + 2))
theorem tsum_eq_tsum_primes_add_tsum_primes_of_support_subset_prime_powers {α : Type u_1} [AddCommGroup α] [UniformSpace α] [UniformAddGroup α] [CompleteSpace α] [T0Space α] {f : α} (hfm : Summable f) (hf : Function.support f {n : | IsPrimePow n}) :
∑' (n : ), f n = ∑' (p : Nat.Primes), f p + ∑' (p : Nat.Primes) (k : ), f (p ^ (k + 2))

The L-series of the von Mangoldt function restricted to a residue class #

@[reducible, inline]
noncomputable abbrev ArithmeticFunction.vonMangoldt.residueClass {q : } (a : ZMod q) :

The von Mangoldt function restricted to the residue class a mod q.

Equations
Instances For
    theorem ArithmeticFunction.vonMangoldt.support_residueClass_prime_div {q : } (a : ZMod q) :
    (Function.support fun (n : ) => (if Nat.Prime n then residueClass a n else 0) / n) = {p : | Nat.Prime p p = a}

    The set we are interested in (prime numbers in the residue class a) is the same as the support of ArithmeticFunction.vonMangoldt.residueClass restricted to primes (and divided by n; this is how this result is used later).

    The function n ↦ Λ n / n, restricted to non-primes in a residue class, is summable. This is used to convert results on ArithmeticFunction.vonMangoldt.residueClass to results on primes in an arithmetic progression.

    theorem ArithmeticFunction.vonMangoldt.residueClass_apply {q : } {a : ZMod q} [NeZero q] (ha : IsUnit a) (n : ) :
    (residueClass a n) = (↑q.totient)⁻¹ * χ : DirichletCharacter q, χ a⁻¹ * χ n * (vonMangoldt n)

    We can express ArithmeticFunction.vonMangoldt.residueClass as a linear combination of twists of the von Mangoldt function by Dirichlet characters.

    theorem ArithmeticFunction.vonMangoldt.residueClass_eq {q : } {a : ZMod q} [NeZero q] (ha : IsUnit a) :
    (fun (n : ) => (residueClass a n)) = (↑q.totient)⁻¹ χ : DirichletCharacter q, χ a⁻¹ fun (n : ) => χ n * (vonMangoldt n)

    We can express ArithmeticFunction.vonMangoldt.residueClass as a linear combination of twists of the von Mangoldt function by Dirichlet characters.

    theorem ArithmeticFunction.vonMangoldt.LSeries_residueClass_eq {q : } {a : ZMod q} [NeZero q] (ha : IsUnit a) {s : } (hs : 1 < s.re) :
    LSeries (fun (n : ) => (residueClass a n)) s = -(↑q.totient)⁻¹ * χ : DirichletCharacter q, χ a⁻¹ * (deriv (DirichletCharacter.LFunction χ) s / DirichletCharacter.LFunction χ s)

    The L-series of the von Mangoldt function restricted to the residue class a mod q with a invertible in ZMod q is a linear combination of logarithmic derivatives of L-functions of the Dirichlet characters mod q (on re s > 1).

    @[reducible, inline]
    noncomputable abbrev ArithmeticFunction.vonMangoldt.LFunctionResidueClassAux {q : } (a : ZMod q) [NeZero q] (s : ) :

    The auxiliary function used, e.g., with the Wiener-Ikehara Theorem to prove Dirichlet's Theorem. On re s > 1, it agrees with the L-series of the von Mangoldt function restricted to the residue class a : ZMod q minus the principal part (q.totient)⁻¹/(s-1) of the pole at s = 1; see ArithmeticFunction.vonMangoldt.eqOn_LFunctionResidueClassAux.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For

      The auxiliary function is continuous away from the zeros of the L-functions of the Dirichlet characters mod q (including at s = 1).

      The L-series of the von Mangoldt function restricted to the prime residue class a mod q is continuous on re s ≥ 1 except for a simple pole at s = 1 with residue (q.totient)⁻¹. The statement as given here in terms of ArithmeticFunction.vonMangoldt.LFunctionResidueClassAux is equivalent.

      theorem ArithmeticFunction.vonMangoldt.eqOn_LFunctionResidueClassAux {q : } {a : ZMod q} [NeZero q] (ha : IsUnit a) :
      Set.EqOn (LFunctionResidueClassAux a) (fun (s : ) => LSeries (fun (n : ) => (residueClass a n)) s - (↑q.totient)⁻¹ / (s - 1)) {s : | 1 < s.re}

      The auxiliary function agrees on re s > 1 with the L-series of the von Mangoldt function restricted to the residue class a : ZMod q minus the principal part (q.totient)⁻¹/(s-1) of its pole at s = 1.

      The auxiliary function takes real values for real arguments x > 1.

      theorem ArithmeticFunction.vonMangoldt.LSeries_residueClass_lower_bound {q : } [NeZero q] {a : ZMod q} (ha : IsUnit a) :
      ∃ (C : ), ∀ {x : }, x Set.Ioc 1 2(↑q.totient)⁻¹ / (x - 1) - C ∑' (n : ), residueClass a n / n ^ x

      As x approaches 1 from the right along the real axis, the L-series of ArithmeticFunction.vonMangoldt.residueClass is bounded below by (q.totient)⁻¹/(x-1) - C.

      theorem ArithmeticFunction.vonMangoldt.not_summable_residueClass_prime_div {q : } [NeZero q] {a : ZMod q} (ha : IsUnit a) :
      ¬Summable fun (n : ) => (if Nat.Prime n then residueClass a n else 0) / n

      The function n ↦ Λ n / n restricted to primes in an invertible residue class is not summable. This then implies that there must be infinitely many such primes.

      Dirichlet's Theorem #

      theorem Nat.setOf_prime_and_eq_mod_infinite {q : } [NeZero q] {a : ZMod q} (ha : IsUnit a) :
      {p : | Prime p p = a}.Infinite

      Dirichlet's Theorem on primes in arithmetic progression: if q is a positive integer and a : ZMod q is a unit, then there are infintely many prime numbers p such that (p : ZMod q) = a.

      theorem Nat.forall_exists_prime_gt_and_eq_mod {q : } [NeZero q] {a : ZMod q} (ha : IsUnit a) (n : ) :
      p > n, Prime p p = a

      Dirichlet's Theorem on primes in arithmetic progression: if q is a positive integer and a : ZMod q is a unit, then there are infintely many prime numbers p such that (p : ZMod q) = a.