Documentation

Mathlib.Tactic.NormNum.LegendreSymbol

A norm_num extension for Jacobi and Legendre symbols #

We extend the norm_num tactic so that it can be used to provably compute the value of the Jacobi symbol J(a | b) or the Legendre symbol legendreSym p a when the arguments are numerals.

Implementation notes #

We use the Law of Quadratic Reciprocity for the Jacobi symbol to compute the value of J(a | b) efficiently, roughly comparable in effort with the euclidean algorithm for the computation of the gcd of a and b. More precisely, the computation is done in the following steps.

We provide customized versions of these results for the various reduction steps, where we encode the residue classes mod 2, mod 4, or mod 8 by using hypotheses like a % n = b. In this way, the only divisions we have to compute and prove are the ones occurring in the use of QR above.

The Jacobi symbol restricted to natural numbers in both arguments.

Equations
Instances For

    API Lemmas #

    We repeat part of the API for jacobiSym with NormNum.jacobiSymNat and without implicit arguments, in a form that is suitable for constructing proofs in norm_num.

    Base cases: b = 0, b = 1, a = 0, a = 1.

    Turn a Legendre symbol into a Jacobi symbol.

    theorem Mathlib.Meta.NormNum.JacobiSym.mod_left (a : ) (b : ) (ab' : ) (ab : ) (r : ) (b' : ) (hb' : b = b') (hab : a % b' = ab) (h : ab' = ab) (hr : Mathlib.Meta.NormNum.jacobiSymNat ab' b = r) :
    jacobiSym a b = r

    The value depends only on the residue class of a mod b.

    theorem Mathlib.Meta.NormNum.jacobiSymNat.even_even (a : ) (b : ) (hb₀ : Nat.beq (b / 2) 0 = false) (ha : a % 2 = 0) (hb₁ : b % 2 = 0) :

    The symbol vanishes when both entries are even (and b / 2 ≠ 0).

    theorem Mathlib.Meta.NormNum.jacobiSymNat.odd_even (a : ) (b : ) (c : ) (r : ) (ha : a % 2 = 1) (hb : b % 2 = 0) (hc : b / 2 = c) (hr : Mathlib.Meta.NormNum.jacobiSymNat a c = r) :

    When a is odd and b is even, we can replace b by b / 2.

    theorem Mathlib.Meta.NormNum.jacobiSymNat.double_even (a : ) (b : ) (c : ) (r : ) (ha : a % 4 = 0) (hb : b % 2 = 1) (hc : a / 4 = c) (hr : Mathlib.Meta.NormNum.jacobiSymNat c b = r) :

    If a is divisible by 4 and b is odd, then we can remove the factor 4 from a.

    theorem Mathlib.Meta.NormNum.jacobiSymNat.even_odd₁ (a : ) (b : ) (c : ) (r : ) (ha : a % 2 = 0) (hb : b % 8 = 1) (hc : a / 2 = c) (hr : Mathlib.Meta.NormNum.jacobiSymNat c b = r) :

    If a is even and b is odd, then we can remove a factor 2 from a, but we may have to change the sign, depending on b % 8. We give one version for each of the four odd residue classes mod 8.

    theorem Mathlib.Meta.NormNum.jacobiSymNat.even_odd₇ (a : ) (b : ) (c : ) (r : ) (ha : a % 2 = 0) (hb : b % 8 = 7) (hc : a / 2 = c) (hr : Mathlib.Meta.NormNum.jacobiSymNat c b = r) :
    theorem Mathlib.Meta.NormNum.jacobiSymNat.even_odd₃ (a : ) (b : ) (c : ) (r : ) (ha : a % 2 = 0) (hb : b % 8 = 3) (hc : a / 2 = c) (hr : Mathlib.Meta.NormNum.jacobiSymNat c b = r) :
    theorem Mathlib.Meta.NormNum.jacobiSymNat.even_odd₅ (a : ) (b : ) (c : ) (r : ) (ha : a % 2 = 0) (hb : b % 8 = 5) (hc : a / 2 = c) (hr : Mathlib.Meta.NormNum.jacobiSymNat c b = r) :
    theorem Mathlib.Meta.NormNum.jacobiSymNat.qr₁ (a : ) (b : ) (r : ) (ha : a % 4 = 1) (hb : b % 2 = 1) (hr : Mathlib.Meta.NormNum.jacobiSymNat b a = r) :

    Use quadratic reciproity to reduce to smaller b.

    theorem Mathlib.Meta.NormNum.jacobiSymNat.qr₁_mod (a : ) (b : ) (ab : ) (r : ) (ha : a % 4 = 1) (hb : b % 2 = 1) (hab : b % a = ab) (hr : Mathlib.Meta.NormNum.jacobiSymNat ab a = r) :
    theorem Mathlib.Meta.NormNum.jacobiSymNat.qr₁'_mod (a : ) (b : ) (ab : ) (r : ) (ha : a % 2 = 1) (hb : b % 4 = 1) (hab : b % a = ab) (hr : Mathlib.Meta.NormNum.jacobiSymNat ab a = r) :
    theorem Mathlib.Meta.NormNum.jacobiSymNat.qr₃_mod (a : ) (b : ) (ab : ) (r : ) (ha : a % 4 = 3) (hb : b % 4 = 3) (hab : b % a = ab) (hr : Mathlib.Meta.NormNum.jacobiSymNat ab a = r) :

    Certified evaluation of the Jacobi symbol #

    The following functions recursively evaluate a Jacobi symbol and construct the corresponding proof term.

    partial def Mathlib.Meta.NormNum.proveJacobiSymOdd (ea : Q()) (eb : Q()) :
    (er : Q()) × Q(Mathlib.Meta.NormNum.jacobiSymNat «$ea» «$eb» = «$er»)

    This evaluates r := jacobiSymNat a b recursively using quadratic reciprocity and produces a proof term for the equality, assuming that a < b and b is odd.

    def Mathlib.Meta.NormNum.proveJacobiSymNat (ea : Q()) (eb : Q()) :
    (er : Q()) × Q(Mathlib.Meta.NormNum.jacobiSymNat «$ea» «$eb» = «$er»)

    This evaluates r := jacobiSymNat a b and produces a proof term for the equality by removing powers of 2 from b and then calling proveJacobiSymOdd.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      def Mathlib.Meta.NormNum.proveJacobiSym (ea : Q()) (eb : Q()) :
      (er : Q()) × Q(jacobiSym «$ea» «$eb» = «$er»)

      This evaluates r := jacobiSym a b and produces a proof term for the equality. This is done by reducing to r := jacobiSymNat (a % b) b.

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

        The norm_num plug-in #

        This is the norm_num plug-in that evaluates Jacobi symbols.

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

          This is the norm_num plug-in that evaluates Jacobi symbols on natural numbers.

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

            This is the norm_num plug-in that evaluates Legendre symbols.

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