From ec75790d0e54e750599df871561b899d861c88fb Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Fri, 24 Nov 2017 12:57:48 +0100 Subject: [PATCH] Bump Iris. --- opam | 2 +- theories/lang/lang.v | 17 ++++++++--------- theories/lang/lifting.v | 12 +----------- theories/lang/proofmode.v | 5 ++--- theories/lang/races.v | 1 - 5 files changed, 12 insertions(+), 25 deletions(-) diff --git a/opam b/opam index 10a0bd29..c2acece8 100644 --- a/opam +++ b/opam @@ -11,5 +11,5 @@ build: [make "-j%{jobs}%"] install: [make "install"] remove: [ "sh" "-c" "rm -rf '%{lib}%/coq/user-contrib/lrust'" ] depends: [ - "coq-iris" { (= "dev.2017-11-14.0") | (= "dev") } + "coq-iris" { (= "dev.2017-11-24.0") | (= "dev") } ] diff --git a/theories/lang/lang.v b/theories/lang/lang.v index 5a750fb2..eebf275c 100644 --- a/theories/lang/lang.v +++ b/theories/lang/lang.v @@ -622,15 +622,14 @@ Canonical Structure valC := leibnizC val. Canonical Structure exprC := leibnizC expr. (** Language *) -Program Instance lrust_ectxi_lang : EctxiLanguage expr val ectx_item state := - {| ectxi_language.of_val := of_val; - ectxi_language.to_val := to_val; - ectxi_language.fill_item := fill_item; - ectxi_language.head_step := head_step |}. -Solve Obligations with eauto using to_of_val, of_to_val, - val_stuck, fill_item_val, fill_item_no_val_inj, head_ctx_step_val. - -Canonical Structure lrust_lang := ectx_lang expr. +Lemma lrust_lang_mixin : EctxiLanguageMixin of_val to_val fill_item head_step. +Proof. + split; apply _ || eauto using to_of_val, of_to_val, + val_stuck, fill_item_val, fill_item_no_val_inj, head_ctx_step_val. +Qed. +Canonical Structure lrust_ectxi_lang := EctxiLanguage lrust_lang_mixin. +Canonical Structure lrust_ectx_lang := EctxLanguageOfEctxi lrust_ectxi_lang. +Canonical Structure lrust_lang := LanguageOfEctx lrust_ectx_lang. (* Lemmas about the language. *) Lemma stuck_irreducible K σ : irreducible (fill K stuck_term) σ. diff --git a/theories/lang/lifting.v b/theories/lang/lifting.v index 7761b6ed..00bd52a0 100644 --- a/theories/lang/lifting.v +++ b/theories/lang/lifting.v @@ -65,16 +65,6 @@ Implicit Types P Q : iProp Σ. Implicit Types e : expr. Implicit Types ef : option expr. -(** Bind. This bundles some arguments that wp_ectx_bind leaves as indices. *) -Lemma wp_bind {E e} K Φ : - WP e @ E {{ v, WP fill K (of_val v) @ E {{ Φ }} }} -∗ WP fill K e @ E {{ Φ }}. -Proof. exact: wp_ectx_bind. Qed. - -Lemma wp_bindi {E e} Ki Φ : - WP e @ E {{ v, WP fill_item Ki (of_val v) @ E {{ Φ }} }} -∗ - WP fill_item Ki e @ E {{ Φ }}. -Proof. exact: weakestpre.wp_bind. Qed. - (** Base axioms for core primitives of the language: Stateless reductions *) Lemma wp_fork E e : {{{ ▷ WP e {{ _, True }} }}} Fork e @ E {{{ RET LitV LitPoison; True }}}. @@ -352,7 +342,7 @@ Proof. - iIntros (Q Ql) "[He Hl] HΦ". assert (App f ((of_val <$> vs) ++ e :: el) = fill_item (AppRCtx vf vs el) e) as -> by rewrite /= (of_to_val f) //. - iApply wp_bindi. iApply (wp_wand with "He"). iIntros (v) "HQ /=". + iApply wp_bind. iApply (wp_wand with "He"). iIntros (v) "HQ /=". rewrite cons_middle (assoc app) -(fmap_app _ _ [v]) (of_to_val f) //. iApply (IH _ _ with "Hl"). iIntros "* Hvl". rewrite -assoc. iApply ("HΦ" $! (v:::vl)). iFrame. diff --git a/theories/lang/proofmode.v b/theories/lang/proofmode.v index e757c6df..2cd7a8a1 100644 --- a/theories/lang/proofmode.v +++ b/theories/lang/proofmode.v @@ -22,8 +22,7 @@ Lemma tac_wp_pure `{lrustG Σ} K Δ Δ' E e1 e2 φ Φ : envs_entails Δ (WP fill K e1 @ E {{ Φ }}). Proof. rewrite /envs_entails=> ??? HΔ'. rewrite into_laterN_env_sound /=. - rewrite -lifting.wp_bind HΔ' -wp_pure_step_later //. - by rewrite -ectx_lifting.wp_ectx_bind_inv. + rewrite -wp_bind HΔ' -wp_pure_step_later //. by rewrite -wp_bind_inv. Qed. Tactic Notation "wp_pure" open_constr(efoc) := @@ -72,7 +71,7 @@ Tactic Notation "wp_case" := wp_pure (Case _ _); try wp_value_head. Lemma tac_wp_bind `{lrustG Σ} K Δ E Φ e : envs_entails Δ (WP e @ E {{ v, WP fill K (of_val v) @ E {{ Φ }} }})%I → envs_entails Δ (WP fill K e @ E {{ Φ }}). -Proof. rewrite /envs_entails=> ->. by apply wp_bind. Qed. +Proof. rewrite /envs_entails=> ->. apply: wp_bind. Qed. Ltac wp_bind_core K := lazymatch eval hnf in K with diff --git a/theories/lang/races.v b/theories/lang/races.v index bc7399c1..0d8c1043 100644 --- a/theories/lang/races.v +++ b/theories/lang/races.v @@ -192,7 +192,6 @@ Theorem safe_nonracing el σ : e' ∈ el' → to_val e' = None → reducible e' σ') → nonracing_threadpool el σ. Proof. - change to_val with ectxi_language.to_val. intros Hsafe l a1 a2 (t1&?&t2&?&t3&->&(K1&e1&Ha1&->)&(K2&e2&Ha2&->)). assert (to_val e1 = None). by destruct Ha1. -- GitLab