41959
|
1 |
(* Title: HOL/Hoare/HeapSyntaxAbort.thy
|
13875
|
2 |
Author: Tobias Nipkow
|
|
3 |
Copyright 2002 TUM
|
|
4 |
*)
|
|
5 |
|
72990
|
6 |
section \<open>Heap syntax (abort)\<close>
|
|
7 |
|
|
8 |
theory HeapSyntaxAbort
|
|
9 |
imports Hoare_Logic_Abort Heap
|
|
10 |
begin
|
13875
|
11 |
|
|
12 |
subsection "Field access and update"
|
|
13 |
|
69597
|
14 |
text\<open>Heap update \<open>p^.h := e\<close> is now guarded against \<^term>\<open>p\<close>
|
|
15 |
being Null. However, \<^term>\<open>p\<close> may still be illegal,
|
13875
|
16 |
e.g. uninitialized or dangling. To guard against that, one needs a
|
|
17 |
more detailed model of the heap where allocated and free addresses are
|
|
18 |
distinguished, e.g. by making the heap a map, or by carrying the set
|
|
19 |
of free addresses around. This is needed anyway as soon as we want to
|
62042
|
20 |
reason about storage allocation/deallocation.\<close>
|
13875
|
21 |
|
|
22 |
syntax
|
35101
|
23 |
"_refupdate" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a ref \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)"
|
13875
|
24 |
("_/'((_ \<rightarrow> _)')" [1000,0] 900)
|
35101
|
25 |
"_fassign" :: "'a ref => id => 'v => 's com"
|
13875
|
26 |
("(2_^._ :=/ _)" [70,1000,65] 61)
|
35101
|
27 |
"_faccess" :: "'a ref => ('a ref \<Rightarrow> 'v) => 'v"
|
13875
|
28 |
("_^._" [65,1000] 65)
|
|
29 |
translations
|
35113
|
30 |
"_refupdate f r v" == "f(CONST addr r := v)"
|
|
31 |
"p^.f := e" => "(p \<noteq> CONST Null) \<rightarrow> (f := _refupdate f p e)"
|
|
32 |
"p^.f" => "f(CONST addr p)"
|
13875
|
33 |
|
|
34 |
|
|
35 |
declare fun_upd_apply[simp del] fun_upd_same[simp] fun_upd_other[simp]
|
|
36 |
|
|
37 |
|
|
38 |
text "An example due to Suzuki:"
|
|
39 |
|
|
40 |
lemma "VARS v n
|
|
41 |
{w = Ref w0 & x = Ref x0 & y = Ref y0 & z = Ref z0 &
|
|
42 |
distinct[w0,x0,y0,z0]}
|
|
43 |
w^.v := (1::int); w^.n := x;
|
|
44 |
x^.v := 2; x^.n := y;
|
|
45 |
y^.v := 3; y^.n := z;
|
|
46 |
z^.v := 4; x^.n := z
|
|
47 |
{w^.n^.n^.v = 4}"
|
|
48 |
by vcg_simp
|
|
49 |
|
|
50 |
end
|