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 #
ArithmeticFunction Rconsists of functionsf : ℕ → Rsuch thatf 0 = 0.- An arithmetic function
fIsMultiplicativewhenx.Coprime y → f (x * y) = f x * f y. - Multiplication and power instances on
ArithmeticFunction R, are defined using Dirichlet convolution.
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
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
- ArithmeticFunction R = ZeroHom ℕ R
Instances For
Equations
- ArithmeticFunction.zero R = inferInstanceAs (Zero (ZeroHom ℕ R))
Equations
Equations
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.
Instances For
Equations
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.
Instances For
Equations
Equations
- ArithmeticFunction.add = { add := fun (f g : ArithmeticFunction R) => { toFun := fun (n : ℕ) => f n + g n, map_zero' := ⋯ } }
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
- ArithmeticFunction.instAddCommMonoid = { toAddMonoid := ArithmeticFunction.instAddMonoid, add_comm := ⋯ }
Equations
- ArithmeticFunction.instNeg = { neg := fun (f : ArithmeticFunction R) => { toFun := fun (n : ℕ) => -f n, map_zero' := ⋯ } }
Equations
- One or more equations did not get rendered due to their size.
Equations
- ArithmeticFunction.instAddCommGroup = { toAddGroup := have this := inferInstance; this, add_comm := ⋯ }
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.
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
- ArithmeticFunction.instMul = { mul := fun (x1 x2 : ArithmeticFunction R) => x1 • x2 }
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
- ArithmeticFunction.instCommSemiring = { toSemiring := ArithmeticFunction.instSemiring, mul_comm := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- ArithmeticFunction.instModule = { toSMul := ArithmeticFunction.instSMul, mul_smul := ⋯, one_smul := ⋯, smul_zero := ⋯, smul_add := ⋯, add_smul := ⋯, zero_smul := ⋯ }
Multiplicative functions
Instances For
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
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