Akra-Bazzi theorem: the sum transform #
We develop further required preliminaries for the theorem, up to the sum transform.
Main definitions and results #
AkraBazziRecurrence T g a b r
: the predicate stating thatT : ℕ → ℝ
satisfies an Akra-Bazzi recurrence with parametersg
,a
,b
andr
as above.GrowsPolynomially
: The growth condition thatg
must satisfy for the theorem to apply. It roughly states thatc₁ g(n) ≤ g(u) ≤ c₂ g(n)
, for u between b*n and n for any constantb ∈ (0,1)
.sumTransform
: The transformation which turns a functiong
inton^p * ∑ u ∈ Finset.Ico n₀ n, g u / u^(p+1)
.
References #
- Mohamad Akra and Louay Bazzi, On the solution of linear recurrence equations
- Tom Leighton, Notes on better master theorems for divide-and-conquer recurrences
- Manuel Eberl, Asymptotic reasoning in a proof assistant
Definition of Akra-Bazzi recurrences #
This section defines the predicate AkraBazziRecurrence T g a b r
which states that T
satisfies the recurrence
T(n) = ∑_{i=0}^{k-1} a_i T(r_i(n)) + g(n)
with appropriate conditions on the various parameters.
An Akra-Bazzi recurrence is a function that satisfies the recurrence
T n = (∑ i, a i * T (r i n)) + g n
.
- n₀ : ℕ
Point below which the recurrence is in the base case
n₀
is always> 0
The
a
's are nonzeroThe
b
's are nonzeroThe b's are less than 1
g
is nonnegative- g_grows_poly : GrowsPolynomially g
g
grows polynomially The actual recurrence
Base case:
T(n) > 0
whenevern < n₀
The
r
's always reducen
- dist_r_b (i : α) : (fun (n : ℕ) => ↑(r i n) - b i * ↑n) =o[Filter.atTop] fun (n : ℕ) => ↑n / Real.log ↑n ^ 2
The
r
's approximate theb
's
Instances For
Smallest b i
Equations
Instances For
Largest b i
Equations
Instances For
Smoothing function #
We define ε
as the "smoothing function" fun n => 1 / log n
, which will be used in the form of a
factor of 1 ± ε n
needed to make the induction step go through.
This is its own definition to make it easier to switch to a different smoothing function.
For example, choosing 1 / log n ^ δ
for a suitable choice of δ
leads to a slightly tighter
theorem at the price of a more complicated proof.
This part of the file then proves several properties of this function that will be needed later in the proof.
The "smoothing function" is defined as 1 / log n
. This is defined as an ℝ → ℝ
function
as opposed to ℕ → ℝ
since this is more convenient for the proof, where we need to e.g. take
derivatives.
Equations
Instances For
Akra-Bazzi exponent p
#
Every Akra-Bazzi recurrence has an associated exponent, denoted by p : ℝ
, such that
∑ a_i b_i^p = 1
. This section shows the existence and uniqueness of this exponent p
for any
R : AkraBazziRecurrence
, and defines R.asympBound
to be the asymptotic bound satisfied by R
,
namely n^p (1 + ∑_{u < n} g(u) / u^(p+1))
.
The function x ↦ ∑ a_i b_i^x is injective. This implies the uniqueness of p
.
The exponent p
associated with a particular Akra-Bazzi recurrence.
Equations
- AkraBazziRecurrence.p a b = Function.invFun (fun (p : ℝ) => ∑ i : α, a i * b i ^ p) 1
Instances For
The sum transform #
This section defines the "sum transform" of a function g
as
∑ u ∈ Finset.Ico n₀ n, g u / u^(p+1)
,
and uses it to define asympBound
as the bound satisfied by an Akra-Bazzi recurrence.
Several properties of the sum transform are then proven.
The transformation which turns a function g
into
n^p * ∑ u ∈ Finset.Ico n₀ n, g u / u^(p+1)
.
Equations
- AkraBazziRecurrence.sumTransform p g n₀ n = ↑n ^ p * ∑ u ∈ Finset.Ico n₀ n, g ↑u / ↑u ^ (p + 1)
Instances For
The asymptotic bound satisfied by an Akra-Bazzi recurrence, namely
n^p (1 + ∑_{u < n} g(u) / u^(p+1))
.
Equations
- AkraBazziRecurrence.asympBound g a b n = ↑n ^ AkraBazziRecurrence.p a b + AkraBazziRecurrence.sumTransform (AkraBazziRecurrence.p a b) g 0 n