13857
|
1 |
theory Separation = HoareAbort:
|
|
2 |
|
|
3 |
types heap = "(nat \<Rightarrow> nat option)"
|
|
4 |
|
|
5 |
|
|
6 |
text{* The semantic definition of a few connectives: *}
|
|
7 |
|
|
8 |
constdefs
|
|
9 |
R:: "heap \<Rightarrow> heap \<Rightarrow> heap \<Rightarrow> bool"
|
|
10 |
"R h h1 h2 == dom h1 \<inter> dom h2 = {} \<and> h = h1 ++ h2"
|
|
11 |
|
|
12 |
star:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)"
|
|
13 |
"star P Q == \<lambda>h. \<exists>h1 h2. R h h1 h2 \<and> P h1 \<and> Q h2"
|
|
14 |
|
|
15 |
singl:: "heap \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
|
|
16 |
"singl h x y == dom h = {x} & h x = Some y"
|
|
17 |
|
|
18 |
lemma "VARS x y z w h
|
|
19 |
{star (%h. singl h x y) (%h. singl h z w) h}
|
|
20 |
SKIP
|
|
21 |
{x \<noteq> z}"
|
|
22 |
apply vcg
|
|
23 |
apply(auto simp:star_def R_def singl_def)
|
|
24 |
done
|
|
25 |
|
|
26 |
text{* To suppress the heap parameter of the connectives, we assume it
|
|
27 |
is always called H and add/remove it upon parsing/printing. Thus
|
|
28 |
every pointer program needs to have a program variable H, and
|
|
29 |
assertions should not contain any locally bound Hs - otherwise they
|
|
30 |
may bind the implicit H. *}
|
|
31 |
|
|
32 |
text{* Nice input syntax: *}
|
|
33 |
|
|
34 |
syntax
|
|
35 |
"@singl" :: "nat \<Rightarrow> nat \<Rightarrow> bool" ("[_ \<mapsto> _]")
|
|
36 |
"@star" :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "**" 60)
|
|
37 |
|
|
38 |
ML{*
|
|
39 |
fun singl_tr [p,q] = Syntax.const "singl" $ Syntax.free "H" $ p $ q
|
|
40 |
| singl_tr ts = raise TERM ("singl_tr", ts);
|
|
41 |
fun star_tr [P,Q] = Syntax.const "star" $
|
|
42 |
absfree("H",dummyT,P) $ absfree("H",dummyT,Q) $ Syntax.free "H"
|
|
43 |
| star_tr ts = raise TERM ("star_tr", ts);
|
|
44 |
*}
|
|
45 |
|
|
46 |
parse_translation {* [("@singl", singl_tr),("@star", star_tr)] *}
|
|
47 |
|
|
48 |
lemma "VARS H x y z w
|
|
49 |
{[x\<mapsto>y] ** [z\<mapsto>w]}
|
|
50 |
SKIP
|
|
51 |
{x \<noteq> z}"
|
|
52 |
apply vcg
|
|
53 |
apply(auto simp:star_def R_def singl_def)
|
|
54 |
done
|
|
55 |
|
|
56 |
text{* Nice output syntax: *}
|
|
57 |
|
|
58 |
ML{*
|
|
59 |
fun singl_tr' [_,p,q] = Syntax.const "@singl" $ p $ q
|
|
60 |
fun star_tr' [Abs(_,_,P),Abs(_,_,Q),_] = Syntax.const "@star" $ P $ Q
|
|
61 |
*}
|
|
62 |
|
|
63 |
print_translation {* [("singl", singl_tr'),("star", star_tr')] *}
|
|
64 |
|
|
65 |
lemma "VARS H x y z w
|
|
66 |
{[x\<mapsto>y] ** [z\<mapsto>w]}
|
13867
|
67 |
y := w
|
13857
|
68 |
{x \<noteq> z}"
|
|
69 |
apply vcg
|
|
70 |
apply(auto simp:star_def R_def singl_def)
|
|
71 |
done
|
|
72 |
|
|
73 |
|
|
74 |
consts llist :: "(heap * nat)set"
|
|
75 |
inductive llist
|
|
76 |
intros
|
|
77 |
empty: "(%n. None, 0) : llist"
|
|
78 |
cons: "\<lbrakk> R h h1 h2; pto h1 p q; (h2,q):llist \<rbrakk> \<Longrightarrow> (h,p):llist"
|
|
79 |
|
|
80 |
lemma "VARS p q h
|
|
81 |
{(h,p) : llist}
|
|
82 |
h := h(q \<mapsto> p)
|
|
83 |
{(h,q) : llist}"
|
|
84 |
apply vcg
|
|
85 |
apply(rule_tac "h1.0" = "%n. if n=q then Some p else None" in llist.cons)
|
|
86 |
prefer 3 apply assumption
|
|
87 |
prefer 2 apply(simp add:singl_def dom_def)
|
|
88 |
apply(simp add:R_def dom_def)
|
|
89 |
|
|
90 |
|