| 13771 |      1 | (*  Title:      HOL/Hoare/Pointers.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Tobias Nipkow
 | 
|  |      4 |     Copyright   2002 TUM
 | 
|  |      5 | 
 | 
|  |      6 | This is like Pointers.thy, but instead of a type constructor 'a ref
 | 
|  |      7 | that adjoins Null to a type, Null is simply a distinguished element of
 | 
|  |      8 | the address type. This avoids the Ref constructor and thus simplifies
 | 
|  |      9 | specifications (a bit). However, the proofs don't seem to get simpler
 | 
|  |     10 | - in fact in some case they appear to get (a bit) more complicated.
 | 
|  |     11 | *)
 | 
|  |     12 | 
 | 
| 16417 |     13 | theory Pointers0 imports Hoare begin
 | 
| 13771 |     14 | 
 | 
|  |     15 | subsection "References"
 | 
|  |     16 | 
 | 
|  |     17 | axclass ref < type
 | 
|  |     18 | consts Null :: "'a::ref"
 | 
|  |     19 | 
 | 
|  |     20 | subsection "Field access and update"
 | 
|  |     21 | 
 | 
|  |     22 | syntax
 | 
|  |     23 |   "@fassign"  :: "'a::ref => id => 'v => 's com"
 | 
|  |     24 |    ("(2_^._ :=/ _)" [70,1000,65] 61)
 | 
|  |     25 |   "@faccess"  :: "'a::ref => ('a::ref \<Rightarrow> 'v) => 'v"
 | 
|  |     26 |    ("_^._" [65,1000] 65)
 | 
|  |     27 | translations
 | 
|  |     28 |   "p^.f := e"  =>  "f := fun_upd f p e"
 | 
|  |     29 |   "p^.f"       =>  "f p"
 | 
|  |     30 | 
 | 
|  |     31 | 
 | 
|  |     32 | text "An example due to Suzuki:"
 | 
|  |     33 | 
 | 
|  |     34 | lemma "VARS v n
 | 
|  |     35 |   {distinct[w,x,y,z]}
 | 
|  |     36 |   w^.v := (1::int); w^.n := x;
 | 
|  |     37 |   x^.v := 2; x^.n := y;
 | 
|  |     38 |   y^.v := 3; y^.n := z;
 | 
|  |     39 |   z^.v := 4; x^.n := z
 | 
|  |     40 |   {w^.n^.n^.v = 4}"
 | 
|  |     41 | by vcg_simp
 | 
|  |     42 | 
 | 
|  |     43 | 
 | 
|  |     44 | section "The heap"
 | 
|  |     45 | 
 | 
|  |     46 | subsection "Paths in the heap"
 | 
|  |     47 | 
 | 
|  |     48 | consts
 | 
|  |     49 |  Path :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a \<Rightarrow> bool"
 | 
|  |     50 | primrec
 | 
|  |     51 | "Path h x [] y = (x = y)"
 | 
|  |     52 | "Path h x (a#as) y = (x \<noteq> Null \<and> x = a \<and> Path h (h a) as y)"
 | 
|  |     53 | 
 | 
|  |     54 | lemma [iff]: "Path h Null xs y = (xs = [] \<and> y = Null)"
 | 
|  |     55 | apply(case_tac xs)
 | 
|  |     56 | apply fastsimp
 | 
|  |     57 | apply fastsimp
 | 
|  |     58 | done
 | 
|  |     59 | 
 | 
|  |     60 | lemma [simp]: "a \<noteq> Null \<Longrightarrow> Path h a as z =
 | 
