liminfs and limsups of functions and filters #
Defines the liminf/limsup of a function taking values in a conditionally complete lattice, with respect to an arbitrary filter.
We define limsSup f (limsInf f) where f is a filter taking values in a conditionally complete
lattice. limsSup f is the smallest element a such that, eventually, u ≤ a (and vice versa for
limsInf f). To work with the Limsup along a function u use limsSup (map u f).
Usually, one defines the Limsup as inf (sup s) where the Inf is taken over all sets in the filter.
For instance, in ℕ along a function u, this is inf_n (sup_{k ≥ n} u k) (and the latter quantity
decreases with n, so this is in fact a limit.). There is however a difficulty: it is well possible
that u is not bounded on the whole space, only eventually (think of limsup (fun x ↦ 1/x) on ℝ.
Then there is no guarantee that the quantity above really decreases (the value of the sup
beforehand is not really well defined, as one can not use ∞), so that the Inf could be anything.
So one can not use this inf sup ... definition in conditionally complete lattices, and one has
to use a less tractable definition.
In conditionally complete lattices, the definition is only useful for filters which are eventually bounded above (otherwise, the Limsup would morally be +∞, which does not belong to the space) and which are frequently bounded below (otherwise, the Limsup would morally be -∞, which is not in the space either). We start with definitions of these concepts for arbitrary filters, before turning to the definitions of Limsup and Liminf.
In complete lattices, however, it coincides with the Inf Sup definition.
f.IsBounded (≺): the filter f is eventually bounded w.r.t. the relation ≺, i.e.
eventually, it is bounded by some uniform bound.
r will be usually instantiated with ≤ or ≥.
Equations
- Filter.IsBounded r f = ∃ (b : α), ∀ᶠ (x : α) in f, r x b
Instances For
f.IsBoundedUnder (≺) u: the image of the filter f under u is eventually bounded w.r.t.
the relation ≺, i.e. eventually, it is bounded by some uniform bound.
Equations
- Filter.IsBoundedUnder r f u = Filter.IsBounded r (Filter.map u f)
Instances For
f is eventually bounded if and only if, there exists an admissible set on which it is
bounded.
A bounded function u is in particular eventually bounded.
A bounded above function u is in particular eventually bounded above.
A bounded below function u is in particular eventually bounded below.
IsCobounded (≺) f states that the filter f does not tend to infinity w.r.t. ≺. This is
also called frequently bounded. Will be usually instantiated with ≤ or ≥.
There is a subtlety in this definition: we want f.IsCobounded to hold for any f in the case of
complete lattices. This will be relevant to deduce theorems on complete lattices from their
versions on conditionally complete lattices with additional assumptions. We have to be careful in
the edge case of the trivial filter containing the empty set: the other natural definition
¬ ∀ a, ∀ᶠ n in f, a ≤ n
would not work as well in this case.
Equations
- Filter.IsCobounded r f = ∃ (b : α), ∀ (a : α), (∀ᶠ (x : α) in f, r x a) → r b a
Instances For
IsCoboundedUnder (≺) f u states that the image of the filter f under the map u does not
tend to infinity w.r.t. ≺. This is also called frequently bounded. Will be usually instantiated
with ≤ or ≥.
Equations
- Filter.IsCoboundedUnder r f u = Filter.IsCobounded r (Filter.map u f)
Instances For
To check that a filter is frequently bounded, it suffices to have a witness
which bounds f at some point for every admissible set.
This is only an implication, as the other direction is wrong for the trivial filter.
A filter which is eventually bounded is in particular frequently bounded (in the opposite direction). At least if the filter is not trivial.
Filters are automatically bounded or cobounded in complete lattices. To use the same statements
in complete and conditionally complete lattices but let automation fill automatically the
boundedness proofs in complete lattices, we use the tactic isBoundedDefault in the statements,
in the form (hf : f.IsBounded (≥) := by isBoundedDefault).
Equations
- Filter.tacticIsBoundedDefault = Lean.ParserDescr.node `Filter.tacticIsBoundedDefault 1024 (Lean.ParserDescr.nonReservedSymbol "isBoundedDefault" false)
Instances For
The limsSup of a filter f is the infimum of the a such that, eventually for f,
holds x ≤ a.
Instances For
The limsInf of a filter f is the supremum of the a such that, eventually for f,
holds x ≥ a.
Instances For
The limsup of a function u along a filter f is the infimum of the a such that,
eventually for f, holds u x ≤ a.
Equations
- Filter.limsup u f = (Filter.map u f).limsSup
Instances For
The liminf of a function u along a filter f is the supremum of the a such that,
eventually for f, holds u x ≥ a.
Equations
- Filter.liminf u f = (Filter.map u f).limsInf
Instances For
The blimsup of a function u along a filter f, bounded by a predicate p, is the infimum
of the a such that, eventually for f, u x ≤ a whenever p x holds.
Equations
- Filter.blimsup u f p = sInf {a : α | ∀ᶠ (x : β) in f, p x → u x ≤ a}
Instances For
The bliminf of a function u along a filter f, bounded by a predicate p, is the supremum
of the a such that, eventually for f, a ≤ u x whenever p x holds.
Equations
- Filter.bliminf u f p = sSup {a : α | ∀ᶠ (x : β) in f, p x → a ≤ u x}
Instances For
Same as limsup_const applied to ⊥ but without the NeBot f assumption
Same as limsup_const applied to ⊤ but without the NeBot f assumption
In a complete lattice, the limsup of a function is the infimum over sets s in the filter
of the supremum of the function over s
In a complete lattice, the liminf of a function is the infimum over sets s in the filter
of the supremum of the function over s
If f : α → α is a morphism of complete lattices, then the limsup of its iterates of any
a : α is a fixed point.
If f : α → α is a morphism of complete lattices, then the liminf of its iterates of any
a : α is a fixed point.
See also Filter.blimsup_or_eq_sup.
See also Filter.bliminf_or_eq_inf.
In other words, limsup cofinite s is the set of elements lying inside the family s
infinitely often.
In other words, liminf cofinite s is the set of elements lying outside the family s
finitely often.
Given an indexed family of sets s j over j : Subtype p and a function f, then
liminf_reparam j is equal to j if f is bounded below on s j, and otherwise to some
index k such that f is bounded below on s k (if there exists one).
To ensure good measurability behavior, this index k is chosen as the minimal suitable index.
This function is used to write down a liminf in a measurable way,
in Filter.HasBasis.liminf_eq_ciSup_ciInf and Filter.HasBasis.liminf_eq_ite.
Equations
Instances For
Writing a liminf as a supremum of infimum, in a (possibly non-complete) conditionally complete linear order. A reparametrization trick is needed to avoid taking the infimum of sets which are not bounded below.
Writing a liminf as a supremum of infimum, in a (possibly non-complete) conditionally complete linear order. A reparametrization trick is needed to avoid taking the infimum of sets which are not bounded below.
Given an indexed family of sets s j and a function f, then limsup_reparam j is equal
to j if f is bounded above on s j, and otherwise to some index k such that f is bounded
above on s k (if there exists one). To ensure good measurability behavior, this index k is
chosen as the minimal suitable index. This function is used to write down a limsup in a measurable
way, in Filter.HasBasis.limsup_eq_ciInf_ciSup and Filter.HasBasis.limsup_eq_ite.
Equations
- Filter.limsup_reparam f s p j = Filter.liminf_reparam f s p j
Instances For
Writing a limsup as an infimum of supremum, in a (possibly non-complete) conditionally complete linear order. A reparametrization trick is needed to avoid taking the supremum of sets which are not bounded above.
Writing a limsup as an infimum of supremum, in a (possibly non-complete) conditionally complete linear order. A reparametrization trick is needed to avoid taking the supremum of sets which are not bounded below.