Skip to content
Snippets Groups Projects
Commit 9e429c46 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Make sure that the fixpoints are non-expansive.

In order to do that, we need to quantify over non-expansive
predicates instead of arbitrary predicates.
parent c01717fb
No related branches found
No related tags found
No related merge requests found
...@@ -5,20 +5,26 @@ Import uPred. ...@@ -5,20 +5,26 @@ Import uPred.
(** Least and greatest fixpoint of a monotone function, defined entirely inside (** Least and greatest fixpoint of a monotone function, defined entirely inside
the logic. *) the logic. *)
Class BIMonoPred {M A} (F : (A uPred M) (A uPred M)) := Class BIMonoPred {M} {A : ofeT} (F : (A uPred M) (A uPred M)) := {
bi_mono_pred Φ Ψ : (( x, Φ x -∗ Ψ x) x, F Φ x -∗ F Ψ x)%I. bi_mono_pred Φ Ψ : (( x, Φ x -∗ Ψ x) x, F Φ x -∗ F Ψ x)%I;
bi_mono_pred_ne Φ : NonExpansive Φ NonExpansive (F Φ)
}.
Arguments bi_mono_pred {_ _ _ _} _ _. Arguments bi_mono_pred {_ _ _ _} _ _.
Local Existing Instance bi_mono_pred_ne.
Definition uPred_least_fixpoint {M A} (F : (A uPred M) (A uPred M)) Definition uPred_least_fixpoint {M} {A : ofeT}
(x : A) : uPred M := (F : (A uPred M) (A uPred M)) (x : A) : uPred M :=
( Φ, ( x, F Φ x Φ x) Φ x)%I. ( Φ : A -n> uPredC M, ( x, F Φ x Φ x) Φ x)%I.
Definition uPred_greatest_fixpoint {M A} (F : (A uPred M) (A uPred M)) Definition uPred_greatest_fixpoint {M} {A : ofeT}
(x : A) : uPred M := (F : (A uPred M) (A uPred M)) (x : A) : uPred M :=
( Φ, ( x, Φ x F Φ x) Φ x)%I. ( Φ : A -n> uPredC M, ( x, Φ x F Φ x) Φ x)%I.
Section least. Section least.
Context {M A} (F : (A uPred M) (A uPred M)) `{!BIMonoPred F}. Context {M} {A : ofeT} (F : (A uPred M) (A uPred M)) `{!BIMonoPred F}.
Global Instance least_fixpoint_ne : NonExpansive (uPred_least_fixpoint F).
Proof. solve_proper. Qed.
Lemma least_fixpoint_unfold_2 x : F (uPred_least_fixpoint F) x uPred_least_fixpoint F x. Lemma least_fixpoint_unfold_2 x : F (uPred_least_fixpoint F) x uPred_least_fixpoint F x.
Proof. Proof.
...@@ -30,8 +36,8 @@ Section least. ...@@ -30,8 +36,8 @@ Section least.
Lemma least_fixpoint_unfold_1 x : Lemma least_fixpoint_unfold_1 x :
uPred_least_fixpoint F x F (uPred_least_fixpoint F) x. uPred_least_fixpoint F x F (uPred_least_fixpoint F) x.
Proof. Proof.
iIntros "HF". iApply "HF". iIntros "!#" (y) "Hy". iIntros "HF". iApply ("HF" $! (CofeMor (F (uPred_least_fixpoint F))) with "[#]").
iApply bi_mono_pred; last done. iIntros "!#" (z) "?". iIntros "!#" (y) "Hy". iApply bi_mono_pred; last done. iIntros "!#" (z) "?".
by iApply least_fixpoint_unfold_2. by iApply least_fixpoint_unfold_2.
Qed. Qed.
...@@ -41,13 +47,18 @@ Section least. ...@@ -41,13 +47,18 @@ Section least.
apply (anti_symm _); auto using least_fixpoint_unfold_1, least_fixpoint_unfold_2. apply (anti_symm _); auto using least_fixpoint_unfold_1, least_fixpoint_unfold_2.
Qed. Qed.
Lemma least_fixpoint_ind (Φ : A uPred M) : Lemma least_fixpoint_ind (Φ : A uPred M) `{!NonExpansive Φ} :
( y, F Φ y Φ y) x, uPred_least_fixpoint F x Φ x. ( y, F Φ y Φ y) x, uPred_least_fixpoint F x Φ x.
Proof. iIntros "#HΦ" (x) "HF". iApply "HF". done. Qed. Proof.
iIntros "#HΦ" (x) "HF". by iApply ("HF" $! (CofeMor Φ) with "[#]").
Qed.
End least. End least.
Section greatest. Section greatest.
Context {M A} (F : (A uPred M) (A uPred M)) `{!BIMonoPred F}. Context {M} {A : ofeT} (F : (A uPred M) (A uPred M)) `{!BIMonoPred F}.
Global Instance greatest_fixpoint_ne : NonExpansive (uPred_greatest_fixpoint F).
Proof. solve_proper. Qed.
Lemma greatest_fixpoint_unfold_1 x : Lemma greatest_fixpoint_unfold_1 x :
uPred_greatest_fixpoint F x F (uPred_greatest_fixpoint F) x. uPred_greatest_fixpoint F x F (uPred_greatest_fixpoint F) x.
...@@ -61,7 +72,7 @@ Section greatest. ...@@ -61,7 +72,7 @@ Section greatest.
Lemma greatest_fixpoint_unfold_2 x : Lemma greatest_fixpoint_unfold_2 x :
F (uPred_greatest_fixpoint F) x uPred_greatest_fixpoint F x. F (uPred_greatest_fixpoint F) x uPred_greatest_fixpoint F x.
Proof. Proof.
iIntros "HF". iExists (F (uPred_greatest_fixpoint F)). iIntros "HF". iExists (CofeMor (F (uPred_greatest_fixpoint F))).
iIntros "{$HF} !#" (y) "Hy". iApply (bi_mono_pred with "[] Hy"). iIntros "{$HF} !#" (y) "Hy". iApply (bi_mono_pred with "[] Hy").
iIntros "!#" (z) "?". by iApply greatest_fixpoint_unfold_1. iIntros "!#" (z) "?". by iApply greatest_fixpoint_unfold_1.
Qed. Qed.
...@@ -72,7 +83,7 @@ Section greatest. ...@@ -72,7 +83,7 @@ Section greatest.
apply (anti_symm _); auto using greatest_fixpoint_unfold_1, greatest_fixpoint_unfold_2. apply (anti_symm _); auto using greatest_fixpoint_unfold_1, greatest_fixpoint_unfold_2.
Qed. Qed.
Lemma greatest_fixpoint_coind (Φ : A uPred M) : Lemma greatest_fixpoint_coind (Φ : A uPred M) `{!NonExpansive Φ} :
( y, Φ y F Φ y) x, Φ x uPred_greatest_fixpoint F x. ( y, Φ y F Φ y) x, Φ x uPred_greatest_fixpoint F x.
Proof. iIntros "#HΦ" (x) "Hx". iExists Φ. by iIntros "{$Hx} !#". Qed. Proof. iIntros "#HΦ" (x) "Hx". iExists (CofeMor Φ). by iIntros "{$Hx} !#". Qed.
End greatest. End greatest.
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