From bd9a61e5f85eed0c6702f17cf915ef1ab809999c Mon Sep 17 00:00:00 2001 From: Ralf Jung <jung@mpi-sws.org> Date: Thu, 14 May 2020 12:08:19 +0200 Subject: [PATCH] make inv_heapG part of heapG --- tests/heap_lang.v | 4 ++-- tests/one_shot.v | 2 +- tests/one_shot_once.v | 2 +- theories/base_logic/lib/gen_inv_heap.v | 2 +- theories/heap_lang/adequacy.v | 8 +++++--- theories/heap_lang/lifting.v | 2 +- theories/heap_lang/total_adequacy.v | 5 +++-- 7 files changed, 14 insertions(+), 11 deletions(-) diff --git a/tests/heap_lang.v b/tests/heap_lang.v index 913a9d8c0..9d011cd8f 100644 --- a/tests/heap_lang.v +++ b/tests/heap_lang.v @@ -204,7 +204,7 @@ Section tests. End tests. Section notation_tests. - Context `{!heapG Σ, !inv_heapG Σ}. + Context `{!heapG Σ}. (* Make sure these parse and type-check. *) Lemma inv_mapsto_own_test (l : loc) : ⊢ l ↦_(λ _, True) #5. Abort. @@ -274,4 +274,4 @@ Section error_tests. End error_tests. Lemma heap_e_adequate σ : adequate NotStuck heap_e σ (λ v _, v = #2). -Proof. eapply (heap_adequacy heapΣ)=> ?. by apply heap_e_spec. Qed. +Proof. eapply (heap_adequacy heapΣ)=> ?. iIntros "_". by iApply heap_e_spec. Qed. diff --git a/tests/one_shot.v b/tests/one_shot.v index 756b8f406..349deb199 100644 --- a/tests/one_shot.v +++ b/tests/one_shot.v @@ -124,7 +124,7 @@ Definition clientΣ : gFunctors := #[ heapΣ; one_shotΣ; spawnΣ ]. (** This lemma implicitly shows that these functors are enough to meet all library assumptions. *) Lemma client_adequate σ : adequate NotStuck client σ (λ _ _, True). -Proof. apply (heap_adequacy clientΣ)=> ?. apply client_safe. Qed. +Proof. apply (heap_adequacy clientΣ)=> ?. iIntros "_". iApply client_safe. Qed. (* Since we check the output of the test files, this means our test suite will fail if we ever accidentally add an axiom diff --git a/tests/one_shot_once.v b/tests/one_shot_once.v index f11b93903..b7cfef5df 100644 --- a/tests/one_shot_once.v +++ b/tests/one_shot_once.v @@ -142,4 +142,4 @@ Definition clientΣ : gFunctors := #[ heapΣ; one_shotΣ; spawnΣ ]. (** This lemma implicitly shows that these functors are enough to meet all library assumptions. *) Lemma client_adequate σ : adequate NotStuck client σ (λ _ _, True). -Proof. apply (heap_adequacy clientΣ)=> ?. apply client_safe. Qed. +Proof. apply (heap_adequacy clientΣ)=> ?. iIntros "_". iApply client_safe. Qed. diff --git a/theories/base_logic/lib/gen_inv_heap.v b/theories/base_logic/lib/gen_inv_heap.v index a6fd87779..20311ce97 100644 --- a/theories/base_logic/lib/gen_inv_heap.v +++ b/theories/base_logic/lib/gen_inv_heap.v @@ -113,7 +113,7 @@ Section to_inv_heap. Qed. End to_inv_heap. -Lemma inv_heap_init {L V : Type} `{Countable L, !invG Σ, !gen_heapG L V Σ, !inv_heapPreG L V Σ} E: +Lemma inv_heap_init (L V : Type) `{Countable L, !invG Σ, !gen_heapG L V Σ, !inv_heapPreG L V Σ} E : ⊢ |==> ∃ _ : inv_heapG L V Σ, |={E}=> inv_heap_inv L V. Proof. iMod (own_alloc (â— (to_inv_heap ∅))) as (γ) "Hâ—". diff --git a/theories/heap_lang/adequacy.v b/theories/heap_lang/adequacy.v index d3fc86c70..2b6366bb2 100644 --- a/theories/heap_lang/adequacy.v +++ b/theories/heap_lang/adequacy.v @@ -7,23 +7,25 @@ Set Default Proof Using "Type". Class heapPreG Σ := HeapPreG { heap_preG_iris :> invPreG Σ; heap_preG_heap :> gen_heapPreG loc val Σ; + heap_preG_inv_heap :> inv_heapPreG loc val Σ; heap_preG_proph :> proph_mapPreG proph_id (val * val) Σ; }. Definition heapΣ : gFunctors := - #[invΣ; gen_heapΣ loc val; proph_mapΣ proph_id (val * val)]. + #[invΣ; gen_heapΣ loc val; inv_heapΣ loc val; proph_mapΣ proph_id (val * val)]. Instance subG_heapPreG {Σ} : subG heapΣ Σ → heapPreG Σ. Proof. solve_inG. Qed. Definition heap_adequacy Σ `{!heapPreG Σ} s e σ φ : - (∀ `{!heapG Σ}, ⊢ WP e @ s; ⊤ {{ v, ⌜φ v⌠}}) → + (∀ `{!heapG Σ}, ⊢ inv_heap_inv -∗ WP e @ s; ⊤ {{ v, ⌜φ v⌠}}) → adequate s e σ (λ v _, φ v). Proof. intros Hwp; eapply (wp_adequacy _ _); iIntros (??) "". iMod (gen_heap_init σ.(heap)) as (?) "Hh". + iMod (inv_heap_init loc val) as (?) ">Hi". iMod (proph_map_init κs σ.(used_proph_id)) as (?) "Hp". iModIntro. iExists (λ σ κs, (gen_heap_ctx σ.(heap) ∗ proph_map_ctx κs σ.(used_proph_id))%I), (λ _, True%I). - iFrame. iApply (Hwp (HeapG _ _ _ _)). + iFrame. iApply (Hwp (HeapG _ _ _ _ _)). done. Qed. diff --git a/theories/heap_lang/lifting.v b/theories/heap_lang/lifting.v index 3bbdc425a..584d4a340 100644 --- a/theories/heap_lang/lifting.v +++ b/theories/heap_lang/lifting.v @@ -11,6 +11,7 @@ Set Default Proof Using "Type". Class heapG Σ := HeapG { heapG_invG : invG Σ; heapG_gen_heapG :> gen_heapG loc val Σ; + heapG_inv_heapG :> inv_heapG loc val Σ; heapG_proph_mapG :> proph_mapG proph_id (val * val) Σ; }. @@ -34,7 +35,6 @@ Notation "l ↦□ I" := (inv_mapsto (L:=loc) (V:=val) l I%stdpp%type) (at level 20, format "l ↦□ I") : bi_scope. Notation "l ↦_ I v" := (inv_mapsto_own (L:=loc) (V:=val) l v I%stdpp%type) (at level 20, I at level 9, format "l ↦_ I v") : bi_scope. -Notation inv_heapG := (inv_heapG loc val). Notation inv_heap_inv := (inv_heap_inv loc val). (** The tactic [inv_head_step] performs inversion on hypotheses of the shape diff --git a/theories/heap_lang/total_adequacy.v b/theories/heap_lang/total_adequacy.v index 2194d6df0..346b024d2 100644 --- a/theories/heap_lang/total_adequacy.v +++ b/theories/heap_lang/total_adequacy.v @@ -5,15 +5,16 @@ From iris.heap_lang Require Import proofmode notation. Set Default Proof Using "Type". Definition heap_total Σ `{!heapPreG Σ} s e σ φ : - (∀ `{!heapG Σ}, ⊢ WP e @ s; ⊤ [{ v, ⌜φ v⌠}]) → + (∀ `{!heapG Σ}, ⊢ inv_heap_inv -∗ WP e @ s; ⊤ [{ v, ⌜φ v⌠}]) → sn erased_step ([e], σ). Proof. intros Hwp; eapply (twp_total _ _); iIntros (?) "". iMod (gen_heap_init σ.(heap)) as (?) "Hh". + iMod (inv_heap_init loc val) as (?) ">Hi". iMod (proph_map_init [] σ.(used_proph_id)) as (?) "Hp". iModIntro. iExists (λ σ κs _, (gen_heap_ctx σ.(heap) ∗ proph_map_ctx κs σ.(used_proph_id))%I), (λ _, True%I); iFrame. - iApply (Hwp (HeapG _ _ _ _)). + iApply (Hwp (HeapG _ _ _ _ _)). done. Qed. -- GitLab