Documentation

Mathlib.NumberTheory.ArithmeticFunction.Misc

Miscellaneous arithmetic Functions #

This file defines some simple examples of arithmetic functions (functions ℕ → R vanishing at 0, considered as a ring under Dirichlet convolution). Note that the Von Mangoldt and Möbius functions are in separate files.

Main Definitions #

Notation #

The arithmetic functions σ, ω and Ω have Greek letter names. This notation is scoped to the separate locales ArithmeticFunction.sigma for σ, ArithmeticFunction.omega for ω and ArithmeticFunction.Omega for Ω, to allow for selective access.

Tags #

arithmetic functions, dirichlet convolution, divisors

The map $n \mapsto \prod_{p \mid n} f(p)$ as an arithmetic function

Equations
Instances For

    ∏ᵖ p ∣ n, f p is custom notation for prodPrimeFactors f n

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem ArithmeticFunction.prodPrimeFactors_apply {R : Type u_1} [CommMonoidWithZero R] {f : R} {n : } (hn : n 0) :
      (prodPrimeFactors fun (p : ) => f p) n = pn.primeFactors, f p

      The identity on as an ArithmeticFunction.

      Equations
      Instances For
        @[simp]
        theorem ArithmeticFunction.id_apply {x : } :
        id x = x

        pow k n = n ^ k, except pow 0 0 = 0.

        Equations
        Instances For
          @[simp]
          theorem ArithmeticFunction.pow_apply {k n : } :
          (pow k) n = if k = 0 n = 0 then 0 else n ^ k

          σ k n is the sum of the kth powers of the divisors of n

          Equations
          Instances For

            σ k n is the sum of the kth powers of the divisors of n

            Equations
            Instances For
              theorem ArithmeticFunction.sigma_apply {k n : } :
              (sigma k) n = dn.divisors, d ^ k
              @[simp]
              theorem ArithmeticFunction.sigma_eq_zero {k n : } :
              (sigma k) n = 0 n = 0
              @[simp]
              theorem ArithmeticFunction.sigma_pos_iff {k n : } :
              0 < (sigma k) n 0 < n
              theorem ArithmeticFunction.sigma_apply_prime_pow {k p i : } (hp : Nat.Prime p) :
              (sigma k) (p ^ i) = jFinset.range (i + 1), p ^ (j * k)
              theorem ArithmeticFunction.sigma_one_apply (n : ) :
              (sigma 1) n = dn.divisors, d
              theorem ArithmeticFunction.sigma_one_apply_prime_pow {p i : } (hp : Nat.Prime p) :
              (sigma 1) (p ^ i) = kFinset.range (i + 1), p ^ k
              theorem ArithmeticFunction.sigma_eq_sum_div (k n : ) :
              (sigma k) n = dn.divisors, (n / d) ^ k
              @[simp]
              theorem ArithmeticFunction.sigma_one (k : ) :
              (sigma k) 1 = 1
              theorem ArithmeticFunction.sigma_pos (k n : ) (hn0 : n 0) :
              0 < (sigma k) n
              theorem ArithmeticFunction.sigma_mono (k k' n : ) (hk : k k') :
              (sigma k) n (sigma k') n
              theorem Nat.card_divisors {n : } (hn : n 0) :
              n.divisors.card = xn.primeFactors, (n.factorization x + 1)
              @[simp]
              @[simp]
              theorem ArithmeticFunction.sigma_eq_one_iff (k n : ) :
              (sigma k) n = 1 n = 1
              theorem Nat.sum_divisors {n : } (hn : n 0) :
              dn.divisors, d = pn.primeFactors, kFinset.range (n.factorization p + 1), p ^ k

              Ω n is the number of prime factors of n.

              Equations
              Instances For

                Ω n is the number of prime factors of n.

                Equations
                Instances For
                  theorem ArithmeticFunction.cardFactors_mul {m n : } (m0 : m 0) (n0 : n 0) :

                  ω n is the number of distinct prime factors of n.

                  Equations
                  Instances For

                    ω n is the number of distinct prime factors of n.

                    Equations
                    Instances For
                      theorem ArithmeticFunction.cardDistinctFactors_prod {ι : Type u_2} {s : Finset ι} {f : ι} (h : (↑s).Pairwise (Function.onFun Nat.Coprime f)) :
                      cardDistinctFactors (∏ is, f i) = is, cardDistinctFactors (f i)
                      theorem ArithmeticFunction.sum_Ioc_zeta (N : ) :
                      nFinset.Ioc 0 N, zeta n = N
                      theorem ArithmeticFunction.sum_Ioc_mul_eq_sum_prod_filter {R : Type u_2} [Semiring R] (f g : ArithmeticFunction R) (N : ) :
                      nFinset.Ioc 0 N, (f * g) n = xFinset.Ioc 0 N ×ˢ Finset.Ioc 0 N with x.1 * x.2 N, f x.1 * g x.2
                      theorem ArithmeticFunction.sum_Ioc_mul_eq_sum_sum {R : Type u_2} [Semiring R] (f g : ArithmeticFunction R) (N : ) :
                      nFinset.Ioc 0 N, (f * g) n = nFinset.Ioc 0 N, f n * mFinset.Ioc 0 (N / n), g m
                      theorem ArithmeticFunction.sum_Ioc_mul_zeta_eq_sum {R : Type u_2} [Semiring R] (f : ArithmeticFunction R) (N : ) :
                      nFinset.Ioc 0 N, (f * zeta) n = nFinset.Ioc 0 N, f n * ↑(N / n)
                      theorem ArithmeticFunction.sum_Ioc_sigma0_eq_sum_div (N : ) :
                      nFinset.Ioc 0 N, (sigma 0) n = nFinset.Ioc 0 N, N / n

                      An O(N) formula for the sum of the number of divisors function.

                      theorem Nat.Coprime.sum_divisors_mul {m n : } (hmn : m.Coprime n) :
                      d(m * n).divisors, d = (∑ dm.divisors, d) * dn.divisors, d

                      Extension for ArithmeticFunction.sigma.

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