Super-Polynomial Function Decay #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file defines a predicate asymptotics.superpolynomial_decay f for a function satisfying
one of following equivalent definitions (The definition is in terms of the first condition):
x ^ n * ftends to𝓝 0for all (or sufficiently large) naturalsn|x ^ n * f|tends to𝓝 0for all naturalsn(superpolynomial_decay_iff_abs_tendsto_zero)|x ^ n * f|is bounded for all naturalsn(superpolynomial_decay_iff_abs_is_bounded_under)fiso(x ^ c)for all integersc(superpolynomial_decay_iff_is_o)fisO(x ^ c)for all integersc(superpolynomial_decay_iff_is_O)
These conditions are all equivalent to conditions in terms of polynomials, replacing x ^ c with
p(x) or p(x)⁻¹ as appropriate, since asymptotically p(x) behaves like X ^ p.nat_degree.
These further equivalences are not proven in mathlib but would be good future projects.
The definition of superpolynomial decay for f : α → β is relative to a parameter k : α → β.
Super-polynomial decay then means f x decays faster than (k x) ^ c for all integers c.
Equivalently f x decays faster than p.eval (k x) for all polynomials p : β[X].
The definition is also relative to a filter l : filter α where the decay rate is compared.
When the map k is given by n ↦ ↑n : ℕ → ℝ this defines negligible functions:
https://en.wikipedia.org/wiki/Negligible_function
When the map k is given by (r₁,...,rₙ) ↦ r₁*...*rₙ : ℝⁿ → ℝ this is equivalent
to the definition of rapidly decreasing functions given here:
https://ncatlab.org/nlab/show/rapidly+decreasing+function
Main Theorems #
superpolynomial_decay.polynomial_mulsays that iff(x)is negligible, then so isp(x) * f(x)for any polynomialp.superpolynomial_decay_iff_zpow_tendsto_zerogives an equivalence between definitions in terms of decaying faster thank(x) ^ nfor all naturalsnork(x) ^ cfor all integerc.
f has superpolynomial decay in parameter k along filter l if
k ^ n * f tends to zero at l for all naturals n
Equations
- asymptotics.superpolynomial_decay l k f = ∀ (n : ℕ), filter.tendsto (λ (a : α), k a ^ n * f a) l (nhds 0)