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

Simplify and clean up wp_fix.

Shorten proofs and name variables that we introduce.
parent 1ae6f1c1
No related branches found
No related tags found
No related merge requests found
From Coq Require Import Wf_nat. From Coq Require Import Wf_nat.
From iris.program_logic Require Import weakestpre wsat. From iris.program_logic Require Import weakestpre wsat.
Local Hint Extern 10 (_ _) => omega.
Local Hint Extern 10 (@eq coPset _ _) => set_solver.
Local Hint Extern 10 ({_} _) =>
repeat match goal with
| H : wsat _ _ _ _ |- _ => apply wsat_valid in H; last omega
end; solve_validN.
(** This files provides an alternative definition of wp in terms of a (** This files provides an alternative definition of wp in terms of a fixpoint
fixpoint of a contractive function, rather than a CoInductive type. of a contractive function, rather than a CoInductive type. This is how we define
This is how we define wp on paper. wp on paper. We show that the two versions are equivalent. *)
We show that the two versions are equivalent. *)
Section def. Section def.
Context {Λ : language} {Σ : iFunctor}. Context {Λ : language} {Σ : iFunctor}.
Local Notation iProp := (iProp Λ Σ). Local Notation iProp := (iProp Λ Σ).
Local Notation coPsetC := (leibnizC (coPset)).
Definition coPsetC := leibnizC (coPset).
Definition pre_wp_def (wp : coPsetC -n> exprC Λ -n> (valC Λ -n> iProp) -n> iProp) Program Definition wp_pre
(E : coPset) (e1 : expr Λ) (Φ : valC Λ -n> iProp) (wp : coPsetC -n> exprC Λ -n> (valC Λ -n> iProp) -n> iProp)
(n : nat) (r1 : iRes Λ Σ) : Prop := (E : coPset) (e1 : expr Λ) (Φ : valC Λ -n> iProp) : iProp :=
rf k Ef σ1, 0 k < n E Ef = {| uPred_holds n r1 := rf k Ef σ1,
wsat (S k) (E Ef) σ1 (r1 rf) 0 k < n E Ef =
( v, to_val e1 = Some v wsat (S k) (E Ef) σ1 (r1 rf)
r2, Φ v (S k) r2 wsat (S k) (E Ef) σ1 (r2 rf)) ( v, to_val e1 = Some v
(to_val e1 = None 0 < k r2, Φ v (S k) r2 wsat (S k) (E Ef) σ1 (r2 rf))
reducible e1 σ1 (to_val e1 = None 0 < k
( e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef reducible e1 σ1
r2 r2', ( e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef
wsat k (E Ef) σ2 (r2 r2' rf) r2 r2',
wp E e2 Φ k r2 wsat k (E Ef) σ2 (r2 r2' rf)
default True ef (λ ef, wp ef (cconst True%I) k r2'))). wp E e2 Φ k r2
default True ef (λ ef, wp ef (cconst True%I) k r2'))) |}.
Next Obligation.
intros wp E e1 Φ n r1 r1' Hwp Hr1 rf k Ef σ1 ???; simpl in *.
destruct (Hwp rf k Ef σ1); auto.
by rewrite (dist_le _ _ _ _ Hr1); last omega.
Qed.
Next Obligation.
intros wp E e1 Φ n1 n2 r1 ? Hwp [r2 ?] ?? rf k Ef σ1 ???; setoid_subst.
destruct (Hwp (r2 rf) k Ef σ1) as [Hval Hstep]; rewrite ?assoc; auto.
split.
- intros v Hv. destruct (Hval v Hv) as [r3 [??]].
exists (r3 r2). rewrite -assoc. eauto using uPred_weaken, cmra_included_l.
- intros ??. destruct Hstep as [Hred Hpstep]; auto.
split; [done|]=> e2 σ2 ef ?.
edestruct Hpstep as (r3&r3'&?&?&?); eauto.
exists r3, (r3' r2); split_and?; auto.
+ by rewrite assoc -assoc.
+ destruct ef; simpl in *; eauto using uPred_weaken, cmra_included_l.
Qed.
Program Definition pre_wp Lemma wp_pre_contractive' n E e Φ1 Φ2 r
(wp : coPsetC -n> exprC Λ -n> (valC Λ -n> iProp) -n> iProp) (wp1 wp2 : coPsetC -n> exprC Λ -n> (valC Λ -n> iProp) -n> iProp) :
(E : coPset) (e1 : expr Λ) (Φ : valC Λ -n> iProp) : iProp := ( i : nat, i < n wp1 {i} wp2) Φ1 {n} Φ2
{| uPred_holds := pre_wp_def wp E e1 Φ |}. wp_pre wp1 E e Φ1 n r wp_pre wp2 E e Φ2 n r.
Next Obligation. Proof.
intros ????? r1 r1' Hwp EQr. intros HI Hwp rf k Ef σ1 ???.
intros ???? Hk ??. edestruct Hwp as [Hval Hstep]; [done..| |]. destruct (Hwp rf k Ef σ1) as [Hval Hstep]; auto.
{ eapply wsat_ne'; last done. apply (dist_le n); last omega. split.
by rewrite EQr. } { intros v ?. destruct (Hval v) as (r2&?&?); auto.
split; done. exists r2. split; [apply |]; auto. }
Qed. intros ??. destruct Hstep as [Hred Hpstep]; auto.
Next Obligation. split; [done|]=> e2 σ2 ef ?.
intros ???? n1 n2 r1 r2 Hwp [r3 Hr] Hn Hvalid. destruct (Hpstep e2 σ2 ef) as (r2&r2'&?&?&?); [done..|].
intros ???? Hk ??. edestruct (Hwp (r3 rf) k) as [Hval Hstep]; [|done..| |]. exists r2, r2'; split_and?; auto.
{ omega. } - apply HI with k; auto.
{ eapply wsat_ne'; last done. by rewrite Hr assoc. } assert (wp1 E e2 Φ2 {n} wp1 E e2 Φ1) as HwpΦ by (by rewrite ).
split. apply HwpΦ; auto.
- intros v Hv. destruct (Hval v Hv) as [r4 [ Hw]]. - destruct ef as [ef|]; simpl in *; last done.
exists (r4 r3). split. apply HI with k; auto.
+ eapply uPred_weaken; first exact: ; eauto. Qed.
* by eexists. Instance wp_pre_ne n wp E e : Proper (dist n ==> dist n) (wp_pre wp E e).
* apply wsat_valid in Hw; last done. solve_validN. Proof.
+ by rewrite -assoc. split; split;
- intros ??. edestruct Hstep as [Hred Hpstep]; [done..|]. eapply wp_pre_contractive'; eauto using dist_le, (symmetry (R:=dist _)).
split; first done. intros ????. Qed.
edestruct Hpstep as (r4 & r4' & Hw & He2 & Hef); [done..|].
exists r4, (r4' r3). split_and?.
+ move: Hw. by rewrite -!assoc.
+ done.
+ destruct ef; last done.
eapply uPred_weaken; first exact: Hef; eauto.
* by eexists.
* apply wsat_valid in Hw; last omega. solve_validN.
Qed.
Local Instance pre_wp_ne n wp E e : Definition wp_preC
Proper (dist n ==> dist n) (pre_wp wp E e). (wp : coPsetC -n> exprC Λ -n> (valC Λ -n> iProp) -n> iProp) :
Proof. coPsetC -n> exprC Λ -n> (valC Λ -n> iProp) -n> iProp :=
cut ( n Φ1 Φ2, Φ1 {n} Φ2 r, CofeMor (λ E : coPsetC, CofeMor (λ e : exprC Λ, CofeMor (wp_pre wp E e))).
pre_wp wp E e Φ1 n r pre_wp wp E e Φ2 n r).
{ intros H Φ1 Φ2 EQΦ. split; split; eapply H.
- eauto using dist_le.
- symmetry. eauto using dist_le. }
clear n. intros n Φ1 Φ2 EQΦ r Hwp.
intros ???? Hk ??. edestruct Hwp as [Hval Hstep]; [done..|].
split.
- intros ??. edestruct Hval as [r2 [ Hw]]; [done..|].
exists r2. split; last done. apply EQΦ, .
+ omega.
+ apply wsat_valid in Hw; last omega. solve_validN.
- intros ??. edestruct Hstep as [Hred Hpstep]; [done..|].
split; first done. intros ????.
edestruct Hpstep as (r2 & r2' & Hw & He2 & Hef); [done..|].
exists r2, r2'. split_and?; try done.
eapply uPred_holds_ne, He2.
+ apply (dist_le n); last omega. by rewrite -EQΦ.
+ apply wsat_valid in Hw; last omega. solve_validN.
Qed.
Definition pre_wp_mor wp : coPsetC -n> exprC Λ -n> (valC Λ -n> iProp) -n> iProp := Local Instance pre_wp_contractive : Contractive wp_preC.
CofeMor (λ E : coPsetC, CofeMor (λ e : exprC Λ, CofeMor (pre_wp wp E e))). Proof.
split; split; eapply wp_pre_contractive'; auto using (symmetry (R:=dist _)).
Qed.
Local Instance pre_wp_contractive : Contractive pre_wp_mor. Definition wp_fix : coPsetC -n> exprC Λ -n> (valC Λ -n> iProp) -n> iProp :=
Proof. fixpoint wp_preC.
cut ( n (wp1 wp2 : coPsetC -n> exprC Λ -n> (valC Λ -n> iProp) -n> iProp),
( i : nat, i < n wp1 {i} wp2) E e Φ r,
pre_wp wp1 E e Φ n r pre_wp wp2 E e Φ n r).
{ intros H n wp1 wp2 HI. split; split; eapply H; intros.
- apply HI. omega.
- symmetry. apply HI. omega. }
intros n wp1 wp2 HI E e1 Φ r1 Hwp.
intros ???? Hk ??. edestruct Hwp as [Hval Hstep]; [done..|].
split; first done.
intros ??. edestruct Hstep as [Hred Hpstep]; [done..|].
split; first done. intros ????.
edestruct Hpstep as (r2 & r2' & Hw & He2 & Hef); [done..|].
exists r2, r2'. split_and?; try done.
- eapply uPred_holds_ne, He2.
+ apply HI. omega.
+ apply wsat_valid in Hw; last omega. solve_validN.
- destruct ef; last done. eapply uPred_holds_ne, Hef.
+ apply HI. omega.
+ apply wsat_valid in Hw; last omega. solve_validN.
Qed.
Definition wp_fix : coPsetC -n> exprC Λ -n> (valC Λ -n> iProp) -n> iProp := Lemma wp_fix_unfold E e Φ : wp_fix E e Φ ⊣⊢ wp_preC wp_fix E e Φ.
fixpoint pre_wp_mor. Proof. by rewrite /wp_fix -fixpoint_unfold. Qed.
Lemma wp_fix_unfold E e Φ : pre_wp_mor wp_fix E e Φ ⊣⊢ wp_fix E e Φ. Lemma wp_fix_correct E e (Φ : valC Λ -n> iProp) : wp_fix E e Φ ⊣⊢ wp E e Φ.
Proof. rewrite -fixpoint_unfold. done. Qed. Proof.
split=> n r Hr. rewrite wp_eq /wp_def {2}/uPred_holds.
Lemma wp_fix_sound (E : coPset) (e : expr Λ) (Φ : val Λ -> iProp) split; revert r E e Φ Hr.
(Hproper : n, Proper (dist n ==> dist n) (Φ : valC Λ -> iProp)) : - induction n as [n IH] using lt_wf_ind=> r1 E e Φ ? Hwp.
wp_fix E e (@CofeMor _ _ _ Hproper) wp E e Φ.
Proof.
split. rewrite wp_eq /wp_def {2}/uPred_holds.
intros n. revert E e Φ Hproper.
induction n as [n IH] using lt_wf_ind=> E e Φ Hproper r1 Hr1 Hwp.
case EQ: (to_val e)=>[v|]. case EQ: (to_val e)=>[v|].
- rewrite -(of_to_val _ _ EQ) {IH}. constructor. rewrite pvs_eq. + rewrite -(of_to_val _ _ EQ) {IH}. constructor. rewrite pvs_eq.
intros rf k Ef σ ???. destruct k; first (exfalso; omega). intros rf [|k] Ef σ ???; first omega.
apply wp_fix_unfold in Hwp; last done. apply wp_fix_unfold in Hwp; last done.
edestruct (Hwp rf k Ef σ); eauto with omega; set_solver. destruct (Hwp rf k Ef σ); auto.
- constructor; first done. intros ???? Hk ??. + constructor; [done|]=> rf k Ef σ1 ???.
apply wp_fix_unfold in Hwp; last done. apply wp_fix_unfold in Hwp; last done.
edestruct Hwp as [Hval Hstep]; [|done..|]; first omega. destruct (Hwp rf k Ef σ1) as [Hval [Hred Hpstep]]; auto.
edestruct Hstep as [Hred Hpstep]; [done|omega|]. clear Hstep. split; [done|]=> e2 σ2 ef ?.
split; first done. intros ????. destruct (Hpstep e2 σ2 ef) as (r2&r2'&?&?&?); [done..|].
edestruct Hpstep as (r2 & r2' & Hw & He2 & Hef); [done..|]. exists r2, r2'; split_and?; auto.
exists r2, r2'. split_and?; first done. intros ? ->. change (weakestpre.wp_pre (cconst True%I) e' k r2'); eauto.
+ apply: IH; last done; first omega. - induction n as [n IH] using lt_wf_ind=> r1 E e Φ ? Hwp.
apply wsat_valid in Hw; last omega. solve_validN. apply wp_fix_unfold; [done|]=> rf k Ef σ1 ???. split.
+ intros e' ?. subst ef. apply: IH; last done; first omega. + intros v Hval.
apply wsat_valid in Hw; last omega. solve_validN. destruct Hwp as [??? Hpvs|]; rewrite ?to_of_val in Hval; simplify_eq/=.
Qed.
Lemma wp_fix_complete (E : coPset) (e : expr Λ) (Φ : val Λ -> iProp)
(Hproper : n, Proper (dist n ==> dist n) (Φ : valC Λ -> iProp)) :
wp E e Φ wp_fix E e (@CofeMor _ _ _ Hproper).
Proof.
split. rewrite wp_eq /wp_def {1}/uPred_holds.
intros n. revert E e Φ Hproper.
induction n as [n IH] using lt_wf_ind=> E e Φ Hproper r1 Hr1 Hwp.
apply wp_fix_unfold; first done.
intros rf k Ef σ1 ???. split.
- intros ? Hval. destruct Hwp as [??? Hpvs|]; last by destruct (to_val e1).
rewrite pvs_eq in Hpvs. rewrite pvs_eq in Hpvs.
edestruct (Hpvs rf (S k) Ef σ1) as (r2 & ? & ?); [omega|set_solver|done|]. destruct (Hpvs rf (S k) Ef σ1) as (r2&?&?); eauto.
exists r2. split; last done. rewrite to_of_val in Hval. + intros Hval ?.
by simplify_option_eq. destruct Hwp as [|???? Hwp]; rewrite ?to_of_val in Hval; simplify_eq/=.
- intros Hval ?. destruct Hwp as [|??? _ Hwp]. edestruct (Hwp rf k Ef σ1) as [? Hpstep]; auto.
{ by rewrite to_of_val in Hval. } split; [done|]=> e2 σ2 ef ?.
edestruct Hwp as [? Hpstep]; try done; first omega; []. destruct (Hpstep e2 σ2 ef) as (r2&r2'&?&?&?); auto.
split; first done. exists r2, r2'. destruct ef; simpl; auto.
intros ????. edestruct Hpstep as (r2 & r2' & Hw & He2 & Hef); [done..|]. Qed.
exists r2, r2'. split_and?; first done.
+ apply IH, He2; first omega.
apply wsat_valid in Hw; last omega. solve_validN.
+ destruct ef; last done.
apply IH, Hef; first omega; last done.
apply wsat_valid in Hw; last omega. solve_validN.
Qed.
End def. End def.
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