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