diff --git a/theories/heap_lang/lang.v b/theories/heap_lang/lang.v index 6acb7414fb136c4bde0dfd8799dbbec381febc9a..cf19a3c5ff5b830987b8bd7827bf166c25c2b610 100644 --- a/theories/heap_lang/lang.v +++ b/theories/heap_lang/lang.v @@ -420,7 +420,7 @@ Lemma is_closed_of_val X v : is_closed X (of_val v). Proof. apply is_closed_weaken_nil. induction v; simpl; auto. Qed. Lemma is_closed_to_val X e v : to_val e = Some v → is_closed X e. -Proof. intros Hev; rewrite -(of_to_val e v Hev); apply is_closed_of_val. Qed. +Proof. intros <-%of_to_val. apply is_closed_of_val. Qed. Lemma is_closed_subst X e x es : is_closed [] es → is_closed (x :: X) e → is_closed X (subst x es e). diff --git a/theories/heap_lang/lifting.v b/theories/heap_lang/lifting.v index 644becba6f550e5bcfb5e59cd6657cc93254bf7f..d1e410afd1fc7e21987f81d025765a9821f78dc9 100644 --- a/theories/heap_lang/lifting.v +++ b/theories/heap_lang/lifting.v @@ -83,14 +83,12 @@ Qed. Local Ltac solve_exec_safe := intros; subst; do 3 eexists; econstructor; eauto. Local Ltac solve_exec_puredet := simpl; intros; by inv_head_step. Local Ltac solve_pure_exec := - repeat lazymatch goal with - | H: IntoVal ?e _ |- _ => rewrite -(of_to_val e _ into_val); clear H - | H: AsRec _ _ _ _ |- _ => rewrite H; clear H - end; + unfold AsRec, IntoVal, AsVal in *; subst; + repeat match goal with H : is_Some _ |- _ => destruct H as [??] end; apply det_head_step_pure_exec; [ solve_exec_safe | solve_exec_puredet ]. -Global Instance pure_rec f x (erec e1 e2 : expr) (v2 : val) - `{!IntoVal e2 v2, AsRec e1 f x erec, Closed (f :b: x :b: []) erec} : +Global Instance pure_rec f x (erec e1 e2 : expr) + `{!AsVal e2, AsRec e1 f x erec, Closed (f :b: x :b: []) erec} : PureExec True (App e1 e2) (subst' x e2 (subst' f e1 erec)). Proof. solve_pure_exec. Qed. @@ -110,11 +108,11 @@ Global Instance pure_if_false e1 e2 : PureExec True (If (Lit (LitBool false)) e1 e2) e2. Proof. solve_pure_exec. Qed. -Global Instance pure_fst e1 e2 v1 v2 `{!IntoVal e1 v1, !IntoVal e2 v2} : +Global Instance pure_fst e1 e2 v1 `{!IntoVal e1 v1, !AsVal e2} : PureExec True (Fst (Pair e1 e2)) e1. Proof. solve_pure_exec. Qed. -Global Instance pure_snd e1 e2 v1 v2 `{!IntoVal e1 v1, !IntoVal e2 v2} : +Global Instance pure_snd e1 e2 v2 `{!AsVal e1, !IntoVal e2 v2} : PureExec True (Snd (Pair e1 e2)) e2. Proof. solve_pure_exec. Qed. @@ -128,7 +126,7 @@ Proof. solve_pure_exec. Qed. (** Heap *) Lemma wp_alloc E e v : - to_val e = Some v → + IntoVal e v → {{{ True }}} Alloc e @ E {{{ l, RET LitV (LitLoc l); l ↦ v }}}. Proof. iIntros (<-%of_to_val Φ) "_ HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. @@ -149,7 +147,7 @@ Proof. Qed. Lemma wp_store E l v' e v : - to_val e = Some v → + IntoVal e v → {{{ ▷ l ↦ v' }}} Store (Lit (LitLoc l)) e @ E {{{ RET LitV LitUnit; l ↦ v }}}. Proof. iIntros (<-%of_to_val Φ) ">Hl HΦ". @@ -160,12 +158,12 @@ Proof. iModIntro. iSplit=>//. by iApply "HΦ". Qed. -Lemma wp_cas_fail E l q v' e1 v1 e2 v2 : - to_val e1 = Some v1 → to_val e2 = Some v2 → v' ≠v1 → +Lemma wp_cas_fail E l q v' e1 v1 e2 : + IntoVal e1 v1 → AsVal e2 → v' ≠v1 → {{{ ▷ l ↦{q} v' }}} CAS (Lit (LitLoc l)) e1 e2 @ E {{{ RET LitV (LitBool false); l ↦{q} v' }}}. Proof. - iIntros (<-%of_to_val <-%of_to_val ? Φ) ">Hl HΦ". + iIntros (<-%of_to_val [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. @@ -173,7 +171,7 @@ Proof. Qed. Lemma wp_cas_suc E l e1 v1 e2 v2 : - to_val e1 = Some v1 → to_val e2 = Some v2 → + IntoVal e1 v1 → IntoVal e2 v2 → {{{ ▷ l ↦ v1 }}} CAS (Lit (LitLoc l)) e1 e2 @ E {{{ RET LitV (LitBool true); l ↦ v2 }}}. Proof. diff --git a/theories/heap_lang/proofmode.v b/theories/heap_lang/proofmode.v index 5801ae991352c2165e474ce612f38f3b3cd22cae..503b21286bfa620ba037c1bd3eb9987fdf9d2cc3 100644 --- a/theories/heap_lang/proofmode.v +++ b/theories/heap_lang/proofmode.v @@ -108,8 +108,8 @@ Proof. rewrite right_id. by apply later_mono, sep_mono_r, wand_mono. Qed. -Lemma tac_wp_cas_fail Δ Δ' E i K l q v e1 v1 e2 v2 Φ : - IntoVal e1 v1 → IntoVal e2 v2 → +Lemma tac_wp_cas_fail Δ Δ' E i K l q v e1 v1 e2 Φ : + IntoVal e1 v1 → AsVal e2 → IntoLaterNEnvs 1 Δ Δ' → envs_lookup i Δ' = Some (false, l ↦{q} v)%I → v ≠v1 → (Δ' ⊢ WP fill K (Lit (LitBool false)) @ E {{ Φ }}) → diff --git a/theories/heap_lang/tactics.v b/theories/heap_lang/tactics.v index 5602dd515d6a8a5cc5de44760605a95e337b73c8..badd134390805e08c15665997d270e592e59f711 100644 --- a/theories/heap_lang/tactics.v +++ b/theories/heap_lang/tactics.v @@ -210,7 +210,7 @@ Ltac solve_closed := Hint Extern 0 (Closed _ _) => solve_closed : typeclass_instances. Ltac solve_to_val := - rewrite /IntoVal; + rewrite /AsVal /IntoVal; try match goal with | |- context E [language.to_val ?e] => let X := context E [to_val e] in change X @@ -224,6 +224,7 @@ Ltac solve_to_val := apply W.to_val_is_Some, (bool_decide_unpack _); vm_compute; exact I end. Hint Extern 10 (IntoVal _ _) => solve_to_val : typeclass_instances. +Hint Extern 10 (AsVal _) => solve_to_val : typeclass_instances. Ltac solve_atomic := match goal with diff --git a/theories/program_logic/language.v b/theories/program_logic/language.v index 51318c3fde85092e66630ff179ac6ce340ef7e6f..f2a9b9e8bb590682e803094e8d8e6af0f9d71ccb 100644 --- a/theories/program_logic/language.v +++ b/theories/program_logic/language.v @@ -126,4 +126,13 @@ Section language. (* This is a family of frequent assumptions for PureExec *) Class IntoVal (e : expr Λ) (v : val Λ) := into_val : to_val e = Some v. + + Class AsVal (e : expr Λ) := as_val : is_Some (to_val e). + (* There is no instance [IntoVal → AsVal] as often one can solve [AsVal] more + efficiently since no witness has to be computed. *) + Global Instance as_vals_of_val vs : TCForall AsVal (of_val <$> vs). + Proof. + apply TCForall_Forall, Forall_fmap, Forall_true=> v. + rewrite /AsVal /= to_of_val; eauto. + Qed. End language.