diff --git a/CHANGELOG.md b/CHANGELOG.md index 77e1cf45fbac7c8a87e7b66172bf569005ab0e7e..1f9d478f9d27f5ba747a6f62950a810dcf14dc53 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -32,6 +32,7 @@ Changes in Coq: - `cmra_opM_assoc_L` → `cmra_op_opM_assoc_L` - `cmra_opM_assoc'` → `cmra_opM_opM_assoc` * `namespaces` has been moved to std++. +* Changed `IntoVal` to be directly usable for rewriting `e` into `of_val v`. ## Iris 3.1.0 (released 2017-12-19) diff --git a/theories/heap_lang/lib/atomic_heap.v b/theories/heap_lang/lib/atomic_heap.v index fcacab5936ea0a6408f101a0968e638190fa27cf..bc4160c9b620dcf3bd952ad6fa490e4f6e781f0f 100644 --- a/theories/heap_lang/lib/atomic_heap.v +++ b/theories/heap_lang/lib/atomic_heap.v @@ -82,7 +82,7 @@ Section proof. (λ v (_:()), l ↦ w)%I (λ _ _, #()%V). Proof. - iIntros (<-%of_to_val Q Φ) "? AU". wp_let. wp_proj. wp_proj. + iIntros (<- Q Φ) "? AU". wp_let. wp_proj. wp_proj. iMod (aupd_acc with "AU") as (v) "[H↦ [_ Hclose]]"; first solve_ndisj. wp_store. iMod ("Hclose" $! () with "H↦") as "HΦ". by iApply "HΦ". Qed. @@ -95,7 +95,7 @@ Section proof. (λ v (_:()), if decide (v = w1) then l ↦ w2 else l ↦ v)%I (λ v _, #(if decide (v = w1) then true else false)%V). Proof. - iIntros (<-%of_to_val <-%of_to_val Q Φ) "? AU". wp_let. repeat wp_proj. + iIntros (<- <- Q Φ) "? AU". wp_let. repeat wp_proj. iMod (aupd_acc with "AU") as (v) "[H↦ [_ Hclose]]"; first solve_ndisj. destruct (decide (v = w1)) as [Hv|Hv]; [wp_cas_suc|wp_cas_fail]; iMod ("Hclose" $! () with "H↦") as "HΦ"; by iApply "HΦ". diff --git a/theories/heap_lang/lib/par.v b/theories/heap_lang/lib/par.v index f37cf4ebd65845d13360a69f0bb912fcf498f839..eaf7d05cb6ac5e2b1ef327b887238573bed16a1e 100644 --- a/theories/heap_lang/lib/par.v +++ b/theories/heap_lang/lib/par.v @@ -21,13 +21,13 @@ Context `{!heapG Σ, !spawnG Σ}. brought together. That is strictly stronger than first stripping a later and then merging them, as demonstrated by [tests/joining_existentials.v]. This is why these are not Texan triples. *) -Lemma par_spec (Ψ1 Ψ2 : val → iProp Σ) e (f1 f2 : val) (Φ : val → iProp Σ) - `{Hef : !IntoVal e (f1,f2)} : +Lemma par_spec (Ψ1 Ψ2 : val → iProp Σ) e (f1 f2 : val) (Φ : val → iProp Σ) : + IntoVal e (f1,f2) → WP f1 #() {{ Ψ1 }} -∗ WP f2 #() {{ Ψ2 }} -∗ (▷ ∀ v1 v2, Ψ1 v1 ∗ Ψ2 v2 -∗ ▷ Φ (v1,v2)%V) -∗ WP par e {{ Φ }}. Proof. - apply of_to_val in Hef as <-. iIntros "Hf1 Hf2 HΦ". + iIntros (<-) "Hf1 Hf2 HΦ". rewrite /par /=. wp_let. wp_proj. wp_apply (spawn_spec parN with "Hf1"). iIntros (l) "Hl". wp_let. wp_proj. wp_bind (f2 _). diff --git a/theories/heap_lang/lib/spawn.v b/theories/heap_lang/lib/spawn.v index 79d6c669426af9f85cb3250c2c4f0398608d266e..75175fe82e3fbda945f08ce00ef0a5594e4807e6 100644 --- a/theories/heap_lang/lib/spawn.v +++ b/theories/heap_lang/lib/spawn.v @@ -44,10 +44,11 @@ Global Instance join_handle_ne n l : Proof. solve_proper. Qed. (** The main proofs. *) -Lemma spawn_spec (Ψ : val → iProp Σ) e (f : val) `{Hef : !IntoVal e f} : +Lemma spawn_spec (Ψ : val → iProp Σ) e (f : val) : + IntoVal e f → {{{ WP f #() {{ Ψ }} }}} spawn e {{{ l, RET #l; join_handle l Ψ }}}. Proof. - apply of_to_val in Hef as <-. iIntros (Φ) "Hf HΦ". rewrite /spawn /=. + iIntros (<- Φ) "Hf HΦ". rewrite /spawn /=. wp_let. wp_alloc l as "Hl". wp_let. iMod (own_alloc (Excl ())) as (γ) "Hγ"; first done. iMod (inv_alloc N _ (spawn_inv γ l Ψ) with "[Hl]") as "#?". diff --git a/theories/heap_lang/lifting.v b/theories/heap_lang/lifting.v index 7d00192b0b79aa01a3cc1fa0fae6c4ba1d8fe9cf..68ef2f8e178c7e2717a83ae017605cd6c9cb4ae1 100644 --- a/theories/heap_lang/lifting.v +++ b/theories/heap_lang/lifting.v @@ -125,7 +125,7 @@ Lemma wp_alloc s E e v : IntoVal e v → {{{ True }}} Alloc e @ s; E {{{ l, RET LitV (LitLoc l); l ↦ v }}}. Proof. - iIntros (<-%of_to_val Φ) "_ HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. + iIntros (<- Φ) "_ HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>"; iSplit; first by auto. iNext; iIntros (v2 σ2 efs Hstep); inv_head_step. iMod (@gen_heap_alloc with "Hσ") as "[Hσ Hl]"; first done. @@ -135,7 +135,7 @@ Lemma twp_alloc s E e v : IntoVal e v → [[{ True }]] Alloc e @ s; E [[{ l, RET LitV (LitLoc l); l ↦ v }]]. Proof. - iIntros (<-%of_to_val Φ) "_ HΦ". iApply twp_lift_atomic_head_step_no_fork; auto. + iIntros (<- Φ) "_ HΦ". iApply twp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>"; iSplit; first by auto. iIntros (v2 σ2 efs Hstep); inv_head_step. iMod (@gen_heap_alloc with "Hσ") as "[Hσ Hl]"; first done. @@ -165,7 +165,7 @@ Lemma wp_store s E l v' e v : IntoVal e v → {{{ ▷ l ↦ v' }}} Store (Lit (LitLoc l)) e @ s; E {{{ RET LitV LitUnit; l ↦ v }}}. Proof. - iIntros (<-%of_to_val Φ) ">Hl HΦ". + iIntros (<- Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iNext; iIntros (v2 σ2 efs Hstep); inv_head_step. @@ -176,7 +176,7 @@ Lemma twp_store s E l v' e v : IntoVal e v → [[{ l ↦ v' }]] Store (Lit (LitLoc l)) e @ s; E [[{ RET LitV LitUnit; l ↦ v }]]. Proof. - iIntros (<-%of_to_val Φ) "Hl HΦ". + iIntros (<- Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iIntros (v2 σ2 efs Hstep); inv_head_step. @@ -189,7 +189,7 @@ Lemma wp_cas_fail s E l q v' e1 v1 e2 : {{{ ▷ l ↦{q} v' }}} CAS (Lit (LitLoc l)) e1 e2 @ s; E {{{ RET LitV (LitBool false); l ↦{q} v' }}}. Proof. - iIntros (<-%of_to_val [v2 <-%of_to_val] ? Φ) ">Hl HΦ". + iIntros (<- [v2 <-%of_to_val] ? Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step. @@ -200,7 +200,7 @@ Lemma twp_cas_fail s E l q v' e1 v1 e2 : [[{ l ↦{q} v' }]] CAS (Lit (LitLoc l)) e1 e2 @ s; E [[{ RET LitV (LitBool false); l ↦{q} v' }]]. Proof. - iIntros (<-%of_to_val [v2 <-%of_to_val] ? Φ) "Hl HΦ". + iIntros (<- [v2 <-%of_to_val] ? Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iIntros (v2' σ2 efs Hstep); inv_head_step. @@ -212,7 +212,7 @@ Lemma wp_cas_suc s E l e1 v1 e2 v2 : {{{ ▷ l ↦ v1 }}} CAS (Lit (LitLoc l)) e1 e2 @ s; E {{{ RET LitV (LitBool true); l ↦ v2 }}}. Proof. - iIntros (<-%of_to_val <-%of_to_val Φ) ">Hl HΦ". + iIntros (<- <- Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step. @@ -224,7 +224,7 @@ Lemma twp_cas_suc s E l e1 v1 e2 v2 : [[{ l ↦ v1 }]] CAS (Lit (LitLoc l)) e1 e2 @ s; E [[{ RET LitV (LitBool true); l ↦ v2 }]]. Proof. - iIntros (<-%of_to_val <-%of_to_val Φ) "Hl HΦ". + iIntros (<- <- Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iIntros (v2' σ2 efs Hstep); inv_head_step. @@ -237,7 +237,7 @@ Lemma wp_faa s E l i1 e2 i2 : {{{ ▷ l ↦ LitV (LitInt i1) }}} FAA (Lit (LitLoc l)) e2 @ s; E {{{ RET LitV (LitInt i1); l ↦ LitV (LitInt (i1 + i2)) }}}. Proof. - iIntros (<-%of_to_val Φ) ">Hl HΦ". + iIntros (<- Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step. @@ -249,7 +249,7 @@ Lemma twp_faa s E l i1 e2 i2 : [[{ l ↦ LitV (LitInt i1) }]] FAA (Lit (LitLoc l)) e2 @ s; E [[{ RET LitV (LitInt i1); l ↦ LitV (LitInt (i1 + i2)) }]]. Proof. - iIntros (<-%of_to_val Φ) "Hl HΦ". + iIntros (<- Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iIntros (v2' σ2 efs Hstep); inv_head_step. diff --git a/theories/heap_lang/tactics.v b/theories/heap_lang/tactics.v index 1789ca57080da7b48708345567d2701eafcaed5f..0c5c751dadf82d63bc533df86c45ed02d48d4f41 100644 --- a/theories/heap_lang/tactics.v +++ b/theories/heap_lang/tactics.v @@ -147,6 +147,9 @@ Qed. Lemma to_val_is_Some e : is_Some (to_val e) → is_Some (heap_lang.to_val (to_expr e)). Proof. intros [v ?]; exists v; eauto using to_val_Some. Qed. +Lemma expr_of_val e v : + to_val e = Some v → of_val v = W.to_expr e. +Proof. intros ?. apply of_to_val, to_val_Some. done. Qed. Fixpoint subst (x : string) (es : expr) (e : expr) : expr := match e with @@ -217,8 +220,8 @@ Hint Extern 0 (Closed _ _) => solve_closed : typeclass_instances. Ltac solve_into_val := match goal with | |- IntoVal ?e ?v => - let e' := W.of_expr e in change (to_val (W.to_expr e') = Some v); - apply W.to_val_Some; simpl; unfold W.to_expr; reflexivity + let e' := W.of_expr e in change (of_val v = W.to_expr e'); + apply W.expr_of_val; simpl; unfold W.to_expr; reflexivity end. Hint Extern 10 (IntoVal _ _) => solve_into_val : typeclass_instances. diff --git a/theories/program_logic/language.v b/theories/program_logic/language.v index 707f9c671861cf99218208bf490f352c8637849f..fb541089779ae76e042d49a91f5c6941e111c4bc 100644 --- a/theories/program_logic/language.v +++ b/theories/program_logic/language.v @@ -160,7 +160,7 @@ Section language. (* This is a family of frequent assumptions for PureExec *) Class IntoVal (e : expr Λ) (v : val Λ) := - into_val : to_val e = Some v. + into_val : of_val v = e. Class AsVal (e : expr Λ) := as_val : is_Some (to_val e). (* There is no instance [IntoVal → AsVal] as often one can solve [AsVal] more diff --git a/theories/program_logic/lifting.v b/theories/program_logic/lifting.v index 0abcd876db8c93bdd07759b757011b3ae327845a..8cc97b7e2d283d145ee44a2ec3551bcadc3f03c0 100644 --- a/theories/program_logic/lifting.v +++ b/theories/program_logic/lifting.v @@ -94,7 +94,7 @@ Proof. iMod (fupd_intro_mask' E2 ∅) as "Hclose"; [set_solver|]. iIntros "!> !>". iMod "Hclose" as "_". iMod "H" as "($ & HΦ & $)". destruct (to_val e2) eqn:?; last by iExFalso. - by iApply wp_value. + iApply wp_value; last done. by apply of_to_val. Qed. Lemma wp_lift_atomic_step {s E Φ} e1 : diff --git a/theories/program_logic/ownp.v b/theories/program_logic/ownp.v index 2601d7ae71d009637dacf187136199780c868ee2..7fa9d42b1570e04e98e009009ea3c1e478a9e0f8 100644 --- a/theories/program_logic/ownp.v +++ b/theories/program_logic/ownp.v @@ -147,7 +147,7 @@ Section lifting. iNext; iIntros (e2 σ2 efs) "% Hσ". iDestruct ("H" $! e2 σ2 efs with "[] [Hσ]") as "[HΦ $]"; [by eauto..|]. destruct (to_val e2) eqn:?; last by iExFalso. - by iMod "Hclose"; iApply wp_value; auto using to_of_val. + iMod "Hclose"; iApply wp_value; last done. by apply of_to_val. Qed. Lemma ownP_lift_atomic_det_step {s E Φ e1} σ1 v2 σ2 efs : diff --git a/theories/program_logic/total_lifting.v b/theories/program_logic/total_lifting.v index b9c62e0714dd6c9e4afcb544c676b8ced2c6c33d..57117988947aae18b0a58e5f8c8aebf0ff1e033f 100644 --- a/theories/program_logic/total_lifting.v +++ b/theories/program_logic/total_lifting.v @@ -54,7 +54,7 @@ Proof. iIntros "!>" (e2 σ2 efs) "%". iMod "Hclose" as "_". iMod ("H" $! e2 σ2 efs with "[#]") as "($ & HΦ & $)"; first by eauto. destruct (to_val e2) eqn:?; last by iExFalso. - by iApply twp_value. + iApply twp_value; last done. by apply of_to_val. Qed. Lemma twp_lift_pure_det_step `{Inhabited (state Λ)} {s E Φ} e1 e2 efs : diff --git a/theories/program_logic/total_weakestpre.v b/theories/program_logic/total_weakestpre.v index d782a858abeac9401ef567ef09e177e370350656..920a8043a2f0b92235cb6f097c1c89fd751eede0 100644 --- a/theories/program_logic/total_weakestpre.v +++ b/theories/program_logic/total_weakestpre.v @@ -342,14 +342,14 @@ Global Instance twp_mono' s E e : Proper (pointwise_relation _ (⊢) ==> (⊢)) (@twp Λ Σ _ s E e). Proof. by intros Φ Φ' ?; apply twp_mono. Qed. -Lemma twp_value s E Φ e v `{!IntoVal e v} : Φ v -∗ WP e @ s; E [{ Φ }]. -Proof. intros; rewrite -(of_to_val e v) //; by apply twp_value'. Qed. +Lemma twp_value s E Φ e v : IntoVal e v → Φ v -∗ WP e @ s; E [{ Φ }]. +Proof. intros <-. by apply twp_value'. Qed. Lemma twp_value_fupd' s E Φ v : (|={E}=> Φ v) -∗ WP of_val v @ s; E [{ Φ }]. Proof. intros. by rewrite -twp_fupd -twp_value'. Qed. -Lemma twp_value_fupd s E Φ e v `{!IntoVal e v} : (|={E}=> Φ v) -∗ WP e @ s; E [{ Φ }]. -Proof. intros. rewrite -twp_fupd -twp_value //. Qed. -Lemma twp_value_inv s E Φ e v `{!IntoVal e v} : WP e @ s; E [{ Φ }] ={E}=∗ Φ v. -Proof. intros; rewrite -(of_to_val e v) //; by apply twp_value_inv'. Qed. +Lemma twp_value_fupd s E Φ e v : IntoVal e v → (|={E}=> Φ v) -∗ WP e @ s; E [{ Φ }]. +Proof. intros ?. rewrite -twp_fupd -twp_value //. Qed. +Lemma twp_value_inv s E Φ e v : IntoVal e v → WP e @ s; E [{ Φ }] ={E}=∗ Φ v. +Proof. intros <-. by apply twp_value_inv'. Qed. Lemma twp_frame_l s E e Φ R : R ∗ WP e @ s; E [{ Φ }] -∗ WP e @ s; E [{ v, R ∗ Φ v }]. Proof. iIntros "[? H]". iApply (twp_strong_mono with "H"); auto with iFrame. Qed. diff --git a/theories/program_logic/weakestpre.v b/theories/program_logic/weakestpre.v index 9e1ff0076adb8f8c952b0fff7f18bc18e2f0cdb3..90ca3af2de55fc66de25ee432814aa42579b04b6 100644 --- a/theories/program_logic/weakestpre.v +++ b/theories/program_logic/weakestpre.v @@ -292,15 +292,15 @@ Global Instance wp_mono' s E e : Proper (pointwise_relation _ (⊢) ==> (⊢)) (@wp Λ Σ _ s E e). Proof. by intros Φ Φ' ?; apply wp_mono. Qed. -Lemma wp_value s E Φ e v `{!IntoVal e v} : Φ v ⊢ WP e @ s; E {{ Φ }}. -Proof. intros; rewrite -(of_to_val e v) //; by apply wp_value'. Qed. +Lemma wp_value s E Φ e v : IntoVal e v → Φ v ⊢ WP e @ s; E {{ Φ }}. +Proof. intros <-. by apply wp_value'. Qed. Lemma wp_value_fupd' s E Φ v : (|={E}=> Φ v) ⊢ WP of_val v @ s; E {{ Φ }}. Proof. intros. by rewrite -wp_fupd -wp_value'. Qed. Lemma wp_value_fupd s E Φ e v `{!IntoVal e v} : (|={E}=> Φ v) ⊢ WP e @ s; E {{ Φ }}. Proof. intros. rewrite -wp_fupd -wp_value //. Qed. -Lemma wp_value_inv s E Φ e v `{!IntoVal e v} : WP e @ s; E {{ Φ }} ={E}=∗ Φ v. -Proof. intros; rewrite -(of_to_val e v) //; by apply wp_value_inv'. Qed. +Lemma wp_value_inv s E Φ e v : IntoVal e v → WP e @ s; E {{ Φ }} ={E}=∗ Φ v. +Proof. intros <-. by apply wp_value_inv'. Qed. Lemma wp_frame_l s E e Φ R : R ∗ WP e @ s; E {{ Φ }} ⊢ WP e @ s; E {{ v, R ∗ Φ v }}. Proof. iIntros "[? H]". iApply (wp_strong_mono with "H"); auto with iFrame. Qed.