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

Merge branch 'ralf/wp_bind_inv' into 'master'

do not make use of wp_bind_inv unnecessarily

See merge request FP/iris-coq!87
parents eba7e9e4 8c6c9404
No related branches found
No related tags found
No related merge requests found
...@@ -5,15 +5,15 @@ From iris.heap_lang Require Export tactics lifting. ...@@ -5,15 +5,15 @@ From iris.heap_lang Require Export tactics lifting.
Set Default Proof Using "Type". Set Default Proof Using "Type".
Import uPred. Import uPred.
Lemma tac_wp_pure `{heapG Σ} K Δ Δ' E e1 e2 φ Φ : Lemma tac_wp_pure `{heapG Σ} Δ Δ' E e1 e2 φ Φ :
PureExec φ e1 e2 PureExec φ e1 e2
φ φ
IntoLaterNEnvs 1 Δ Δ' IntoLaterNEnvs 1 Δ Δ'
envs_entails Δ' (WP fill K e2 @ E {{ Φ }}) envs_entails Δ' (WP e2 @ E {{ Φ }})
envs_entails Δ (WP fill K e1 @ E {{ Φ }}). envs_entails Δ (WP e1 @ E {{ Φ }}).
Proof. Proof.
rewrite /envs_entails=> ??? HΔ'. rewrite into_laterN_env_sound /=. rewrite /envs_entails=> ??? HΔ'. rewrite into_laterN_env_sound /=.
rewrite -wp_bind HΔ' -wp_pure_step_later //. by rewrite -wp_bind_inv. rewrite HΔ' -wp_pure_step_later //.
Qed. Qed.
Lemma tac_wp_value `{heapG Σ} Δ E Φ e v : Lemma tac_wp_value `{heapG Σ} Δ E Φ e v :
...@@ -26,13 +26,16 @@ Ltac wp_value_head := eapply tac_wp_value; [apply _|lazy beta]. ...@@ -26,13 +26,16 @@ Ltac wp_value_head := eapply tac_wp_value; [apply _|lazy beta].
Tactic Notation "wp_pure" open_constr(efoc) := Tactic Notation "wp_pure" open_constr(efoc) :=
iStartProof; iStartProof;
lazymatch goal with lazymatch goal with
| |- envs_entails _ (wp ?E ?e ?Q) => reshape_expr e ltac:(fun K e' => | |- envs_entails _ (wp ?E ?e ?Q) =>
unify e' efoc; let e := eval simpl in e in
eapply (tac_wp_pure K); reshape_expr e ltac:(fun K e' =>
[simpl; apply _ (* PureExec *) unify e' efoc;
|try fast_done (* The pure condition for PureExec *) eapply (tac_wp_pure _ _ _ (fill K e'));
|apply _ (* IntoLaters *) [apply _ (* PureExec *)
|simpl_subst; try wp_value_head (* new goal *)]) |try fast_done (* The pure condition for PureExec *)
|apply _ (* IntoLaters *)
|simpl_subst; try wp_value_head (* new goal *)
])
|| fail "wp_pure: cannot find" efoc "in" e "or" efoc "is not a reduct" || fail "wp_pure: cannot find" efoc "in" e "or" efoc "is not a reduct"
| _ => fail "wp_pure: not a 'wp'" | _ => fail "wp_pure: not a 'wp'"
end. end.
......
...@@ -205,6 +205,12 @@ Section ectx_language. ...@@ -205,6 +205,12 @@ Section ectx_language.
eexists e2', σ2, efs. by apply head_prim_step. eexists e2', σ2, efs. by apply head_prim_step.
- intros σ1 e2' σ2 efs ? ?%head_reducible_prim_step; eauto. - intros σ1 e2' σ2 efs ? ?%head_reducible_prim_step; eauto.
Qed. Qed.
Global Instance pure_exec_fill K e1 e2 φ :
PureExec φ e1 e2
PureExec φ (fill K e1) (fill K e2).
Proof. apply: pure_exec_ctx. Qed.
End ectx_language. End ectx_language.
Arguments ectx_lang : clear implicits. Arguments ectx_lang : clear implicits.
......
...@@ -140,6 +140,19 @@ Section language. ...@@ -140,6 +140,19 @@ Section language.
PureExec P e1 e2. PureExec P e1 e2.
Proof. intros HPE. split; intros; eapply HPE; eauto. Qed. Proof. intros HPE. split; intros; eapply HPE; eauto. Qed.
(* We do not make this an instance because it is awfully general. *)
Lemma pure_exec_ctx K `{LanguageCtx Λ K} e1 e2 φ :
PureExec φ e1 e2
PureExec φ (K e1) (K e2).
Proof.
intros [Hred Hstep]. split.
- unfold reducible in *. naive_solver eauto using fill_step.
- intros σ1 e2' σ2 efs ? Hpstep.
destruct (fill_step_inv e1 σ1 e2' σ2 efs) as (e2'' & -> & ?); [|exact Hpstep|].
+ destruct (Hred σ1) as (? & ? & ? & ?); eauto using val_stuck.
+ edestruct (Hstep σ1 e2'' σ2 efs) as (-> & -> & ->); auto.
Qed.
(* This is a family of frequent assumptions for PureExec *) (* This is a family of frequent assumptions for PureExec *)
Class IntoVal (e : expr Λ) (v : val Λ) := Class IntoVal (e : expr Λ) (v : val Λ) :=
into_val : to_val e = Some v. into_val : to_val e = Some v.
......
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