diff --git a/README b/README index 3a3dd3dcd71ad25e0b6e9a93cfe0abaeb6d5cb0c..53ccfc1bb8d08577edb61b5243af406f8a5acae1 100644 --- a/README +++ b/README @@ -4,9 +4,15 @@ PREREQUISITES This version is known to compile with: - Coq 8.5 - - Ssreflect >= 3cc0493 + - Ssreflect 1.6 - Autosubst 1.4 +For development, better make sure you have a version of Ssreflect that includes +commit be724937 (no such version has been released so far, you will have to +fetch the development branch yourself). Iris compiles fine even without this +patch, but proof bullets will only be in 'strict' (enforcing) mode with the +fixed version of Ssreflect. + BUILDING INSTRUCTIONS --------------------- diff --git a/heap_lang/lifting.v b/heap_lang/lifting.v index 883d642f611cadf65a992eafa2d8ad5cbb4b415c..71116edeca8be5a019f84c233d02939df0d758c4 100644 --- a/heap_lang/lifting.v +++ b/heap_lang/lifting.v @@ -21,8 +21,9 @@ Lemma wp_alloc_pst E σ e v Q : (ownP σ ★ ▷ (∀ l, ■(σ !! l = None) ∧ ownP (<[l:=v]>σ) -★ Q (LocV l))) ⊑ wp E (Alloc e) Q. Proof. - intros; rewrite -(wp_lift_atomic_step (Alloc e) (λ v' σ' ef, - ∃ l, ef = None ∧ v' = LocV l ∧ σ' = <[l:=v]>σ ∧ σ !! l = None) σ) //; + (* TODO RJ: This works around ssreflect bug #22. *) + intros. set (φ v' σ' ef := ∃ l, ef = @None expr ∧ v' = LocV l ∧ σ' = <[l:=v]>σ ∧ σ !! l = None). + rewrite -(wp_lift_atomic_step (Alloc e) φ σ) // /φ; last by intros; inv_step; eauto 8. apply sep_mono, later_mono; first done. apply forall_intro=>e2; apply forall_intro=>σ2; apply forall_intro=>ef.