Skip to content
Snippets Groups Projects
Commit 5758b8cd authored by Ralf Jung's avatar Ralf Jung
Browse files

Merge branch 'master' of gitlab.mpi-sws.org:FP/iris-coq

parents 16240e35 fe82f3a4
No related branches found
No related tags found
No related merge requests found
......@@ -188,12 +188,21 @@ Definition fixpoint_eq : @fixpoint = @fixpoint_def := proj2_sig fixpoint_aux.
Section fixpoint.
Context {A : cofeT} `{Inhabited A} (f : A A) `{!Contractive f}.
Lemma fixpoint_unfold : fixpoint f f (fixpoint f).
Proof.
apply equiv_dist=>n.
rewrite fixpoint_eq /fixpoint_def (conv_compl n (fixpoint_chain f)) //.
induction n as [|n IH]; simpl; eauto using contractive_0, contractive_S.
Qed.
Lemma fixpoint_unique (x : A) : x f x x fixpoint f.
Proof.
rewrite !equiv_dist=> Hx n. induction n as [|n IH]; simpl in *.
- rewrite Hx fixpoint_unfold; eauto using contractive_0.
- rewrite Hx fixpoint_unfold. apply (contractive_S _), IH.
Qed.
Lemma fixpoint_ne (g : A A) `{!Contractive g} n :
( z, f z {n} g z) fixpoint f {n} fixpoint g.
Proof.
......
......@@ -39,7 +39,7 @@ In order to solve the recursive domain equation in \Sref{sec:model} it is also e
\end{defn}
Intuitively, applying a non-expansive function to some data will not suddenly introduce differences between seemingly equal data.
Elements that cannot be distinguished by programs within $n$ steps remain indistinguishable after applying $f$.
The reason that contractive functions are interesting is that for every contractive $f : \cofe \to \cofe$ with $\cofe$ inhabited, there exists a \emph{unique}\footnote{Uniqueness is not proven in Coq.} fixed-point $\fix(f)$ such that $\fix(f) = f(\fix(f))$.
The reason that contractive functions are interesting is that for every contractive $f : \cofe \to \cofe$ with $\cofe$ inhabited, there exists a \emph{unique} fixed-point $\fix(f)$ such that $\fix(f) = f(\fix(f))$.
\begin{defn}
The category $\COFEs$ consists of COFEs as objects, and non-expansive functions as arrows.
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment