author | haftmann |
Tue, 23 Feb 2010 10:11:15 +0100 | |
changeset 35316 | 870dfea4f9c0 |
parent 35113 | 1a0c129bb2e0 |
child 35416 | d8d7d1b785af |
permissions | -rw-r--r-- |
13903 | 1 |
(* Title: HOL/Hoare/Separation.thy |
2 |
Author: Tobias Nipkow |
|
3 |
Copyright 2003 TUM |
|
4 |
||
5 |
A first attempt at a nice syntactic embedding of separation logic. |
|
14074 | 6 |
Already builds on the theory for list abstractions. |
7 |
||
8 |
If we suppress the H parameter for "List", we have to hardwired this |
|
9 |
into parser and pretty printer, which is not very modular. |
|
10 |
Alternative: some syntax like <P> which stands for P H. No more |
|
11 |
compact, but avoids the funny H. |
|
12 |
||
13903 | 13 |
*) |
14 |
||
35316
870dfea4f9c0
dropped axclass; dropped Id; session theory Hoare.thy
haftmann
parents:
35113
diff
changeset
|
15 |
theory Separation imports Hoare_Logic_Abort SepLogHeap begin |
13857 | 16 |
|
17 |
text{* The semantic definition of a few connectives: *} |
|
18 |
||
19 |
constdefs |
|
13875 | 20 |
ortho:: "heap \<Rightarrow> heap \<Rightarrow> bool" (infix "\<bottom>" 55) |
21 |
"h1 \<bottom> h2 == dom h1 \<inter> dom h2 = {}" |
|
13857 | 22 |
|
13875 | 23 |
is_empty :: "heap \<Rightarrow> bool" |
24 |
"is_empty h == h = empty" |
|
13857 | 25 |
|
26 |
singl:: "heap \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" |
|
27 |
"singl h x y == dom h = {x} & h x = Some y" |
|
28 |
||
13875 | 29 |
star:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)" |
30 |
"star P Q == \<lambda>h. \<exists>h1 h2. h = h1++h2 \<and> h1 \<bottom> h2 \<and> P h1 \<and> Q h2" |
|
31 |
||
32 |
wand:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)" |
|
33 |
"wand P Q == \<lambda>h. \<forall>h'. h' \<bottom> h \<and> P h' \<longrightarrow> Q(h++h')" |
|
34 |
||
13903 | 35 |
text{*This is what assertions look like without any syntactic sugar: *} |
36 |
||
13857 | 37 |
lemma "VARS x y z w h |
38 |
{star (%h. singl h x y) (%h. singl h z w) h} |
|
39 |
SKIP |
|
40 |
{x \<noteq> z}" |
|
41 |
apply vcg |
|
13875 | 42 |
apply(auto simp:star_def ortho_def singl_def) |
13857 | 43 |
done |
44 |
||
13903 | 45 |
text{* Now we add nice input syntax. To suppress the heap parameter |
46 |
of the connectives, we assume it is always called H and add/remove it |
|
47 |
upon parsing/printing. Thus every pointer program needs to have a |
|
48 |
program variable H, and assertions should not contain any locally |
|
49 |
bound Hs - otherwise they may bind the implicit H. *} |
|
13857 | 50 |
|
51 |
syntax |
|
35101 | 52 |
"_emp" :: "bool" ("emp") |
53 |
"_singl" :: "nat \<Rightarrow> nat \<Rightarrow> bool" ("[_ \<mapsto> _]") |
|
54 |
"_star" :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "**" 60) |
|
55 |
"_wand" :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "-*" 60) |
|
13857 | 56 |
|
17781 | 57 |
(* FIXME does not handle "_idtdummy" *) |
13857 | 58 |
ML{* |
13875 | 59 |
(* free_tr takes care of free vars in the scope of sep. logic connectives: |
60 |
they are implicitly applied to the heap *) |
|
61 |
fun free_tr(t as Free _) = t $ Syntax.free "H" |
|
14074 | 62 |
(* |
63 |
| free_tr((list as Free("List",_))$ p $ ps) = list $ Syntax.free "H" $ p $ ps |
|
64 |
*) |
|
13875 | 65 |
| free_tr t = t |
66 |
||
35113 | 67 |
fun emp_tr [] = Syntax.const @{const_syntax is_empty} $ Syntax.free "H" |
13875 | 68 |
| emp_tr ts = raise TERM ("emp_tr", ts); |
35113 | 69 |
fun singl_tr [p, q] = Syntax.const @{const_syntax singl} $ Syntax.free "H" $ p $ q |
13857 | 70 |
| singl_tr ts = raise TERM ("singl_tr", ts); |
35113 | 71 |
fun star_tr [P,Q] = Syntax.const @{const_syntax star} $ |
72 |
absfree ("H", dummyT, free_tr P) $ absfree ("H", dummyT, free_tr Q) $ |
|
13875 | 73 |
Syntax.free "H" |
74 |
| star_tr ts = raise TERM ("star_tr", ts); |
|
35113 | 75 |
fun wand_tr [P, Q] = Syntax.const @{const_syntax wand} $ |
76 |
absfree ("H", dummyT, P) $ absfree ("H", dummyT, Q) $ Syntax.free "H" |
|
13875 | 77 |
| wand_tr ts = raise TERM ("wand_tr", ts); |
13857 | 78 |
*} |
79 |
||
35113 | 80 |
parse_translation {* |
81 |
[(@{syntax_const "_emp"}, emp_tr), |
|
82 |
(@{syntax_const "_singl"}, singl_tr), |
|
83 |
(@{syntax_const "_star"}, star_tr), |
|
84 |
(@{syntax_const "_wand"}, wand_tr)] |
|
85 |
*} |
|
13857 | 86 |
|
13903 | 87 |
text{* Now it looks much better: *} |
88 |
||
13857 | 89 |
lemma "VARS H x y z w |
90 |
{[x\<mapsto>y] ** [z\<mapsto>w]} |
|
91 |
SKIP |
|
92 |
{x \<noteq> z}" |
|
93 |
apply vcg |
|
13875 | 94 |
apply(auto simp:star_def ortho_def singl_def) |
95 |
done |
|
96 |
||
97 |
lemma "VARS H x y z w |
|
98 |
{emp ** emp} |
|
99 |
SKIP |
|
100 |
{emp}" |
|
101 |
apply vcg |
|
102 |
apply(auto simp:star_def ortho_def is_empty_def) |
|
13857 | 103 |
done |
104 |
||
13903 | 105 |
text{* But the output is still unreadable. Thus we also strip the heap |
106 |
parameters upon output: *} |
|
107 |
||
35113 | 108 |
ML {* |
109 |
local |
|
13857 | 110 |
|
13903 | 111 |
fun strip (Abs(_,_,(t as Const("_free",_) $ Free _) $ Bound 0)) = t |
112 |
| strip (Abs(_,_,(t as Free _) $ Bound 0)) = t |
|
14074 | 113 |
(* |
114 |
| strip (Abs(_,_,((list as Const("List",_))$ Bound 0 $ p $ ps))) = list$p$ps |
|
115 |
*) |
|
13903 | 116 |
| strip (Abs(_,_,(t as Const("_var",_) $ Var _) $ Bound 0)) = t |
13875 | 117 |
| strip (Abs(_,_,P)) = P |
35113 | 118 |
| strip (Const(@{const_syntax is_empty},_)) = Syntax.const @{syntax_const "_emp"} |
13875 | 119 |
| strip t = t; |
35113 | 120 |
|
13875 | 121 |
in |
35113 | 122 |
|
123 |
fun is_empty_tr' [_] = Syntax.const @{syntax_const "_emp"} |
|
124 |
fun singl_tr' [_,p,q] = Syntax.const @{syntax_const "_singl"} $ p $ q |
|
125 |
fun star_tr' [P,Q,_] = Syntax.const @{syntax_const "_star"} $ strip P $ strip Q |
|
126 |
fun wand_tr' [P,Q,_] = Syntax.const @{syntax_const "_wand"} $ strip P $ strip Q |
|
127 |
||
13875 | 128 |
end |
13857 | 129 |
*} |
130 |
||
35113 | 131 |
print_translation {* |
132 |
[(@{const_syntax is_empty}, is_empty_tr'), |
|
133 |
(@{const_syntax singl}, singl_tr'), |
|
134 |
(@{const_syntax star}, star_tr'), |
|
135 |
(@{const_syntax wand}, wand_tr')] |
|
136 |
*} |
|
13903 | 137 |
|
138 |
text{* Now the intermediate proof states are also readable: *} |
|
13857 | 139 |
|
140 |
lemma "VARS H x y z w |
|
141 |
{[x\<mapsto>y] ** [z\<mapsto>w]} |
|
13867 | 142 |
y := w |
13857 | 143 |
{x \<noteq> z}" |
144 |
apply vcg |
|
13875 | 145 |
apply(auto simp:star_def ortho_def singl_def) |
146 |
done |
|
147 |
||
148 |
lemma "VARS H x y z w |
|
149 |
{emp ** emp} |
|
150 |
SKIP |
|
151 |
{emp}" |
|
152 |
apply vcg |
|
153 |
apply(auto simp:star_def ortho_def is_empty_def) |
|
154 |
done |
|
155 |
||
13903 | 156 |
text{* So far we have unfolded the separation logic connectives in |
157 |
proofs. Here comes a simple example of a program proof that uses a law |
|
158 |
of separation logic instead. *} |
|
159 |
||
13875 | 160 |
(* a law of separation logic *) |
161 |
lemma star_comm: "P ** Q = Q ** P" |
|
18447 | 162 |
by(auto simp add:star_def ortho_def dest: map_add_comm) |
13875 | 163 |
|
164 |
lemma "VARS H x y z w |
|
165 |
{P ** Q} |
|
166 |
SKIP |
|
167 |
{Q ** P}" |
|
168 |
apply vcg |
|
169 |
apply(simp add: star_comm) |
|
170 |
done |
|
171 |
||
14074 | 172 |
|
173 |
lemma "VARS H |
|
174 |
{p\<noteq>0 \<and> [p \<mapsto> x] ** List H q qs} |
|
175 |
H := H(p \<mapsto> q) |
|
176 |
{List H p (p#qs)}" |
|
177 |
apply vcg |
|
178 |
apply(simp add: star_def ortho_def singl_def) |
|
179 |
apply clarify |
|
180 |
apply(subgoal_tac "p \<notin> set qs") |
|
181 |
prefer 2 |
|
182 |
apply(blast dest:list_in_heap) |
|
183 |
apply simp |
|
184 |
done |
|
185 |
||
186 |
lemma "VARS H p q r |
|
187 |
{List H p Ps ** List H q Qs} |
|
188 |
WHILE p \<noteq> 0 |
|
189 |
INV {\<exists>ps qs. (List H p ps ** List H q qs) \<and> rev ps @ qs = rev Ps @ Qs} |
|
190 |
DO r := p; p := the(H p); H := H(r \<mapsto> q); q := r OD |
|
191 |
{List H q (rev Ps @ Qs)}" |
|
192 |
apply vcg |
|
193 |
apply(simp_all add: star_def ortho_def singl_def) |
|
194 |
||
195 |
apply fastsimp |
|
196 |
||
197 |
apply (clarsimp simp add:List_non_null) |
|
198 |
apply(rename_tac ps') |
|
199 |
apply(rule_tac x = ps' in exI) |
|
200 |
apply(rule_tac x = "p#qs" in exI) |
|
201 |
apply simp |
|
202 |
apply(rule_tac x = "h1(p:=None)" in exI) |
|
203 |
apply(rule_tac x = "h2(p\<mapsto>q)" in exI) |
|
204 |
apply simp |
|
205 |
apply(rule conjI) |
|
206 |
apply(rule ext) |
|
207 |
apply(simp add:map_add_def split:option.split) |
|
208 |
apply(rule conjI) |
|
209 |
apply blast |
|
210 |
apply(simp add:map_add_def split:option.split) |
|
211 |
apply(rule conjI) |
|
212 |
apply(subgoal_tac "p \<notin> set qs") |
|
213 |
prefer 2 |
|
214 |
apply(blast dest:list_in_heap) |
|
215 |
apply(simp) |
|
216 |
apply fast |
|
217 |
||
218 |
apply(fastsimp) |
|
219 |
done |
|
220 |
||
13875 | 221 |
end |