Skip to content
Snippets Groups Projects
Commit d903a458 authored by Ralf Jung's avatar Ralf Jung
Browse files

show that allocating a single ref is like normal map insertion

parent e9a19f54
No related branches found
No related tags found
No related merge requests found
......@@ -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.
......
......@@ -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 σ :
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment