Contracting maps
A Lipschitz continuous self-map with Lipschitz constant K < 1
is called a contracting map.
In this file we prove the Banach fixed point theorem, some explicit estimates on the rate
of convergence, and some properties of the map sending a contracting map to its fixed point.
Main definitions
contracting_with K f
: a Lipschitz continuous self-map withK < 1
;efixed_point
: given a contracting mapf
on a complete emetric space and a pointx
such thatedist x (f x) < ∞
,efixed_point f hf x hx
is the unique fixed point off
inemetric.ball x ∞
;fixed_point
: the unique fixed point of a contracting map on a complete nonempty metric space.
Tags
contracting map, fixed point, Banach fixed point theorem
A map is said to be contracting_with K
, if K < 1
and f
is lipschitz_with K
.
Equations
- contracting_with K f = (K < 1 ∧ lipschitz_with K f)
If a map f
is contracting_with K
, and s
is a forward-invariant set, then
restriction of f
to s
is contracting_with K
as well.
Banach fixed-point theorem, contraction mapping theorem, emetric_space
version.
A contracting map on a complete metric space has a fixed point.
We include more conclusions in this theorem to avoid proving them again later.
The main API for this theorem are the functions efixed_point
and fixed_point
,
and lemmas about these functions.
Let x
be a point of a complete emetric space. Suppose that f
is a contracting map,
and edist x (f x) < ∞
. Then efixed_point
is the unique fixed point of f
in emetric.ball x ∞
.
Equations
- contracting_with.efixed_point f hf x hx = classical.some _
Banach fixed-point theorem for maps contracting on a complete subset.
Let s
be a complete forward-invariant set of a self-map f
. If f
contracts on s
and x ∈ s
satisfies edist x (f x) < ⊤
, then efixed_point'
is the unique fixed point
of the restriction of f
to s ∩ emetric.ball x ⊤
.
Equations
- contracting_with.efixed_point' f hsc hsf hf x hxs hx = classical.some _
If a globally contracting map f
has two complete forward-invariant sets s
, t
,
and x ∈ s
is at a finite distance from y ∈ t
, then the efixed_point'
constructed by x
is the same as the efixed_point'
constructed by y
.
This lemma takes additional arguments stating that f
contracts on s
and t
because this way
it can be used to prove the desired equality with non-trivial proofs of these facts.
Let f
be a contracting map with constant K
; let g
be another map uniformly
C
-close to f
. If x
and y
are their fixed points, then dist x y ≤ C / (1 - K)
.
The unique fixed point of a contracting map in a nonempty complete metric space.
Equations
- contracting_with.fixed_point f hf = contracting_with.efixed_point f hf (classical.choice _inst_2) _
The point provided by contracting_with.fixed_point
is actually a fixed point.
Aposteriori estimates on the convergence of iterates to the fixed point.