Library prosa.util.fixpoint
From mathcomp Require Import ssreflect ssrbool eqtype ssrfun ssrnat seq fintype bigop path.
Require Export prosa.util.rel.
Require Export prosa.util.list.
Require Export prosa.util.minmax.
Require Export prosa.util.rel.
Require Export prosa.util.list.
Require Export prosa.util.minmax.
Fixpoint Search
Least Positive Fixpoint
Fixpoint find_fixpoint_from (f : nat → nat) (x h fuel: nat): option nat :=
if fuel is S fuel' then
if f x == x then
Some x
else
if (f x) ≤ h then
find_fixpoint_from f (f x) h fuel'
else None
else None.
if fuel is S fuel' then
if f x == x then
Some x
else
if (f x) ≤ h then
find_fixpoint_from f (f x) h fuel'
else None
else None.
Given a horizon h and a monotonic function f, find the least positive
fixpoint of f no larger than h, if any.
In the following, we show that the fixpoint search works as expected.
Consider any function f ...
... and any given horizon.
We show that the result of find_fixpoint_from is indeed a fixpoint.
Using the above fact, we observe that the result of find_fixpoint is a
fixpoint.
Next, we establish properties that rely on the monotonic properties of the
function.
Suppose f is monotonically increasing ...
... and not zero at 1.
It follows that find_fixpoint_from finds the least fixpoint larger
than s.
Lemma ffpf_finds_least_fixpoint :
∀ y s fuel,
find_fixpoint_from f s h fuel = Some y →
∀ x,
s ≤ x < y →
x != f x.
∀ y s fuel,
find_fixpoint_from f s h fuel = Some y →
∀ x,
s ≤ x < y →
x != f x.
Hence find_fixpoint finds the least positive fixpoint of f.
Furthermore, we observe that find_fixpoint_from finds positive fixpoints
when provided with a positive starting point s.
Lemma ffpf_finds_positive_fixpoint :
∀ s fuel x,
Some x = find_fixpoint_from f s h fuel →
0 < s →
0 < x.
∀ s fuel x,
Some x = find_fixpoint_from f s h fuel →
0 < s →
0 < x.
Hence finds_fixpoint finds only positive fixpoints.
If find_fixpoint_from finds no fixpoint between s and h,
there is none.
Lemma ffpf_finds_none :
∀ s fuel,
s ≤ f s →
fuel ≥ h - s →
find_fixpoint_from f s h fuel = None →
∀ x,
s ≤ x < h →
x != f x.
∀ s fuel,
s ≤ f s →
fuel ≥ h - s →
find_fixpoint_from f s h fuel = None →
∀ x,
s ≤ x < h →
x != f x.
Hence, if find_fixpoint finds no positive fixpoint less than h, there
is none.
Lemma ffp_finds_none :
find_fixpoint f h = None →
∀ x,
0 < x < h →
x != f x.
End MonotonicFunction.
End FixpointSearch.
find_fixpoint f h = None →
∀ x,
0 < x < h →
x != f x.
End MonotonicFunction.
End FixpointSearch.
Maximal Fixpoint across a Search Space
Definition find_max_fixpoint_of_seq (f : nat → nat → nat)
(sp : seq nat) (h : nat) : option nat :=
let is_some opt := opt != None in
let fixpoints := [seq find_fixpoint (f s) h | s <- sp] in
let max := \max_(fp <- fixpoints | is_some fp) (odflt 0) fp in
if all is_some fixpoints then Some max else None.
(sp : seq nat) (h : nat) : option nat :=
let is_some opt := opt != None in
let fixpoints := [seq find_fixpoint (f s) h | s <- sp] in
let max := \max_(fp <- fixpoints | is_some fp) (odflt 0) fp in
if all is_some fixpoints then Some max else None.
We prove that the result of the find_max_fixpoint_of_seq is a fixpoint of
the function for an element of the search space.
Lemma fmfs_finds_fixpoint :
∀ (f : nat → nat → nat) (sp : seq nat) (h x : nat),
~~ nilp sp →
find_max_fixpoint_of_seq f sp h = Some x →
exists2 a, a \in sp & x = f a x.
∀ (f : nat → nat → nat) (sp : seq nat) (h x : nat),
~~ nilp sp →
find_max_fixpoint_of_seq f sp h = Some x →
exists2 a, a \in sp & x = f a x.
Furthermore, we show that the result is the maximum of all fixpoints, with
regard to the search space.
Lemma fmfs_is_maximum :
∀ (f : nat → nat → nat) (sp : seq nat) (h s r: nat),
Some r = find_max_fixpoint_of_seq f sp h →
s \in sp →
exists2 v,
Some v = find_fixpoint (f s) h & r ≥ v.
∀ (f : nat → nat → nat) (sp : seq nat) (h s r: nat),
Some r = find_max_fixpoint_of_seq f sp h →
s \in sp →
exists2 v,
Some v = find_fixpoint (f s) h & r ≥ v.
Consider a limit L ...
... and any predicate P on numbers.
We create a derive predicate that defines the search space as all values
less than L that satisfy P.
This is used to create a corresponding sequence of points in the search
space.
Based on this conversion, we define a function to find the maximum of all
fixpoints within the search space.
Definition find_max_fixpoint (f : nat → nat → nat) (h : nat) :=
if (has P (iota 0 L)) then find_max_fixpoint_of_seq f sp h else None.
if (has P (iota 0 L)) then find_max_fixpoint_of_seq f sp h else None.
The result of find_max_fixpoint is a fixpoint of the function f for a
an element of the search space.
Corollary fmf_finds_fixpoint :
∀ (f : nat → nat → nat) (h x: nat),
find_max_fixpoint f h = Some x →
exists2 a, is_in_search_space a & x = f a x.
∀ (f : nat → nat → nat) (h x: nat),
find_max_fixpoint f h = Some x →
exists2 a, is_in_search_space a & x = f a x.
Finally, we observe that there is no fixpoint in the search space larger
than the result of find_maximum_fixpoint.