1(*  Title:      HOL/Imperative_HOL/ex/Linked_Lists.thy
2    Author:     Lukas Bulwahn, TU Muenchen
3*)
4
5section \<open>Linked Lists by ML references\<close>
6
7theory Linked_Lists
8imports "../Imperative_HOL" "HOL-Library.Code_Target_Int"
9begin
10
11section \<open>Definition of Linked Lists\<close>
12
13setup \<open>Sign.add_const_constraint (\<^const_name>\<open>Ref\<close>, SOME \<^typ>\<open>nat \<Rightarrow> 'a::type ref\<close>)\<close>
14datatype 'a node = Empty | Node 'a "'a node ref"
15
16primrec
17  node_encode :: "'a::countable node \<Rightarrow> nat"
18where
19  "node_encode Empty = 0"
20  | "node_encode (Node x r) = Suc (to_nat (x, r))"
21
22instance node :: (countable) countable
23proof (rule countable_classI [of "node_encode"])
24  fix x y :: "'a::countable node"
25  show "node_encode x = node_encode y \<Longrightarrow> x = y"
26  by (induct x, auto, induct y, auto, induct y, auto)
27qed
28
29instance node :: (heap) heap ..
30
31primrec make_llist :: "'a::heap list \<Rightarrow> 'a node Heap"
32where 
33  [simp del]: "make_llist []     = return Empty"
34            | "make_llist (x#xs) = do { tl \<leftarrow> make_llist xs;
35                                        next \<leftarrow> ref tl;
36                                        return (Node x next)
37                                   }"
38
39
40partial_function (heap) traverse :: "'a::heap node \<Rightarrow> 'a list Heap"
41where
42  [code del]: "traverse l =
43    (case l of Empty \<Rightarrow> return []
44     | Node x r \<Rightarrow> do { tl \<leftarrow> Ref.lookup r;
45                              xs \<leftarrow> traverse tl;
46                              return (x#xs)
47                         })"
48
49lemma traverse_simps[code, simp]:
50  "traverse Empty      = return []"
51  "traverse (Node x r) = do { tl \<leftarrow> Ref.lookup r;
52                              xs \<leftarrow> traverse tl;
53                              return (x#xs)
54                         }"
55by (simp_all add: traverse.simps[of "Empty"] traverse.simps[of "Node x r"])
56
57
58section \<open>Proving correctness with relational abstraction\<close>
59
60subsection \<open>Definition of list_of, list_of', refs_of and refs_of'\<close>
61
62primrec list_of :: "heap \<Rightarrow> ('a::heap) node \<Rightarrow> 'a list \<Rightarrow> bool"
63where
64  "list_of h r [] = (r = Empty)"
65| "list_of h r (a#as) = (case r of Empty \<Rightarrow> False | Node b bs \<Rightarrow> (a = b \<and> list_of h (Ref.get h bs) as))"
66 
67definition list_of' :: "heap \<Rightarrow> ('a::heap) node ref \<Rightarrow> 'a list \<Rightarrow> bool"
68where
69  "list_of' h r xs = list_of h (Ref.get h r) xs"
70
71primrec refs_of :: "heap \<Rightarrow> ('a::heap) node \<Rightarrow> 'a node ref list \<Rightarrow> bool"
72where
73  "refs_of h r [] = (r = Empty)"
74| "refs_of h r (x#xs) = (case r of Empty \<Rightarrow> False | Node b bs \<Rightarrow> (x = bs) \<and> refs_of h (Ref.get h bs) xs)"
75
76primrec refs_of' :: "heap \<Rightarrow> ('a::heap) node ref \<Rightarrow> 'a node ref list \<Rightarrow> bool"
77where
78  "refs_of' h r [] = False"
79| "refs_of' h r (x#xs) = ((x = r) \<and> refs_of h (Ref.get h x) xs)"
80
81
82subsection \<open>Properties of these definitions\<close>
83
84lemma list_of_Empty[simp]: "list_of h Empty xs = (xs = [])"
85by (cases xs, auto)
86
87lemma list_of_Node[simp]: "list_of h (Node x ps) xs = (\<exists>xs'. (xs = x # xs') \<and> list_of h (Ref.get h ps) xs')"
88by (cases xs, auto)
89
90lemma list_of'_Empty[simp]: "Ref.get h q = Empty \<Longrightarrow> list_of' h q xs = (xs = [])"
91unfolding list_of'_def by simp
92
93lemma list_of'_Node[simp]: "Ref.get h q = Node x ps \<Longrightarrow> list_of' h q xs = (\<exists>xs'. (xs = x # xs') \<and> list_of' h ps xs')"
94unfolding list_of'_def by simp
95
96lemma list_of'_Nil: "list_of' h q [] \<Longrightarrow> Ref.get h q = Empty"
97unfolding list_of'_def by simp
98
99lemma list_of'_Cons: 
100assumes "list_of' h q (x#xs)"
101obtains n where "Ref.get h q = Node x n" and "list_of' h n xs"
102using assms unfolding list_of'_def by (auto split: node.split_asm)
103
104lemma refs_of_Empty[simp] : "refs_of h Empty xs = (xs = [])"
105  by (cases xs, auto)
106
107lemma refs_of_Node[simp]: "refs_of h (Node x ps) xs = (\<exists>prs. xs = ps # prs \<and> refs_of h (Ref.get h ps) prs)"
108  by (cases xs, auto)
109
110lemma refs_of'_def': "refs_of' h p ps = (\<exists>prs. (ps = (p # prs)) \<and> refs_of h (Ref.get h p) prs)"
111by (cases ps, auto)
112
113lemma refs_of'_Node:
114  assumes "refs_of' h p xs"
115  assumes "Ref.get h p = Node x pn"
116  obtains pnrs
117  where "xs = p # pnrs" and "refs_of' h pn pnrs"
118using assms
119unfolding refs_of'_def' by auto
120
121lemma list_of_is_fun: "\<lbrakk> list_of h n xs; list_of h n ys\<rbrakk> \<Longrightarrow> xs = ys"
122proof (induct xs arbitrary: ys n)
123  case Nil thus ?case by auto
124next
125  case (Cons x xs')
126  thus ?case
127    by (cases ys,  auto split: node.split_asm)
128qed
129
130lemma refs_of_is_fun: "\<lbrakk> refs_of h n xs; refs_of h n ys\<rbrakk> \<Longrightarrow> xs = ys"
131proof (induct xs arbitrary: ys n)
132  case Nil thus ?case by auto
133next
134  case (Cons x xs')
135  thus ?case
136    by (cases ys,  auto split: node.split_asm)
137qed
138
139lemma refs_of'_is_fun: "\<lbrakk> refs_of' h p as; refs_of' h p bs \<rbrakk> \<Longrightarrow> as = bs"
140unfolding refs_of'_def' by (auto dest: refs_of_is_fun)
141
142
143lemma list_of_refs_of_HOL:
144  assumes "list_of h r xs"
145  shows "\<exists>rs. refs_of h r rs"
146using assms
147proof (induct xs arbitrary: r)
148  case Nil thus ?case by auto
149next
150  case (Cons x xs')
151  thus ?case
152    by (cases r, auto)
153qed
154    
155lemma list_of_refs_of:
156  assumes "list_of h r xs"
157  obtains rs where "refs_of h r rs"
158using list_of_refs_of_HOL[OF assms]
159by auto
160
161lemma list_of'_refs_of'_HOL:
162  assumes "list_of' h r xs"
163  shows "\<exists>rs. refs_of' h r rs"
164proof -
165  from assms obtain rs' where "refs_of h (Ref.get h r) rs'"
166    unfolding list_of'_def by (rule list_of_refs_of)
167  thus ?thesis unfolding refs_of'_def' by auto
168qed
169
170lemma list_of'_refs_of':
171  assumes "list_of' h r xs"
172  obtains rs where "refs_of' h r rs"
173using list_of'_refs_of'_HOL[OF assms]
174by auto
175
176lemma refs_of_list_of_HOL:
177  assumes "refs_of h r rs"
178  shows "\<exists>xs. list_of h r xs"
179using assms
180proof (induct rs arbitrary: r)
181  case Nil thus ?case by auto
182next
183  case (Cons r rs')
184  thus ?case
185    by (cases r, auto)
186qed
187
188lemma refs_of_list_of:
189  assumes "refs_of h r rs"
190  obtains xs where "list_of h r xs"
191using refs_of_list_of_HOL[OF assms]
192by auto
193
194lemma refs_of'_list_of'_HOL:
195  assumes "refs_of' h r rs"
196  shows "\<exists>xs. list_of' h r xs"
197using assms
198unfolding list_of'_def refs_of'_def'
199by (auto intro: refs_of_list_of)
200
201
202lemma refs_of'_list_of':
203  assumes "refs_of' h r rs"
204  obtains xs where "list_of' h r xs"
205using refs_of'_list_of'_HOL[OF assms]
206by auto
207
208lemma refs_of'E: "refs_of' h q rs \<Longrightarrow> q \<in> set rs"
209unfolding refs_of'_def' by auto
210
211lemma list_of'_refs_of'2:
212  assumes "list_of' h r xs"
213  shows "\<exists>rs'. refs_of' h r (r#rs')"
214proof -
215  from assms obtain rs where "refs_of' h r rs" by (rule list_of'_refs_of')
216  thus ?thesis by (auto simp add: refs_of'_def')
217qed
218
219subsection \<open>More complicated properties of these predicates\<close>
220
221lemma list_of_append:
222  "list_of h n (as @ bs) \<Longrightarrow> \<exists>m. list_of h m bs"
223apply (induct as arbitrary: n)
224apply auto
225apply (case_tac n)
226apply auto
227done
228
229lemma refs_of_append: "refs_of h n (as @ bs) \<Longrightarrow> \<exists>m. refs_of h m bs"
230apply (induct as arbitrary: n)
231apply auto
232apply (case_tac n)
233apply auto
234done
235
236lemma refs_of_next:
237assumes "refs_of h (Ref.get h p) rs"
238  shows "p \<notin> set rs"
239proof (rule ccontr)
240  assume a: "\<not> (p \<notin> set rs)"
241  from this obtain as bs where split:"rs = as @ p # bs" by (fastforce dest: split_list)
242  with assms obtain q where "refs_of h q (p # bs)" by (fast dest: refs_of_append)
243  with assms split show "False"
244    by (cases q,auto dest: refs_of_is_fun)
245qed
246
247lemma refs_of_distinct: "refs_of h p rs \<Longrightarrow> distinct rs"
248proof (induct rs arbitrary: p)
249  case Nil thus ?case by simp
250next
251  case (Cons r rs')
252  thus ?case
253    by (cases p, auto simp add: refs_of_next)
254qed
255
256lemma refs_of'_distinct: "refs_of' h p rs \<Longrightarrow> distinct rs"
257  unfolding refs_of'_def'
258  by (fastforce simp add: refs_of_distinct refs_of_next)
259
260
261subsection \<open>Interaction of these predicates with our heap transitions\<close>
262
263lemma list_of_set_ref: "refs_of h q rs \<Longrightarrow> p \<notin> set rs \<Longrightarrow> list_of (Ref.set p v h) q as = list_of h q as"
264proof (induct as arbitrary: q rs)
265  case Nil thus ?case by simp
266next
267  case (Cons x xs)
268  thus ?case
269  proof (cases q)
270    case Empty thus ?thesis by auto
271  next
272    case (Node a ref)
273    from Cons(2) Node obtain rs' where 1: "refs_of h (Ref.get h ref) rs'" and rs_rs': "rs = ref # rs'" by auto
274    from Cons(3) rs_rs' have "ref \<noteq> p" by fastforce
275    hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq)
276    from rs_rs' Cons(3) have 2: "p \<notin> set rs'" by simp
277    from Cons.hyps[OF 1 2] Node ref_eq show ?thesis by simp
278  qed
279qed
280
281lemma refs_of_set_ref: "refs_of h q rs \<Longrightarrow> p \<notin> set rs \<Longrightarrow> refs_of (Ref.set p v h) q as = refs_of h q as"
282proof (induct as arbitrary: q rs)
283  case Nil thus ?case by simp
284next
285  case (Cons x xs)
286  thus ?case
287  proof (cases q)
288    case Empty thus ?thesis by auto
289  next
290    case (Node a ref)
291    from Cons(2) Node obtain rs' where 1: "refs_of h (Ref.get h ref) rs'" and rs_rs': "rs = ref # rs'" by auto
292    from Cons(3) rs_rs' have "ref \<noteq> p" by fastforce
293    hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq)
294    from rs_rs' Cons(3) have 2: "p \<notin> set rs'" by simp
295    from Cons.hyps[OF 1 2] Node ref_eq show ?thesis by auto
296  qed
297qed
298
299lemma refs_of_set_ref2: "refs_of (Ref.set p v h) q rs \<Longrightarrow> p \<notin> set rs \<Longrightarrow> refs_of (Ref.set p v h) q rs = refs_of h q rs"
300proof (induct rs arbitrary: q)
301  case Nil thus ?case by simp
302next
303  case (Cons x xs)
304  thus ?case
305  proof (cases q)
306    case Empty thus ?thesis by auto
307  next
308    case (Node a ref)
309    from Cons(2) Node have 1:"refs_of (Ref.set p v h) (Ref.get (Ref.set p v h) ref) xs" and x_ref: "x = ref" by auto
310    from Cons(3) this have "ref \<noteq> p" by fastforce
311    hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq)
312    from Cons(3) have 2: "p \<notin> set xs" by simp
313    with Cons.hyps 1 2 Node ref_eq show ?thesis
314      by simp
315  qed
316qed
317
318lemma list_of'_set_ref:
319  assumes "refs_of' h q rs"
320  assumes "p \<notin> set rs"
321  shows "list_of' (Ref.set p v h) q as = list_of' h q as"
322proof -
323  from assms have "q \<noteq> p" by (auto simp only: dest!: refs_of'E)
324  with assms show ?thesis
325    unfolding list_of'_def refs_of'_def'
326    by (auto simp add: list_of_set_ref)
327qed
328
329lemma list_of'_set_next_ref_Node[simp]:
330  assumes "list_of' h r xs"
331  assumes "Ref.get h p = Node x r'"
332  assumes "refs_of' h r rs"
333  assumes "p \<notin> set rs"
334  shows "list_of' (Ref.set p (Node x r) h) p (x#xs) = list_of' h r xs"
335using assms
336unfolding list_of'_def refs_of'_def'
337by (auto simp add: list_of_set_ref Ref.noteq_sym)
338
339lemma refs_of'_set_ref:
340  assumes "refs_of' h q rs"
341  assumes "p \<notin> set rs"
342  shows "refs_of' (Ref.set p v h) q as = refs_of' h q as"
343using assms
344proof -
345  from assms have "q \<noteq> p" by (auto simp only: dest!: refs_of'E)
346  with assms show ?thesis
347    unfolding refs_of'_def'
348    by (auto simp add: refs_of_set_ref)
349qed
350
351lemma refs_of'_set_ref2:
352  assumes "refs_of' (Ref.set p v h) q rs"
353  assumes "p \<notin> set rs"
354  shows "refs_of' (Ref.set p v h) q as = refs_of' h q as"
355using assms
356proof -
357  from assms have "q \<noteq> p" by (auto simp only: dest!: refs_of'E)
358  with assms show ?thesis
359    unfolding refs_of'_def'
360    apply auto
361    apply (subgoal_tac "prs = prsa")
362    apply (insert refs_of_set_ref2[of p v h "Ref.get h q"])
363    apply (erule_tac x="prs" in meta_allE)
364    apply auto
365    apply (auto dest: refs_of_is_fun)
366    done
367qed
368
369lemma refs_of'_set_next_ref:
370assumes "Ref.get h1 p = Node x pn"
371assumes "refs_of' (Ref.set p (Node x r1) h1) p rs"
372obtains r1s where "rs = (p#r1s)" and "refs_of' h1 r1 r1s"
373proof -
374  from assms refs_of'_distinct[OF assms(2)] have "\<exists> r1s. rs = (p # r1s) \<and> refs_of' h1 r1 r1s"
375    apply -
376    unfolding refs_of'_def'[of _ p]
377    apply (auto, frule refs_of_set_ref2) by (auto dest: Ref.noteq_sym)
378  with assms that show thesis by auto
379qed
380
381section \<open>Proving make_llist and traverse correct\<close>
382
383lemma refs_of_invariant:
384  assumes "refs_of h (r::('a::heap) node) xs"
385  assumes "\<forall>refs. refs_of h r refs \<longrightarrow> (\<forall>ref \<in> set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)"
386  shows "refs_of h' r xs"
387using assms
388proof (induct xs arbitrary: r)
389  case Nil thus ?case by simp
390next
391  case (Cons x xs')
392  from Cons(2) obtain v where Node: "r = Node v x" by (cases r, auto)
393  from Cons(2) Node have refs_of_next: "refs_of h (Ref.get h x) xs'" by simp
394  from Cons(2-3) Node have ref_eq: "Ref.get h x = Ref.get h' x" by auto
395  from ref_eq refs_of_next have 1: "refs_of h (Ref.get h' x) xs'" by simp
396  from Cons(2) Cons(3) have "\<forall>ref \<in> set xs'. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref"
397    by fastforce
398  with Cons(3) 1 have 2: "\<forall>refs. refs_of h (Ref.get h' x) refs \<longrightarrow> (\<forall>ref \<in> set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)"
399    by (fastforce dest: refs_of_is_fun)
400  from Cons.hyps[OF 1 2] have "refs_of h' (Ref.get h' x) xs'" .
401  with Node show ?case by simp
402qed
403
404lemma refs_of'_invariant:
405  assumes "refs_of' h r xs"
406  assumes "\<forall>refs. refs_of' h r refs \<longrightarrow> (\<forall>ref \<in> set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)"
407  shows "refs_of' h' r xs"
408using assms
409proof -
410  from assms obtain prs where refs:"refs_of h (Ref.get h r) prs" and xs_def: "xs = r # prs"
411    unfolding refs_of'_def' by auto
412  from xs_def assms have x_eq: "Ref.get h r = Ref.get h' r" by fastforce
413  from refs assms xs_def have 2: "\<forall>refs. refs_of h (Ref.get h r) refs \<longrightarrow>
414     (\<forall>ref\<in>set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)" 
415    by (fastforce dest: refs_of_is_fun)
416  from refs_of_invariant [OF refs 2] xs_def x_eq show ?thesis
417    unfolding refs_of'_def' by auto
418qed
419
420lemma list_of_invariant:
421  assumes "list_of h (r::('a::heap) node) xs"
422  assumes "\<forall>refs. refs_of h r refs \<longrightarrow> (\<forall>ref \<in> set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)"
423  shows "list_of h' r xs"
424using assms
425proof (induct xs arbitrary: r)
426  case Nil thus ?case by simp
427next
428  case (Cons x xs')
429
430  from Cons(2) obtain ref where Node: "r = Node x ref"
431    by (cases r, auto)
432  from Cons(2) obtain rs where rs_def: "refs_of h r rs" by (rule list_of_refs_of)
433  from Node rs_def obtain rss where refs_of: "refs_of h r (ref#rss)" and rss_def: "rs = ref#rss" by auto
434  from Cons(3) Node refs_of have ref_eq: "Ref.get h ref = Ref.get h' ref"
435    by auto
436  from Cons(2) ref_eq Node have 1: "list_of h (Ref.get h' ref) xs'" by simp
437  from refs_of Node ref_eq have refs_of_ref: "refs_of h (Ref.get h' ref) rss" by simp
438  from Cons(3) rs_def have rs_heap_eq: "\<forall>ref\<in>set rs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref" by simp
439  from refs_of_ref rs_heap_eq rss_def have 2: "\<forall>refs. refs_of h (Ref.get h' ref) refs \<longrightarrow>
440          (\<forall>ref\<in>set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)"
441    by (auto dest: refs_of_is_fun)
442  from Cons(1)[OF 1 2]
443  have "list_of h' (Ref.get h' ref) xs'" .
444  with Node show ?case
445    unfolding list_of'_def
446    by simp
447qed
448
449lemma effect_ref:
450  assumes "effect (ref v) h h' x"
451  obtains "Ref.get h' x = v"
452  and "\<not> Ref.present h x"
453  and "Ref.present h' x"
454  and "\<forall>y. Ref.present h y \<longrightarrow> Ref.get h y = Ref.get h' y"
455 (* and "lim h' = Suc (lim h)" *)
456  and "\<forall>y. Ref.present h y \<longrightarrow> Ref.present h' y"
457  using assms
458  unfolding Ref.ref_def
459  apply (elim effect_heapE)
460  unfolding Ref.alloc_def
461  apply (simp add: Let_def)
462  unfolding Ref.present_def
463  apply auto
464  unfolding Ref.get_def Ref.set_def
465  apply auto
466  done
467
468lemma make_llist:
469assumes "effect (make_llist xs) h h' r"
470shows "list_of h' r xs \<and> (\<forall>rs. refs_of h' r rs \<longrightarrow> (\<forall>ref \<in> (set rs). Ref.present h' ref))"
471using assms 
472proof (induct xs arbitrary: h h' r)
473  case Nil thus ?case by (auto elim: effect_returnE simp add: make_llist.simps)
474next
475  case (Cons x xs')
476  from Cons.prems obtain h1 r1 r' where make_llist: "effect (make_llist xs') h h1 r1"
477    and effect_refnew:"effect (ref r1) h1 h' r'" and Node: "r = Node x r'"
478    unfolding make_llist.simps
479    by (auto elim!: effect_bindE effect_returnE)
480  from Cons.hyps[OF make_llist] have list_of_h1: "list_of h1 r1 xs'" ..
481  from Cons.hyps[OF make_llist] obtain rs' where rs'_def: "refs_of h1 r1 rs'" by (auto intro: list_of_refs_of)
482  from Cons.hyps[OF make_llist] rs'_def have refs_present: "\<forall>ref\<in>set rs'. Ref.present h1 ref" by simp
483  from effect_refnew rs'_def refs_present have refs_unchanged: "\<forall>refs. refs_of h1 r1 refs \<longrightarrow>
484         (\<forall>ref\<in>set refs. Ref.present h1 ref \<and> Ref.present h' ref \<and> Ref.get h1 ref = Ref.get h' ref)"
485    by (auto elim!: effect_ref dest: refs_of_is_fun)
486  with list_of_invariant[OF list_of_h1 refs_unchanged] Node effect_refnew have fstgoal: "list_of h' r (x # xs')"
487    unfolding list_of.simps
488    by (auto elim!: effect_refE)
489  from refs_unchanged rs'_def have refs_still_present: "\<forall>ref\<in>set rs'. Ref.present h' ref" by auto
490  from refs_of_invariant[OF rs'_def refs_unchanged] refs_unchanged Node effect_refnew refs_still_present
491  have sndgoal: "\<forall>rs. refs_of h' r rs \<longrightarrow> (\<forall>ref\<in>set rs. Ref.present h' ref)"
492    by (fastforce elim!: effect_refE dest: refs_of_is_fun)
493  from fstgoal sndgoal show ?case ..
494qed
495
496lemma traverse: "list_of h n r \<Longrightarrow> effect (traverse n) h h r"
497proof (induct r arbitrary: n)
498  case Nil
499  thus ?case
500    by (auto intro: effect_returnI)
501next
502  case (Cons x xs)
503  thus ?case
504  apply (cases n, auto)
505  by (auto intro!: effect_bindI effect_returnI effect_lookupI)
506qed
507
508lemma traverse_make_llist':
509  assumes effect: "effect (make_llist xs \<bind> traverse) h h' r"
510  shows "r = xs"
511proof -
512  from effect obtain h1 r1
513    where makell: "effect (make_llist xs) h h1 r1"
514    and trav: "effect (traverse r1) h1 h' r"
515    by (auto elim!: effect_bindE)
516  from make_llist[OF makell] have "list_of h1 r1 xs" ..
517  from traverse [OF this] trav show ?thesis
518    using effect_deterministic by fastforce
519qed
520
521section \<open>Proving correctness of in-place reversal\<close>
522
523subsection \<open>Definition of in-place reversal\<close>
524
525partial_function (heap) rev' :: "('a::heap) node ref \<Rightarrow> 'a node ref \<Rightarrow> 'a node ref Heap"
526where
527  [code]: "rev' q p =
528   do {
529     v \<leftarrow> !p;
530     (case v of
531        Empty \<Rightarrow> return q
532      | Node x next \<Rightarrow>
533        do {
534          p := Node x q;
535          rev' p next
536        })
537    }"
538  
539primrec rev :: "('a:: heap) node \<Rightarrow> 'a node Heap" 
540where
541  "rev Empty = return Empty"
542| "rev (Node x n) = do { q \<leftarrow> ref Empty; p \<leftarrow> ref (Node x n); v \<leftarrow> rev' q p; !v }"
543
544subsection \<open>Correctness Proof\<close>
545
546lemma rev'_invariant:
547  assumes "effect (rev' q p) h h' v"
548  assumes "list_of' h q qs"
549  assumes "list_of' h p ps"
550  assumes "\<forall>qrs prs. refs_of' h q qrs \<and> refs_of' h p prs \<longrightarrow> set prs \<inter> set qrs = {}"
551  shows "\<exists>vs. list_of' h' v vs \<and> vs = (List.rev ps) @ qs"
552using assms
553proof (induct ps arbitrary: qs p q h)
554  case Nil
555  thus ?case
556    unfolding rev'.simps[of q p] list_of'_def
557    by (auto elim!: effect_bindE effect_lookupE effect_returnE)
558next
559  case (Cons x xs)
560  (*"LinkedList.list_of h' (get_ref v h') (List.rev xs @ x # qsa)"*)
561  from Cons(4) obtain ref where 
562    p_is_Node: "Ref.get h p = Node x ref"
563    (*and "ref_present ref h"*)
564    and list_of'_ref: "list_of' h ref xs"
565    unfolding list_of'_def by (cases "Ref.get h p", auto)
566  from p_is_Node Cons(2) have effect_rev': "effect (rev' p ref) (Ref.set p (Node x q) h) h' v"
567    by (auto simp add: rev'.simps [of q p] elim!: effect_bindE effect_lookupE effect_updateE)
568  from Cons(3) obtain qrs where qrs_def: "refs_of' h q qrs" by (elim list_of'_refs_of')
569  from Cons(4) obtain prs where prs_def: "refs_of' h p prs" by (elim list_of'_refs_of')
570  from qrs_def prs_def Cons(5) have distinct_pointers: "set qrs \<inter> set prs = {}" by fastforce
571  from qrs_def prs_def distinct_pointers refs_of'E have p_notin_qrs: "p \<notin> set qrs" by fastforce
572  from Cons(3) qrs_def this have 1: "list_of' (Ref.set p (Node x q) h) p (x#qs)"
573    unfolding list_of'_def  
574    apply (simp)
575    unfolding list_of'_def[symmetric]
576    by (simp add: list_of'_set_ref)
577  from list_of'_refs_of'2[OF Cons(4)] p_is_Node prs_def obtain refs where refs_def: "refs_of' h ref refs" and prs_refs: "prs = p # refs"
578    unfolding refs_of'_def' by auto
579  from prs_refs prs_def have p_not_in_refs: "p \<notin> set refs"
580    by (fastforce dest!: refs_of'_distinct)
581  with refs_def p_is_Node list_of'_ref have 2: "list_of' (Ref.set p (Node x q) h) ref xs"
582    by (auto simp add: list_of'_set_ref)
583  from p_notin_qrs qrs_def have refs_of1: "refs_of' (Ref.set p (Node x q) h) p (p#qrs)"
584    unfolding refs_of'_def'
585    apply (simp)
586    unfolding refs_of'_def'[symmetric]
587    by (simp add: refs_of'_set_ref)
588  from p_not_in_refs p_is_Node refs_def have refs_of2: "refs_of' (Ref.set p (Node x q) h) ref refs"
589    by (simp add: refs_of'_set_ref)
590  from p_not_in_refs refs_of1 refs_of2 distinct_pointers prs_refs have 3: "\<forall>qrs prs. refs_of' (Ref.set p (Node x q) h) p qrs \<and> refs_of' (Ref.set p (Node x q) h) ref prs \<longrightarrow> set prs \<inter> set qrs = {}"
591    apply - apply (rule allI)+ apply (rule impI) apply (erule conjE)
592    apply (drule refs_of'_is_fun) back back apply assumption
593    apply (drule refs_of'_is_fun) back back apply assumption
594    apply auto done
595  from Cons.hyps [OF effect_rev' 1 2 3] show ?case by simp
596qed
597
598
599lemma rev_correctness:
600  assumes list_of_h: "list_of h r xs"
601  assumes validHeap: "\<forall>refs. refs_of h r refs \<longrightarrow> (\<forall>r \<in> set refs. Ref.present h r)"
602  assumes effect_rev: "effect (rev r) h h' r'"
603  shows "list_of h' r' (List.rev xs)"
604using assms
605proof (cases r)
606  case Empty
607  with list_of_h effect_rev show ?thesis
608    by (auto simp add: list_of_Empty elim!: effect_returnE)
609next
610  case (Node x ps)
611  with effect_rev obtain p q h1 h2 h3 v where
612    init: "effect (ref Empty) h h1 q"
613    "effect (ref (Node x ps)) h1 h2 p"
614    and effect_rev':"effect (rev' q p) h2 h3 v"
615    and lookup: "effect (!v) h3 h' r'"
616    using rev.simps
617    by (auto elim!: effect_bindE)
618  from init have a1:"list_of' h2 q []"
619    unfolding list_of'_def
620    by (auto elim!: effect_ref)
621  from list_of_h obtain refs where refs_def: "refs_of h r refs" by (rule list_of_refs_of)
622  from validHeap init refs_def have heap_eq: "\<forall>refs. refs_of h r refs \<longrightarrow> (\<forall>ref\<in>set refs. Ref.present h ref \<and> Ref.present h2 ref \<and> Ref.get h ref = Ref.get h2 ref)"
623    by (fastforce elim!: effect_ref dest: refs_of_is_fun)
624  from list_of_invariant[OF list_of_h heap_eq] have "list_of h2 r xs" .
625  from init this Node have a2: "list_of' h2 p xs"
626    apply -
627    unfolding list_of'_def
628    apply (auto elim!: effect_refE)
629    done
630  from init have refs_of_q: "refs_of' h2 q [q]"
631    by (auto elim!: effect_ref)
632  from refs_def Node have refs_of'_ps: "refs_of' h ps refs"
633    by (auto simp add: refs_of'_def'[symmetric])
634  from validHeap refs_def have all_ref_present: "\<forall>r\<in>set refs. Ref.present h r" by simp
635  from init refs_of'_ps this
636    have heap_eq: "\<forall>refs. refs_of' h ps refs \<longrightarrow> (\<forall>ref\<in>set refs. Ref.present h ref \<and> Ref.present h2 ref \<and> Ref.get h ref = Ref.get h2 ref)"
637    by (auto elim!: effect_ref [where ?'a="'a node", where ?'b="'a node", where ?'c="'a node"] dest: refs_of'_is_fun)
638  from refs_of'_invariant[OF refs_of'_ps this] have "refs_of' h2 ps refs" .
639  with init have refs_of_p: "refs_of' h2 p (p#refs)"
640    by (auto elim!: effect_refE simp add: refs_of'_def')
641  with init all_ref_present have q_is_new: "q \<notin> set (p#refs)"
642    by (auto elim!: effect_refE intro!: Ref.noteq_I)
643  from refs_of_p refs_of_q q_is_new have a3: "\<forall>qrs prs. refs_of' h2 q qrs \<and> refs_of' h2 p prs \<longrightarrow> set prs \<inter> set qrs = {}"
644    by (fastforce simp only: list.set dest: refs_of'_is_fun)
645  from rev'_invariant [OF effect_rev' a1 a2 a3] have "list_of h3 (Ref.get h3 v) (List.rev xs)" 
646    unfolding list_of'_def by auto
647  with lookup show ?thesis
648    by (auto elim: effect_lookupE)
649qed
650
651
652section \<open>The merge function on Linked Lists\<close>
653text \<open>We also prove merge correct\<close>
654
655text\<open>First, we define merge on lists in a natural way.\<close>
656
657fun Lmerge :: "('a::ord) list \<Rightarrow> 'a list \<Rightarrow> 'a list"
658where
659  "Lmerge (x#xs) (y#ys) =
660     (if x \<le> y then x # Lmerge xs (y#ys) else y # Lmerge (x#xs) ys)"
661| "Lmerge [] ys = ys"
662| "Lmerge xs [] = xs"
663
664subsection \<open>Definition of merge function\<close>
665
666partial_function (heap) merge :: "('a::{heap, ord}) node ref \<Rightarrow> 'a node ref \<Rightarrow> 'a node ref Heap"
667where
668[code]: "merge p q = (do { v \<leftarrow> !p; w \<leftarrow> !q;
669  (case v of Empty \<Rightarrow> return q
670          | Node valp np \<Rightarrow>
671            (case w of Empty \<Rightarrow> return p
672                     | Node valq nq \<Rightarrow>
673                       if (valp \<le> valq) then do {
674                         npq \<leftarrow> merge np q;
675                         p := Node valp npq;
676                         return p }
677                       else do {
678                         pnq \<leftarrow> merge p nq;
679                         q := Node valq pnq;
680                         return q }))})"
681
682
683lemma if_return: "(if P then return x else return y) = return (if P then x else y)"
684by auto
685
686lemma if_distrib_App: "(if P then f else g) x = (if P then f x else g x)"
687by auto
688lemma redundant_if: "(if P then (if P then x else z) else y) = (if P then x else y)"
689  "(if P then x else (if P then z else y)) = (if P then x else y)"
690by auto
691
692
693
694lemma sum_distrib: "case_sum fl fr (case x of Empty \<Rightarrow> y | Node v n \<Rightarrow> (z v n)) = (case x of Empty \<Rightarrow> case_sum fl fr y | Node v n \<Rightarrow> case_sum fl fr (z v n))"
695by (cases x) auto
696
697subsection \<open>Induction refinement by applying the abstraction function to our induct rule\<close>
698
699text \<open>From our original induction rule Lmerge.induct, we derive a new rule with our list_of' predicate\<close>
700
701lemma merge_induct2:
702  assumes "list_of' h (p::'a::{heap, ord} node ref) xs"
703  assumes "list_of' h q ys"
704  assumes "\<And> ys p q. \<lbrakk> list_of' h p []; list_of' h q ys; Ref.get h p = Empty \<rbrakk> \<Longrightarrow> P p q [] ys"
705  assumes "\<And> x xs' p q pn. \<lbrakk> list_of' h p (x#xs'); list_of' h q []; Ref.get h p = Node x pn; Ref.get h q = Empty \<rbrakk> \<Longrightarrow> P p q (x#xs') []"
706  assumes "\<And> x xs' y ys' p q pn qn.
707  \<lbrakk> list_of' h p (x#xs'); list_of' h q (y#ys'); Ref.get h p = Node x pn; Ref.get h q = Node y qn;
708  x \<le> y; P pn q xs' (y#ys') \<rbrakk>
709  \<Longrightarrow> P p q (x#xs') (y#ys')"
710  assumes "\<And> x xs' y ys' p q pn qn.
711  \<lbrakk> list_of' h p (x#xs'); list_of' h q (y#ys'); Ref.get h p = Node x pn; Ref.get h q = Node y qn;
712  \<not> x \<le> y; P p qn (x#xs') ys'\<rbrakk>
713  \<Longrightarrow> P p q (x#xs') (y#ys')"
714  shows "P p q xs ys"
715using assms(1-2)
716proof (induct xs ys arbitrary: p q rule: Lmerge.induct)
717  case (2 ys)
718  from 2(1) have "Ref.get h p = Empty" unfolding list_of'_def by simp
719  with 2(1-2) assms(3) show ?case by blast
720next
721  case (3 x xs')
722  from 3(1) obtain pn where Node: "Ref.get h p = Node x pn" by (rule list_of'_Cons)
723  from 3(2) have "Ref.get h q = Empty" unfolding list_of'_def by simp
724  with Node 3(1-2) assms(4) show ?case by blast
725next
726  case (1 x xs' y ys')
727  from 1(3) obtain pn where pNode:"Ref.get h p = Node x pn"
728    and list_of'_pn: "list_of' h pn xs'" by (rule list_of'_Cons)
729  from 1(4) obtain qn where qNode:"Ref.get h q = Node y qn"
730    and  list_of'_qn: "list_of' h qn ys'" by (rule list_of'_Cons)
731  show ?case
732  proof (cases "x \<le> y")
733    case True
734    from 1(1)[OF True list_of'_pn 1(4)] assms(5) 1(3-4) pNode qNode True
735    show ?thesis by blast
736  next
737    case False
738    from 1(2)[OF False 1(3) list_of'_qn] assms(6) 1(3-4) pNode qNode False
739    show ?thesis by blast
740  qed
741qed
742
743
744text \<open>secondly, we add the effect statement in the premise, and derive the effect statements for the single cases which we then eliminate with our effect elim rules.\<close>
745  
746lemma merge_induct3: 
747assumes  "list_of' h p xs"
748assumes  "list_of' h q ys"
749assumes  "effect (merge p q) h h' r"
750assumes  "\<And> ys p q. \<lbrakk> list_of' h p []; list_of' h q ys; Ref.get h p = Empty \<rbrakk> \<Longrightarrow> P p q h h q [] ys"
751assumes  "\<And> x xs' p q pn. \<lbrakk> list_of' h p (x#xs'); list_of' h q []; Ref.get h p = Node x pn; Ref.get h q = Empty \<rbrakk> \<Longrightarrow> P p q h h p (x#xs') []"
752assumes  "\<And> x xs' y ys' p q pn qn h1 r1 h'.
753  \<lbrakk> list_of' h p (x#xs'); list_of' h q (y#ys');Ref.get h p = Node x pn; Ref.get h q = Node y qn;
754  x \<le> y; effect (merge pn q) h h1 r1 ; P pn q h h1 r1 xs' (y#ys'); h' = Ref.set p (Node x r1) h1 \<rbrakk>
755  \<Longrightarrow> P p q h h' p (x#xs') (y#ys')"
756assumes "\<And> x xs' y ys' p q pn qn h1 r1 h'.
757  \<lbrakk> list_of' h p (x#xs'); list_of' h q (y#ys'); Ref.get h p = Node x pn; Ref.get h q = Node y qn;
758  \<not> x \<le> y; effect (merge p qn) h h1 r1; P p qn h h1 r1 (x#xs') ys'; h' = Ref.set q (Node y r1) h1 \<rbrakk>
759  \<Longrightarrow> P p q h h' q (x#xs') (y#ys')"
760shows "P p q h h' r xs ys"
761using assms(3)
762proof (induct arbitrary: h' r rule: merge_induct2[OF assms(1) assms(2)])
763  case (1 ys p q)
764  from 1(3-4) have "h = h' \<and> r = q"
765    unfolding merge.simps[of p q]
766    by (auto elim!: effect_lookupE effect_bindE effect_returnE)
767  with assms(4)[OF 1(1) 1(2) 1(3)] show ?case by simp
768next
769  case (2 x xs' p q pn)
770  from 2(3-5) have "h = h' \<and> r = p"
771    unfolding merge.simps[of p q]
772    by (auto elim!: effect_lookupE effect_bindE effect_returnE)
773  with assms(5)[OF 2(1-4)] show ?case by simp
774next
775  case (3 x xs' y ys' p q pn qn)
776  from 3(3-5) 3(7) obtain h1 r1 where
777    1: "effect (merge pn q) h h1 r1" 
778    and 2: "h' = Ref.set p (Node x r1) h1 \<and> r = p"
779    unfolding merge.simps[of p q]
780    by (auto elim!: effect_lookupE effect_bindE effect_returnE effect_ifE effect_updateE)
781  from 3(6)[OF 1] assms(6) [OF 3(1-5)] 1 2 show ?case by simp
782next
783  case (4 x xs' y ys' p q pn qn)
784  from 4(3-5) 4(7) obtain h1 r1 where
785    1: "effect (merge p qn) h h1 r1" 
786    and 2: "h' = Ref.set q (Node y r1) h1 \<and> r = q"
787    unfolding merge.simps[of p q]
788    by (auto elim!: effect_lookupE effect_bindE effect_returnE effect_ifE effect_updateE)
789  from 4(6)[OF 1] assms(7) [OF 4(1-5)] 1 2 show ?case by simp
790qed
791
792
793subsection \<open>Proving merge correct\<close>
794
795text \<open>As many parts of the following three proofs are identical, we could actually move the
796same reasoning into an extended induction rule\<close>
797 
798lemma merge_unchanged:
799  assumes "refs_of' h p xs"
800  assumes "refs_of' h q ys"  
801  assumes "effect (merge p q) h h' r'"
802  assumes "set xs \<inter> set ys = {}"
803  assumes "r \<notin> set xs \<union> set ys"
804  shows "Ref.get h r = Ref.get h' r"
805proof -
806  from assms(1) obtain ps where ps_def: "list_of' h p ps" by (rule refs_of'_list_of')
807  from assms(2) obtain qs where qs_def: "list_of' h q qs" by (rule refs_of'_list_of')
808  show ?thesis using assms(1) assms(2) assms(4) assms(5)
809  proof (induct arbitrary: xs ys r rule: merge_induct3[OF ps_def qs_def assms(3)])
810    case 1 thus ?case by simp
811  next
812    case 2 thus ?case by simp
813  next
814    case (3 x xs' y ys' p q pn qn h1 r1 h' xs ys r)
815    from 3(9) 3(3) obtain pnrs
816      where pnrs_def: "xs = p#pnrs"
817      and refs_of'_pn: "refs_of' h pn pnrs"
818      by (rule refs_of'_Node)
819    with 3(12) have r_in: "r \<notin> set pnrs \<union> set ys" by auto
820    from pnrs_def 3(12) have "r \<noteq> p" by auto
821    with 3(11) 3(12) pnrs_def refs_of'_distinct[OF 3(9)] have p_in: "p \<notin> set pnrs \<union> set ys" by auto
822    from 3(11) pnrs_def have no_inter: "set pnrs \<inter> set ys = {}" by auto
823    from 3(7)[OF refs_of'_pn 3(10) this p_in] 3(3) have p_is_Node: "Ref.get h1 p = Node x pn"
824      by simp
825    from 3(7)[OF refs_of'_pn 3(10) no_inter r_in] 3(8) \<open>r \<noteq> p\<close> show ?case
826      by simp
827  next
828    case (4 x xs' y ys' p q pn qn h1 r1 h' xs ys r)
829    from 4(10) 4(4) obtain qnrs
830      where qnrs_def: "ys = q#qnrs"
831      and refs_of'_qn: "refs_of' h qn qnrs"
832      by (rule refs_of'_Node)
833    with 4(12) have r_in: "r \<notin> set xs \<union> set qnrs" by auto
834    from qnrs_def 4(12) have "r \<noteq> q" by auto
835    with 4(11) 4(12) qnrs_def refs_of'_distinct[OF 4(10)] have q_in: "q \<notin> set xs \<union> set qnrs" by auto
836    from 4(11) qnrs_def have no_inter: "set xs \<inter> set qnrs = {}" by auto
837    from 4(7)[OF 4(9) refs_of'_qn this q_in] 4(4) have q_is_Node: "Ref.get h1 q = Node y qn" by simp
838    from 4(7)[OF 4(9) refs_of'_qn no_inter r_in] 4(8) \<open>r \<noteq> q\<close> show ?case
839      by simp
840  qed
841qed
842
843lemma refs_of'_merge:
844  assumes "refs_of' h p xs"
845  assumes "refs_of' h q ys"
846  assumes "effect (merge p q) h h' r"
847  assumes "set xs \<inter> set ys = {}"
848  assumes "refs_of' h' r rs"
849  shows "set rs \<subseteq> set xs \<union> set ys"
850proof -
851  from assms(1) obtain ps where ps_def: "list_of' h p ps" by (rule refs_of'_list_of')
852  from assms(2) obtain qs where qs_def: "list_of' h q qs" by (rule refs_of'_list_of')
853  show ?thesis using assms(1) assms(2) assms(4) assms(5)
854  proof (induct arbitrary: xs ys rs rule: merge_induct3[OF ps_def qs_def assms(3)])
855    case 1
856    from 1(5) 1(7) have "rs = ys" by (fastforce simp add: refs_of'_is_fun)
857    thus ?case by auto
858  next
859    case 2
860    from 2(5) 2(8) have "rs = xs" by (auto simp add: refs_of'_is_fun)
861    thus ?case by auto
862  next
863    case (3 x xs' y ys' p q pn qn h1 r1 h' xs ys rs)
864    from 3(9) 3(3) obtain pnrs
865      where pnrs_def: "xs = p#pnrs"
866      and refs_of'_pn: "refs_of' h pn pnrs"
867      by (rule refs_of'_Node)
868    from 3(10) 3(9) 3(11) pnrs_def refs_of'_distinct[OF 3(9)] have p_in: "p \<notin> set pnrs \<union> set ys" by auto
869    from 3(11) pnrs_def have no_inter: "set pnrs \<inter> set ys = {}" by auto
870    from merge_unchanged[OF refs_of'_pn 3(10) 3(6) no_inter p_in] have p_stays: "Ref.get h1 p = Ref.get h p" ..
871    from 3 p_stays obtain r1s
872      where rs_def: "rs = p#r1s" and refs_of'_r1:"refs_of' h1 r1 r1s"
873      by (auto elim: refs_of'_set_next_ref)
874    from 3(7)[OF refs_of'_pn 3(10) no_inter refs_of'_r1] rs_def pnrs_def show ?case by auto
875  next
876    case (4 x xs' y ys' p q pn qn h1 r1 h' xs ys rs)
877    from 4(10) 4(4) obtain qnrs
878      where qnrs_def: "ys = q#qnrs"
879      and refs_of'_qn: "refs_of' h qn qnrs"
880      by (rule refs_of'_Node)
881    from 4(10) 4(9) 4(11) qnrs_def refs_of'_distinct[OF 4(10)] have q_in: "q \<notin> set xs \<union> set qnrs" by auto
882    from 4(11) qnrs_def have no_inter: "set xs \<inter> set qnrs = {}" by auto
883    from merge_unchanged[OF 4(9) refs_of'_qn 4(6) no_inter q_in] have q_stays: "Ref.get h1 q = Ref.get h q" ..
884    from 4 q_stays obtain r1s
885      where rs_def: "rs = q#r1s" and refs_of'_r1:"refs_of' h1 r1 r1s"
886      by (auto elim: refs_of'_set_next_ref)
887    from 4(7)[OF 4(9) refs_of'_qn no_inter refs_of'_r1] rs_def qnrs_def show ?case by auto
888  qed
889qed
890
891lemma
892  assumes "list_of' h p xs"
893  assumes "list_of' h q ys"
894  assumes "effect (merge p q) h h' r"
895  assumes "\<forall>qrs prs. refs_of' h q qrs \<and> refs_of' h p prs \<longrightarrow> set prs \<inter> set qrs = {}"
896  shows "list_of' h' r (Lmerge xs ys)"
897using assms(4)
898proof (induct rule: merge_induct3[OF assms(1-3)])
899  case 1
900  thus ?case by simp
901next
902  case 2
903  thus ?case by simp
904next
905  case (3 x xs' y ys' p q pn qn h1 r1 h')
906  from 3(1) obtain prs where prs_def: "refs_of' h p prs" by (rule list_of'_refs_of')
907  from 3(2) obtain qrs where qrs_def: "refs_of' h q qrs" by (rule list_of'_refs_of')
908  from prs_def 3(3) obtain pnrs
909    where pnrs_def: "prs = p#pnrs"
910    and refs_of'_pn: "refs_of' h pn pnrs"
911    by (rule refs_of'_Node)
912  from prs_def qrs_def 3(9) pnrs_def refs_of'_distinct[OF prs_def] have p_in: "p \<notin> set pnrs \<union> set qrs" by fastforce
913  from prs_def qrs_def 3(9) pnrs_def have no_inter: "set pnrs \<inter> set qrs = {}" by fastforce
914  from no_inter refs_of'_pn qrs_def have no_inter2: "\<forall>qrs prs. refs_of' h q qrs \<and> refs_of' h pn prs \<longrightarrow> set prs \<inter> set qrs = {}"
915    by (fastforce dest: refs_of'_is_fun)
916  from merge_unchanged[OF refs_of'_pn qrs_def 3(6) no_inter p_in] have p_stays: "Ref.get h1 p = Ref.get h p" ..
917  from 3(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of')
918  from refs_of'_merge[OF refs_of'_pn qrs_def 3(6) no_inter this] p_in have p_rs: "p \<notin> set rs" by auto
919  with 3(7)[OF no_inter2] 3(1-5) 3(8) p_rs rs_def p_stays
920  show ?case by (auto simp: list_of'_set_ref)
921next
922  case (4 x xs' y ys' p q pn qn h1 r1 h')
923  from 4(1) obtain prs where prs_def: "refs_of' h p prs" by (rule list_of'_refs_of')
924  from 4(2) obtain qrs where qrs_def: "refs_of' h q qrs" by (rule list_of'_refs_of')
925  from qrs_def 4(4) obtain qnrs
926    where qnrs_def: "qrs = q#qnrs"
927    and refs_of'_qn: "refs_of' h qn qnrs"
928    by (rule refs_of'_Node)
929  from prs_def qrs_def 4(9) qnrs_def refs_of'_distinct[OF qrs_def] have q_in: "q \<notin> set prs \<union> set qnrs" by fastforce
930  from prs_def qrs_def 4(9) qnrs_def have no_inter: "set prs \<inter> set qnrs = {}" by fastforce
931  from no_inter refs_of'_qn prs_def have no_inter2: "\<forall>qrs prs. refs_of' h qn qrs \<and> refs_of' h p prs \<longrightarrow> set prs \<inter> set qrs = {}"
932    by (fastforce dest: refs_of'_is_fun)
933  from merge_unchanged[OF prs_def refs_of'_qn 4(6) no_inter q_in] have q_stays: "Ref.get h1 q = Ref.get h q" ..
934  from 4(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of')
935  from refs_of'_merge[OF prs_def refs_of'_qn 4(6) no_inter this] q_in have q_rs: "q \<notin> set rs" by auto
936  with 4(7)[OF no_inter2] 4(1-5) 4(8) q_rs rs_def q_stays
937  show ?case by (auto simp: list_of'_set_ref)
938qed
939
940section \<open>Code generation\<close>
941
942text \<open>A simple example program\<close>
943
944definition test_1 where "test_1 = (do { ll_xs \<leftarrow> make_llist [1..(15::int)]; xs \<leftarrow> traverse ll_xs; return xs })" 
945definition test_2 where "test_2 = (do { ll_xs \<leftarrow> make_llist [1..(15::int)]; ll_ys \<leftarrow> rev ll_xs; ys \<leftarrow> traverse ll_ys; return ys })"
946
947definition test_3 where "test_3 =
948  (do {
949    ll_xs \<leftarrow> make_llist (filter (%n. n mod 2 = 0) [2..8]);
950    ll_ys \<leftarrow> make_llist (filter (%n. n mod 2 = 1) [5..11]);
951    r \<leftarrow> ref ll_xs;
952    q \<leftarrow> ref ll_ys;
953    p \<leftarrow> merge r q;
954    ll_zs \<leftarrow> !p;
955    zs \<leftarrow> traverse ll_zs;
956    return zs
957  })"
958
959code_reserved SML upto
960
961ML_val \<open>@{code test_1} ()\<close>
962ML_val \<open>@{code test_2} ()\<close>
963ML_val \<open>@{code test_3} ()\<close>
964
965export_code test_1 test_2 test_3 checking SML SML_imp OCaml? OCaml_imp? Haskell? Scala Scala_imp
966
967end
968