Asymptotic cone of a set #
This file defines the asymptotic cone of a set in a topological affine space.
Implementation details #
The asymptotic cone of a set $A$ is usually defined as the set of points $v$ for which there exist
sequences $t_n > 0$ and $x_n \in A$ such that $t_n \to 0$ and $t_n x_n \to v$. We take a different
approach here using filters: we define the asymptotic cone of s
as the set of vectors v
such
that āį¶ p in Filter.atTop ⢠š v, p ā s
holds.
Main definitions #
AffineSpace.asymptoticNhds
: the filter of neighborhoods at infinity in some direction.asymptoticCone
: the asymptotic cone of a subset of a topological affine space.
Main statements #
Convex.smul_vadd_mem_of_isClosed_of_mem_asymptoticCone
: ifv
is in the asymptotic cone of a closed convex sets
, then every ray of directionv
starting froms
is contained ins
.Convex.smul_vadd_mem_of_mem_nhds_of_mem_asymptoticCone
: ifv
is in the asymptotic cone of a convex sets
, then every ray of directionv
starting from the interior ofs
is contained ins
.
In a topological affine space P
over k
, AffineSpace.asymptoticNhds k P v
is the filter of
neighborhoods at infinity in directions near v
. In a topological vector space, this is the filter
Filter.atTop ⢠š v
. To support affine spaces, the actual definition is different and should be
considered an implementation detail. Use AffineSpace.asymptoticNhds_eq_smul
or
AffineSpace.asymptoticNhds_eq_smul_vadd
for unfolding.
Equations
- AffineSpace.asymptoticNhds k P v = ⨠(p : P), Filter.atTop ⢠nhds v +ᵄ pure p
Instances For
The set of directions v
for which the set has points arbitrarily far in directions near v
.
Instances For
If a closed set s
is star-convex at p
and v
is in the asymptotic cone of s
, then the ray
of direction v
starting from p
is contained in s
.
If v
is in the asymptotic cone of a closed convex set s
, then for every p ā s
, the ray of
direction v
starting from p
is contained in s
.
If v
is in the asymptotic cone of a convex set s
, then for every interior point p
, the ray
of direction v
starting from p
is contained in s
.