Documentation

Mathlib.NumberTheory.ArithmeticFunction.Defs

Arithmetic Functions and Dirichlet Convolution #

This file defines arithmetic functions, which are functions from to a specified type that map 0 to 0. In the literature, they are often instead defined as functions from ℕ+. These arithmetic functions are endowed with a multiplication, given by Dirichlet convolution, and pointwise addition, to form the Dirichlet ring.

Main Definitions #

Further examples of arithmetic functions, such as the Möbius function μ, are available in other files in the Mathlib.NumberTheory.ArithmeticFunction directory.

Tags #

arithmetic functions, dirichlet convolution, divisors

def ArithmeticFunction (R : Type u_1) [Zero R] :
Type u_1

An arithmetic function is a function from that maps 0 to 0. In the literature, they are often instead defined as functions from ℕ+. Multiplication on ArithmeticFunctions is by Dirichlet convolution.

Equations
Instances For
    @[simp]
    theorem ArithmeticFunction.toFun_eq {R : Type u_1} [Zero R] (f : ArithmeticFunction R) :
    f.toFun = f
    @[simp]
    theorem ArithmeticFunction.coe_mk {R : Type u_1} [Zero R] (f : R) (hf : f 0 = 0) :
    { toFun := f, map_zero' := hf } = f
    @[simp]
    theorem ArithmeticFunction.map_zero {R : Type u_1} [Zero R] {f : ArithmeticFunction R} :
    f 0 = 0
    theorem ArithmeticFunction.coe_inj {R : Type u_1} [Zero R] {f g : ArithmeticFunction R} :
    f = g f = g
    @[simp]
    theorem ArithmeticFunction.zero_apply {R : Type u_1} [Zero R] {x : } :
    0 x = 0
    theorem ArithmeticFunction.ext {R : Type u_1} [Zero R] f g : ArithmeticFunction R (h : ∀ (x : ), f x = g x) :
    f = g
    theorem ArithmeticFunction.ext_iff {R : Type u_1} [Zero R] {f g : ArithmeticFunction R} :
    f = g ∀ (x : ), f x = g x
    instance ArithmeticFunction.one {R : Type u_1} [Zero R] [One R] :
    Equations
    theorem ArithmeticFunction.one_apply {R : Type u_1} [Zero R] [One R] {x : } :
    1 x = if x = 1 then 1 else 0
    @[simp]
    theorem ArithmeticFunction.one_one {R : Type u_1} [Zero R] [One R] :
    1 1 = 1
    @[simp]
    theorem ArithmeticFunction.one_apply_ne {R : Type u_1} [Zero R] [One R] {x : } (h : x 1) :
    1 x = 0

    Coerce an arithmetic function with values in to one with values in R. We cannot inline this in natCoe because it gets unfolded too much.

    Equations
    • f = { toFun := fun (n : ) => (f n), map_zero' := }
    Instances For
      @[simp]
      theorem ArithmeticFunction.natCoe_apply {R : Type u_1} [AddMonoidWithOne R] {f : ArithmeticFunction } {x : } :
      f x = (f x)

      Coerce an arithmetic function with values in to one with values in R. We cannot inline this in intCoe because it gets unfolded too much.

      Equations
      • f = { toFun := fun (n : ) => (f n), map_zero' := }
      Instances For
        @[simp]
        theorem ArithmeticFunction.intCoe_apply {R : Type u_1} [AddGroupWithOne R] {f : ArithmeticFunction } {x : } :
        f x = (f x)
        @[simp]
        theorem ArithmeticFunction.coe_coe {R : Type u_1} [AddGroupWithOne R] {f : ArithmeticFunction } :
        f = f
        @[simp]
        @[simp]
        theorem ArithmeticFunction.intCoe_one {R : Type u_1} [AddGroupWithOne R] :
        1 = 1
        Equations
        @[simp]
        theorem ArithmeticFunction.add_apply {R : Type u_1} [AddMonoid R] {f g : ArithmeticFunction R} {n : } :
        (f + g) n = f n + g n
        Equations
        • One or more equations did not get rendered due to their size.
        Equations
        • One or more equations did not get rendered due to their size.
        Equations
        Equations
        • One or more equations did not get rendered due to their size.
        Equations

        The Dirichlet convolution of two arithmetic functions f and g is another arithmetic function such that (f * g) n is the sum of f x * g y over all (x,y) such that x * y = n.

        Equations
        • One or more equations did not get rendered due to their size.
        @[simp]
        theorem ArithmeticFunction.smul_apply {R : Type u_1} {M : Type u_2} [Zero R] [AddCommMonoid M] [SMul R M] {f : ArithmeticFunction R} {g : ArithmeticFunction M} {n : } :
        (f g) n = xn.divisorsAntidiagonal, f x.1 g x.2

        The Dirichlet convolution of two arithmetic functions f and g is another arithmetic function such that (f * g) n is the sum of f x * g y over all (x,y) such that x * y = n.

        Equations
        @[simp]
        theorem ArithmeticFunction.mul_apply {R : Type u_1} [Semiring R] {f g : ArithmeticFunction R} {n : } :
        (f * g) n = xn.divisorsAntidiagonal, f x.1 * g x.2
        theorem ArithmeticFunction.mul_apply_one {R : Type u_1} [Semiring R] {f g : ArithmeticFunction R} :
        (f * g) 1 = f 1 * g 1
        @[simp]
        theorem ArithmeticFunction.natCoe_mul {R : Type u_1} [Semiring R] {f g : ArithmeticFunction } :
        ↑(f * g) = f * g
        @[simp]
        theorem ArithmeticFunction.intCoe_mul {R : Type u_1} [Ring R] {f g : ArithmeticFunction } :
        ↑(f * g) = f * g
        theorem ArithmeticFunction.mul_smul' {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] (f g : ArithmeticFunction R) (h : ArithmeticFunction M) :
        (f * g) h = f g h
        theorem ArithmeticFunction.one_smul' {R : Type u_1} {M : Type u_2} [Semiring R] [AddCommMonoid M] [Module R M] (b : ArithmeticFunction M) :
        1 b = b
        Equations
        • One or more equations did not get rendered due to their size.
        Equations
        • One or more equations did not get rendered due to their size.
        Equations
        • One or more equations did not get rendered due to their size.
        Equations

        Multiplicative functions

        Equations
        Instances For
          @[simp]
          theorem ArithmeticFunction.IsMultiplicative.map_mul_of_coprime {R : Type u_1} [MonoidWithZero R] {f : ArithmeticFunction R} (hf : f.IsMultiplicative) {m n : } (h : m.gcd n = 1) :
          f (m * n) = f m * f n
          theorem ArithmeticFunction.IsMultiplicative.map_prod {R : Type u_1} {ι : Type u_2} [CommMonoidWithZero R] (g : ι) {f : ArithmeticFunction R} (hf : f.IsMultiplicative) (s : Finset ι) (hs : (↑s).Pairwise (Function.onFun Nat.Coprime g)) :
          f (∏ is, g i) = is, f (g i)
          theorem ArithmeticFunction.IsMultiplicative.map_prod_of_prime {R : Type u_1} [CommMonoidWithZero R] {f : ArithmeticFunction R} (h_mult : f.IsMultiplicative) (t : Finset ) (ht : pt, Nat.Prime p) :
          f (∏ at, a) = at, f a
          theorem ArithmeticFunction.IsMultiplicative.map_prod_of_subset_primeFactors {R : Type u_1} [CommMonoidWithZero R] {f : ArithmeticFunction R} (h_mult : f.IsMultiplicative) (l : ) (t : Finset ) (ht : t l.primeFactors) :
          f (∏ at, a) = at, f a
          theorem ArithmeticFunction.IsMultiplicative.map_div_of_coprime {R : Type u_1} [GroupWithZero R] {f : ArithmeticFunction R} (hf : f.IsMultiplicative) {l d : } (hdl : d l) (hl : (l / d).Coprime d) (hd : f d 0) :
          f (l / d) = f l / f d

          For any multiplicative function f and any n > 0, we can evaluate f n by evaluating f at p ^ k over the factorization of n

          theorem ArithmeticFunction.IsMultiplicative.iff_ne_zero {R : Type u_1} [MonoidWithZero R] {f : ArithmeticFunction R} :
          f.IsMultiplicative f 1 = 1 ∀ {m n : }, m 0n 0m.Coprime nf (m * n) = f m * f n

          A recapitulation of the definition of multiplicative that is simpler for proofs

          Two multiplicative functions f and g are equal if and only if they agree on prime powers

          theorem ArithmeticFunction.IsMultiplicative.map_gcd {R : Type u_1} [CommGroupWithZero R] {f : ArithmeticFunction R} (hf : f.IsMultiplicative) {x y : } (hf_lcm : f (x.lcm y) 0) :
          f (x.gcd y) = f x * f y / f (x.lcm y)
          theorem ArithmeticFunction.IsMultiplicative.map_lcm {R : Type u_1} [CommGroupWithZero R] {f : ArithmeticFunction R} (hf : f.IsMultiplicative) {x y : } (hf_gcd : f (x.gcd y) 0) :
          f (x.lcm y) = f x * f y / f (x.gcd y)
          theorem ArithmeticFunction.IsMultiplicative.eq_zero_of_squarefree_of_dvd_eq_zero {R : Type u_1} [MonoidWithZero R] {f : ArithmeticFunction R} (hf : f.IsMultiplicative) {m n : } (hn : Squarefree n) (hmn : m n) (h_zero : f m = 0) :
          f n = 0