author | wenzelm |
Sun, 31 Dec 2023 19:24:37 +0100 | |
changeset 79409 | e1895596e1b9 |
parent 72990 | db8f94656024 |
child 80914 | d97fdabd9e2b |
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. |
|
13903 | 12 |
*) |
13 |
||
72990 | 14 |
section \<open>Separation logic\<close> |
15 |
||
16 |
theory Separation |
|
17 |
imports Hoare_Logic_Abort SepLogHeap |
|
18 |
begin |
|
13857 | 19 |
|
62042 | 20 |
text\<open>The semantic definition of a few connectives:\<close> |
13857 | 21 |
|
38353 | 22 |
definition ortho :: "heap \<Rightarrow> heap \<Rightarrow> bool" (infix "\<bottom>" 55) |
23 |
where "h1 \<bottom> h2 \<longleftrightarrow> dom h1 \<inter> dom h2 = {}" |
|
13857 | 24 |
|
38353 | 25 |
definition is_empty :: "heap \<Rightarrow> bool" |
68451 | 26 |
where "is_empty h \<longleftrightarrow> h = Map.empty" |
13857 | 27 |
|
38353 | 28 |
definition singl:: "heap \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" |
29 |
where "singl h x y \<longleftrightarrow> dom h = {x} & h x = Some y" |
|
13857 | 30 |
|
38353 | 31 |
definition star:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)" |
32 |
where "star P Q = (\<lambda>h. \<exists>h1 h2. h = h1++h2 \<and> h1 \<bottom> h2 \<and> P h1 \<and> Q h2)" |
|
13875 | 33 |
|
38353 | 34 |
definition wand:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)" |
35 |
where "wand P Q = (\<lambda>h. \<forall>h'. h' \<bottom> h \<and> P h' \<longrightarrow> Q(h++h'))" |
|
13875 | 36 |
|
62042 | 37 |
text\<open>This is what assertions look like without any syntactic sugar:\<close> |
13903 | 38 |
|
13857 | 39 |
lemma "VARS x y z w h |
40 |
{star (%h. singl h x y) (%h. singl h z w) h} |
|
41 |
SKIP |
|
42 |
{x \<noteq> z}" |
|
43 |
apply vcg |
|
13875 | 44 |
apply(auto simp:star_def ortho_def singl_def) |
13857 | 45 |
done |
46 |
||
62042 | 47 |
text\<open>Now we add nice input syntax. To suppress the heap parameter |
13903 | 48 |
of the connectives, we assume it is always called H and add/remove it |
49 |
upon parsing/printing. Thus every pointer program needs to have a |
|
50 |
program variable H, and assertions should not contain any locally |
|
62042 | 51 |
bound Hs - otherwise they may bind the implicit H.\<close> |
13857 | 52 |
|
53 |
syntax |
|
35101 | 54 |
"_emp" :: "bool" ("emp") |
55 |
"_singl" :: "nat \<Rightarrow> nat \<Rightarrow> bool" ("[_ \<mapsto> _]") |
|
56 |
"_star" :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "**" 60) |
|
57 |
"_wand" :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "-*" 60) |
|
13857 | 58 |
|
17781 | 59 |
(* FIXME does not handle "_idtdummy" *) |
67444 | 60 |
ML \<open> |
61 |
\<comment> \<open>\<open>free_tr\<close> takes care of free vars in the scope of separation logic connectives: |
|
62 |
they are implicitly applied to the heap\<close> |
|
13875 | 63 |
fun free_tr(t as Free _) = t $ Syntax.free "H" |
67444 | 64 |
\<^cancel>\<open>| free_tr((list as Free("List",_))$ p $ ps) = list $ Syntax.free "H" $ p $ ps\<close> |
13875 | 65 |
| free_tr t = t |
66 |
||
69597 | 67 |
fun emp_tr [] = Syntax.const \<^const_syntax>\<open>is_empty\<close> $ Syntax.free "H" |
13875 | 68 |
| emp_tr ts = raise TERM ("emp_tr", ts); |
69597 | 69 |
fun singl_tr [p, q] = Syntax.const \<^const_syntax>\<open>singl\<close> $ Syntax.free "H" $ p $ q |
13857 | 70 |
| singl_tr ts = raise TERM ("singl_tr", ts); |
69597 | 71 |
fun star_tr [P,Q] = Syntax.const \<^const_syntax>\<open>star\<close> $ |
44241 | 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); |
|
69597 | 75 |
fun wand_tr [P, Q] = Syntax.const \<^const_syntax>\<open>wand\<close> $ |
44241 | 76 |
absfree ("H", dummyT) P $ absfree ("H", dummyT) Q $ Syntax.free "H" |
13875 | 77 |
| wand_tr ts = raise TERM ("wand_tr", ts); |
62042 | 78 |
\<close> |
13857 | 79 |
|
62042 | 80 |
parse_translation \<open> |
69597 | 81 |
[(\<^syntax_const>\<open>_emp\<close>, K emp_tr), |
82 |
(\<^syntax_const>\<open>_singl\<close>, K singl_tr), |
|
83 |
(\<^syntax_const>\<open>_star\<close>, K star_tr), |
|
84 |
(\<^syntax_const>\<open>_wand\<close>, K wand_tr)] |
|
62042 | 85 |
\<close> |
13857 | 86 |
|
62042 | 87 |
text\<open>Now it looks much better:\<close> |
13903 | 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 |
||
62042 | 105 |
text\<open>But the output is still unreadable. Thus we also strip the heap |
106 |
parameters upon output:\<close> |
|
13903 | 107 |
|
62042 | 108 |
ML \<open> |
35113 | 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 |
|
67444 | 113 |
\<^cancel>\<open>| strip (Abs(_,_,((list as Const("List",_))$ Bound 0 $ p $ ps))) = list$p$ps\<close> |
13903 | 114 |
| strip (Abs(_,_,(t as Const("_var",_) $ Var _) $ Bound 0)) = t |
13875 | 115 |
| strip (Abs(_,_,P)) = P |
69597 | 116 |
| strip (Const(\<^const_syntax>\<open>is_empty\<close>,_)) = Syntax.const \<^syntax_const>\<open>_emp\<close> |
13875 | 117 |
| strip t = t; |
35113 | 118 |
|
13875 | 119 |
in |
35113 | 120 |
|
69597 | 121 |
fun is_empty_tr' [_] = Syntax.const \<^syntax_const>\<open>_emp\<close> |
122 |
fun singl_tr' [_,p,q] = Syntax.const \<^syntax_const>\<open>_singl\<close> $ p $ q |
|
123 |
fun star_tr' [P,Q,_] = Syntax.const \<^syntax_const>\<open>_star\<close> $ strip P $ strip Q |
|
124 |
fun wand_tr' [P,Q,_] = Syntax.const \<^syntax_const>\<open>_wand\<close> $ strip P $ strip Q |
|
35113 | 125 |
|
13875 | 126 |
end |
62042 | 127 |
\<close> |
13857 | 128 |
|
62042 | 129 |
print_translation \<open> |
69597 | 130 |
[(\<^const_syntax>\<open>is_empty\<close>, K is_empty_tr'), |
131 |
(\<^const_syntax>\<open>singl\<close>, K singl_tr'), |
|
132 |
(\<^const_syntax>\<open>star\<close>, K star_tr'), |
|
133 |
(\<^const_syntax>\<open>wand\<close>, K wand_tr')] |
|
62042 | 134 |
\<close> |
13903 | 135 |
|
62042 | 136 |
text\<open>Now the intermediate proof states are also readable:\<close> |
13857 | 137 |
|
138 |
lemma "VARS H x y z w |
|
139 |
{[x\<mapsto>y] ** [z\<mapsto>w]} |
|
13867 | 140 |
y := w |
13857 | 141 |
{x \<noteq> z}" |
142 |
apply vcg |
|
13875 | 143 |
apply(auto simp:star_def ortho_def singl_def) |
144 |
done |
|
145 |
||
146 |
lemma "VARS H x y z w |
|
147 |
{emp ** emp} |
|
148 |
SKIP |
|
149 |
{emp}" |
|
150 |
apply vcg |
|
151 |
apply(auto simp:star_def ortho_def is_empty_def) |
|
152 |
done |
|
153 |
||
62042 | 154 |
text\<open>So far we have unfolded the separation logic connectives in |
13903 | 155 |
proofs. Here comes a simple example of a program proof that uses a law |
62042 | 156 |
of separation logic instead.\<close> |
13903 | 157 |
|
67410 | 158 |
\<comment> \<open>a law of separation logic\<close> |
13875 | 159 |
lemma star_comm: "P ** Q = Q ** P" |
18447 | 160 |
by(auto simp add:star_def ortho_def dest: map_add_comm) |
13875 | 161 |
|
162 |
lemma "VARS H x y z w |
|
163 |
{P ** Q} |
|
164 |
SKIP |
|
165 |
{Q ** P}" |
|
166 |
apply vcg |
|
167 |
apply(simp add: star_comm) |
|
168 |
done |
|
169 |
||
14074 | 170 |
|
171 |
lemma "VARS H |
|
172 |
{p\<noteq>0 \<and> [p \<mapsto> x] ** List H q qs} |
|
173 |
H := H(p \<mapsto> q) |
|
174 |
{List H p (p#qs)}" |
|
175 |
apply vcg |
|
176 |
apply(simp add: star_def ortho_def singl_def) |
|
177 |
apply clarify |
|
178 |
apply(subgoal_tac "p \<notin> set qs") |
|
179 |
prefer 2 |
|
180 |
apply(blast dest:list_in_heap) |
|
181 |
apply simp |
|
182 |
done |
|
183 |
||
184 |
lemma "VARS H p q r |
|
185 |
{List H p Ps ** List H q Qs} |
|
186 |
WHILE p \<noteq> 0 |
|
187 |
INV {\<exists>ps qs. (List H p ps ** List H q qs) \<and> rev ps @ qs = rev Ps @ Qs} |
|
188 |
DO r := p; p := the(H p); H := H(r \<mapsto> q); q := r OD |
|
189 |
{List H q (rev Ps @ Qs)}" |
|
190 |
apply vcg |
|
191 |
apply(simp_all add: star_def ortho_def singl_def) |
|
192 |
||
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44241
diff
changeset
|
193 |
apply fastforce |
14074 | 194 |
|
195 |
apply (clarsimp simp add:List_non_null) |
|
196 |
apply(rename_tac ps') |
|
197 |
apply(rule_tac x = ps' in exI) |
|
198 |
apply(rule_tac x = "p#qs" in exI) |
|
199 |
apply simp |
|
200 |
apply(rule_tac x = "h1(p:=None)" in exI) |
|
201 |
apply(rule_tac x = "h2(p\<mapsto>q)" in exI) |
|
202 |
apply simp |
|
203 |
apply(rule conjI) |
|
204 |
apply(rule ext) |
|
205 |
apply(simp add:map_add_def split:option.split) |
|
206 |
apply(rule conjI) |
|
207 |
apply blast |
|
208 |
apply(simp add:map_add_def split:option.split) |
|
209 |
apply(rule conjI) |
|
210 |
apply(subgoal_tac "p \<notin> set qs") |
|
211 |
prefer 2 |
|
212 |
apply(blast dest:list_in_heap) |
|
213 |
apply(simp) |
|
214 |
apply fast |
|
215 |
||
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44241
diff
changeset
|
216 |
apply(fastforce) |
14074 | 217 |
done |
218 |
||
13875 | 219 |
end |