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