diff --git a/theories/base_logic/lib/gen_heap.v b/theories/base_logic/lib/gen_heap.v index 50b2f7310e43423dcfc658376f171dbd8ac13cf0..9f64f3cb7bae2bba26b95edf411a07a76b2e9790 100644 --- a/theories/base_logic/lib/gen_heap.v +++ b/theories/base_logic/lib/gen_heap.v @@ -143,17 +143,17 @@ Section gen_heap. Qed. Lemma gen_heap_alloc_gen σ σ' : - σ ##ₘ σ' → gen_heap_ctx σ ==∗ gen_heap_ctx (σ ∪ σ') ∗ [∗ map] l ↦ v ∈ σ', l ↦ v. + σ ##ₘ σ' → gen_heap_ctx σ ==∗ gen_heap_ctx (σ' ∪ σ) ∗ [∗ map] l ↦ v ∈ σ', l ↦ v. Proof. revert σ; induction σ' as [| l v σ' Hl IHσ'] using map_ind; iIntros (σ Hσdisj) "Hσ". - - by rewrite right_id big_opM_empty; iFrame. + - by rewrite left_id big_opM_empty; iFrame. - iMod (IHσ' with "Hσ") as "[Hσ m]"; first by eapply map_disjoint_insert_r. rewrite big_opM_insert //; iFrame. assert (σ !! l = None). { eapply map_disjoint_Some_r; first by eauto. rewrite lookup_insert //. } - rewrite -insert_union_r //. + rewrite -insert_union_l //. iMod (gen_heap_alloc with "Hσ") as "[$ $]"; last done. apply lookup_union_None; split; auto. Qed. diff --git a/theories/heap_lang/lang.v b/theories/heap_lang/lang.v index 9edced2d4ae526a2ddcf74f5a9071049affe9658..653bd26bbd4b67626f573ebe7a4051ea5a29ce6b 100644 --- a/theories/heap_lang/lang.v +++ b/theories/heap_lang/lang.v @@ -513,7 +513,14 @@ Proof. Qed. Definition state_init_heap (l : loc) (n : Z) (v : val) (σ : state) : state := - state_upd_heap (λ h, h ∪ heap_array l (replicate (Z.to_nat n) v)) σ. + state_upd_heap (λ h, heap_array l (replicate (Z.to_nat n) v) ∪ h) σ. + +Lemma state_init_heap_singleton l v σ : + state_init_heap l 1 v σ = state_upd_heap <[l:=v]> σ. +Proof. + destruct σ as [h p]. rewrite /state_init_heap /=. f_equiv. + rewrite right_id insert_union_singleton_l. done. +Qed. Inductive head_step : expr → state → list observation → expr → state → list expr → Prop := | RecS f x e σ :