|  |     61 |  (as = [] \<and> z = a  \<or>  (\<exists>bs. as = a#bs \<and> Path h (h a) bs z))"
 | 
|  |     62 | apply(case_tac as)
 | 
|  |     63 | apply fastsimp
 | 
|  |     64 | apply fastsimp
 | 
|  |     65 | done
 | 
|  |     66 | 
 | 
|  |     67 | lemma [simp]: "\<And>x. Path f x (as@bs) z = (\<exists>y. Path f x as y \<and> Path f y bs z)"
 | 
|  |     68 | by(induct as, simp+)
 | 
|  |     69 | 
 | 
|  |     70 | lemma [simp]: "\<And>x. u \<notin> set as \<Longrightarrow> Path (f(u := v)) x as y = Path f x as y"
 | 
|  |     71 | by(induct as, simp, simp add:eq_sym_conv)
 | 
|  |     72 | 
 | 
|  |     73 | subsection "Lists on the heap"
 | 
|  |     74 | 
 | 
|  |     75 | subsubsection "Relational abstraction"
 | 
|  |     76 | 
 | 
|  |     77 | constdefs
 | 
|  |     78 |  List :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> bool"
 | 
|  |     79 | "List h x as == Path h x as Null"
 | 
|  |     80 | 
 | 
|  |     81 | lemma [simp]: "List h x [] = (x = Null)"
 | 
|  |     82 | by(simp add:List_def)
 | 
|  |     83 | 
 | 
|  |     84 | lemma [simp]: "List h x (a#as) = (x \<noteq> Null \<and> x = a \<and> List h (h a) as)"
 | 
|  |     85 | by(simp add:List_def)
 | 
|  |     86 | 
 | 
|  |     87 | lemma [simp]: "List h Null as = (as = [])"
 | 
|  |     88 | by(case_tac as, simp_all)
 | 
|  |     89 | 
 | 
|  |     90 | lemma List_Ref[simp]:
 | 
|  |     91 |  "a \<noteq> Null \<Longrightarrow> List h a as = (\<exists>bs. as = a#bs \<and> List h (h a) bs)"
 | 
|  |     92 | by(case_tac as, simp_all, fast)
 | 
|  |     93 | 
 | 
|  |     94 | theorem notin_List_update[simp]:
 | 
|  |     95 |  "\<And>x. a \<notin> set as \<Longrightarrow> List (h(a := y)) x as = List h x as"
 | 
|  |     96 | apply(induct as)
 | 
|  |     97 | apply simp
 | 
|  |     98 | apply(clarsimp simp add:fun_upd_apply)
 | 
|  |     99 | done
 | 
|  |    100 | 
 | 
|  |    101 | 
 | 
|  |    102 | declare fun_upd_apply[simp del]fun_upd_same[simp] fun_upd_other[simp]
 | 
|  |    103 | 
 | 
|  |    104 | lemma List_unique: "\<And>x bs. List h x as \<Longrightarrow> List h x bs \<Longrightarrow> as = bs"
 | 
|  |    105 | by(induct as, simp, clarsimp)
 | 
|  |    106 | 
 | 
|  |    107 | lemma List_unique1: "List h p as \<Longrightarrow> \<exists>!as. List h p as"
 | 
|  |    108 | by(blast intro:List_unique)
 | 
|  |    109 | 
 | 
|  |    110 | lemma List_app: "\<And>x. List h x (as@bs) = (\<exists>y. Path h x as y \<and> List h y bs)"
 | 
|  |    111 | by(induct as, simp, clarsimp)
 | 
|  |    112 | 
 | 
|  |    113 | lemma List_hd_not_in_tl[simp]: "List h (h a) as \<Longrightarrow> a \<notin> set as"
 | 
|  |    114 | apply (clarsimp simp add:in_set_conv_decomp)
 | 
|  |    115 | apply(frule List_app[THEN iffD1])
 | 
|  |    116 | apply(fastsimp dest: List_unique)
 | 
|  |    117 | done
 | 
|  |    118 | 
 | 
|  |    119 | lemma List_distinct[simp]: "\<And>x. List h x as \<Longrightarrow> distinct as"
 | 
|  |    120 | apply(induct as, simp)
 | 
|  |    121 | apply(fastsimp dest:List_hd_not_in_tl)
 | 
|  |    122 | done
 | 
|  |    123 | 
 | 
|  |    124 | subsection "Functional abstraction"
 | 
|  |    125 | 
 | 
|  |    126 | constdefs
 | 
|  |    127 |  islist :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool"
 | 
|  |    128 | "islist h p == \<exists>as. List h p as"
 | 
|  |    129 |  list :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list"
 | 
|  |    130 | "list h p == SOME as. List h p as"
 | 
|  |    131 | 
 | 
|  |    132 | lemma List_conv_islist_list: "List h p as = (islist h p \<and> as = list h p)"
 | 
|  |    133 | apply(simp add:islist_def list_def)
 | 
|  |    134 | apply(rule iffI)
 | 
|  |    135 | apply(rule conjI)
 | 
|  |    136 | apply blast
 | 
|  |    137 | apply(subst some1_equality)
 | 
|  |    138 |   apply(erule List_unique1)
 | 
|  |    139 |  apply assumption
 | 
|  |    140 | apply(rule refl)
 | 
|  |    141 | apply simp
 | 
|  |    142 | apply(rule someI_ex)
 | 
|  |    143 | apply fast
 | 
|  |    144 | done
 | 
|  |    145 | 
 | 
|  |    146 | lemma [simp]: "islist h Null"
 | 
|  |    147 | by(simp add:islist_def)
 | 
|  |    148 | 
 | 
|  |    149 | lemma [simp]: "a \<noteq> Null \<Longrightarrow> islist h a = islist h (h a)"
 | 
|  |    150 | by(simp add:islist_def)
 | 
|  |    151 | 
 | 
|  |    152 | lemma [simp]: "list h Null = []"
 | 
|  |    153 | by(simp add:list_def)
 | 
|  |    154 | 
 | 
|  |    155 | lemma list_Ref_conv[simp]:
 | 
|  |    156 |  "\<lbrakk> a \<noteq> Null; islist h (h a) \<rbrakk> \<Longrightarrow> list h a = a # list h (h a)"
 | 
|  |    157 | apply(insert List_Ref[of _ h])
 | 
|  |    158 | apply(fastsimp simp:List_conv_islist_list)
 | 
|  |    159 | done
 | 
|  |    160 | 
 | 
|  |    161 | lemma [simp]: "islist h (h a) \<Longrightarrow> a \<notin> set(list h (h a))"
 | 
|  |    162 | apply(insert List_hd_not_in_tl[of h])
 | 
|  |    163 | apply(simp add:List_conv_islist_list)
 | 
|  |    164 | done
 | 
|  |    165 | 
 | 
|  |    166 | lemma list_upd_conv[simp]:
 | 
|  |    167 |  "islist h p \<Longrightarrow> y \<notin> set(list h p) \<Longrightarrow> list (h(y := q)) p = list h p"
 | 
|  |    168 | apply(drule notin_List_update[of _ _ h q p])
 | 
|  |    169 | apply(simp add:List_conv_islist_list)
 | 
|  |    170 | done
 | 
|  |    171 | 
 | 
|  |    172 | lemma islist_upd[simp]:
 | 
|  |    173 |  "islist h p \<Longrightarrow> y \<notin> set(list h p) \<Longrightarrow> islist (h(y := q)) p"
 | 
|  |    174 | apply(frule notin_List_update[of _ _ h q p])
 | 
|  |    175 | apply(simp add:List_conv_islist_list)
 | 
|  |    176 | done
 | 
|  |    177 | 
 | 
|  |    178 | 
 | 
|  |    179 | section "Verifications"
 | 
|  |    180 | 
 | 
|  |    181 | subsection "List reversal"
 | 
|  |    182 | 
 | 
|  |    183 | text "A short but unreadable proof:"
 | 
|  |    184 | 
 | 
|  |    185 | lemma "VARS tl p q r
 | 
|  |    186 |   {List tl p Ps \<and> List tl q Qs \<and> set Ps \<inter> set Qs = {}}
 | 
|  |    187 |   WHILE p \<noteq> Null
 | 
|  |    188 |   INV {\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
 | 
|  |    189 |                  rev ps @ qs = rev Ps @ Qs}
 | 
|  |    190 |   DO r := p; p := p^.tl; r^.tl := q; q := r OD
 | 
|  |    191 |   {List tl q (rev Ps @ Qs)}"
 | 
|  |    192 | apply vcg_simp
 | 
|  |    193 |   apply fastsimp
 | 
|  |    194 |  apply(fastsimp intro:notin_List_update[THEN iffD2])
 | 
|  |    195 | (* explicily:
 | 
|  |    196 |  apply clarify
 | 
|  |    197 |  apply(rename_tac ps qs)
 | 
|  |    198 |  apply clarsimp
 | 
|  |    199 |  apply(rename_tac ps')
 | 
|  |    200 |  apply(rule_tac x = ps' in exI)
 | 
|  |    201 |  apply simp
 | 
|  |    202 |  apply(rule_tac x = "p#qs" in exI)
 | 
|  |    203 |  apply simp
 | 
|  |    204 | *)
 | 
|  |    205 | apply fastsimp
 | 
|  |    206 | done
 | 
|  |    207 | 
 | 
|  |    208 | 
 | 
|  |    209 | text "A longer readable version:"
 | 
|  |    210 | 
 | 
|  |    211 | lemma "VARS tl p q r
 | 
|  |    212 |   {List tl p Ps \<and> List tl q Qs \<and> set Ps \<inter> set Qs = {}}
 | 
|  |    213 |   WHILE p \<noteq> Null
 | 
|  |    214 |   INV {\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
 | 
|  |    215 |                rev ps @ qs = rev Ps @ Qs}
 | 
|  |    216 |   DO r := p; p := p^.tl; r^.tl := q; q := r OD
 | 
|  |    217 |   {List tl q (rev Ps @ Qs)}"
 | 
|  |    218 | proof vcg
 | 
|  |    219 |   fix tl p q r
 | 
|  |    220 |   assume "List tl p Ps \<and> List tl q Qs \<and> set Ps \<inter> set Qs = {}"
 | 
|  |    221 |   thus "\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
 | 
|  |    222 |                 rev ps @ qs = rev Ps @ Qs" by fastsimp
 | 
|  |    223 | next
 | 
|  |    224 |   fix tl p q r
 | 
|  |    225 |   assume "(\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
 | 
|  |    226 |                    rev ps @ qs = rev Ps @ Qs) \<and> p \<noteq> Null"
 | 
|  |    227 |          (is "(\<exists>ps qs. ?I ps qs) \<and> _")
 | 
|  |    228 |   then obtain ps qs where I: "?I ps qs \<and> p \<noteq> Null" by fast
 | 
|  |    229 |   then obtain ps' where "ps = p # ps'" by fastsimp
 | 
|  |    230 |   hence "List (tl(p := q)) (p^.tl) ps' \<and>
 | 
|  |    231 |          List (tl(p := q)) p       (p#qs) \<and>
 | 
|  |    232 |          set ps' \<inter> set (p#qs) = {} \<and>
 | 
|  |    233 |          rev ps' @ (p#qs) = rev Ps @ Qs"
 | 
|  |    234 |     using I by fastsimp
 | 
|  |    235 |   thus "\<exists>ps' qs'. List (tl(p := q)) (p^.tl) ps' \<and>
 | 
|  |    236 |                   List (tl(p := q)) p       qs' \<and>
 | 
|  |    237 |                   set ps' \<inter> set qs' = {} \<and>
 | 
|  |    238 |                   rev ps' @ qs' = rev Ps @ Qs" by fast
 | 
|  |    239 | next
 | 
|  |    240 |   fix tl p q r
 | 
|  |    241 |   assume "(\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
 | 
|  |    242 |                    rev ps @ qs = rev Ps @ Qs) \<and> \<not> p \<noteq> Null"
 | 
|  |    243 |   thus "List tl q (rev Ps @ Qs)" by fastsimp
 | 
|  |    244 | qed
 | 
|  |    245 | 
 | 
|  |    246 | 
 | 
|  |    247 | text{* Finaly, the functional version. A bit more verbose, but automatic! *}
 | 
|  |    248 | 
 | 
|  |    249 | lemma "VARS tl p q r
 | 
|  |    250 |   {islist tl p \<and> islist tl q \<and>
 | 
|  |    251 |    Ps = list tl p \<and> Qs = list tl q \<and> set Ps \<inter> set Qs = {}}
 | 
|  |    252 |   WHILE p \<noteq> Null
 | 
|  |    253 |   INV {islist tl p \<and> islist tl q \<and>
 | 
|  |    254 |        set(list tl p) \<inter> set(list tl q) = {} \<and>
 | 
|  |    255 |        rev(list tl p) @ (list tl q) = rev Ps @ Qs}
 | 
|  |    256 |   DO r := p; p := p^.tl; r^.tl := q; q := r OD
 | 
|  |    257 |   {islist tl q \<and> list tl q = rev Ps @ Qs}"
 | 
|  |    258 | apply vcg_simp
 | 
|  |    259 |   apply clarsimp
 | 
|  |    260 |  apply clarsimp
 | 
|  |    261 | apply clarsimp
 | 
|  |    262 | done
 | 
|  |    263 | 
 | 
|  |    264 | 
 | 
|  |    265 | subsection "Searching in a list"
 | 
|  |    266 | 
 | 
|  |    267 | text{*What follows is a sequence of successively more intelligent proofs that
 | 
|  |    268 | a simple loop finds an element in a linked list.
 | 
|  |    269 | 
 | 
|  |    270 | We start with a proof based on the @{term List} predicate. This means it only
 | 
|  |    271 | works for acyclic lists. *}
 | 
|  |    272 | 
 | 
|  |    273 | lemma "VARS tl p
 | 
|  |    274 |   {List tl p Ps \<and> X \<in> set Ps}
 | 
|  |    275 |   WHILE p \<noteq> Null \<and> p \<noteq> X
 | 
|  |    276 |   INV {p \<noteq> Null \<and> (\<exists>ps. List tl p ps \<and> X \<in> set ps)}
 | 
|  |    277 |   DO p := p^.tl OD
 | 
|  |    278 |   {p = X}"
 | 
|  |    279 | apply vcg_simp
 | 
|  |    280 |   apply(case_tac "p = Null")
 | 
|  |    281 |    apply clarsimp
 | 
|  |    282 |   apply fastsimp
 | 
|  |    283 |  apply clarsimp
 | 
|  |    284 |  apply fastsimp
 | 
|  |    285 | apply clarsimp
 | 
|  |    286 | done
 | 
|  |    287 | 
 | 
|  |    288 | text{*Using @{term Path} instead of @{term List} generalizes the correctness
 | 
|  |    289 | statement to cyclic lists as well: *}
 | 
|  |    290 | 
 | 
|  |    291 | lemma "VARS tl p
 | 
|  |    292 |   {Path tl p Ps X}
 | 
|  |    293 |   WHILE p \<noteq> Null \<and> p \<noteq> X
 | 
|  |    294 |   INV {\<exists>ps. Path tl p ps X}
 | 
|  |    295 |   DO p := p^.tl OD
 | 
|  |    296 |   {p = X}"
 | 
|  |    297 | apply vcg_simp
 | 
|  |    298 |   apply blast
 | 
|  |    299 |  apply fastsimp
 | 
|  |    300 | apply clarsimp
 | 
|  |    301 | done
 | 
|  |    302 | 
 | 
|  |    303 | text{*Now it dawns on us that we do not need the list witness at all --- it
 | 
|  |    304 | suffices to talk about reachability, i.e.\ we can use relations directly. *}
 | 
|  |    305 | 
 | 
|  |    306 | lemma "VARS tl p
 | 
|  |    307 |   {(p,X) \<in> {(x,y). y = tl x & x \<noteq> Null}^*}
 | 
|  |    308 |   WHILE p \<noteq> Null \<and> p \<noteq> X
 | 
|  |    309 |   INV {(p,X) \<in> {(x,y). y = tl x & x \<noteq> Null}^*}
 | 
|  |    310 |   DO p := p^.tl OD
 | 
|  |    311 |   {p = X}"
 | 
|  |    312 | apply vcg_simp
 | 
|  |    313 |  apply clarsimp
 | 
|  |    314 |  apply(erule converse_rtranclE)
 | 
|  |    315 |   apply simp
 | 
|  |    316 |  apply(simp)
 | 
|  |    317 | apply(fastsimp elim:converse_rtranclE)
 | 
|  |    318 | done
 | 
|  |    319 | 
 | 
|  |    320 | 
 | 
|  |    321 | subsection "Merging two lists"
 | 
|  |    322 | 
 | 
|  |    323 | text"This is still a bit rough, especially the proof."
 | 
|  |    324 | 
 | 
|  |    325 | consts merge :: "'a list * 'a list * ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list"
 | 
|  |    326 | 
 | 
|  |    327 | recdef merge "measure(%(xs,ys,f). size xs + size ys)"
 | 
|  |    328 | "merge(x#xs,y#ys,f) = (if f x y then x # merge(xs,y#ys,f)
 | 
|  |    329 |                                 else y # merge(x#xs,ys,f))"
 | 
|  |    330 | "merge(x#xs,[],f) = x # merge(xs,[],f)"
 | 
|  |    331 | "merge([],y#ys,f) = y # merge([],ys,f)"
 | 
|  |    332 | "merge([],[],f) = []"
 | 
|  |    333 | 
 | 
|  |    334 | lemma imp_disjCL: "(P|Q \<longrightarrow> R) = ((P \<longrightarrow> R) \<and> (~P \<longrightarrow> Q \<longrightarrow> R))"
 | 
|  |    335 | by blast
 | 
|  |    336 | 
 | 
|  |    337 | declare disj_not1[simp del] imp_disjL[simp del] imp_disjCL[simp]
 | 
|  |    338 | 
 | 
|  |    339 | lemma "VARS hd tl p q r s
 | 
|  |    340 |  {List tl p Ps \<and> List tl q Qs \<and> set Ps \<inter> set Qs = {} \<and>
 | 
|  |    341 |   (p \<noteq> Null \<or> q \<noteq> Null)}
 | 
|  |    342 |  IF if q = Null then True else p ~= Null & p^.hd \<le> q^.hd
 | 
|  |    343 |  THEN r := p; p := p^.tl ELSE r := q; q := q^.tl FI;
 | 
|  |    344 |  s := r;
 | 
|  |    345 |  WHILE p \<noteq> Null \<or> q \<noteq> Null
 | 
|  |    346 |  INV {EX rs ps qs. Path tl r rs s \<and> List tl p ps \<and> List tl q qs \<and>
 | 
|  |    347 |       distinct(s # ps @ qs @ rs) \<and> s \<noteq> Null \<and>
 | 
|  |    348 |       merge(Ps,Qs,\<lambda>x y. hd x \<le> hd y) =
 | 
|  |    349 |       rs @ s # merge(ps,qs,\<lambda>x y. hd x \<le> hd y) \<and>
 | 
|  |    350 |       (tl s = p \<or> tl s = q)}
 | 
|  |    351 |  DO IF if q = Null then True else p \<noteq> Null \<and> p^.hd \<le> q^.hd
 | 
|  |    352 |     THEN s^.tl := p; p := p^.tl ELSE s^.tl := q; q := q^.tl FI;
 | 
|  |    353 |     s := s^.tl
 | 
|  |    354 |  OD
 | 
|  |    355 |  {List tl r (merge(Ps,Qs,\<lambda>x y. hd x \<le> hd y))}"
 | 
|  |    356 | apply vcg_simp
 | 
|  |    357 | 
 | 
|  |    358 | apply (fastsimp)
 | 
|  |    359 | 
 | 
|  |    360 | apply clarsimp
 | 
|  |    361 | apply(rule conjI)
 | 
|  |    362 | apply clarsimp
 | 
|  |    363 | apply(simp add:eq_sym_conv)
 | 
|  |    364 | apply(rule_tac x = "rs @ [s]" in exI)
 | 
|  |    365 | apply simp
 | 
|  |    366 | apply(rule_tac x = "bs" in exI)
 | 
|  |    367 | apply (fastsimp simp:eq_sym_conv)
 | 
|  |    368 | 
 | 
|  |    369 | apply clarsimp
 | 
|  |    370 | apply(rule conjI)
 | 
|  |    371 | apply clarsimp
 | 
|  |    372 | apply(rule_tac x = "rs @ [s]" in exI)
 | 
|  |    373 | apply simp
 | 
|  |    374 | apply(rule_tac x = "bsa" in exI)
 | 
|  |    375 | apply(rule conjI)
 | 
|  |    376 | apply (simp add:eq_sym_conv)
 | 
|  |    377 | apply(rule exI)
 | 
|  |    378 | apply(rule conjI)
 | 
|  |    379 | apply(rule_tac x = bs in exI)
 | 
|  |    380 | apply(rule conjI)
 | 
|  |    381 | apply(rule refl)
 | 
|  |    382 | apply (simp add:eq_sym_conv)
 | 
|  |    383 | apply (simp add:eq_sym_conv)
 | 
|  |    384 | 
 | 
|  |    385 | apply(rule conjI)
 | 
|  |    386 | apply clarsimp
 | 
|  |    387 | apply(rule_tac x = "rs @ [s]" in exI)
 | 
|  |    388 | apply simp
 | 
|  |    389 | apply(rule_tac x = bs in exI)
 | 
|  |    390 | apply (simp add:eq_sym_conv)
 | 
|  |    391 | apply clarsimp
 | 
|  |    392 | apply(rule_tac x = "rs @ [s]" in exI)
 | 
|  |    393 | apply (simp add:eq_sym_conv)
 | 
|  |    394 | apply(rule exI)
 | 
|  |    395 | apply(rule conjI)
 | 
|  |    396 | apply(rule_tac x = bsa in exI)
 | 
|  |    397 | apply(rule conjI)
 | 
|  |    398 | apply(rule refl)
 | 
|  |    399 | apply (simp add:eq_sym_conv)
 | 
|  |    400 | apply(rule_tac x = bs in exI)
 | 
|  |    401 | apply (simp add:eq_sym_conv)
 | 
|  |    402 | 
 | 
|  |    403 | apply(clarsimp simp add:List_app)
 | 
|  |    404 | done
 | 
|  |    405 | 
 | 
|  |    406 | (* TODO: merging with islist/list instead of List: an improvement?
 | 
|  |    407 |    needs (is)path, which is not so easy to prove either. *)
 | 
|  |    408 | 
 | 
|  |    409 | subsection "Storage allocation"
 | 
|  |    410 | 
 | 
|  |    411 | constdefs new :: "'a set \<Rightarrow> 'a::ref"
 | 
|  |    412 | "new A == SOME a. a \<notin> A & a \<noteq> Null"
 | 
|  |    413 | 
 | 
|  |    414 | 
 | 
|  |    415 | lemma new_notin:
 | 
|  |    416 |  "\<lbrakk> ~finite(UNIV::('a::ref)set); finite(A::'a set); B \<subseteq> A \<rbrakk> \<Longrightarrow>
 | 
|  |    417 |   new (A) \<notin> B & new A \<noteq> Null"
 | 
|  |    418 | apply(unfold new_def)
 | 
|  |    419 | apply(rule someI2_ex)
 | 
|  |    420 |  apply (fast dest:ex_new_if_finite[of "insert Null A"])
 | 
|  |    421 | apply (fast)
 | 
|  |    422 | done
 | 
|  |    423 | 
 | 
|  |    424 | lemma "~finite(UNIV::('a::ref)set) \<Longrightarrow>
 | 
|  |    425 |   VARS xs elem next alloc p q
 | 
|  |    426 |   {Xs = xs \<and> p = (Null::'a)}
 | 
|  |    427 |   WHILE xs \<noteq> []
 | 
|  |    428 |   INV {islist next p \<and> set(list next p) \<subseteq> set alloc \<and>
 | 
|  |    429 |        map elem (rev(list next p)) @ xs = Xs}
 | 
|  |    430 |   DO q := new(set alloc); alloc := q#alloc;
 | 
|  |    431 |      q^.next := p; q^.elem := hd xs; xs := tl xs; p := q
 | 
|  |    432 |   OD
 | 
|  |    433 |   {islist next p \<and> map elem (rev(list next p)) = Xs}"
 | 
|  |    434 | apply vcg_simp
 | 
|  |    435 |  apply (clarsimp simp: subset_insert_iff neq_Nil_conv fun_upd_apply new_notin)
 | 
|  |    436 | apply fastsimp
 | 
|  |    437 | done
 | 
|  |    438 | 
 | 
|  |    439 | 
 | 
|  |    440 | end
 |