diff --git a/opam b/opam index c72a62183d425b32aa528abe3388f7d15940a1c5..ca3257db3cf4f28c634d16c948512bdaae8513b0 100644 --- a/opam +++ b/opam @@ -11,5 +11,5 @@ install: [make "install"] remove: ["rm" "-rf" "%{lib}%/coq/user-contrib/iris"] depends: [ "coq" { (>= "8.7.0" & < "8.9~") | (= "dev") } - "coq-stdpp" { (= "dev.2018-05-23.0.a8f65af5") | (= "dev") } + "coq-stdpp" { (= "dev.2018-05-24.0.de797b31") | (= "dev") } ] diff --git a/theories/program_logic/ectx_lifting.v b/theories/program_logic/ectx_lifting.v index 1105a2e11104e202ca16c4760167de1627328f83..e0ec8c01d5c1037e3f07498161beb5a7c7f8cb61 100644 --- a/theories/program_logic/ectx_lifting.v +++ b/theories/program_logic/ectx_lifting.v @@ -68,7 +68,7 @@ Lemma wp_lift_atomic_head_step {s E Φ} e1 : ⌜head_reducible e1 σ1⌠∗ ▷ ∀ e2 σ2 efs, ⌜head_step e1 σ1 e2 σ2 efs⌠={E}=∗ state_interp σ2 ∗ - default False (to_val e2) Φ ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }}) + from_option Φ False (to_val e2) ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }}) ⊢ WP e1 @ s; E {{ Φ }}. Proof. iIntros (?) "H". iApply wp_lift_atomic_step; eauto. @@ -82,7 +82,7 @@ Lemma wp_lift_atomic_head_step_no_fork {s E Φ} e1 : (∀ σ1, state_interp σ1 ={E}=∗ ⌜head_reducible e1 σ1⌠∗ ▷ ∀ e2 σ2 efs, ⌜head_step e1 σ1 e2 σ2 efs⌠={E}=∗ - ⌜efs = []⌠∗ state_interp σ2 ∗ default False (to_val e2) Φ) + ⌜efs = []⌠∗ state_interp σ2 ∗ from_option Φ False (to_val e2)) ⊢ WP e1 @ s; E {{ Φ }}. Proof. iIntros (?) "H". iApply wp_lift_atomic_head_step; eauto. diff --git a/theories/program_logic/lifting.v b/theories/program_logic/lifting.v index f1299cee628ffbf79759ea943f47c145bc1576d4..38f6e3a803cdef266ff9cb8d6f849ca2b1b546ab 100644 --- a/theories/program_logic/lifting.v +++ b/theories/program_logic/lifting.v @@ -72,7 +72,7 @@ Lemma wp_lift_atomic_step {s E Φ} e1 : ⌜if s is NotStuck then reducible e1 σ1 else True⌠∗ ▷ ∀ e2 σ2 efs, ⌜prim_step e1 σ1 e2 σ2 efs⌠={E}=∗ state_interp σ2 ∗ - default False (to_val e2) Φ ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }}) + from_option Φ False (to_val e2) ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }}) ⊢ WP e1 @ s; E {{ Φ }}. Proof. iIntros (?) "H". iApply (wp_lift_step s E _ e1)=>//; iIntros (σ1) "Hσ1". diff --git a/theories/program_logic/ownp.v b/theories/program_logic/ownp.v index abd3aa7da9fcff5683e152cf3facf8b9f26a2530..3927d60e4cfb8ece04289ecad857fd1c1594d369 100644 --- a/theories/program_logic/ownp.v +++ b/theories/program_logic/ownp.v @@ -138,7 +138,7 @@ Section lifting. Lemma ownP_lift_atomic_step {s E Φ} e1 σ1 : (if s is NotStuck then reducible e1 σ1 else to_val e1 = None) → (▷ ownP σ1 ∗ ▷ ∀ e2 σ2 efs, ⌜prim_step e1 σ1 e2 σ2 efs⌠-∗ ownP σ2 -∗ - default False (to_val e2) Φ ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }}) + from_option Φ False (to_val e2) ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }}) ⊢ WP e1 @ s; E {{ Φ }}. Proof. iIntros (?) "[Hσ H]"; iApply ownP_lift_step. @@ -240,7 +240,7 @@ Section ectx_lifting. head_reducible e1 σ1 → ▷ ownP σ1 ∗ ▷ (∀ e2 σ2 efs, ⌜head_step e1 σ1 e2 σ2 efs⌠-∗ ownP σ2 -∗ - default False (to_val e2) Φ ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }}) + from_option Φ False (to_val e2) ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }}) ⊢ WP e1 @ s; E {{ Φ }}. Proof. iIntros (?) "[? H]". iApply ownP_lift_atomic_step; eauto. diff --git a/theories/program_logic/total_ectx_lifting.v b/theories/program_logic/total_ectx_lifting.v index a78740646b2970d56b138c81114236f3251a4c26..e5d1af45e6fc89383ef3199de3dda66078d0532e 100644 --- a/theories/program_logic/total_ectx_lifting.v +++ b/theories/program_logic/total_ectx_lifting.v @@ -43,7 +43,7 @@ Lemma twp_lift_atomic_head_step {s E Φ} e1 : ⌜head_reducible e1 σ1⌠∗ ∀ e2 σ2 efs, ⌜head_step e1 σ1 e2 σ2 efs⌠={E}=∗ state_interp σ2 ∗ - default False (to_val e2) Φ ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ [{ _, True }]) + from_option Φ False (to_val e2) ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ [{ _, True }]) ⊢ WP e1 @ s; E [{ Φ }]. Proof. iIntros (?) "H". iApply twp_lift_atomic_step; eauto. @@ -56,7 +56,7 @@ Lemma twp_lift_atomic_head_step_no_fork {s E Φ} e1 : (∀ σ1, state_interp σ1 ={E}=∗ ⌜head_reducible e1 σ1⌠∗ ∀ e2 σ2 efs, ⌜head_step e1 σ1 e2 σ2 efs⌠={E}=∗ - ⌜efs = []⌠∗ state_interp σ2 ∗ default False (to_val e2) Φ) + ⌜efs = []⌠∗ state_interp σ2 ∗ from_option Φ False (to_val e2)) ⊢ WP e1 @ s; E [{ Φ }]. Proof. iIntros (?) "H". iApply twp_lift_atomic_head_step; eauto. diff --git a/theories/program_logic/total_lifting.v b/theories/program_logic/total_lifting.v index fe690efd28a84f22cad9044be7714b72467debb2..0c64555d0e5ed6b6dbcc8c5ad2612bfc3de3994d 100644 --- a/theories/program_logic/total_lifting.v +++ b/theories/program_logic/total_lifting.v @@ -45,7 +45,7 @@ Lemma twp_lift_atomic_step {s E Φ} e1 : ⌜if s is NotStuck then reducible e1 σ1 else True⌠∗ ∀ e2 σ2 efs, ⌜prim_step e1 σ1 e2 σ2 efs⌠={E}=∗ state_interp σ2 ∗ - default False (to_val e2) Φ ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ [{ _, True }]) + from_option Φ False (to_val e2) ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ [{ _, True }]) ⊢ WP e1 @ s; E [{ Φ }]. Proof. iIntros (?) "H". iApply (twp_lift_step _ E _ e1)=>//; iIntros (σ1) "Hσ1".