Documentation

Mathlib.NumberTheory.LucasLehmer

The Lucas-Lehmer test for Mersenne primes #

We define lucasLehmerResidue : Π p : ℕ, ZMod (2^p - 1), and prove lucasLehmerResidue p = 0 ↔ Prime (mersenne p).

We construct a norm_num extension to calculate this residue to certify primality of Mersenne primes using lucas_lehmer_sufficiency.

TODO #

History #

This development began as a student project by Ainsley Pahljina, and was then cleaned up for mathlib by Kim Morrison. The tactic for certified computation of Lucas-Lehmer residues was provided by Mario Carneiro. This tactic was ported by Thomas Murrills to Lean 4, and then it was converted to a norm_num extension and made to use kernel reductions by Kyle Miller.

def mersenne (p : ) :

The Mersenne numbers, 2^p - 1.

Equations
    Instances For
      @[simp]
      theorem mersenne_lt_mersenne {p q : } :
      theorem GCongr.mersenne_lt_mersenne {p q : } :
      p < qmersenne p < mersenne q

      Alias of the reverse direction of mersenne_lt_mersenne.

      @[simp]

      Alias of the reverse direction of mersenne_le_mersenne.

      @[simp]
      theorem mersenne_zero :
      @[simp]
      theorem mersenne_odd {p : } :
      @[simp]
      theorem mersenne_pos {p : } :
      0 < mersenne p 0 < p
      theorem mersenne_succ (n : ) :
      mersenne (n + 1) = 2 * mersenne n + 1
      theorem Nat.Prime.of_mersenne {p : } (h : Prime (mersenne p)) :

      If 2 ^ p - 1 is prime, then p is prime.

      Alias of the reverse direction of mersenne_pos.

      Extension for the positivity tactic: mersenne.

      Equations
        Instances For
          @[simp]
          theorem one_lt_mersenne {p : } :
          1 < mersenne p 1 < p
          @[simp]
          theorem succ_mersenne (k : ) :
          mersenne k + 1 = 2 ^ k
          theorem mersenne_mod_four {n : } (h : 2 n) :
          mersenne n % 4 = 3
          theorem mersenne_mod_three {n : } (odd : Odd n) (h : 3 n) :
          mersenne n % 3 = 1
          theorem mersenne_mod_eight {n : } (h : 3 n) :
          mersenne n % 8 = 7
          theorem legendreSym_mersenne_two {p : } [Fact (Nat.Prime (mersenne p))] (hp : 3 p) :

          If 2^p - 1 is prime then 2 is a square mod 2^p - 1.

          theorem legendreSym_mersenne_three {p : } [Fact (Nat.Prime (mersenne p))] (hp : 3 p) (odd : Odd p) :

          If 2^p - 1 is prime then 3 is not a square mod 2^p - 1.

          We now define three(!) different versions of the recurrence s (i+1) = (s i)^2 - 2.

          These versions take values either in , in ZMod (2^p - 1), or in but applying % (2^p - 1) at each step.

          They are each useful at different points in the proof, so we take a moment setting up the lemmas relating them.

          The recurrence s (i+1) = (s i)^2 - 2 in .

          Equations
            Instances For
              def LucasLehmer.sZMod (p a✝ : ) :
              ZMod (2 ^ p - 1)

              The recurrence s (i+1) = (s i)^2 - 2 in ZMod (2^p - 1).

              Equations
                Instances For
                  def LucasLehmer.sMod (p : ) :

                  The recurrence s (i+1) = ((s i)^2 - 2) % (2^p - 1) in .

                  Equations
                    Instances For
                      theorem LucasLehmer.mersenne_int_pos {p : } (hp : p 0) :
                      0 < 2 ^ p - 1
                      theorem LucasLehmer.mersenne_int_ne_zero (p : ) (hp : p 0) :
                      2 ^ p - 1 0
                      theorem LucasLehmer.sMod_nonneg (p : ) (hp : p 0) (i : ) :
                      0 sMod p i
                      theorem LucasLehmer.sMod_mod (p i : ) :
                      sMod p i % (2 ^ p - 1) = sMod p i
                      theorem LucasLehmer.sMod_lt (p : ) (hp : p 0) (i : ) :
                      sMod p i < 2 ^ p - 1
                      theorem LucasLehmer.sZMod_eq_s (p' i : ) :
                      sZMod (p' + 2) i = (s i)
                      theorem LucasLehmer.Int.natCast_pow_pred (b p : ) (w : 0 < b) :
                      ↑(b ^ p - 1) = b ^ p - 1
                      theorem LucasLehmer.Int.coe_nat_two_pow_pred (p : ) :
                      ↑(2 ^ p - 1) = 2 ^ p - 1
                      theorem LucasLehmer.sZMod_eq_sMod (p i : ) :
                      sZMod p i = (sMod p i)

                      The Lucas-Lehmer residue is s p (p-2) in ZMod (2^p - 1).

                      Equations
                        Instances For

                          Lucas-Lehmer Test: a Mersenne number 2^p-1 is prime if and only if the Lucas-Lehmer residue s p (p-2) % (2^p - 1) is zero.

                          Equations
                            Instances For

                              q is defined as the minimum factor of mersenne p, bundled as an ℕ+.

                              Equations
                                Instances For

                                  We construct the ring X q as ℤ/qℤ + √3 ℤ/qℤ.

                                  Equations
                                    Instances For
                                      Equations
                                        theorem LucasLehmer.X.ext {q : } {x y : X q} (h₁ : x.1 = y.1) (h₂ : x.2 = y.2) :
                                        x = y
                                        theorem LucasLehmer.X.ext_iff {q : } {x y : X q} :
                                        x = y x.1 = y.1 x.2 = y.2
                                        @[simp]
                                        theorem LucasLehmer.X.zero_fst {q : } :
                                        0.1 = 0
                                        @[simp]
                                        theorem LucasLehmer.X.zero_snd {q : } :
                                        0.2 = 0
                                        @[simp]
                                        theorem LucasLehmer.X.add_fst {q : } (x y : X q) :
                                        (x + y).1 = x.1 + y.1
                                        @[simp]
                                        theorem LucasLehmer.X.add_snd {q : } (x y : X q) :
                                        (x + y).2 = x.2 + y.2
                                        @[simp]
                                        theorem LucasLehmer.X.neg_fst {q : } (x : X q) :
                                        (-x).1 = -x.1
                                        @[simp]
                                        theorem LucasLehmer.X.neg_snd {q : } (x : X q) :
                                        (-x).2 = -x.2
                                        instance LucasLehmer.X.instMul {q : } :
                                        Mul (X q)
                                        Equations
                                          @[simp]
                                          theorem LucasLehmer.X.mul_fst {q : } (x y : X q) :
                                          (x * y).1 = x.1 * y.1 + 3 * x.2 * y.2
                                          @[simp]
                                          theorem LucasLehmer.X.mul_snd {q : } (x y : X q) :
                                          (x * y).2 = x.1 * y.2 + x.2 * y.1
                                          instance LucasLehmer.X.instOne {q : } :
                                          One (X q)
                                          Equations
                                            @[simp]
                                            theorem LucasLehmer.X.one_fst {q : } :
                                            1.1 = 1
                                            @[simp]
                                            theorem LucasLehmer.X.one_snd {q : } :
                                            1.2 = 0
                                            instance LucasLehmer.X.instMonoid {q : } :
                                            Monoid (X q)
                                            Equations
                                              Equations
                                                @[simp]
                                                theorem LucasLehmer.X.fst_natCast {q : } (n : ) :
                                                (↑n).1 = n
                                                @[simp]
                                                theorem LucasLehmer.X.snd_natCast {q : } (n : ) :
                                                (↑n).2 = 0
                                                @[simp]
                                                @[simp]
                                                theorem LucasLehmer.X.ofNat_snd {q : } (n : ) [n.AtLeastTwo] :
                                                (OfNat.ofNat n).2 = 0
                                                theorem LucasLehmer.X.left_distrib {q : } (x y z : X q) :
                                                x * (y + z) = x * y + x * z
                                                theorem LucasLehmer.X.right_distrib {q : } (x y z : X q) :
                                                (x + y) * z = x * z + y * z
                                                instance LucasLehmer.X.instRing {q : } :
                                                Ring (X q)
                                                Equations
                                                  Equations
                                                    @[simp]
                                                    theorem LucasLehmer.X.fst_intCast {q : } (n : ) :
                                                    (↑n).1 = n
                                                    @[simp]
                                                    theorem LucasLehmer.X.snd_intCast {q : } (n : ) :
                                                    (↑n).2 = 0
                                                    theorem LucasLehmer.X.coe_mul {q : } (n m : ) :
                                                    ↑(n * m) = n * m
                                                    theorem LucasLehmer.X.coe_natCast {q : } (n : ) :
                                                    n = n
                                                    def LucasLehmer.X.ω {q : } :
                                                    X q

                                                    We define ω = 2 + √3.

                                                    Equations
                                                      Instances For
                                                        def LucasLehmer.X.ωb {q : } :
                                                        X q

                                                        We define ωb = 2 - √3, which is the inverse of ω.

                                                        Equations
                                                          Instances For
                                                            theorem LucasLehmer.X.closed_form {q : } (i : ) :
                                                            (s i) = ω ^ 2 ^ i + ωb ^ 2 ^ i

                                                            A closed form for the recurrence relation.

                                                            def LucasLehmer.X.α {q : } :
                                                            X q

                                                            We define α = √3.

                                                            Equations
                                                              Instances For
                                                                @[simp]
                                                                theorem LucasLehmer.X.α_sq {q : } :
                                                                α ^ 2 = 3
                                                                @[simp]
                                                                theorem LucasLehmer.X.one_add_α_sq {q : } :
                                                                (1 + α) ^ 2 = 2 * ω
                                                                theorem LucasLehmer.X.α_pow {q : } (i : ) :
                                                                α ^ (2 * i + 1) = 3 ^ i * α

                                                                We show that X q has characteristic q, so that we can apply the binomial theorem.

                                                                instance LucasLehmer.X.instCoeZMod {q : } :
                                                                Coe (ZMod q) (X q)
                                                                Equations
                                                                  theorem LucasLehmer.X.one_add_α_pow_q {q : } [Fact (Nat.Prime q)] (odd : Odd q) (leg3 : legendreSym q 3 = -1) :
                                                                  (1 + α) ^ q = 1 - α

                                                                  If 3 is not a square mod q then (1 + α) ^ q = 1 - α

                                                                  theorem LucasLehmer.X.one_add_α_pow_q_succ {q : } [Fact (Nat.Prime q)] (odd : Odd q) (leg3 : legendreSym q 3 = -1) :
                                                                  (1 + α) ^ (q + 1) = -2

                                                                  If 3 is not a square then (1 + α) ^ (q + 1) = -2.

                                                                  theorem LucasLehmer.X.two_mul_ω_pow {q : } [Fact (Nat.Prime q)] (odd : Odd q) (leg3 : legendreSym q 3 = -1) :
                                                                  (2 * ω) ^ ((q + 1) / 2) = -2

                                                                  If 3 is not a square then (2 * ω) ^ ((q + 1) / 2) = -2.

                                                                  theorem LucasLehmer.X.pow_ω {q : } [Fact (Nat.Prime q)] (odd : Odd q) (leg3 : legendreSym q 3 = -1) (leg2 : legendreSym q 2 = 1) :
                                                                  ω ^ ((q + 1) / 2) = -1

                                                                  If 3 is not a square and 2 is square then $\omega^{(q+1)/2}=-1$.

                                                                  theorem LucasLehmer.X.ω_pow_trace {q : } [Fact (Nat.Prime q)] (odd : Odd q) (leg3 : legendreSym q 3 = -1) (leg2 : legendreSym q 2 = 1) (hq4 : 4 q + 1) :
                                                                  ω ^ ((q + 1) / 4) + ωb ^ ((q + 1) / 4) = 0

                                                                  The final evaluation needed to establish the Lucas-Lehmer necessity.

                                                                  instance LucasLehmer.X.instFintype {q : } [NeZero q] :
                                                                  Equations
                                                                    theorem LucasLehmer.X.card_eq {q : } [NeZero q] :
                                                                    Fintype.card (X q) = q ^ 2

                                                                    The cardinality of X is q^2.

                                                                    theorem LucasLehmer.X.card_units_lt {q : } [NeZero q] (w : 1 < q) :

                                                                    There are strictly fewer than q^2 units, since 0 is not a unit.

                                                                    Here and below, we introduce p' = p - 2, in order to avoid using subtraction in .

                                                                    theorem LucasLehmer.two_lt_q (p' : ) :
                                                                    2 < q (p' + 2)

                                                                    If 1 < p, then q p, the smallest prime factor of mersenne p, is more than 2.

                                                                    theorem LucasLehmer.ω_pow_formula (p' : ) (h : lucasLehmerResidue (p' + 2) = 0) :
                                                                    ∃ (k : ), X.ω ^ 2 ^ (p' + 1) = k * (mersenne (p' + 2)) * X.ω ^ 2 ^ p' - 1
                                                                    theorem LucasLehmer.mersenne_coe_X (p : ) :
                                                                    (mersenne p) = 0

                                                                    q is the minimum factor of mersenne p, so M p = 0 in X q.

                                                                    theorem LucasLehmer.ω_pow_eq_neg_one (p' : ) (h : lucasLehmerResidue (p' + 2) = 0) :
                                                                    X.ω ^ 2 ^ (p' + 1) = -1
                                                                    theorem LucasLehmer.ω_pow_eq_one (p' : ) (h : lucasLehmerResidue (p' + 2) = 0) :
                                                                    X.ω ^ 2 ^ (p' + 2) = 1
                                                                    def LucasLehmer.ωUnit (p : ) :
                                                                    (X (q p))ˣ

                                                                    ω as an element of the group of units.

                                                                    Equations
                                                                      Instances For
                                                                        @[simp]
                                                                        theorem LucasLehmer.ωUnit_coe (p : ) :
                                                                        (ωUnit p) = X.ω
                                                                        theorem LucasLehmer.order_ω (p' : ) (h : lucasLehmerResidue (p' + 2) = 0) :
                                                                        orderOf (ωUnit (p' + 2)) = 2 ^ (p' + 2)

                                                                        The order of ω in the unit group is exactly 2^p.

                                                                        theorem LucasLehmer.order_ineq (p' : ) (h : lucasLehmerResidue (p' + 2) = 0) :
                                                                        2 ^ (p' + 2) < (q (p' + 2)) ^ 2
                                                                        theorem lucas_lehmer_necessity (p : ) (w : 3 p) (hp : Nat.Prime (mersenne p)) :

                                                                        If 2^p-1 is prime then the Lucas-Lehmer test holds, `s(p-2) % (2^p-1) = 0.

                                                                        norm_num extension #

                                                                        Next we define a norm_num extension that calculates LucasLehmerTest p for 1 < p. It makes use of a version of sMod that is specifically written to be reducible by the Lean 4 kernel, which has the capability of efficiently reducing natural number expressions. With this reduction in hand, it's a simple matter of applying the lemma LucasLehmer.residue_eq_zero_iff_sMod_eq_zero.

                                                                        See [Archive/Examples/MersennePrimes.lean] for certifications of all Mersenne primes up through mersenne 4423.

                                                                        Version of sMod that is -valued. One should have q = 2 ^ p - 1. This can be reduced by the kernel.

                                                                        Equations
                                                                          Instances For
                                                                            theorem LucasLehmer.norm_num_ext.sModNat_eq_sMod (p k : ) (hp : 2 p) :
                                                                            (sModNat (2 ^ p - 1) k) = sMod p k

                                                                            Tail-recursive version of sModNat.

                                                                            Equations
                                                                              Instances For

                                                                                Helper function for sMod''.

                                                                                Equations
                                                                                  Instances For

                                                                                    Generalization of sModNat with arbitrary base case, useful for proving sModNatTR and sModNat agree.

                                                                                    Equations
                                                                                      Instances For
                                                                                        theorem LucasLehmer.norm_num_ext.testTrueHelper (p : ) (hp : Nat.blt 1 p = true) (h : sModNatTR (2 ^ p - 1) (p - 2) = 0) :

                                                                                        Calculate LucasLehmer.LucasLehmerTest p for 2 ≤ p by using kernel reduction for the sMod' function.

                                                                                        Equations
                                                                                          Instances For

                                                                                            This implementation works successfully to prove (2^4423 - 1).Prime, and all the Mersenne primes up to this point appear in [Archive/Examples/MersennePrimes.lean]. These can be calculated nearly instantly, and (2^9689 - 1).Prime only fails due to deep recursion.

                                                                                            (Note by kmill: the following notes were for the Lean 3 version. They seem like they could still be useful, so I'm leaving them here.)

                                                                                            There's still low hanging fruit available to do faster computations based on the formula

                                                                                            n ≡ (n % 2^p) + (n / 2^p) [MOD 2^p - 1]
                                                                                            

                                                                                            and the fact that % 2^p and / 2^p can be very efficient on the binary representation. Someone should do this, too!

                                                                                            theorem modEq_mersenne (n k : ) :
                                                                                            k k / 2 ^ n + k % 2 ^ n [MOD 2 ^ n - 1]