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.
- Define
ArithmeticFunction.vonMangoldt.residueClass a
fora : ZMod q
, which is a functionℕ → ℝ
taking the value zero when(n : ℤMod q) ≠ a
andΛ n
else (whereΛ
is the von Mangoldt functionArithmeticFunction.vonMangoldt
; we haveΛ (p^k) = log p
for prime powers andΛ n = 0
otherwise.) - Show that this function can be written as a linear combination of functions
of the form
χ * Λ
(pointwise product) with Dirichlet charactersχ
modq
. SeeArithmeticFunction.vonMangoldt.residueClass_eq
. - This implies that the L-series of
ArithmeticFunction.vonMangoldt.residueClass a
agrees (onre s > 1
) with the corresponding linear combination of negative logarithmic derivatives of Dirichlet L-functions. SeeArithmeticFunction.vonMangoldt.LSeries_residueClass_eq
. - 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 ats = 1
. SeeArithmeticFunction.vonMangoldt.eqOn_LFunctionResidueClassAux
for the statement that the auxiliary function agrees with the L-series ofArithmeticFunction.vonMangoldt.residueClass
up to the term(q.totient)⁻¹/(s-1)
. - Show that the auxiliary function is continuous on
re s ≥ 1
; seeArithmeticFunction.vonMangoldt.continuousOn_LFunctionResidueClassAux
. This relies heavily on the non-vanishing of Dirichlet L-functions on the closed half-planere 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 ats = 1
); seeDirichletCharacter.LFunction_eq_LSeries
(contributed by David Loeffler). - Show that the sum of
Λ n / n
over any residue class, but excluding the primes, converges. SeeArithmeticFunction.vonMangoldt.summable_residueClass_non_primes_div
. - Combining these ingredients, we can deduce that the sum of
Λ n / n
over the primes in a residue class must diverge. SeeArithmeticFunction.vonMangoldt.not_summable_residueClass_prime_div
. - This finally easily implies that there must be infinitely many primes in the residue class.
Definitions #
ArithmeticFunction.vonMangoldt.residueClass a
(see above).ArithmeticFunction.vonMangoldt.continuousOn_LFunctionResidueClassAux
(see above).
Main Result #
We give two versions of Dirichlet's Theorem:
Nat.setOf_prime_and_eq_mod_infinite
states that the set of primesp
such that(p : ZMod q) = a
is infinite (whena
is invertible inZMod q
).Nat.forall_exists_prime_gt_and_eq_mod
states that for any natural numbern
there is a primep > n
such that(p : ZMod q) = a
.
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.
The L-series of the von Mangoldt function restricted to a residue class #
The von Mangoldt function restricted to the residue class a
mod q
.
Equations
- ArithmeticFunction.vonMangoldt.residueClass a = {n : ℕ | ↑n = a}.indicator fun (x : ℕ) => ArithmeticFunction.vonMangoldt x
Instances For
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
, restriced 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.
We can express ArithmeticFunction.vonMangoldt.residueClass
as a linear combination
of twists of the von Mangoldt function by Dirichlet charaters.
We can express ArithmeticFunction.vonMangoldt.residueClass
as a linear combination
of twists of the von Mangoldt function by Dirichlet charaters.
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
).
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.
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
.
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
.
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.