merged
authorwenzelm
Tue, 06 Jul 2010 10:02:24 +0200
changeset 37727 8244558af8a5
parent 37726 17b05b104390 (diff)
parent 37713 c82cf6e11669 (current diff)
child 37728 5d2b3e827371
child 37733 bf6c1216db43
merged
--- a/src/HOL/Imperative_HOL/Array.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/Array.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -8,24 +8,132 @@
 imports Heap_Monad
 begin
 
+subsection {* Primitive layer *}
+
+definition 
+  array_present :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> bool" where
+  "array_present a h \<longleftrightarrow> addr_of_array a < lim h"
+
+definition
+  get_array :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> 'a list" where
+  "get_array a h = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))"
+
+definition
+  set_array :: "'a\<Colon>heap array \<Rightarrow> 'a list \<Rightarrow> heap \<Rightarrow> heap" where
+  "set_array a x = 
+  arrays_update (\<lambda>h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))"
+
+definition array :: "'a list \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where
+  "array xs h = (let
+     l = lim h;
+     r = Array l;
+     h'' = set_array r xs (h\<lparr>lim := l + 1\<rparr>)
+   in (r, h''))"
+
+definition length :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> nat" where
+  "length a h = List.length (get_array a h)"
+  
+definition change :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
+  "change a i x h = set_array a ((get_array a h)[i:=x]) h"
+
+text {* Properties of imperative arrays *}
+
+text {* FIXME: Does there exist a "canonical" array axiomatisation in
+the literature?  *}
+
+definition noteq_arrs :: "('a\<Colon>heap) array \<Rightarrow> ('b\<Colon>heap) array \<Rightarrow> bool" (infix "=!!=" 70) where
+  "r =!!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_array r \<noteq> addr_of_array s"
+
+lemma noteq_arrs_sym: "a =!!= b \<Longrightarrow> b =!!= a"
+  and unequal_arrs [simp]: "a \<noteq> a' \<longleftrightarrow> a =!!= a'"
+  unfolding noteq_arrs_def by auto
+
+lemma noteq_arrs_irrefl: "r =!!= r \<Longrightarrow> False"
+  unfolding noteq_arrs_def by auto
+
+lemma present_new_arr: "array_present a h \<Longrightarrow> a =!!= fst (array xs h)"
+  by (simp add: array_present_def noteq_arrs_def array_def Let_def)
+
+lemma array_get_set_eq [simp]: "get_array r (set_array r x h) = x"
+  by (simp add: get_array_def set_array_def o_def)
+
+lemma array_get_set_neq [simp]: "r =!!= s \<Longrightarrow> get_array r (set_array s x h) = get_array r h"
+  by (simp add: noteq_arrs_def get_array_def set_array_def)
+
+lemma set_array_same [simp]:
+  "set_array r x (set_array r y h) = set_array r x h"
+  by (simp add: set_array_def)
+
+lemma array_set_set_swap:
+  "r =!!= r' \<Longrightarrow> set_array r x (set_array r' x' h) = set_array r' x' (set_array r x h)"
+  by (simp add: Let_def expand_fun_eq noteq_arrs_def set_array_def)
+
+lemma get_array_change_eq [simp]:
+  "get_array a (change a i v h) = (get_array a h) [i := v]"
+  by (simp add: change_def)
+
+lemma nth_change_array_neq_array [simp]:
+  "a =!!= b \<Longrightarrow> get_array a (change b j v h) ! i = get_array a h ! i"
+  by (simp add: change_def noteq_arrs_def)
+
+lemma get_arry_array_change_elem_neqIndex [simp]:
+  "i \<noteq> j \<Longrightarrow> get_array a (change a j v h) ! i = get_array a h ! i"
+  by simp
+
+lemma length_change [simp]: 
+  "length a (change b i v h) = length a h"
+  by (simp add: change_def length_def set_array_def get_array_def)
+
+lemma change_swap_neqArray:
+  "a =!!= a' \<Longrightarrow> 
+  change a i v (change a' i' v' h) 
+  = change a' i' v' (change a i v h)"
+apply (unfold change_def)
+apply simp
+apply (subst array_set_set_swap, assumption)
+apply (subst array_get_set_neq)
+apply (erule noteq_arrs_sym)
+apply (simp)
+done
+
+lemma change_swap_neqIndex:
+  "\<lbrakk> i \<noteq> i' \<rbrakk> \<Longrightarrow> change a i v (change a i' v' h) = change a i' v' (change a i v h)"
+  by (auto simp add: change_def array_set_set_swap list_update_swap)
+
+lemma get_array_init_array_list:
+  "get_array (fst (array ls h)) (snd (array ls' h)) = ls'"
+  by (simp add: Let_def split_def array_def)
+
+lemma set_array:
+  "set_array (fst (array ls h))
+     new_ls (snd (array ls h))
+       = snd (array new_ls h)"
+  by (simp add: Let_def split_def array_def)
+
+lemma array_present_change [simp]: 
+  "array_present a (change b i v h) = array_present a h"
+  by (simp add: change_def array_present_def set_array_def get_array_def)
+
+
+
 subsection {* Primitives *}
 
 definition
   new :: "nat \<Rightarrow> 'a\<Colon>heap \<Rightarrow> 'a array Heap" where
-  [code del]: "new n x = Heap_Monad.heap (Heap.array n x)"
+  [code del]: "new n x = Heap_Monad.heap (Array.array (replicate n x))"
 
 definition
   of_list :: "'a\<Colon>heap list \<Rightarrow> 'a array Heap" where
-  [code del]: "of_list xs = Heap_Monad.heap (Heap.array_of_list xs)"
+  [code del]: "of_list xs = Heap_Monad.heap (Array.array xs)"
 
 definition
-  length :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where
-  [code del]: "length arr = Heap_Monad.heap (\<lambda>h. (Heap.length arr h, h))"
+  len :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where
+  [code del]: "len arr = Heap_Monad.heap (\<lambda>h. (Array.length arr h, h))"
 
 definition
   nth :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a Heap"
 where
-  [code del]: "nth a i = (do len \<leftarrow> length a;
+  [code del]: "nth a i = (do len \<leftarrow> len a;
                  (if i < len
                      then Heap_Monad.heap (\<lambda>h. (get_array a h ! i, h))
                      else raise ''array lookup: index out of range'')
@@ -34,9 +142,9 @@
 definition
   upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a\<Colon>heap array Heap"
 where
-  [code del]: "upd i x a = (do len \<leftarrow> length a;
+  [code del]: "upd i x a = (do len \<leftarrow> len a;
                       (if i < len
-                           then Heap_Monad.heap (\<lambda>h. (a, Heap.upd a i x h))
+                           then Heap_Monad.heap (\<lambda>h. (a, change a i x h))
                            else raise ''array update: index out of range'')
                    done)" 
 
@@ -73,7 +181,7 @@
   freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap"
 where
   "freeze a = (do
-     n \<leftarrow> length a;
+     n \<leftarrow> len a;
      mapM (nth a) [0..<n]
    done)"
 
@@ -81,19 +189,18 @@
    map :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
 where
   "map f a = (do
-     n \<leftarrow> length a;
+     n \<leftarrow> len a;
      mapM (\<lambda>n. map_entry n f a) [0..<n];
      return a
    done)"
 
-hide_const (open) new map -- {* avoid clashed with some popular names *}
 
 
 subsection {* Properties *}
 
 lemma array_make [code]:
   "Array.new n x = make n (\<lambda>_. x)"
-  by (rule Heap_eqI) (simp add: make_def new_def array_of_list_replicate map_replicate_trivial of_list_def)
+  by (rule Heap_eqI) (simp add: make_def new_def map_replicate_trivial of_list_def)
 
 lemma array_of_list_make [code]:
   "of_list xs = make (List.length xs) (\<lambda>n. xs ! n)"
@@ -125,12 +232,12 @@
   "Array.make n f = Array.make' (Code_Numeral.of_nat n) (f o Code_Numeral.nat_of)"
   by (simp add: make'_def o_def)
 
-definition length' where
-  [code del]: "length' a = Array.length a \<guillemotright>= (\<lambda>n. return (Code_Numeral.of_nat n))"
-hide_const (open) length'
+definition len' where
+  [code del]: "len' a = Array.len a \<guillemotright>= (\<lambda>n. return (Code_Numeral.of_nat n))"
+hide_const (open) len'
 lemma [code]:
-  "Array.length a = Array.length' a \<guillemotright>= (\<lambda>i. return (Code_Numeral.nat_of i))"
-  by (simp add: length'_def)
+  "Array.len a = Array.len' a \<guillemotright>= (\<lambda>i. return (Code_Numeral.nat_of i))"
+  by (simp add: len'_def)
 
 definition nth' where
   [code del]: "nth' a = Array.nth a o Code_Numeral.nat_of"
@@ -154,7 +261,7 @@
 code_const Array.new' (SML "(fn/ ()/ =>/ Array.array/ ((_),/ (_)))")
 code_const Array.of_list' (SML "(fn/ ()/ =>/ Array.fromList/ _)")
 code_const Array.make' (SML "(fn/ ()/ =>/ Array.tabulate/ ((_),/ (_)))")
-code_const Array.length' (SML "(fn/ ()/ =>/ Array.length/ _)")
+code_const Array.len' (SML "(fn/ ()/ =>/ Array.length/ _)")
 code_const Array.nth' (SML "(fn/ ()/ =>/ Array.sub/ ((_),/ (_)))")
 code_const Array.upd' (SML "(fn/ ()/ =>/ Array.update/ ((_),/ (_),/ (_)))")
 
@@ -167,7 +274,7 @@
 code_const Array (OCaml "failwith/ \"bare Array\"")
 code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ (Big'_int.int'_of'_big'_int/ _)/ _)")
 code_const Array.of_list' (OCaml "(fun/ ()/ ->/ Array.of'_list/ _)")
-code_const Array.length' (OCaml "(fun/ ()/ ->/ Big'_int.big'_int'_of'_int/ (Array.length/ _))")
+code_const Array.len' (OCaml "(fun/ ()/ ->/ Big'_int.big'_int'_of'_int/ (Array.length/ _))")
 code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ (Big'_int.int'_of'_big'_int/ _))")
 code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ (Big'_int.int'_of'_big'_int/ _)/ _)")
 
@@ -180,8 +287,10 @@
 code_const Array (Haskell "error/ \"bare Array\"")
 code_const Array.new' (Haskell "Heap.newArray/ (0,/ _)")
 code_const Array.of_list' (Haskell "Heap.newListArray/ (0,/ _)")
-code_const Array.length' (Haskell "Heap.lengthArray")
+code_const Array.len' (Haskell "Heap.lengthArray")
 code_const Array.nth' (Haskell "Heap.readArray")
 code_const Array.upd' (Haskell "Heap.writeArray")
 
+hide_const (open) new map -- {* avoid clashed with some popular names *}
+
 end
--- a/src/HOL/Imperative_HOL/Heap.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/Heap.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -1,4 +1,4 @@
-(*  Title:      HOL/Library/Heap.thy
+(*  Title:      HOL/Imperative_HOL/Heap.thy
     Author:     John Matthews, Galois Connections; Alexander Krauss, TU Muenchen
 *)
 
@@ -14,8 +14,6 @@
 
 class heap = typerep + countable
 
-text {* Instances for common HOL types *}
-
 instance nat :: heap ..
 
 instance prod :: (heap, heap) heap ..
@@ -34,47 +32,26 @@
 
 instance String.literal :: heap ..
 
-text {* Reflected types themselves are heap-representable *}
-
-instantiation typerep :: countable
-begin
-
-fun to_nat_typerep :: "typerep \<Rightarrow> nat" where
-  "to_nat_typerep (Typerep.Typerep c ts) = to_nat (to_nat c, to_nat (map to_nat_typerep ts))"
-
-instance
-proof (rule countable_classI)
-  fix t t' :: typerep and ts
-  have "(\<forall>t'. to_nat_typerep t = to_nat_typerep t' \<longrightarrow> t = t')
-    \<and> (\<forall>ts'. map to_nat_typerep ts = map to_nat_typerep ts' \<longrightarrow> ts = ts')"
-  proof (induct rule: typerep.induct)
-    case (Typerep c ts) show ?case
-    proof (rule allI, rule impI)
-      fix t'
-      assume hyp: "to_nat_typerep (Typerep.Typerep c ts) = to_nat_typerep t'"
-      then obtain c' ts' where t': "t' = (Typerep.Typerep c' ts')"
-        by (cases t') auto
-      with Typerep hyp have "c = c'" and "ts = ts'" by simp_all
-      with t' show "Typerep.Typerep c ts = t'" by simp
-    qed
-  next
-    case Nil_typerep then show ?case by simp
-  next
-    case (Cons_typerep t ts) then show ?case by auto
-  qed
-  then have "to_nat_typerep t = to_nat_typerep t' \<Longrightarrow> t = t'" by auto
-  moreover assume "to_nat_typerep t = to_nat_typerep t'"
-  ultimately show "t = t'" by simp
-qed
-
-end
-
 instance typerep :: heap ..
 
 
 subsection {* A polymorphic heap with dynamic arrays and references *}
 
+text {*
+  References and arrays are developed in parallel,
+  but keeping them separate makes some later proofs simpler.
+*}
+
 types addr = nat -- "untyped heap references"
+types heap_rep = nat -- "representable values"
+
+record heap =
+  arrays :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep list"
+  refs :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep"
+  lim  :: addr
+
+definition empty :: heap where
+  "empty = \<lparr>arrays = (\<lambda>_ _. []), refs = (\<lambda>_ _. 0), lim = 0\<rparr>"
 
 datatype 'a array = Array addr
 datatype 'a ref = Ref addr -- "note the phantom type 'a "
@@ -99,6 +76,8 @@
 instance ref :: (type) countable
   by (rule countable_classI [of addr_of_ref]) simp
 
+text {* Syntactic convenience *}
+
 setup {*
   Sign.add_const_constraint (@{const_name Array}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap array"})
   #> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap ref"})
@@ -106,335 +85,6 @@
   #> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a\<Colon>heap ref \<Rightarrow> nat"})
 *}
 
-types heap_rep = nat -- "representable values"
-
-record heap =
-  arrays :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep list"
-  refs :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep"
-  lim  :: addr
-
-definition empty :: heap where
-  "empty = \<lparr>arrays = (\<lambda>_. undefined), refs = (\<lambda>_. undefined), lim = 0\<rparr>" -- "why undefined?"
-
-
-subsection {* Imperative references and arrays *}
-
-text {*
-  References and arrays are developed in parallel,
-  but keeping them separate makes some later proofs simpler.
-*}
-
-subsubsection {* Primitive operations *}
-
-definition
-  new_ref :: "heap \<Rightarrow> ('a\<Colon>heap) ref \<times> heap" where
-  "new_ref h = (let l = lim h in (Ref l, h\<lparr>lim := l + 1\<rparr>))"
-
-definition
-  new_array :: "heap \<Rightarrow> ('a\<Colon>heap) array \<times> heap" where
-  "new_array h = (let l = lim h in (Array l, h\<lparr>lim := l + 1\<rparr>))"
-
-definition
-  ref_present :: "'a\<Colon>heap ref \<Rightarrow> heap \<Rightarrow> bool" where
-  "ref_present r h \<longleftrightarrow> addr_of_ref r < lim h"
-
-definition 
-  array_present :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> bool" where
-  "array_present a h \<longleftrightarrow> addr_of_array a < lim h"
-
-definition
-  get_ref :: "'a\<Colon>heap ref \<Rightarrow> heap \<Rightarrow> 'a" where
-  "get_ref r h = from_nat (refs h (TYPEREP('a)) (addr_of_ref r))"
-
-definition
-  get_array :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> 'a list" where
-  "get_array a h = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))"
-
-definition
-  set_ref :: "'a\<Colon>heap ref \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
-  "set_ref r x = 
-  refs_update (\<lambda>h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r:=to_nat x))))"
-
-definition
-  set_array :: "'a\<Colon>heap array \<Rightarrow> 'a list \<Rightarrow> heap \<Rightarrow> heap" where
-  "set_array a x = 
-  arrays_update (\<lambda>h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))"
-
-subsubsection {* Interface operations *}
-
-definition
-  ref :: "'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap ref \<times> heap" where
-  "ref x h = (let (r, h') = new_ref h;
-                   h''    = set_ref r x h'
-         in (r, h''))"
-
-definition
-  array :: "nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where
-  "array n x h = (let (r, h') = new_array h;
-                       h'' = set_array r (replicate n x) h'
-        in (r, h''))"
-
-definition
-  array_of_list :: "'a list \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where
-  "array_of_list xs h = (let (r, h') = new_array h;
-           h'' = set_array r xs h'
-        in (r, h''))"  
-
-definition
-  upd :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
-  "upd a i x h = set_array a ((get_array a h)[i:=x]) h"
-
-definition
-  length :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> nat" where
-  "length a h = size (get_array a h)"
-
-definition
-  array_ran :: "('a\<Colon>heap) option array \<Rightarrow> heap \<Rightarrow> 'a set" where
-  "array_ran a h = {e. Some e \<in> set (get_array a h)}"
-    -- {*FIXME*}
-
-
-subsubsection {* Reference equality *}
-
-text {* 
-  The following relations are useful for comparing arrays and references.
-*}
-
-definition
-  noteq_refs :: "('a\<Colon>heap) ref \<Rightarrow> ('b\<Colon>heap) ref \<Rightarrow> bool" (infix "=!=" 70)
-where
-  "r =!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_ref r \<noteq> addr_of_ref s"
-
-definition
-  noteq_arrs :: "('a\<Colon>heap) array \<Rightarrow> ('b\<Colon>heap) array \<Rightarrow> bool" (infix "=!!=" 70)
-where
-  "r =!!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_array r \<noteq> addr_of_array s"
-
-lemma noteq_refs_sym: "r =!= s \<Longrightarrow> s =!= r"
-  and noteq_arrs_sym: "a =!!= b \<Longrightarrow> b =!!= a"
-  and unequal_refs [simp]: "r \<noteq> r' \<longleftrightarrow> r =!= r'" -- "same types!"
-  and unequal_arrs [simp]: "a \<noteq> a' \<longleftrightarrow> a =!!= a'"
-unfolding noteq_refs_def noteq_arrs_def by auto
-
-lemma noteq_refs_irrefl: "r =!= r \<Longrightarrow> False"
-  unfolding noteq_refs_def by auto
-
-lemma present_new_ref: "ref_present r h \<Longrightarrow> r =!= fst (ref v h)"
-  by (simp add: ref_present_def new_ref_def ref_def Let_def noteq_refs_def)
-
-lemma present_new_arr: "array_present a h \<Longrightarrow> a =!!= fst (array v x h)"
-  by (simp add: array_present_def noteq_arrs_def new_array_def array_def Let_def)
-
-
-subsubsection {* Properties of heap containers *}
-
-text {* Properties of imperative arrays *}
-
-text {* FIXME: Does there exist a "canonical" array axiomatisation in
-the literature?  *}
-
-lemma array_get_set_eq [simp]: "get_array r (set_array r x h) = x"
-  by (simp add: get_array_def set_array_def o_def)
-
-lemma array_get_set_neq [simp]: "r =!!= s \<Longrightarrow> get_array r (set_array s x h) = get_array r h"
-  by (simp add: noteq_arrs_def get_array_def set_array_def)
-
-lemma set_array_same [simp]:
-  "set_array r x (set_array r y h) = set_array r x h"
-  by (simp add: set_array_def)
-
-lemma array_set_set_swap:
-  "r =!!= r' \<Longrightarrow> set_array r x (set_array r' x' h) = set_array r' x' (set_array r x h)"
-  by (simp add: Let_def expand_fun_eq noteq_arrs_def set_array_def)
-
-lemma array_ref_set_set_swap:
-  "set_array r x (set_ref r' x' h) = set_ref r' x' (set_array r x h)"
-  by (simp add: Let_def expand_fun_eq set_array_def set_ref_def)
-
-lemma get_array_upd_eq [simp]:
-  "get_array a (upd a i v h) = (get_array a h) [i := v]"
-  by (simp add: upd_def)
-
-lemma nth_upd_array_neq_array [simp]:
-  "a =!!= b \<Longrightarrow> get_array a (upd b j v h) ! i = get_array a h ! i"
-  by (simp add: upd_def noteq_arrs_def)
-
-lemma get_arry_array_upd_elem_neqIndex [simp]:
-  "i \<noteq> j \<Longrightarrow> get_array a (upd a j v h) ! i = get_array a h ! i"
-  by simp
-
-lemma length_upd_eq [simp]: 
-  "length a (upd a i v h) = length a h" 
-  by (simp add: length_def upd_def)
-
-lemma length_upd_neq [simp]: 
-  "length a (upd b i v h) = length a h"
-  by (simp add: upd_def length_def set_array_def get_array_def)
-
-lemma upd_swap_neqArray:
-  "a =!!= a' \<Longrightarrow> 
-  upd a i v (upd a' i' v' h) 
-  = upd a' i' v' (upd a i v h)"
-apply (unfold upd_def)
-apply simp
-apply (subst array_set_set_swap, assumption)
-apply (subst array_get_set_neq)
-apply (erule noteq_arrs_sym)
-apply (simp)
-done
-
-lemma upd_swap_neqIndex:
-  "\<lbrakk> i \<noteq> i' \<rbrakk> \<Longrightarrow> upd a i v (upd a i' v' h) = upd a i' v' (upd a i v h)"
-by (auto simp add: upd_def array_set_set_swap list_update_swap)
-
-lemma get_array_init_array_list:
-  "get_array (fst (array_of_list ls h)) (snd (array_of_list ls' h)) = ls'"
-  by (simp add: Let_def split_def array_of_list_def)
-
-lemma set_array:
-  "set_array (fst (array_of_list ls h))
-     new_ls (snd (array_of_list ls h))
-       = snd (array_of_list new_ls h)"
-  by (simp add: Let_def split_def array_of_list_def)
-
-lemma array_present_upd [simp]: 
-  "array_present a (upd b i v h) = array_present a h"
-  by (simp add: upd_def array_present_def set_array_def get_array_def)
-
-lemma array_of_list_replicate:
-  "array_of_list (replicate n x) = array n x"
-  by (simp add: expand_fun_eq array_of_list_def array_def)
-
-
-text {* Properties of imperative references *}
-
-lemma next_ref_fresh [simp]:
-  assumes "(r, h') = new_ref h"
-  shows "\<not> ref_present r h"
-  using assms by (cases h) (auto simp add: new_ref_def ref_present_def Let_def)
-
-lemma next_ref_present [simp]:
-  assumes "(r, h') = new_ref h"
-  shows "ref_present r h'"
-  using assms by (cases h) (auto simp add: new_ref_def ref_present_def Let_def)
-
-lemma ref_get_set_eq [simp]: "get_ref r (set_ref r x h) = x"
-  by (simp add: get_ref_def set_ref_def)
-
-lemma ref_get_set_neq [simp]: "r =!= s \<Longrightarrow> get_ref r (set_ref s x h) = get_ref r h"
-  by (simp add: noteq_refs_def get_ref_def set_ref_def)
-
-(* FIXME: We need some infrastructure to infer that locally generated
-  new refs (by new_ref(_no_init), new_array(')) are distinct
-  from all existing refs.
-*)
-
-lemma ref_set_get: "set_ref r (get_ref r h) h = h"
-apply (simp add: set_ref_def get_ref_def)
-oops
-
-lemma set_ref_same[simp]:
-  "set_ref r x (set_ref r y h) = set_ref r x h"
-  by (simp add: set_ref_def)
-
-lemma ref_set_set_swap:
-  "r =!= r' \<Longrightarrow> set_ref r x (set_ref r' x' h) = set_ref r' x' (set_ref r x h)"
-  by (simp add: Let_def expand_fun_eq noteq_refs_def set_ref_def)
-
-lemma ref_new_set: "fst (ref v (set_ref r v' h)) = fst (ref v h)"
-  by (simp add: ref_def new_ref_def set_ref_def Let_def)
-
-lemma ref_get_new [simp]:
-  "get_ref (fst (ref v h)) (snd (ref v' h)) = v'"
-  by (simp add: ref_def Let_def split_def)
-
-lemma ref_set_new [simp]:
-  "set_ref (fst (ref v h)) new_v (snd (ref v h)) = snd (ref new_v h)"
-  by (simp add: ref_def Let_def split_def)
-
-lemma ref_get_new_neq: "r =!= (fst (ref v h)) \<Longrightarrow> 
-  get_ref r (snd (ref v h)) = get_ref r h"
-  by (simp add: get_ref_def set_ref_def ref_def Let_def new_ref_def noteq_refs_def)
-
-lemma lim_set_ref [simp]:
-  "lim (set_ref r v h) = lim h"
-  by (simp add: set_ref_def)
-
-lemma ref_present_new_ref [simp]: 
-  "ref_present r h \<Longrightarrow> ref_present r (snd (ref v h))"
-  by (simp add: new_ref_def ref_present_def ref_def Let_def)
-
-lemma ref_present_set_ref [simp]:
-  "ref_present r (set_ref r' v h) = ref_present r h"
-  by (simp add: set_ref_def ref_present_def)
-
-lemma noteq_refsI: "\<lbrakk> ref_present r h; \<not>ref_present r' h \<rbrakk>  \<Longrightarrow> r =!= r'"
-  unfolding noteq_refs_def ref_present_def
-  by auto
-
-lemma array_ranI: "\<lbrakk> Some b = get_array a h ! i; i < Heap.length a h \<rbrakk> \<Longrightarrow> b \<in> array_ran a h"
-unfolding array_ran_def Heap.length_def by simp
-
-lemma array_ran_upd_array_Some:
-  assumes "cl \<in> array_ran a (Heap.upd a i (Some b) h)"
-  shows "cl \<in> array_ran a h \<or> cl = b"
-proof -
-  have "set (get_array a h[i := Some b]) \<subseteq> insert (Some b) (set (get_array a h))" by (rule set_update_subset_insert)
-  with assms show ?thesis 
-    unfolding array_ran_def Heap.upd_def by fastsimp
-qed
-
-lemma array_ran_upd_array_None:
-  assumes "cl \<in> array_ran a (Heap.upd a i None h)"
-  shows "cl \<in> array_ran a h"
-proof -
-  have "set (get_array a h[i := None]) \<subseteq> insert None (set (get_array a h))" by (rule set_update_subset_insert)
-  with assms show ?thesis
-    unfolding array_ran_def Heap.upd_def by auto
-qed
-
-
-text {* Non-interaction between imperative array and imperative references *}
-
-lemma get_array_set_ref [simp]: "get_array a (set_ref r v h) = get_array a h"
-  by (simp add: get_array_def set_ref_def)
-
-lemma nth_set_ref [simp]: "get_array a (set_ref r v h) ! i = get_array a h ! i"
-  by simp
-
-lemma get_ref_upd [simp]: "get_ref r (upd a i v h) = get_ref r h"
-  by (simp add: get_ref_def set_array_def upd_def)
-
-lemma new_ref_upd: "fst (ref v (upd a i v' h)) = fst (ref v h)"
-  by (simp add: set_array_def get_array_def Let_def ref_new_set upd_def ref_def  new_ref_def)
-
-text {*not actually true ???*}
-lemma upd_set_ref_swap: "upd a i v (set_ref r v' h) = set_ref r v' (upd a i v h)"
-apply (case_tac a)
-apply (simp add: Let_def upd_def)
-apply auto
-oops
-
-lemma length_new_ref[simp]: 
-  "length a (snd (ref v h)) = length a h"
-  by (simp add: get_array_def set_ref_def length_def new_ref_def ref_def Let_def)
-
-lemma get_array_new_ref [simp]: 
-  "get_array a (snd (ref v h)) = get_array a h"
-  by (simp add: new_ref_def ref_def set_ref_def get_array_def Let_def)
-
-lemma ref_present_upd [simp]: 
-  "ref_present r (upd a i v h) = ref_present r h"
-  by (simp add: upd_def ref_present_def set_array_def get_array_def)
-
-lemma array_present_set_ref [simp]:
-  "array_present a (set_ref r v h) = array_present a h"
-  by (simp add: array_present_def set_ref_def)
-
-lemma array_present_new_ref [simp]:
-  "array_present a h \<Longrightarrow> array_present a (snd (ref v h))"
-  by (simp add: array_present_def new_ref_def ref_def Let_def)
-
-hide_const (open) empty array array_of_list upd length ref
+hide_const (open) empty
 
 end
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -349,8 +349,6 @@
 lemmas MREC_rule = mrec.MREC_rule
 lemmas MREC_pinduct = mrec.MREC_pinduct
 
-hide_const (open) heap execute
-
 
 subsection {* Code generator setup *}
 
@@ -365,8 +363,6 @@
 
 code_datatype raise' -- {* avoid @{const "Heap"} formally *}
 
-hide_const (open) raise'
-
 
 subsubsection {* SML and OCaml *}
 
@@ -493,4 +489,6 @@
 code_const return (Haskell "return")
 code_const Heap_Monad.raise' (Haskell "error/ _")
 
+hide_const (open) Heap heap execute raise'
+
 end
--- a/src/HOL/Imperative_HOL/Ref.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/Ref.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -5,7 +5,7 @@
 header {* Monadic references *}
 
 theory Ref
-imports Heap_Monad
+imports Array
 begin
 
 text {*
@@ -14,45 +14,177 @@
   and http://www.smlnj.org/doc/Conversion/top-level-comparison.html
 *}
 
+subsection {* Primitive layer *}
+
+definition present :: "heap \<Rightarrow> 'a\<Colon>heap ref \<Rightarrow> bool" where
+  "present h r \<longleftrightarrow> addr_of_ref r < lim h"
+
+definition get :: "heap \<Rightarrow> 'a\<Colon>heap ref \<Rightarrow> 'a" where
+  "get h = from_nat \<circ> refs h TYPEREP('a) \<circ> addr_of_ref"
+
+definition set :: "'a\<Colon>heap ref \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
+  "set r x = refs_update
+    (\<lambda>h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r := to_nat x))))"
+
+definition alloc :: "'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap ref \<times> heap" where
+  "alloc x h = (let
+     l = lim h;
+     r = Ref l
+   in (r, set r x (h\<lparr>lim := l + 1\<rparr>)))"
+
+definition noteq :: "'a\<Colon>heap ref \<Rightarrow> 'b\<Colon>heap ref \<Rightarrow> bool" (infix "=!=" 70) where
+  "r =!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_ref r \<noteq> addr_of_ref s"
+
+lemma noteq_sym: "r =!= s \<Longrightarrow> s =!= r"
+  and unequal [simp]: "r \<noteq> r' \<longleftrightarrow> r =!= r'" -- "same types!"
+  by (auto simp add: noteq_def)
+
+lemma noteq_irrefl: "r =!= r \<Longrightarrow> False"
+  by (auto simp add: noteq_def)
+
+lemma present_alloc_neq: "present h r \<Longrightarrow> r =!= fst (alloc v h)"
+  by (simp add: present_def alloc_def noteq_def Let_def)
+
+lemma next_fresh [simp]:
+  assumes "(r, h') = alloc x h"
+  shows "\<not> present h r"
+  using assms by (cases h) (auto simp add: alloc_def present_def Let_def)
+
+lemma next_present [simp]:
+  assumes "(r, h') = alloc x h"
+  shows "present h' r"
+  using assms by (cases h) (auto simp add: alloc_def set_def present_def Let_def)
+
+lemma get_set_eq [simp]:
+  "get (set r x h) r = x"
+  by (simp add: get_def set_def)
+
+lemma get_set_neq [simp]:
+  "r =!= s \<Longrightarrow> get (set s x h) r = get h r"
+  by (simp add: noteq_def get_def set_def)
+
+lemma set_same [simp]:
+  "set r x (set r y h) = set r x h"
+  by (simp add: set_def)
+
+lemma set_set_swap:
+  "r =!= r' \<Longrightarrow> set r x (set r' x' h) = set r' x' (set r x h)"
+  by (simp add: noteq_def set_def expand_fun_eq)
+
+lemma alloc_set:
+  "fst (alloc x (set r x' h)) = fst (alloc x h)"
+  by (simp add: alloc_def set_def Let_def)
+
+lemma get_alloc [simp]:
+  "get (snd (alloc x h)) (fst (alloc x' h)) = x"
+  by (simp add: alloc_def Let_def)
+
+lemma set_alloc [simp]:
+  "set (fst (alloc v h)) v' (snd (alloc v h)) = snd (alloc v' h)"
+  by (simp add: alloc_def Let_def)
+
+lemma get_alloc_neq: "r =!= fst (alloc v h) \<Longrightarrow> 
+  get (snd (alloc v h)) r  = get h r"
+  by (simp add: get_def set_def alloc_def Let_def noteq_def)
+
+lemma lim_set [simp]:
+  "lim (set r v h) = lim h"
+  by (simp add: set_def)
+
+lemma present_alloc [simp]: 
+  "present h r \<Longrightarrow> present (snd (alloc v h)) r"
+  by (simp add: present_def alloc_def Let_def)
+
+lemma present_set [simp]:
+  "present (set r v h) = present h"
+  by (simp add: present_def expand_fun_eq)
+
+lemma noteq_I:
+  "present h r \<Longrightarrow> \<not> present h r' \<Longrightarrow> r =!= r'"
+  by (auto simp add: noteq_def present_def)
+
+
 subsection {* Primitives *}
 
-definition
-  new :: "'a\<Colon>heap \<Rightarrow> 'a ref Heap" where
-  [code del]: "new v = Heap_Monad.heap (Heap.ref v)"
+definition ref :: "'a\<Colon>heap \<Rightarrow> 'a ref Heap" where
+  [code del]: "ref v = Heap_Monad.heap (alloc v)"
 
-definition
-  lookup :: "'a\<Colon>heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
-  [code del]: "lookup r = Heap_Monad.heap (\<lambda>h. (get_ref r h, h))"
+definition lookup :: "'a\<Colon>heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
+  [code del]: "lookup r = Heap_Monad.heap (\<lambda>h. (get h r, h))"
 
-definition
-  update :: "'a ref \<Rightarrow> ('a\<Colon>heap) \<Rightarrow> unit Heap" ("_ := _" 62) where
-  [code del]: "update r e = Heap_Monad.heap (\<lambda>h. ((), set_ref r e h))"
+definition update :: "'a ref \<Rightarrow> 'a\<Colon>heap \<Rightarrow> unit Heap" ("_ := _" 62) where
+  [code del]: "update r v = Heap_Monad.heap (\<lambda>h. ((), set r v h))"
 
 
 subsection {* Derivates *}
 
-definition
-  change :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap"
-where
-  "change f r = (do x \<leftarrow> ! r;
-                    let y = f x;
-                    r := y;
-                    return y
-                 done)"
-
-hide_const (open) new lookup update change
+definition change :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap" where
+  "change f r = (do
+     x \<leftarrow> ! r;
+     let y = f x;
+     r := y;
+     return y
+   done)"
 
 
 subsection {* Properties *}
 
 lemma lookup_chain:
   "(!r \<guillemotright> f) = f"
-  by (cases f)
-    (auto simp add: Let_def bindM_def lookup_def expand_fun_eq)
+  by (rule Heap_eqI) (simp add: lookup_def)
 
 lemma update_change [code]:
-  "r := e = Ref.change (\<lambda>_. e) r \<guillemotright> return ()"
-  by (auto simp add: change_def lookup_chain)
+  "r := e = change (\<lambda>_. e) r \<guillemotright> return ()"
+  by (rule Heap_eqI) (simp add: change_def lookup_chain)
+
+
+text {* Non-interaction between imperative array and imperative references *}
+
+lemma get_array_set [simp]:
+  "get_array a (set r v h) = get_array a h"
+  by (simp add: get_array_def set_def)
+
+lemma nth_set [simp]:
+  "get_array a (set r v h) ! i = get_array a h ! i"
+  by simp
+
+lemma get_change [simp]:
+  "get (Array.change a i v h) r  = get h r"
+  by (simp add: get_def Array.change_def set_array_def)
+
+lemma alloc_change:
+  "fst (alloc v (Array.change a i v' h)) = fst (alloc v h)"
+  by (simp add: Array.change_def get_array_def set_array_def alloc_def Let_def)
+
+lemma change_set_swap:
+  "Array.change a i v (set r v' h) = set r v' (Array.change a i v h)"
+  by (simp add: Array.change_def get_array_def set_array_def set_def)
+
+lemma length_alloc [simp]: 
+  "Array.length a (snd (alloc v h)) = Array.length a h"
+  by (simp add: Array.length_def get_array_def alloc_def set_def Let_def)
+
+lemma get_array_alloc [simp]: 
+  "get_array a (snd (alloc v h)) = get_array a h"
+  by (simp add: get_array_def alloc_def set_def Let_def)
+
+lemma present_change [simp]: 
+  "present (Array.change a i v h) = present h"
+  by (simp add: Array.change_def set_array_def expand_fun_eq present_def)
+
+lemma array_present_set [simp]:
+  "array_present a (set r v h) = array_present a h"
+  by (simp add: array_present_def set_def)
+
+lemma array_present_alloc [simp]:
+  "array_present a h \<Longrightarrow> array_present a (snd (alloc v h))"
+  by (simp add: array_present_def alloc_def Let_def)
+
+lemma set_array_set_swap:
+  "set_array a xs (set r x' h) = set r x' (set_array a xs h)"
+  by (simp add: set_array_def set_def)
+
+hide_const (open) present get set alloc lookup update change
 
 
 subsection {* Code generator setup *}
@@ -61,7 +193,7 @@
 
 code_type ref (SML "_/ Unsynchronized.ref")
 code_const Ref (SML "raise/ (Fail/ \"bare Ref\")")
-code_const Ref.new (SML "(fn/ ()/ =>/ Unsynchronized.ref/ _)")
+code_const ref (SML "(fn/ ()/ =>/ Unsynchronized.ref/ _)")
 code_const Ref.lookup (SML "(fn/ ()/ =>/ !/ _)")
 code_const Ref.update (SML "(fn/ ()/ =>/ _/ :=/ _)")
 
@@ -72,7 +204,7 @@
 
 code_type ref (OCaml "_/ ref")
 code_const Ref (OCaml "failwith/ \"bare Ref\")")
-code_const Ref.new (OCaml "(fn/ ()/ =>/ ref/ _)")
+code_const ref (OCaml "(fn/ ()/ =>/ ref/ _)")
 code_const Ref.lookup (OCaml "(fn/ ()/ =>/ !/ _)")
 code_const Ref.update (OCaml "(fn/ ()/ =>/ _/ :=/ _)")
 
@@ -83,7 +215,7 @@
 
 code_type ref (Haskell "Heap.STRef/ Heap.RealWorld/ _")
 code_const Ref (Haskell "error/ \"bare Ref\"")
-code_const Ref.new (Haskell "Heap.newSTRef")
+code_const ref (Haskell "Heap.newSTRef")
 code_const Ref.lookup (Haskell "Heap.readSTRef")
 code_const Ref.update (Haskell "Heap.writeSTRef")
 
--- a/src/HOL/Imperative_HOL/Relational.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/Relational.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -91,29 +91,29 @@
 subsection {* Elimination rules for array commands *}
 
 lemma crel_length:
-  assumes "crel (length a) h h' r"
-  obtains "h = h'" "r = Heap.length a h'"
+  assumes "crel (len a) h h' r"
+  obtains "h = h'" "r = Array.length a h'"
   using assms
-  unfolding length_def
+  unfolding Array.len_def
   by (elim crel_heap) simp
 
 (* Strong version of the lemma for this operation is missing *) 
 lemma crel_new_weak:
   assumes "crel (Array.new n v) h h' r"
   obtains "get_array r h' = List.replicate n v"
-  using assms unfolding  Array.new_def
-  by (elim crel_heap) (auto simp:Heap.array_def Let_def split_def)
+  using assms unfolding Array.new_def
+  by (elim crel_heap) (auto simp: array_def Let_def split_def)
 
 lemma crel_nth[consumes 1]:
   assumes "crel (nth a i) h h' r"
-  obtains "r = (get_array a h) ! i" "h = h'" "i < Heap.length a h"
+  obtains "r = (get_array a h) ! i" "h = h'" "i < Array.length a h"
   using assms
   unfolding nth_def
   by (auto elim!: crelE crel_if crel_raise crel_length crel_heap)
 
 lemma crel_upd[consumes 1]:
   assumes "crel (upd i v a) h h' r"
-  obtains "r = a" "h' = Heap.upd a i v h"
+  obtains "r = a" "h' = Array.change a i v h"
   using assms
   unfolding upd_def
   by (elim crelE crel_if crel_return crel_raise
@@ -129,14 +129,14 @@
 
 lemma crel_map_entry:
   assumes "crel (Array.map_entry i f a) h h' r"
-  obtains "r = a" "h' = Heap.upd a i (f (get_array a h ! i)) h"
+  obtains "r = a" "h' = Array.change a i (f (get_array a h ! i)) h"
   using assms
   unfolding Array.map_entry_def
   by (elim crelE crel_upd crel_nth) auto
 
 lemma crel_swap:
   assumes "crel (Array.swap i x a) h h' r"
-  obtains "r = get_array a h ! i" "h' = Heap.upd a i x h"
+  obtains "r = get_array a h ! i" "h' = Array.change a i x h"
   using assms
   unfolding Array.swap_def
   by (elim crelE crel_upd crel_nth crel_return) auto
@@ -160,25 +160,25 @@
 
 lemma crel_mapM_nth:
   assumes
-  "crel (mapM (Array.nth a) [Heap.length a h - n..<Heap.length a h]) h h' xs"
-  assumes "n \<le> Heap.length a h"
-  shows "h = h' \<and> xs = drop (Heap.length a h - n) (get_array a h)"
+  "crel (mapM (Array.nth a) [Array.length a h - n..<Array.length a h]) h h' xs"
+  assumes "n \<le> Array.length a h"
+  shows "h = h' \<and> xs = drop (Array.length a h - n) (get_array a h)"
 using assms
 proof (induct n arbitrary: xs h h')
   case 0 thus ?case
-    by (auto elim!: crel_return simp add: Heap.length_def)
+    by (auto elim!: crel_return simp add: Array.length_def)
 next
   case (Suc n)
-  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
+  from Suc(3) have "[Array.length a h - Suc n..<Array.length a h] = (Array.length a h - Suc n)#[Array.length a h - n..<Array.length a h]"
     by (simp add: upt_conv_Cons')
   with Suc(2) obtain r where
-    crel_mapM: "crel (mapM (Array.nth a) [Heap.length a h - n..<Heap.length a h]) h h' r"
-    and xs_def: "xs = get_array a h ! (Heap.length a h - Suc n) # r"
+    crel_mapM: "crel (mapM (Array.nth a) [Array.length a h - n..<Array.length a h]) h h' r"
+    and xs_def: "xs = get_array a h ! (Array.length a h - Suc n) # r"
     by (auto elim!: crelE crel_nth crel_return)
-  from Suc(3) have "Heap.length a h - n = Suc (Heap.length a h - Suc n)" 
+  from Suc(3) have "Array.length a h - n = Suc (Array.length a h - Suc n)" 
     by arith
   with Suc.hyps[OF crel_mapM] xs_def show ?case
-    unfolding Heap.length_def
+    unfolding Array.length_def
     by (auto simp add: nth_drop')
 qed
 
@@ -186,17 +186,17 @@
   assumes "crel (Array.freeze a) h h' xs"
   obtains "h = h'" "xs = get_array a h"
 proof 
-  from assms have "crel (mapM (Array.nth a) [0..<Heap.length a h]) h h' xs"
+  from assms have "crel (mapM (Array.nth a) [0..<Array.length a h]) h h' xs"
     unfolding freeze_def
     by (auto elim: crelE crel_length)
-  hence "crel (mapM (Array.nth a) [(Heap.length a h - Heap.length a h)..<Heap.length a h]) h h' xs"
+  hence "crel (mapM (Array.nth a) [(Array.length a h - Array.length a h)..<Array.length a h]) h h' xs"
     by simp
   from crel_mapM_nth[OF this] show "h = h'" and "xs = get_array a h" by auto
 qed
 
 lemma crel_mapM_map_entry_remains:
-  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h h' r"
-  assumes "i < Heap.length a h - n"
+  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) h h' r"
+  assumes "i < Array.length a h - n"
   shows "get_array a h ! i = get_array a h' ! i"
 using assms
 proof (induct n arbitrary: h h' r)
@@ -205,23 +205,23 @@
     by (auto elim: crel_return)
 next
   case (Suc n)
-  let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h"
-  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
+  let ?h1 = "Array.change a (Array.length a h - Suc n) (f (get_array a h ! (Array.length a h - Suc n))) h"
+  from Suc(3) have "[Array.length a h - Suc n..<Array.length a h] = (Array.length a h - Suc n)#[Array.length a h - n..<Array.length a h]"
     by (simp add: upt_conv_Cons')
   from Suc(2) this obtain r where
-    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) ?h1 h' r"
+    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) ?h1 h' r"
     by (auto simp add: elim!: crelE crel_map_entry crel_return)
-  have length_remains: "Heap.length a ?h1 = Heap.length a h" by simp
-  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a ?h1 - n..<Heap.length a ?h1]) ?h1 h' r"
+  have length_remains: "Array.length a ?h1 = Array.length a h" by simp
+  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a ?h1 - n..<Array.length a ?h1]) ?h1 h' r"
     by simp
   from Suc(1)[OF this] length_remains Suc(3) show ?case by simp
 qed
 
 lemma crel_mapM_map_entry_changes:
-  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h h' r"
-  assumes "n \<le> Heap.length a h"  
-  assumes "i \<ge> Heap.length a h - n"
-  assumes "i < Heap.length a h"
+  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) h h' r"
+  assumes "n \<le> Array.length a h"  
+  assumes "i \<ge> Array.length a h - n"
+  assumes "i < Array.length a h"
   shows "get_array a h' ! i = f (get_array a h ! i)"
 using assms
 proof (induct n arbitrary: h h' r)
@@ -230,54 +230,54 @@
     by (auto elim!: crel_return)
 next
   case (Suc n)
-  let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h"
-  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
+  let ?h1 = "Array.change a (Array.length a h - Suc n) (f (get_array a h ! (Array.length a h - Suc n))) h"
+  from Suc(3) have "[Array.length a h - Suc n..<Array.length a h] = (Array.length a h - Suc n)#[Array.length a h - n..<Array.length a h]"
     by (simp add: upt_conv_Cons')
   from Suc(2) this obtain r where
-    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) ?h1 h' r"
+    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) ?h1 h' r"
     by (auto simp add: elim!: crelE crel_map_entry crel_return)
-  have length_remains: "Heap.length a ?h1 = Heap.length a h" by simp
-  from Suc(3) have less: "Heap.length a h - Suc n < Heap.length a h - n" by arith
-  from Suc(3) have less2: "Heap.length a h - Suc n < Heap.length a h" by arith
-  from Suc(4) length_remains have cases: "i = Heap.length a ?h1 - Suc n \<or> i \<ge> Heap.length a ?h1 - n" by arith
-  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a ?h1 - n..<Heap.length a ?h1]) ?h1 h' r"
+  have length_remains: "Array.length a ?h1 = Array.length a h" by simp
+  from Suc(3) have less: "Array.length a h - Suc n < Array.length a h - n" by arith
+  from Suc(3) have less2: "Array.length a h - Suc n < Array.length a h" by arith
+  from Suc(4) length_remains have cases: "i = Array.length a ?h1 - Suc n \<or> i \<ge> Array.length a ?h1 - n" by arith
+  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a ?h1 - n..<Array.length a ?h1]) ?h1 h' r"
     by simp
   from Suc(1)[OF this] cases Suc(3) Suc(5) length_remains
-    crel_mapM_map_entry_remains[OF this, of "Heap.length a h - Suc n", symmetric] less less2
+    crel_mapM_map_entry_remains[OF this, of "Array.length a h - Suc n", symmetric] less less2
   show ?case
-    by (auto simp add: nth_list_update_eq Heap.length_def)
+    by (auto simp add: nth_list_update_eq Array.length_def)
 qed
 
 lemma crel_mapM_map_entry_length:
-  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h h' r"
-  assumes "n \<le> Heap.length a h"
-  shows "Heap.length a h' = Heap.length a h"
+  assumes "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) h h' r"
+  assumes "n \<le> Array.length a h"
+  shows "Array.length a h' = Array.length a h"
 using assms
 proof (induct n arbitrary: h h' r)
   case 0
   thus ?case by (auto elim!: crel_return)
 next
   case (Suc n)
-  let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h"
-  from Suc(3) have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
+  let ?h1 = "Array.change a (Array.length a h - Suc n) (f (get_array a h ! (Array.length a h - Suc n))) h"
+  from Suc(3) have "[Array.length a h - Suc n..<Array.length a h] = (Array.length a h - Suc n)#[Array.length a h - n..<Array.length a h]"
     by (simp add: upt_conv_Cons')
   from Suc(2) this obtain r where 
-    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) ?h1 h' r"
+    crel_mapM: "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) ?h1 h' r"
     by (auto elim!: crelE crel_map_entry crel_return)
-  have length_remains: "Heap.length a ?h1 = Heap.length a h" by simp
-  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a ?h1 - n..<Heap.length a ?h1]) ?h1 h' r"
+  have length_remains: "Array.length a ?h1 = Array.length a h" by simp
+  from crel_mapM have crel_mapM': "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a ?h1 - n..<Array.length a ?h1]) ?h1 h' r"
     by simp
   from Suc(1)[OF this] length_remains Suc(3) show ?case by simp
 qed
 
 lemma crel_mapM_map_entry:
-assumes "crel (mapM (\<lambda>n. map_entry n f a) [0..<Heap.length a h]) h h' r"
+assumes "crel (mapM (\<lambda>n. map_entry n f a) [0..<Array.length a h]) h h' r"
   shows "get_array a h' = List.map f (get_array a h)"
 proof -
-  from assms have "crel (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - Heap.length a h..<Heap.length a h]) h h' r" by simp
+  from assms have "crel (mapM (\<lambda>n. map_entry n f a) [Array.length a h - Array.length a h..<Array.length a h]) h h' r" by simp
   from crel_mapM_map_entry_length[OF this]
   crel_mapM_map_entry_changes[OF this] show ?thesis
-    unfolding Heap.length_def
+    unfolding Array.length_def
     by (auto intro: nth_equalityI)
 qed
 
@@ -311,44 +311,42 @@
 extends h' h x \<Longrightarrow> lim h' = Suc (lim h)
 *)
 
-lemma crel_Ref_new:
-  assumes "crel (Ref.new v) h h' x"
-  obtains "get_ref x h' = v"
-  and "\<not> ref_present x h"
-  and "ref_present x h'"
-  and "\<forall>y. ref_present y h \<longrightarrow> get_ref y h = get_ref y h'"
+lemma crel_ref:
+  assumes "crel (ref v) h h' x"
+  obtains "Ref.get h' x = v"
+  and "\<not> Ref.present h x"
+  and "Ref.present h' x"
+  and "\<forall>y. Ref.present h y \<longrightarrow> Ref.get h y = Ref.get h' y"
  (* and "lim h' = Suc (lim h)" *)
-  and "\<forall>y. ref_present y h \<longrightarrow> ref_present y h'"
+  and "\<forall>y. Ref.present h y \<longrightarrow> Ref.present h' y"
   using assms
-  unfolding Ref.new_def
+  unfolding Ref.ref_def
   apply (elim crel_heap)
-  unfolding Heap.ref_def
+  unfolding Ref.alloc_def
   apply (simp add: Let_def)
-  unfolding Heap.new_ref_def
-  apply (simp add: Let_def)
-  unfolding ref_present_def
+  unfolding Ref.present_def
   apply auto
-  unfolding get_ref_def set_ref_def
+  unfolding Ref.get_def Ref.set_def
   apply auto
   done
 
 lemma crel_lookup:
-  assumes "crel (!ref) h h' r"
-  obtains "h = h'" "r = get_ref ref h"
+  assumes "crel (!r') h h' r"
+  obtains "h = h'" "r = Ref.get h r'"
 using assms
 unfolding Ref.lookup_def
 by (auto elim: crel_heap)
 
 lemma crel_update:
-  assumes "crel (ref := v) h h' r"
-  obtains "h' = set_ref ref v h" "r = ()"
+  assumes "crel (r' := v) h h' r"
+  obtains "h' = Ref.set r' v h" "r = ()"
 using assms
 unfolding Ref.update_def
 by (auto elim: crel_heap)
 
 lemma crel_change:
-  assumes "crel (Ref.change f ref) h h' r"
-  obtains "h' = set_ref ref (f (get_ref ref h)) h" "r = f (get_ref ref h)"
+  assumes "crel (Ref.change f r') h h' r"
+  obtains "h' = Ref.set r' (f (Ref.get h r')) h" "r = f (Ref.get h r')"
 using assms
 unfolding Ref.change_def Let_def
 by (auto elim!: crelE crel_lookup crel_update crel_return)
@@ -433,22 +431,22 @@
 subsection {* Introduction rules for array commands *}
 
 lemma crel_lengthI:
-  shows "crel (length a) h h (Heap.length a h)"
-  unfolding length_def
+  shows "crel (Array.len a) h h (Array.length a h)"
+  unfolding len_def
   by (rule crel_heapI') auto
 
 (* thm crel_newI for Array.new is missing *)
 
 lemma crel_nthI:
-  assumes "i < Heap.length a h"
+  assumes "i < Array.length a h"
   shows "crel (nth a i) h h ((get_array a h) ! i)"
   using assms
   unfolding nth_def
   by (auto intro!: crelI crel_ifI crel_raiseI crel_lengthI crel_heapI')
 
 lemma crel_updI:
-  assumes "i < Heap.length a h"
-  shows "crel (upd i v a) h (Heap.upd a i v h) a"
+  assumes "i < Array.length a h"
+  shows "crel (upd i v a) h (Array.change a i v h) a"
   using assms
   unfolding upd_def
   by (auto intro!: crelI crel_ifI crel_returnI crel_raiseI
@@ -469,15 +467,15 @@
 subsubsection {* Introduction rules for reference commands *}
 
 lemma crel_lookupI:
-  shows "crel (!ref) h h (get_ref ref h)"
+  shows "crel (!r) h h (Ref.get h r)"
   unfolding lookup_def by (auto intro!: crel_heapI')
 
 lemma crel_updateI:
-  shows "crel (ref := v) h (set_ref ref v h) ()"
+  shows "crel (r := v) h (Ref.set r v h) ()"
   unfolding update_def by (auto intro!: crel_heapI')
 
 lemma crel_changeI: 
-  shows "crel (Ref.change f ref) h (set_ref ref (f (get_ref ref h)) h) (f (get_ref ref h))"
+  shows "crel (Ref.change f r) h (Ref.set r (f (Ref.get h r)) h) (f (Ref.get h r))"
 unfolding change_def Let_def by (auto intro!: crelI crel_returnI crel_lookupI crel_updateI)
 
 subsubsection {* Introduction rules for the assert command *}
@@ -589,8 +587,8 @@
 subsection {* Introduction rules for array commands *}
 
 lemma noError_length:
-  shows "noError (Array.length a) h"
-  unfolding length_def
+  shows "noError (Array.len a) h"
+  unfolding len_def
   by (intro noError_heap)
 
 lemma noError_new:
@@ -598,14 +596,14 @@
 unfolding Array.new_def by (intro noError_heap)
 
 lemma noError_upd:
-  assumes "i < Heap.length a h"
+  assumes "i < Array.length a h"
   shows "noError (Array.upd i v a) h"
   using assms
   unfolding upd_def
   by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length)
 
 lemma noError_nth:
-assumes "i < Heap.length a h"
+assumes "i < Array.length a h"
   shows "noError (Array.nth a i) h"
   using assms
   unfolding nth_def
@@ -616,14 +614,14 @@
   unfolding of_list_def by (rule noError_heap)
 
 lemma noError_map_entry:
-  assumes "i < Heap.length a h"
+  assumes "i < Array.length a h"
   shows "noError (map_entry i f a) h"
   using assms
   unfolding map_entry_def
   by (auto elim: crel_nth intro!: noErrorI noError_nth noError_upd)
 
 lemma noError_swap:
-  assumes "i < Heap.length a h"
+  assumes "i < Array.length a h"
   shows "noError (swap i x a) h"
   using assms
   unfolding swap_def
@@ -646,42 +644,42 @@
   noError_nth crel_nthI elim: crel_length)
 
 lemma noError_mapM_map_entry:
-  assumes "n \<le> Heap.length a h"
-  shows "noError (mapM (\<lambda>n. map_entry n f a) [Heap.length a h - n..<Heap.length a h]) h"
+  assumes "n \<le> Array.length a h"
+  shows "noError (mapM (\<lambda>n. map_entry n f a) [Array.length a h - n..<Array.length a h]) h"
 using assms
 proof (induct n arbitrary: h)
   case 0
   thus ?case by (auto intro: noError_return)
 next
   case (Suc n)
-  from Suc.prems have "[Heap.length a h - Suc n..<Heap.length a h] = (Heap.length a h - Suc n)#[Heap.length a h - n..<Heap.length a h]"
+  from Suc.prems have "[Array.length a h - Suc n..<Array.length a h] = (Array.length a h - Suc n)#[Array.length a h - n..<Array.length a h]"
     by (simp add: upt_conv_Cons')
-  with Suc.hyps[of "(Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h)"] Suc.prems show ?case
+  with Suc.hyps[of "(Array.change a (Array.length a h - Suc n) (f (get_array a h ! (Array.length a h - Suc n))) h)"] Suc.prems show ?case
     by (auto simp add: intro!: noErrorI noError_return noError_map_entry elim: crel_map_entry)
 qed
 
 lemma noError_map:
   shows "noError (Array.map f a) h"
-using noError_mapM_map_entry[of "Heap.length a h" a h]
+using noError_mapM_map_entry[of "Array.length a h" a h]
 unfolding Array.map_def
 by (auto intro: noErrorI noError_length noError_return elim!: crel_length)
 
 subsection {* Introduction rules for the reference commands *}
 
 lemma noError_Ref_new:
-  shows "noError (Ref.new v) h"
-unfolding Ref.new_def by (intro noError_heap)
+  shows "noError (ref v) h"
+  unfolding Ref.ref_def by (intro noError_heap)
 
 lemma noError_lookup:
-  shows "noError (!ref) h"
+  shows "noError (!r) h"
   unfolding lookup_def by (intro noError_heap)
 
 lemma noError_update:
-  shows "noError (ref := v) h"
+  shows "noError (r := v) h"
   unfolding update_def by (intro noError_heap)
 
 lemma noError_change:
-  shows "noError (Ref.change f ref) h"
+  shows "noError (Ref.change f r) h"
   unfolding Ref.change_def Let_def by (intro noErrorI noError_lookup noError_update noError_return)
 
 subsection {* Introduction rules for the assert command *}
@@ -698,7 +696,7 @@
 lemmas crel_elim_all =
   crelE crelE' crel_return crel_raise crel_if crel_option_case
   crel_length crel_new_weak crel_nth crel_upd crel_of_list_weak crel_map_entry crel_swap crel_make_weak crel_freeze crel_map_weak
-  crel_Ref_new crel_lookup crel_update crel_change
+  crel_ref crel_lookup crel_update crel_change
   crel_assert
 
 lemmas crel_intro_all =
--- a/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -27,7 +27,7 @@
   = multiset_of (get_array a h)"
   using assms
   unfolding swap_def
-  by (auto simp add: Heap.length_def multiset_of_swap dest: sym [of _ "h'"] elim!: crelE crel_nth crel_return crel_upd)
+  by (auto simp add: Array.length_def multiset_of_swap dest: sym [of _ "h'"] elim!: crelE crel_nth crel_return crel_upd)
 
 function part1 :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
 where
@@ -101,7 +101,7 @@
 
 lemma part_length_remains:
   assumes "crel (part1 a l r p) h h' rs"
-  shows "Heap.length a h = Heap.length a h'"
+  shows "Array.length a h = Array.length a h'"
 using assms
 proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
   case (1 a l r p h h' rs)
@@ -207,7 +207,7 @@
         by (elim crelE crel_nth crel_if crel_return) auto
       from swp False have "get_array a h1 ! r \<ge> p"
         unfolding swap_def
-        by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return)
+        by (auto simp add: Array.length_def elim!: crelE crel_nth crel_upd crel_return)
       with part_outer_remains [OF rec2] lr have a_r: "get_array a h' ! r \<ge> p"
         by fastsimp
       have "\<forall>i. (i \<le> r = (i = r \<or> i \<le> r - 1))" by arith
@@ -243,7 +243,7 @@
 
 lemma partition_length_remains:
   assumes "crel (partition a l r) h h' rs"
-  shows "Heap.length a h = Heap.length a h'"
+  shows "Array.length a h = Array.length a h'"
 proof -
   from assms part_length_remains show ?thesis
     unfolding partition.simps swap_def
@@ -287,14 +287,14 @@
          else middle)"
     unfolding partition.simps
     by (elim crelE crel_return crel_nth crel_if crel_upd) simp
-  from swap have h'_def: "h' = Heap.upd a r (get_array a h1 ! rs)
-    (Heap.upd a rs (get_array a h1 ! r) h1)"
+  from swap have h'_def: "h' = Array.change a r (get_array a h1 ! rs)
+    (Array.change a rs (get_array a h1 ! r) h1)"
     unfolding swap_def
     by (elim crelE crel_return crel_nth crel_upd) simp
-  from swap have in_bounds: "r < Heap.length a h1 \<and> rs < Heap.length a h1"
+  from swap have in_bounds: "r < Array.length a h1 \<and> rs < Array.length a h1"
     unfolding swap_def
     by (elim crelE crel_return crel_nth crel_upd) simp
-  from swap have swap_length_remains: "Heap.length a h1 = Heap.length a h'"
+  from swap have swap_length_remains: "Array.length a h1 = Array.length a h'"
     unfolding swap_def by (elim crelE crel_return crel_nth crel_upd) auto
   from `l < r` have "l \<le> r - 1" by simp 
   note middle_in_bounds = part_returns_index_in_bounds[OF part this]
@@ -304,7 +304,7 @@
   with swap
   have right_remains: "get_array a h ! r = get_array a h' ! rs"
     unfolding swap_def
-    by (auto simp add: Heap.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto)
+    by (auto simp add: Array.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto)
   from part_partitions [OF part]
   show ?thesis
   proof (cases "get_array a h1 ! middle \<le> ?pivot")
@@ -314,12 +314,12 @@
       fix i
       assume i_is_left: "l \<le> i \<and> i < rs"
       with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r`
-      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+      have i_props: "i < Array.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
       from i_is_left rs_equals have "l \<le> i \<and> i < middle \<or> i = middle" by arith
       with part_partitions[OF part] right_remains True
       have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
       with i_props h'_def in_bounds have "get_array a h' ! i \<le> get_array a h' ! rs"
-        unfolding Heap.upd_def Heap.length_def by simp
+        unfolding Array.change_def Array.length_def by simp
     }
     moreover
     {
@@ -331,7 +331,7 @@
       proof
         assume i_is: "rs < i \<and> i \<le> r - 1"
         with swap_length_remains in_bounds middle_in_bounds rs_equals
-        have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+        have i_props: "i < Array.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
         from part_partitions[OF part] rs_equals right_remains i_is
         have "get_array a h' ! rs \<le> get_array a h1 ! i"
           by fastsimp
@@ -345,7 +345,7 @@
           by fastsimp
         with i_is True rs_equals right_remains h'_def
         show ?thesis using in_bounds
-          unfolding Heap.upd_def Heap.length_def
+          unfolding Array.change_def Array.length_def
           by auto
       qed
     }
@@ -357,11 +357,11 @@
       fix i
       assume i_is_left: "l \<le> i \<and> i < rs"
       with swap_length_remains in_bounds middle_in_bounds rs_equals
-      have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+      have i_props: "i < Array.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
       from part_partitions[OF part] rs_equals right_remains i_is_left
       have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
       with i_props h'_def have "get_array a h' ! i \<le> get_array a h' ! rs"
-        unfolding Heap.upd_def by simp
+        unfolding Array.change_def by simp
     }
     moreover
     {
@@ -372,7 +372,7 @@
       proof
         assume i_is: "rs < i \<and> i \<le> r - 1"
         with swap_length_remains in_bounds middle_in_bounds rs_equals
-        have i_props: "i < Heap.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
+        have i_props: "i < Array.length a h'" "i \<noteq> r" "i \<noteq> rs" by auto
         from part_partitions[OF part] rs_equals right_remains i_is
         have "get_array a h' ! rs \<le> get_array a h1 ! i"
           by fastsimp
@@ -381,7 +381,7 @@
         assume i_is: "i = r"
         from i_is False rs_equals right_remains h'_def
         show ?thesis using in_bounds
-          unfolding Heap.upd_def Heap.length_def
+          unfolding Array.change_def Array.length_def
           by auto
       qed
     }
@@ -425,7 +425,7 @@
 
 lemma length_remains:
   assumes "crel (quicksort a l r) h h' rs"
-  shows "Heap.length a h = Heap.length a h'"
+  shows "Array.length a h = Array.length a h'"
 using assms
 proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
   case (1 a l r h h' rs)
@@ -482,7 +482,7 @@
  
 lemma quicksort_sorts:
   assumes "crel (quicksort a l r) h h' rs"
-  assumes l_r_length: "l < Heap.length a h" "r < Heap.length a h" 
+  assumes l_r_length: "l < Array.length a h" "r < Array.length a h" 
   shows "sorted (subarray l (r + 1) a h')"
   using assms
 proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
@@ -524,7 +524,7 @@
       from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True
         length_remains 1(5) pivot multiset_of_sublist [of l p "get_array a h1" "get_array a h2"]
       have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)"
-        unfolding Heap.length_def subarray_def by (cases p, auto)
+        unfolding Array.length_def subarray_def by (cases p, auto)
       with left_subarray_remains part_conds1 pivot_unchanged
       have part_conds2': "\<forall>j. j \<in> set (subarray l p a h') \<longrightarrow> j \<le> get_array a h' ! p"
         by (simp, subst set_of_multiset_of[symmetric], simp)
@@ -535,7 +535,7 @@
       from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True
         length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "get_array a h2" "get_array a h'"]
       have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)"
-        unfolding Heap.length_def subarray_def by auto
+        unfolding Array.length_def subarray_def by auto
       with right_subarray_remains part_conds2 pivot_unchanged
       have part_conds1': "\<forall>j. j \<in> set (subarray (p + 1) (r + 1) a h') \<longrightarrow> get_array a h' ! p \<le> j"
         by (simp, subst set_of_multiset_of[symmetric], simp)
@@ -556,18 +556,18 @@
 
 
 lemma quicksort_is_sort:
-  assumes crel: "crel (quicksort a 0 (Heap.length a h - 1)) h h' rs"
+  assumes crel: "crel (quicksort a 0 (Array.length a h - 1)) h h' rs"
   shows "get_array a h' = sort (get_array a h)"
 proof (cases "get_array a h = []")
   case True
   with quicksort_is_skip[OF crel] show ?thesis
-  unfolding Heap.length_def by simp
+  unfolding Array.length_def by simp
 next
   case False
   from quicksort_sorts [OF crel] False have "sorted (sublist' 0 (List.length (get_array a h)) (get_array a h'))"
-    unfolding Heap.length_def subarray_def by auto
+    unfolding Array.length_def subarray_def by auto
   with length_remains[OF crel] have "sorted (get_array a h')"
-    unfolding Heap.length_def by simp
+    unfolding Array.length_def by simp
   with quicksort_permutes [OF crel] properties_for_sort show ?thesis by fastsimp
 qed
 
@@ -576,7 +576,7 @@
 We will now show that exceptions do not occur. *}
 
 lemma noError_part1: 
-  assumes "l < Heap.length a h" "r < Heap.length a h"
+  assumes "l < Array.length a h" "r < Array.length a h"
   shows "noError (part1 a l r p) h"
   using assms
 proof (induct a l r p arbitrary: h rule: part1.induct)
@@ -587,7 +587,7 @@
 qed
 
 lemma noError_partition:
-  assumes "l < r" "l < Heap.length a h" "r < Heap.length a h"
+  assumes "l < r" "l < Array.length a h" "r < Array.length a h"
   shows "noError (partition a l r) h"
 using assms
 unfolding partition.simps swap_def
@@ -603,7 +603,7 @@
 done
 
 lemma noError_quicksort:
-  assumes "l < Heap.length a h" "r < Heap.length a h"
+  assumes "l < Array.length a h" "r < Array.length a h"
   shows "noError (quicksort a l r) h"
 using assms
 proof (induct a l r arbitrary: h rule: quicksort.induct)
@@ -628,7 +628,7 @@
 subsection {* Example *}
 
 definition "qsort a = do
-    k \<leftarrow> length a;
+    k \<leftarrow> len a;
     quicksort a 0 (k - 1);
     return a
   done"
--- a/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -37,7 +37,7 @@
       else get_array a h ! k)"
 using assms unfolding swap.simps
 by (elim crel_elim_all)
- (auto simp: Heap.length_def)
+ (auto simp: length_def)
 
 lemma rev_pointwise: assumes "crel (rev a i j) h h' r"
   shows "get_array a h' ! k = (if k < i then get_array a h ! k
@@ -69,7 +69,7 @@
 
 lemma rev_length:
   assumes "crel (rev a i j) h h' r"
-  shows "Heap.length a h = Heap.length a h'"
+  shows "Array.length a h = Array.length a h'"
 using assms
 proof (induct a i j arbitrary: h h' rule: rev.induct)
   case (1 a i j h h'')
@@ -93,7 +93,7 @@
 qed
 
 lemma rev2_rev': assumes "crel (rev a i j) h h' u"
-  assumes "j < Heap.length a h"
+  assumes "j < Array.length a h"
   shows "subarray i (j + 1) a h' = List.rev (subarray i (j + 1) a h)"
 proof - 
   {
@@ -103,15 +103,15 @@
       by auto
   } 
   with assms(2) rev_length[OF assms(1)] show ?thesis
-  unfolding subarray_def Heap.length_def
+  unfolding subarray_def Array.length_def
   by (auto simp add: length_sublist' rev_nth min_def nth_sublist' intro!: nth_equalityI)
 qed
 
 lemma rev2_rev: 
-  assumes "crel (rev a 0 (Heap.length a h - 1)) h h' u"
+  assumes "crel (rev a 0 (Array.length a h - 1)) h h' u"
   shows "get_array a h' = List.rev (get_array a h)"
   using rev2_rev'[OF assms] rev_length[OF assms] assms
-    by (cases "Heap.length a h = 0", auto simp add: Heap.length_def
+    by (cases "Array.length a h = 0", auto simp add: Array.length_def
       subarray_def sublist'_all rev.simps[where j=0] elim!: crel_elim_all)
   (drule sym[of "List.length (get_array a h)"], simp)
 
--- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -13,7 +13,7 @@
 setup {* Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>type ref"}) *}
 datatype 'a node = Empty | Node 'a "('a node) ref"
 
-fun
+primrec
   node_encode :: "'a\<Colon>countable node \<Rightarrow> nat"
 where
   "node_encode Empty = 0"
@@ -28,11 +28,11 @@
 
 instance node :: (heap) heap ..
 
-fun make_llist :: "'a\<Colon>heap list \<Rightarrow> 'a node Heap"
+primrec make_llist :: "'a\<Colon>heap list \<Rightarrow> 'a node Heap"
 where 
   [simp del]: "make_llist []     = return Empty"
             | "make_llist (x#xs) = do tl   \<leftarrow> make_llist xs;
-                                      next \<leftarrow> Ref.new tl;
+                                      next \<leftarrow> ref tl;
                                       return (Node x next)
                                    done"
 
@@ -63,24 +63,24 @@
 
 subsection {* Definition of list_of, list_of', refs_of and refs_of' *}
 
-fun list_of :: "heap \<Rightarrow> ('a::heap) node \<Rightarrow> 'a list \<Rightarrow> bool"
+primrec list_of :: "heap \<Rightarrow> ('a::heap) node \<Rightarrow> 'a list \<Rightarrow> bool"
 where
   "list_of h r [] = (r = Empty)"
-| "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))"
+| "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))"
  
 definition list_of' :: "heap \<Rightarrow> ('a::heap) node ref \<Rightarrow> 'a list \<Rightarrow> bool"
 where
-  "list_of' h r xs = list_of h (get_ref r h) xs"
+  "list_of' h r xs = list_of h (Ref.get h r) xs"
 
-fun refs_of :: "heap \<Rightarrow> ('a::heap) node \<Rightarrow> 'a node ref list \<Rightarrow> bool"
+primrec refs_of :: "heap \<Rightarrow> ('a::heap) node \<Rightarrow> 'a node ref list \<Rightarrow> bool"
 where
   "refs_of h r [] = (r = Empty)"
-| "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)"
+| "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)"
 
-fun refs_of' :: "heap \<Rightarrow> ('a::heap) node ref \<Rightarrow> 'a node ref list \<Rightarrow> bool"
+primrec refs_of' :: "heap \<Rightarrow> ('a::heap) node ref \<Rightarrow> 'a node ref list \<Rightarrow> bool"
 where
   "refs_of' h r [] = False"
-| "refs_of' h r (x#xs) = ((x = r) \<and> refs_of h (get_ref x h) xs)"
+| "refs_of' h r (x#xs) = ((x = r) \<and> refs_of h (Ref.get h x) xs)"
 
 
 subsection {* Properties of these definitions *}
@@ -88,35 +88,35 @@
 lemma list_of_Empty[simp]: "list_of h Empty xs = (xs = [])"
 by (cases xs, auto)
 
-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')"
+lemma 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')"
 by (cases xs, auto)
 
-lemma list_of'_Empty[simp]: "get_ref q h = Empty \<Longrightarrow> list_of' h q xs = (xs = [])"
+lemma list_of'_Empty[simp]: "Ref.get h q = Empty \<Longrightarrow> list_of' h q xs = (xs = [])"
 unfolding list_of'_def by simp
 
-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')"
+lemma 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')"
 unfolding list_of'_def by simp
 
-lemma list_of'_Nil: "list_of' h q [] \<Longrightarrow> get_ref q h = Empty"
+lemma list_of'_Nil: "list_of' h q [] \<Longrightarrow> Ref.get h q = Empty"
 unfolding list_of'_def by simp
 
 lemma list_of'_Cons: 
 assumes "list_of' h q (x#xs)"
-obtains n where "get_ref q h = Node x n" and "list_of' h n xs"
+obtains n where "Ref.get h q = Node x n" and "list_of' h n xs"
 using assms unfolding list_of'_def by (auto split: node.split_asm)
 
 lemma refs_of_Empty[simp] : "refs_of h Empty xs = (xs = [])"
   by (cases xs, auto)
 
-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)"
+lemma 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)"
   by (cases xs, auto)
 
-lemma refs_of'_def': "refs_of' h p ps = (\<exists>prs. (ps = (p # prs)) \<and> refs_of h (get_ref p h) prs)"
+lemma refs_of'_def': "refs_of' h p ps = (\<exists>prs. (ps = (p # prs)) \<and> refs_of h (Ref.get h p) prs)"
 by (cases ps, auto)
 
 lemma refs_of'_Node:
   assumes "refs_of' h p xs"
-  assumes "get_ref p h = Node x pn"
+  assumes "Ref.get h p = Node x pn"
   obtains pnrs
   where "xs = p # pnrs" and "refs_of' h pn pnrs"
 using assms
@@ -166,7 +166,7 @@
   assumes "list_of' h r xs"
   shows "\<exists>rs. refs_of' h r rs"
 proof -
-  from assms obtain rs' where "refs_of h (get_ref r h) rs'"
+  from assms obtain rs' where "refs_of h (Ref.get h r) rs'"
     unfolding list_of'_def by (rule list_of_refs_of)
   thus ?thesis unfolding refs_of'_def' by auto
 qed
@@ -238,7 +238,7 @@
 done
 
 lemma refs_of_next:
-assumes "refs_of h (get_ref p h) rs"
+assumes "refs_of h (Ref.get h p) rs"
   shows "p \<notin> set rs"
 proof (rule ccontr)
   assume a: "\<not> (p \<notin> set rs)"
@@ -264,7 +264,7 @@
 
 subsection {* Interaction of these predicates with our heap transitions *}
 
-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"
+lemma 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"
 using assms
 proof (induct as arbitrary: q rs)
   case Nil thus ?case by simp
@@ -275,15 +275,15 @@
     case Empty thus ?thesis by auto
   next
     case (Node a ref)
-    from Cons(2) Node obtain rs' where 1: "refs_of h (get_ref ref h) rs'" and rs_rs': "rs = ref # rs'" by auto
+    from Cons(2) Node obtain rs' where 1: "refs_of h (Ref.get h ref) rs'" and rs_rs': "rs = ref # rs'" by auto
     from Cons(3) rs_rs' have "ref \<noteq> p" by fastsimp
-    hence ref_eq: "get_ref ref (set_ref p v h) = (get_ref ref h)" by (auto simp add: ref_get_set_neq)
+    hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq)
     from rs_rs' Cons(3) have 2: "p \<notin> set rs'" by simp
     from Cons.hyps[OF 1 2] Node ref_eq show ?thesis by simp
   qed
 qed
 
-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"
+lemma 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"
 proof (induct as arbitrary: q rs)
   case Nil thus ?case by simp
 next
@@ -293,15 +293,15 @@
     case Empty thus ?thesis by auto
   next
     case (Node a ref)
-    from Cons(2) Node obtain rs' where 1: "refs_of h (get_ref ref h) rs'" and rs_rs': "rs = ref # rs'" by auto
+    from Cons(2) Node obtain rs' where 1: "refs_of h (Ref.get h ref) rs'" and rs_rs': "rs = ref # rs'" by auto
     from Cons(3) rs_rs' have "ref \<noteq> p" by fastsimp
-    hence ref_eq: "get_ref ref (set_ref p v h) = (get_ref ref h)" by (auto simp add: ref_get_set_neq)
+    hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq)
     from rs_rs' Cons(3) have 2: "p \<notin> set rs'" by simp
     from Cons.hyps[OF 1 2] Node ref_eq show ?thesis by auto
   qed
 qed
 
-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"
+lemma 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"
 proof (induct rs arbitrary: q)
   case Nil thus ?case by simp
 next
@@ -311,9 +311,9 @@
     case Empty thus ?thesis by auto
   next
     case (Node a ref)
-    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
+    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
     from Cons(3) this have "ref \<noteq> p" by fastsimp
-    hence ref_eq: "get_ref ref (set_ref p v h) = (get_ref ref h)" by (auto simp add: ref_get_set_neq)
+    hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq)
     from Cons(3) have 2: "p \<notin> set xs" by simp
     with Cons.hyps 1 2 Node ref_eq show ?thesis
       by simp
@@ -323,7 +323,7 @@
 lemma list_of'_set_ref:
   assumes "refs_of' h q rs"
   assumes "p \<notin> set rs"
-  shows "list_of' (set_ref p v h) q as = list_of' h q as"
+  shows "list_of' (Ref.set p v h) q as = list_of' h q as"
 proof -
   from assms have "q \<noteq> p" by (auto simp only: dest!: refs_of'E)
   with assms show ?thesis
@@ -333,18 +333,18 @@
 
 lemma list_of'_set_next_ref_Node[simp]:
   assumes "list_of' h r xs"
-  assumes "get_ref p h = Node x r'"
+  assumes "Ref.get h p = Node x r'"
   assumes "refs_of' h r rs"
   assumes "p \<notin> set rs"
-  shows "list_of' (set_ref p (Node x r) h) p (x#xs) = list_of' h r xs"
+  shows "list_of' (Ref.set p (Node x r) h) p (x#xs) = list_of' h r xs"
 using assms
 unfolding list_of'_def refs_of'_def'
-by (auto simp add: list_of_set_ref noteq_refs_sym)
+by (auto simp add: list_of_set_ref Ref.noteq_sym)
 
 lemma refs_of'_set_ref:
   assumes "refs_of' h q rs"
   assumes "p \<notin> set rs"
-  shows "refs_of' (set_ref p v h) q as = refs_of' h q as"
+  shows "refs_of' (Ref.set p v h) q as = refs_of' h q as"
 using assms
 proof -
   from assms have "q \<noteq> p" by (auto simp only: dest!: refs_of'E)
@@ -354,9 +354,9 @@
 qed
 
 lemma refs_of'_set_ref2:
-  assumes "refs_of' (set_ref p v h) q rs"
+  assumes "refs_of' (Ref.set p v h) q rs"
   assumes "p \<notin> set rs"
-  shows "refs_of' (set_ref p v h) q as = refs_of' h q as"
+  shows "refs_of' (Ref.set p v h) q as = refs_of' h q as"
 using assms
 proof -
   from assms have "q \<noteq> p" by (auto simp only: dest!: refs_of'E)
@@ -364,7 +364,7 @@
     unfolding refs_of'_def'
     apply auto
     apply (subgoal_tac "prs = prsa")
-    apply (insert refs_of_set_ref2[of p v h "get_ref q h"])
+    apply (insert refs_of_set_ref2[of p v h "Ref.get h q"])
     apply (erule_tac x="prs" in meta_allE)
     apply auto
     apply (auto dest: refs_of_is_fun)
@@ -372,15 +372,15 @@
 qed
 
 lemma refs_of'_set_next_ref:
-assumes "get_ref p h1 = Node x pn"
-assumes "refs_of' (set_ref p (Node x r1) h1) p rs"
+assumes "Ref.get h1 p = Node x pn"
+assumes "refs_of' (Ref.set p (Node x r1) h1) p rs"
 obtains r1s where "rs = (p#r1s)" and "refs_of' h1 r1 r1s"
 using assms
 proof -
   from assms refs_of'_distinct[OF assms(2)] have "\<exists> r1s. rs = (p # r1s) \<and> refs_of' h1 r1 r1s"
     apply -
     unfolding refs_of'_def'[of _ p]
-    apply (auto, frule refs_of_set_ref2) by (auto dest: noteq_refs_sym)
+    apply (auto, frule refs_of_set_ref2) by (auto dest: Ref.noteq_sym)
   with prems show thesis by auto
 qed
 
@@ -388,7 +388,7 @@
 
 lemma refs_of_invariant:
   assumes "refs_of h (r::('a::heap) node) xs"
-  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')"
+  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)"
   shows "refs_of h' r xs"
 using assms
 proof (induct xs arbitrary: r)
@@ -396,28 +396,28 @@
 next
   case (Cons x xs')
   from Cons(2) obtain v where Node: "r = Node v x" by (cases r, auto)
-  from Cons(2) Node have refs_of_next: "refs_of h (get_ref x h) xs'" by simp
-  from Cons(2-3) Node have ref_eq: "get_ref x h = get_ref x h'" by auto
-  from ref_eq refs_of_next have 1: "refs_of h (get_ref x h') xs'" by simp
-  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'"
+  from Cons(2) Node have refs_of_next: "refs_of h (Ref.get h x) xs'" by simp
+  from Cons(2-3) Node have ref_eq: "Ref.get h x = Ref.get h' x" by auto
+  from ref_eq refs_of_next have 1: "refs_of h (Ref.get h' x) xs'" by simp
+  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"
     by fastsimp
-  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')"
+  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)"
     by (fastsimp dest: refs_of_is_fun)
-  from Cons.hyps[OF 1 2] have "refs_of h' (get_ref x h') xs'" .
+  from Cons.hyps[OF 1 2] have "refs_of h' (Ref.get h' x) xs'" .
   with Node show ?case by simp
 qed
 
 lemma refs_of'_invariant:
   assumes "refs_of' h r xs"
-  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')"
+  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)"
   shows "refs_of' h' r xs"
 using assms
 proof -
-  from assms obtain prs where refs:"refs_of h (get_ref r h) prs" and xs_def: "xs = r # prs"
+  from assms obtain prs where refs:"refs_of h (Ref.get h r) prs" and xs_def: "xs = r # prs"
     unfolding refs_of'_def' by auto
-  from xs_def assms have x_eq: "get_ref r h = get_ref r h'" by fastsimp
-  from refs assms xs_def have 2: "\<forall>refs. refs_of h (get_ref r 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')" 
+  from xs_def assms have x_eq: "Ref.get h r = Ref.get h' r" by fastsimp
+  from refs assms xs_def have 2: "\<forall>refs. refs_of h (Ref.get 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)" 
     by (fastsimp dest: refs_of_is_fun)
   from refs_of_invariant [OF refs 2] xs_def x_eq show ?thesis
     unfolding refs_of'_def' by auto
@@ -425,7 +425,7 @@
 
 lemma list_of_invariant:
   assumes "list_of h (r::('a::heap) node) xs"
-  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')"
+  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)"
   shows "list_of h' r xs"
 using assms
 proof (induct xs arbitrary: r)
@@ -437,16 +437,16 @@
     by (cases r, auto)
   from Cons(2) obtain rs where rs_def: "refs_of h r rs" by (rule list_of_refs_of)
   from Node rs_def obtain rss where refs_of: "refs_of h r (ref#rss)" and rss_def: "rs = ref#rss" by auto
-  from Cons(3) Node refs_of have ref_eq: "get_ref ref h = get_ref ref h'"
+  from Cons(3) Node refs_of have ref_eq: "Ref.get h ref = Ref.get h' ref"
     by auto
-  from Cons(2) ref_eq Node have 1: "list_of h (get_ref ref h') xs'" by simp
-  from refs_of Node ref_eq have refs_of_ref: "refs_of h (get_ref ref h') rss" by simp
-  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
-  from refs_of_ref rs_heap_eq rss_def have 2: "\<forall>refs. refs_of h (get_ref ref 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')"
+  from Cons(2) ref_eq Node have 1: "list_of h (Ref.get h' ref) xs'" by simp
+  from refs_of Node ref_eq have refs_of_ref: "refs_of h (Ref.get h' ref) rss" by simp
+  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
+  from refs_of_ref rs_heap_eq rss_def have 2: "\<forall>refs. refs_of h (Ref.get h' ref) 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)"
     by (auto dest: refs_of_is_fun)
   from Cons(1)[OF 1 2]
-  have "list_of h' (get_ref ref h') xs'" .
+  have "list_of h' (Ref.get h' ref) xs'" .
   with Node show ?case
     unfolding list_of'_def
     by simp
@@ -454,29 +454,29 @@
 
 lemma make_llist:
 assumes "crel (make_llist xs) h h' r"
-shows "list_of h' r xs \<and> (\<forall>rs. refs_of h' r rs \<longrightarrow> (\<forall>ref \<in> (set rs). ref_present ref h'))"
+shows "list_of h' r xs \<and> (\<forall>rs. refs_of h' r rs \<longrightarrow> (\<forall>ref \<in> (set rs). Ref.present h' ref))"
 using assms 
 proof (induct xs arbitrary: h h' r)
   case Nil thus ?case by (auto elim: crel_return simp add: make_llist.simps)
 next
   case (Cons x xs')
   from Cons.prems obtain h1 r1 r' where make_llist: "crel (make_llist xs') h h1 r1"
-    and crel_refnew:"crel (Ref.new r1) h1 h' r'" and Node: "r = Node x r'"
+    and crel_refnew:"crel (ref r1) h1 h' r'" and Node: "r = Node x r'"
     unfolding make_llist.simps
     by (auto elim!: crelE crel_return)
   from Cons.hyps[OF make_llist] have list_of_h1: "list_of h1 r1 xs'" ..
   from Cons.hyps[OF make_llist] obtain rs' where rs'_def: "refs_of h1 r1 rs'" by (auto intro: list_of_refs_of)
-  from Cons.hyps[OF make_llist] rs'_def have refs_present: "\<forall>ref\<in>set rs'. ref_present ref h1" by simp
+  from Cons.hyps[OF make_llist] rs'_def have refs_present: "\<forall>ref\<in>set rs'. Ref.present h1 ref" by simp
   from crel_refnew rs'_def refs_present have refs_unchanged: "\<forall>refs. refs_of h1 r1 refs \<longrightarrow>
-         (\<forall>ref\<in>set refs. ref_present ref h1 \<and> ref_present ref h' \<and> get_ref ref h1 = get_ref ref h')"
-    by (auto elim!: crel_Ref_new dest: refs_of_is_fun)
+         (\<forall>ref\<in>set refs. Ref.present h1 ref \<and> Ref.present h' ref \<and> Ref.get h1 ref = Ref.get h' ref)"
+    by (auto elim!: crel_ref dest: refs_of_is_fun)
   with list_of_invariant[OF list_of_h1 refs_unchanged] Node crel_refnew have fstgoal: "list_of h' r (x # xs')"
     unfolding list_of.simps
-    by (auto elim!: crel_Ref_new)
-  from refs_unchanged rs'_def have refs_still_present: "\<forall>ref\<in>set rs'. ref_present ref h'" by auto
+    by (auto elim!: crel_ref)
+  from refs_unchanged rs'_def have refs_still_present: "\<forall>ref\<in>set rs'. Ref.present h' ref" by auto
   from refs_of_invariant[OF rs'_def refs_unchanged] refs_unchanged Node crel_refnew refs_still_present
-  have sndgoal: "\<forall>rs. refs_of h' r rs \<longrightarrow> (\<forall>ref\<in>set rs. ref_present ref h')"
-    by (fastsimp elim!: crel_Ref_new dest: refs_of_is_fun)
+  have sndgoal: "\<forall>rs. refs_of h' r rs \<longrightarrow> (\<forall>ref\<in>set rs. Ref.present h' ref)"
+    by (fastsimp elim!: crel_ref dest: refs_of_is_fun)
   from fstgoal sndgoal show ?case ..
 qed
 
@@ -533,10 +533,10 @@
 thm arg_cong2
   by (auto simp add: expand_fun_eq intro: arg_cong2[where f = "op \<guillemotright>="] split: node.split)
 
-fun rev :: "('a:: heap) node \<Rightarrow> 'a node Heap" 
+primrec rev :: "('a:: heap) node \<Rightarrow> 'a node Heap" 
 where
   "rev Empty = return Empty"
-| "rev (Node x n) = (do q \<leftarrow> Ref.new Empty; p \<leftarrow> Ref.new (Node x n); v \<leftarrow> rev' (q, p); !v done)"
+| "rev (Node x n) = (do q \<leftarrow> ref Empty; p \<leftarrow> ref (Node x n); v \<leftarrow> rev' (q, p); !v done)"
 
 subsection {* Correctness Proof *}
 
@@ -556,17 +556,17 @@
   case (Cons x xs)
   (*"LinkedList.list_of h' (get_ref v h') (List.rev xs @ x # qsa)"*)
   from Cons(4) obtain ref where 
-    p_is_Node: "get_ref p h = Node x ref"
+    p_is_Node: "Ref.get h p = Node x ref"
     (*and "ref_present ref h"*)
     and list_of'_ref: "list_of' h ref xs"
-    unfolding list_of'_def by (cases "get_ref p h", auto)
-  from p_is_Node Cons(2) have crel_rev': "crel (rev' (p, ref)) (set_ref p (Node x q) h) h' v"
+    unfolding list_of'_def by (cases "Ref.get h p", auto)
+  from p_is_Node Cons(2) have crel_rev': "crel (rev' (p, ref)) (Ref.set p (Node x q) h) h' v"
     by (auto simp add: rev'_simps [of q p] elim!: crelE crel_lookup crel_update)
   from Cons(3) obtain qrs where qrs_def: "refs_of' h q qrs" by (elim list_of'_refs_of')
   from Cons(4) obtain prs where prs_def: "refs_of' h p prs" by (elim list_of'_refs_of')
   from qrs_def prs_def Cons(5) have distinct_pointers: "set qrs \<inter> set prs = {}" by fastsimp
   from qrs_def prs_def distinct_pointers refs_of'E have p_notin_qrs: "p \<notin> set qrs" by fastsimp
-  from Cons(3) qrs_def this have 1: "list_of' (set_ref p (Node x q) h) p (x#qs)"
+  from Cons(3) qrs_def this have 1: "list_of' (Ref.set p (Node x q) h) p (x#qs)"
     unfolding list_of'_def  
     apply (simp)
     unfolding list_of'_def[symmetric]
@@ -575,16 +575,16 @@
     unfolding refs_of'_def' by auto
   from prs_refs prs_def have p_not_in_refs: "p \<notin> set refs"
     by (fastsimp dest!: refs_of'_distinct)
-  with refs_def p_is_Node list_of'_ref have 2: "list_of' (set_ref p (Node x q) h) ref xs"
+  with refs_def p_is_Node list_of'_ref have 2: "list_of' (Ref.set p (Node x q) h) ref xs"
     by (auto simp add: list_of'_set_ref)
-  from p_notin_qrs qrs_def have refs_of1: "refs_of' (set_ref p (Node x q) h) p (p#qrs)"
+  from p_notin_qrs qrs_def have refs_of1: "refs_of' (Ref.set p (Node x q) h) p (p#qrs)"
     unfolding refs_of'_def'
     apply (simp)
     unfolding refs_of'_def'[symmetric]
     by (simp add: refs_of'_set_ref)
-  from p_not_in_refs p_is_Node refs_def have refs_of2: "refs_of' (set_ref p (Node x q) h) ref refs"
+  from p_not_in_refs p_is_Node refs_def have refs_of2: "refs_of' (Ref.set p (Node x q) h) ref refs"
     by (simp add: refs_of'_set_ref)
-  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 = {}"
+  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 = {}"
     apply - apply (rule allI)+ apply (rule impI) apply (erule conjE)
     apply (drule refs_of'_is_fun) back back apply assumption
     apply (drule refs_of'_is_fun) back back apply assumption
@@ -595,7 +595,7 @@
 
 lemma rev_correctness:
   assumes list_of_h: "list_of h r xs"
-  assumes validHeap: "\<forall>refs. refs_of h r refs \<longrightarrow> (\<forall>r \<in> set refs. ref_present r h)"
+  assumes validHeap: "\<forall>refs. refs_of h r refs \<longrightarrow> (\<forall>r \<in> set refs. Ref.present h r)"
   assumes crel_rev: "crel (rev r) h h' r'"
   shows "list_of h' r' (List.rev xs)"
 using assms
@@ -606,39 +606,39 @@
 next
   case (Node x ps)
   with crel_rev obtain p q h1 h2 h3 v where
-    init: "crel (Ref.new Empty) h h1 q"
-    "crel (Ref.new (Node x ps)) h1 h2 p"
+    init: "crel (ref Empty) h h1 q"
+    "crel (ref (Node x ps)) h1 h2 p"
     and crel_rev':"crel (rev' (q, p)) h2 h3 v"
     and lookup: "crel (!v) h3 h' r'"
     using rev.simps
     by (auto elim!: crelE)
   from init have a1:"list_of' h2 q []"
     unfolding list_of'_def
-    by (auto elim!: crel_Ref_new)
+    by (auto elim!: crel_ref)
   from list_of_h obtain refs where refs_def: "refs_of h r refs" by (rule list_of_refs_of)
-  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)"
-    by (fastsimp elim!: crel_Ref_new dest: refs_of_is_fun)
+  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)"
+    by (fastsimp elim!: crel_ref dest: refs_of_is_fun)
   from list_of_invariant[OF list_of_h heap_eq] have "list_of h2 r xs" .
   from init this Node have a2: "list_of' h2 p xs"
     apply -
     unfolding list_of'_def
-    apply (auto elim!: crel_Ref_new)
+    apply (auto elim!: crel_ref)
     done
   from init have refs_of_q: "refs_of' h2 q [q]"
-    by (auto elim!: crel_Ref_new)
+    by (auto elim!: crel_ref)
   from refs_def Node have refs_of'_ps: "refs_of' h ps refs"
     by (auto simp add: refs_of'_def'[symmetric])
-  from validHeap refs_def have all_ref_present: "\<forall>r\<in>set refs. ref_present r h" by simp
-  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)"
-    by (fastsimp elim!: crel_Ref_new dest: refs_of'_is_fun)
+  from validHeap refs_def have all_ref_present: "\<forall>r\<in>set refs. Ref.present h r" by simp
+  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 h ref \<and> Ref.present h2 ref \<and> Ref.get h ref = Ref.get h2 ref)"
+    by (fastsimp elim!: crel_ref dest: refs_of'_is_fun)
   from refs_of'_invariant[OF refs_of'_ps this] have "refs_of' h2 ps refs" .
   with init have refs_of_p: "refs_of' h2 p (p#refs)"
-    by (auto elim!: crel_Ref_new simp add: refs_of'_def')
+    by (auto elim!: crel_ref simp add: refs_of'_def')
   with init all_ref_present have q_is_new: "q \<notin> set (p#refs)"
-    by (auto elim!: crel_Ref_new intro!: noteq_refsI)
+    by (auto elim!: crel_ref intro!: Ref.noteq_I)
   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 = {}"
     by (fastsimp simp only: set.simps dest: refs_of'_is_fun)
-  from rev'_invariant [OF crel_rev' a1 a2 a3] have "list_of h3 (get_ref v h3) (List.rev xs)" 
+  from rev'_invariant [OF crel_rev' a1 a2 a3] have "list_of h3 (Ref.get h3 v) (List.rev xs)" 
     unfolding list_of'_def by auto
   with lookup show ?thesis
     by (auto elim: crel_lookup)
@@ -734,32 +734,32 @@
 lemma merge_induct2:
   assumes "list_of' h (p::'a::{heap, ord} node ref) xs"
   assumes "list_of' h q ys"
-  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"
-  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') []"
+  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"
+  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') []"
   assumes "\<And> x xs' y ys' p q pn qn.
-  \<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;
+  \<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;
   x \<le> y; P pn q xs' (y#ys') \<rbrakk>
   \<Longrightarrow> P p q (x#xs') (y#ys')"
   assumes "\<And> x xs' y ys' p q pn qn.
-  \<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;
+  \<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;
   \<not> x \<le> y; P p qn (x#xs') ys'\<rbrakk>
   \<Longrightarrow> P p q (x#xs') (y#ys')"
   shows "P p q xs ys"
 using assms(1-2)
 proof (induct xs ys arbitrary: p q rule: Lmerge.induct)
   case (2 ys)
-  from 2(1) have "get_ref p h = Empty" unfolding list_of'_def by simp
+  from 2(1) have "Ref.get h p = Empty" unfolding list_of'_def by simp
   with 2(1-2) assms(3) show ?case by blast
 next
   case (3 x xs')
-  from 3(1) obtain pn where Node: "get_ref p h = Node x pn" by (rule list_of'_Cons)
-  from 3(2) have "get_ref q h = Empty" unfolding list_of'_def by simp
+  from 3(1) obtain pn where Node: "Ref.get h p = Node x pn" by (rule list_of'_Cons)
+  from 3(2) have "Ref.get h q = Empty" unfolding list_of'_def by simp
   with Node 3(1-2) assms(4) show ?case by blast
 next
   case (1 x xs' y ys')
-  from 1(3) obtain pn where pNode:"get_ref p h = Node x pn"
+  from 1(3) obtain pn where pNode:"Ref.get h p = Node x pn"
     and list_of'_pn: "list_of' h pn xs'" by (rule list_of'_Cons)
-  from 1(4) obtain qn where qNode:"get_ref q h = Node y qn"
+  from 1(4) obtain qn where qNode:"Ref.get h q = Node y qn"
     and  list_of'_qn: "list_of' h qn ys'" by (rule list_of'_Cons)
   show ?case
   proof (cases "x \<le> y")
@@ -780,15 +780,15 @@
 assumes  "list_of' h p xs"
 assumes  "list_of' h q ys"
 assumes  "crel (merge p q) h h' r"
-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"
-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') []"
+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 h h q [] ys"
+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 h h p (x#xs') []"
 assumes  "\<And> x xs' y ys' p q pn qn h1 r1 h'.
-  \<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;
-  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>
+  \<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;
+  x \<le> y; crel (merge pn q) h h1 r1 ; P pn q h h1 r1 xs' (y#ys'); h' = Ref.set p (Node x r1) h1 \<rbrakk>
   \<Longrightarrow> P p q h h' p (x#xs') (y#ys')"
 assumes "\<And> x xs' y ys' p q pn qn h1 r1 h'.
-  \<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;
-  \<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>
+  \<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;
+  \<not> x \<le> y; crel (merge p qn) h h1 r1; P p qn h h1 r1 (x#xs') ys'; h' = Ref.set q (Node y r1) h1 \<rbrakk>
   \<Longrightarrow> P p q h h' q (x#xs') (y#ys')"
 shows "P p q h h' r xs ys"
 using assms(3)
@@ -808,7 +808,7 @@
   case (3 x xs' y ys' p q pn qn)
   from 3(3-5) 3(7) obtain h1 r1 where
     1: "crel (merge pn q) h h1 r1" 
-    and 2: "h' = set_ref p (Node x r1) h1 \<and> r = p"
+    and 2: "h' = Ref.set p (Node x r1) h1 \<and> r = p"
     unfolding merge_simps[of p q]
     by (auto elim!: crel_lookup crelE crel_return crel_if crel_update)
   from 3(6)[OF 1] assms(6) [OF 3(1-5)] 1 2 show ?case by simp
@@ -816,7 +816,7 @@
   case (4 x xs' y ys' p q pn qn)
   from 4(3-5) 4(7) obtain h1 r1 where
     1: "crel (merge p qn) h h1 r1" 
-    and 2: "h' = set_ref q (Node y r1) h1 \<and> r = q"
+    and 2: "h' = Ref.set q (Node y r1) h1 \<and> r = q"
     unfolding merge_simps[of p q]
     by (auto elim!: crel_lookup crelE crel_return crel_if crel_update)
   from 4(6)[OF 1] assms(7) [OF 4(1-5)] 1 2 show ?case by simp
@@ -834,7 +834,7 @@
   assumes "crel (merge p q) h h' r'"
   assumes "set xs \<inter> set ys = {}"
   assumes "r \<notin> set xs \<union> set ys"
-  shows "get_ref r h = get_ref r h'"
+  shows "Ref.get h r = Ref.get h' r"
 proof -
   from assms(1) obtain ps where ps_def: "list_of' h p ps" by (rule refs_of'_list_of')
   from assms(2) obtain qs where qs_def: "list_of' h q qs" by (rule refs_of'_list_of')
@@ -853,7 +853,8 @@
     from pnrs_def 3(12) have "r \<noteq> p" by auto
     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
     from 3(11) pnrs_def have no_inter: "set pnrs \<inter> set ys = {}" by auto
-    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
+    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"
+      by simp
     from 3(7)[OF refs_of'_pn 3(10) no_inter r_in] 3(8) `r \<noteq> p` show ?case
       by simp
   next
@@ -866,7 +867,7 @@
     from qnrs_def 4(12) have "r \<noteq> q" by auto
     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
     from 4(11) qnrs_def have no_inter: "set xs \<inter> set qnrs = {}" by auto
-    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
+    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
     from 4(7)[OF 4(9) refs_of'_qn no_inter r_in] 4(8) `r \<noteq> q` show ?case
       by simp
   qed
@@ -899,7 +900,7 @@
       by (rule refs_of'_Node)
     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
     from 3(11) pnrs_def have no_inter: "set pnrs \<inter> set ys = {}" by auto
-    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" ..
+    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" ..
     from 3 p_stays obtain r1s
       where rs_def: "rs = p#r1s" and refs_of'_r1:"refs_of' h1 r1 r1s"
       by (auto elim: refs_of'_set_next_ref)
@@ -912,7 +913,7 @@
       by (rule refs_of'_Node)
     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
     from 4(11) qnrs_def have no_inter: "set xs \<inter> set qnrs = {}" by auto
-    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" ..
+    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" ..
     from 4 q_stays obtain r1s
       where rs_def: "rs = q#r1s" and refs_of'_r1:"refs_of' h1 r1 r1s"
       by (auto elim: refs_of'_set_next_ref)
@@ -945,7 +946,7 @@
   from prs_def qrs_def 3(9) pnrs_def have no_inter: "set pnrs \<inter> set qrs = {}" by fastsimp
   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 = {}"
     by (fastsimp dest: refs_of'_is_fun)
-  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" ..
+  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" ..
   from 3(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of')
   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
   with 3(7)[OF no_inter2] 3(1-5) 3(8) p_rs rs_def p_stays
@@ -962,7 +963,7 @@
   from prs_def qrs_def 4(9) qnrs_def have no_inter: "set prs \<inter> set qnrs = {}" by fastsimp
   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 = {}"
     by (fastsimp dest: refs_of'_is_fun)
-  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" ..
+  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" ..
   from 4(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of')
   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
   with 4(7)[OF no_inter2] 4(1-5) 4(8) q_rs rs_def q_stays
@@ -984,8 +985,8 @@
   (do
     ll_xs \<leftarrow> make_llist (filter (%n. n mod 2 = 0) [2..8]);
     ll_ys \<leftarrow> make_llist (filter (%n. n mod 2 = 1) [5..11]);
-    r \<leftarrow> Ref.new ll_xs;
-    q \<leftarrow> Ref.new ll_ys;
+    r \<leftarrow> ref ll_xs;
+    q \<leftarrow> ref ll_ys;
     p \<leftarrow> merge r q;
     ll_zs \<leftarrow> !p;
     zs \<leftarrow> traverse ll_zs;
@@ -998,4 +999,4 @@
 ML {* @{code test_2} () *}
 ML {* @{code test_3} () *}
 
-end
\ No newline at end of file
+end
--- a/src/HOL/Imperative_HOL/ex/SatChecker.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/ex/SatChecker.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -118,6 +118,32 @@
 
 text {* Specific definition for derived clauses in the Array *}
 
+definition
+  array_ran :: "('a\<Colon>heap) option array \<Rightarrow> heap \<Rightarrow> 'a set" where
+  "array_ran a h = {e. Some e \<in> set (get_array a h)}"
+    -- {*FIXME*}
+
+lemma array_ranI: "\<lbrakk> Some b = get_array a h ! i; i < Array.length a h \<rbrakk> \<Longrightarrow> b \<in> array_ran a h"
+unfolding array_ran_def Array.length_def by simp
+
+lemma array_ran_upd_array_Some:
+  assumes "cl \<in> array_ran a (Array.change a i (Some b) h)"
+  shows "cl \<in> array_ran a h \<or> cl = b"
+proof -
+  have "set (get_array a h[i := Some b]) \<subseteq> insert (Some b) (set (get_array a h))" by (rule set_update_subset_insert)
+  with assms show ?thesis 
+    unfolding array_ran_def Array.change_def by fastsimp
+qed
+
+lemma array_ran_upd_array_None:
+  assumes "cl \<in> array_ran a (Array.change a i None h)"
+  shows "cl \<in> array_ran a h"
+proof -
+  have "set (get_array a h[i := None]) \<subseteq> insert None (set (get_array a h))" by (rule set_update_subset_insert)
+  with assms show ?thesis
+    unfolding array_ran_def Array.change_def by auto
+qed
+
 definition correctArray :: "Clause list \<Rightarrow> Clause option array \<Rightarrow> heap \<Rightarrow> bool"
 where
   "correctArray rootcls a h =
@@ -126,7 +152,7 @@
 lemma correctArray_update:
   assumes "correctArray rcs a h"
   assumes "correctClause rcs c" "sorted c" "distinct c"
-  shows "correctArray rcs a (Heap.upd a i (Some c) h)"
+  shows "correctArray rcs a (Array.change a i (Some c) h)"
   using assms
   unfolding correctArray_def
   by (auto dest:array_ran_upd_array_Some)
@@ -145,7 +171,7 @@
 
 subsection{* Function definitions *}
 
-fun res_mem :: "Lit \<Rightarrow> Clause \<Rightarrow> Clause Heap"
+primrec res_mem :: "Lit \<Rightarrow> Clause \<Rightarrow> Clause Heap"
 where
   "res_mem l [] = raise ''MiniSatChecked.res_thm: Cannot find literal''"
 | "res_mem l (x#xs) = (if (x = l) then return xs else (do v \<leftarrow> res_mem l xs; return (x # v) done))"
@@ -393,7 +419,7 @@
         done)"
 
 
-fun res_thm2 :: "Clause option array \<Rightarrow> (Lit * ClauseId) \<Rightarrow> Clause \<Rightarrow> Clause Heap"
+primrec res_thm2 :: "Clause option array \<Rightarrow> (Lit * ClauseId) \<Rightarrow> Clause \<Rightarrow> Clause Heap"
 where
   "res_thm2 a (l, j) cli =
   ( if l = 0 then raise(''Illegal literal'')
@@ -445,7 +471,7 @@
     fix clj
     let ?rs = "merge (remove l cli) (remove (compl l) clj)"
     let ?rs' = "merge (remove (compl l) cli) (remove l clj)"
-    assume "h = h'" "Some clj = get_array a h' ! j" "j < Heap.length a h'"
+    assume "h = h'" "Some clj = get_array a h' ! j" "j < Array.length a h'"
     with correct_a have clj: "correctClause r clj" "sorted clj" "distinct clj"
       unfolding correctArray_def by (auto intro: array_ranI)
     with clj l_not_zero correct_cli
@@ -459,7 +485,7 @@
   }
   {
     fix v clj
-    assume "Some clj = get_array a h ! j" "j < Heap.length a h"
+    assume "Some clj = get_array a h ! j" "j < Array.length a h"
     with correct_a have clj: "correctClause r clj \<and> sorted clj \<and> distinct clj"
       unfolding correctArray_def by (auto intro: array_ranI)
     assume "crel (res_thm' l cli clj) h h' rs"
@@ -606,7 +632,7 @@
 
 subsection {* Checker functions *}
 
-fun lres_thm :: "Clause option list \<Rightarrow> (Lit * ClauseId) \<Rightarrow> Clause \<Rightarrow> Clause Heap" 
+primrec lres_thm :: "Clause option list \<Rightarrow> (Lit * ClauseId) \<Rightarrow> Clause \<Rightarrow> Clause Heap" 
 where
   "lres_thm xs (l, j) cli = (if (j < List.length xs) then (case (xs ! j) of
   None \<Rightarrow> raise (''MiniSatChecked.res_thm: No resolvant clause in thms array for Conflict step.'')
@@ -640,7 +666,7 @@
 
 section {* Functional version with RedBlackTrees *}
 
-fun tres_thm :: "(ClauseId, Clause) RBT_Impl.rbt \<Rightarrow> Lit \<times> ClauseId \<Rightarrow> Clause \<Rightarrow> Clause Heap"
+primrec tres_thm :: "(ClauseId, Clause) RBT_Impl.rbt \<Rightarrow> Lit \<times> ClauseId \<Rightarrow> Clause \<Rightarrow> Clause Heap"
 where
   "tres_thm t (l, j) cli =
   (case (RBT_Impl.lookup t j) of 
--- a/src/HOL/Imperative_HOL/ex/Sorted_List.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/ex/Sorted_List.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -33,7 +33,7 @@
 
 text {* The remove function removes an element from a sorted list *}
 
-fun remove :: "('a :: linorder) \<Rightarrow> 'a list \<Rightarrow> 'a list"
+primrec remove :: "('a :: linorder) \<Rightarrow> 'a list \<Rightarrow> 'a list"
 where
   "remove a [] = []"
   |  "remove a (x#xs) = (if a > x then (x # remove a xs) else (if a = x then xs else x#xs))" 
@@ -86,16 +86,13 @@
 apply (auto simp add: sorted_Cons)
 done
 
-subsection {* Efficient member function for sorted lists: smem *}
+subsection {* Efficient member function for sorted lists *}
 
-fun smember :: "('a::linorder) \<Rightarrow> 'a list \<Rightarrow> bool" (infixl "smem" 55)
-where
-  "x smem [] = False"
-|  "x smem (y#ys) = (if x = y then True else (if (x > y) then x smem ys else False))" 
+primrec smember :: "'a list \<Rightarrow> 'a::linorder \<Rightarrow> bool" where
+  "smember [] x \<longleftrightarrow> False"
+| "smember (y#ys) x \<longleftrightarrow> x = y \<or> (x > y \<and> smember ys x)"
 
-lemma "sorted xs \<Longrightarrow> x smem xs = (x \<in> set xs)" 
-apply (induct xs)
-apply (auto simp add: sorted_Cons)
-done
+lemma "sorted xs \<Longrightarrow> smember xs x \<longleftrightarrow> (x \<in> set xs)" 
+  by (induct xs) (auto simp add: sorted_Cons)
 
 end
\ No newline at end of file
--- a/src/HOL/Imperative_HOL/ex/Subarray.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Imperative_HOL/ex/Subarray.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -8,65 +8,64 @@
 imports Array Sublist
 begin
 
-definition subarray :: "nat \<Rightarrow> nat \<Rightarrow> ('a::heap) array \<Rightarrow> heap \<Rightarrow> 'a list"
-where
+definition subarray :: "nat \<Rightarrow> nat \<Rightarrow> ('a::heap) array \<Rightarrow> heap \<Rightarrow> 'a list" where
   "subarray n m a h \<equiv> sublist' n m (get_array a h)"
 
-lemma subarray_upd: "i \<ge> m \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
-apply (simp add: subarray_def Heap.upd_def)
+lemma subarray_upd: "i \<ge> m \<Longrightarrow> subarray n m a (Array.change a i v h) = subarray n m a h"
+apply (simp add: subarray_def Array.change_def)
 apply (simp add: sublist'_update1)
 done
 
-lemma subarray_upd2: " i < n  \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
-apply (simp add: subarray_def Heap.upd_def)
+lemma subarray_upd2: " i < n  \<Longrightarrow> subarray n m a (Array.change a i v h) = subarray n m a h"
+apply (simp add: subarray_def Array.change_def)
 apply (subst sublist'_update2)
 apply fastsimp
 apply simp
 done
 
-lemma subarray_upd3: "\<lbrakk> n \<le> i; i < m\<rbrakk> \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h[i - n := v]"
-unfolding subarray_def Heap.upd_def
+lemma subarray_upd3: "\<lbrakk> n \<le> i; i < m\<rbrakk> \<Longrightarrow> subarray n m a (Array.change a i v h) = subarray n m a h[i - n := v]"
+unfolding subarray_def Array.change_def
 by (simp add: sublist'_update3)
 
 lemma subarray_Nil: "n \<ge> m \<Longrightarrow> subarray n m a h = []"
 by (simp add: subarray_def sublist'_Nil')
 
-lemma subarray_single: "\<lbrakk> n < Heap.length a h \<rbrakk> \<Longrightarrow> subarray n (Suc n) a h = [get_array a h ! n]" 
-by (simp add: subarray_def Heap.length_def sublist'_single)
+lemma subarray_single: "\<lbrakk> n < Array.length a h \<rbrakk> \<Longrightarrow> subarray n (Suc n) a h = [get_array a h ! n]" 
+by (simp add: subarray_def length_def sublist'_single)
 
-lemma length_subarray: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray n m a h) = m - n"
-by (simp add: subarray_def Heap.length_def length_sublist')
+lemma length_subarray: "m \<le> Array.length a h \<Longrightarrow> List.length (subarray n m a h) = m - n"
+by (simp add: subarray_def length_def length_sublist')
 
-lemma length_subarray_0: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray 0 m a h) = m"
+lemma length_subarray_0: "m \<le> Array.length a h \<Longrightarrow> List.length (subarray 0 m a h) = m"
 by (simp add: length_subarray)
 
-lemma subarray_nth_array_Cons: "\<lbrakk> i < Heap.length a h; i < j \<rbrakk> \<Longrightarrow> (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h"
-unfolding Heap.length_def subarray_def
+lemma subarray_nth_array_Cons: "\<lbrakk> i < Array.length a h; i < j \<rbrakk> \<Longrightarrow> (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h"
+unfolding Array.length_def subarray_def
 by (simp add: sublist'_front)
 
-lemma subarray_nth_array_back: "\<lbrakk> i < j; j \<le> Heap.length a h\<rbrakk> \<Longrightarrow> subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]"
-unfolding Heap.length_def subarray_def
+lemma subarray_nth_array_back: "\<lbrakk> i < j; j \<le> Array.length a h\<rbrakk> \<Longrightarrow> subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]"
+unfolding Array.length_def subarray_def
 by (simp add: sublist'_back)
 
 lemma subarray_append: "\<lbrakk> i < j; j < k \<rbrakk> \<Longrightarrow> subarray i j a h @ subarray j k a h = subarray i k a h"
 unfolding subarray_def
 by (simp add: sublist'_append)
 
-lemma subarray_all: "subarray 0 (Heap.length a h) a h = get_array a h"
-unfolding Heap.length_def subarray_def
+lemma subarray_all: "subarray 0 (Array.length a h) a h = get_array a h"
+unfolding Array.length_def subarray_def
 by (simp add: sublist'_all)
 
-lemma nth_subarray: "\<lbrakk> k < j - i; j \<le> Heap.length a h \<rbrakk> \<Longrightarrow> subarray i j a h ! k = get_array a h ! (i + k)"
-unfolding Heap.length_def subarray_def
+lemma nth_subarray: "\<lbrakk> k < j - i; j \<le> Array.length a h \<rbrakk> \<Longrightarrow> subarray i j a h ! k = get_array a h ! (i + k)"
+unfolding Array.length_def subarray_def
 by (simp add: nth_sublist')
 
-lemma subarray_eq_samelength_iff: "Heap.length a h = Heap.length a h' \<Longrightarrow> (subarray i j a h = subarray i j a h') = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> get_array a h ! i' = get_array a h' ! i')"
-unfolding Heap.length_def subarray_def by (rule sublist'_eq_samelength_iff)
+lemma subarray_eq_samelength_iff: "Array.length a h = Array.length a h' \<Longrightarrow> (subarray i j a h = subarray i j a h') = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> get_array a h ! i' = get_array a h' ! i')"
+unfolding Array.length_def subarray_def by (rule sublist'_eq_samelength_iff)
 
-lemma all_in_set_subarray_conv: "(\<forall>j. j \<in> set (subarray l r a h) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
-unfolding subarray_def Heap.length_def by (rule all_in_set_sublist'_conv)
+lemma all_in_set_subarray_conv: "(\<forall>j. j \<in> set (subarray l r a h) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Array.length a h \<longrightarrow> P (get_array a h ! k))"
+unfolding subarray_def Array.length_def by (rule all_in_set_sublist'_conv)
 
-lemma ball_in_set_subarray_conv: "(\<forall>j \<in> set (subarray l r a h). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
-unfolding subarray_def Heap.length_def by (rule ball_in_set_sublist'_conv)
+lemma ball_in_set_subarray_conv: "(\<forall>j \<in> set (subarray l r a h). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Array.length a h \<longrightarrow> P (get_array a h ! k))"
+unfolding subarray_def Array.length_def by (rule ball_in_set_sublist'_conv)
 
 end
\ No newline at end of file
--- a/src/HOL/Library/Countable.thy	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/HOL/Library/Countable.thy	Tue Jul 06 10:02:24 2010 +0200
@@ -110,26 +110,20 @@
   "to_nat_typerep (Typerep.Typerep c ts) = to_nat (to_nat c, to_nat (map to_nat_typerep ts))"
 
 instance proof (rule countable_classI)
-  fix t t' :: typerep and ts
-  have "(\<forall>t'. to_nat_typerep t = to_nat_typerep t' \<longrightarrow> t = t')
-    \<and> (\<forall>ts'. map to_nat_typerep ts = map to_nat_typerep ts' \<longrightarrow> ts = ts')"
-  proof (induct rule: typerep.induct)
-    case (Typerep c ts) show ?case
-    proof (rule allI, rule impI)
-      fix t'
-      assume hyp: "to_nat_typerep (Typerep.Typerep c ts) = to_nat_typerep t'"
-      then obtain c' ts' where t': "t' = (Typerep.Typerep c' ts')"
-        by (cases t') auto
-      with Typerep hyp have "c = c'" and "ts = ts'" by simp_all
-      with t' show "Typerep.Typerep c ts = t'" by simp
-    qed
+  fix t t' :: typerep and ts ts' :: "typerep list"
+  assume "to_nat_typerep t = to_nat_typerep t'"
+  moreover have "to_nat_typerep t = to_nat_typerep t' \<Longrightarrow> t = t'"
+    and "map to_nat_typerep ts = map to_nat_typerep ts' \<Longrightarrow> ts = ts'"
+  proof (induct t and ts arbitrary: t' and ts' rule: typerep.inducts)
+    case (Typerep c ts t')
+    then obtain c' ts' where t': "t' = Typerep.Typerep c' ts'" by (cases t') auto
+    with Typerep have "c = c'" and "ts = ts'" by simp_all
+    with t' show "Typerep.Typerep c ts = t'" by simp
   next
     case Nil_typerep then show ?case by simp
   next
     case (Cons_typerep t ts) then show ?case by auto
   qed
-  then have "to_nat_typerep t = to_nat_typerep t' \<Longrightarrow> t = t'" by auto
-  moreover assume "to_nat_typerep t = to_nat_typerep t'"
   ultimately show "t = t'" by simp
 qed
 
--- a/src/Provers/clasimp.ML	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/Provers/clasimp.ML	Tue Jul 06 10:02:24 2010 +0200
@@ -203,7 +203,7 @@
           (CHANGED o nodup_depth_tac cs' n); (* slower but more general *)
     in  EVERY [ALLGOALS (Simplifier.asm_full_simp_tac ss),
                TRY (Classical.safe_tac cs),
-               REPEAT (FIRSTGOAL maintac),
+               REPEAT_DETERM (FIRSTGOAL maintac),
                TRY (Classical.safe_tac (cs addSss ss)),
                prune_params_tac]
     end;
--- a/src/Pure/unify.ML	Mon Jul 05 23:07:36 2010 +0200
+++ b/src/Pure/unify.ML	Tue Jul 06 10:02:24 2010 +0200
@@ -451,20 +451,23 @@
   end;
 
 
+(*If an argument contains a banned Bound, then it should be deleted.
+  But if the only path is flexible, this is difficult; the code gives up!
+  In  %x y.?a(x) =?= %x y.?b(?c(y)) should we instantiate ?b or ?c *)
+exception CHANGE_FAIL;   (*flexible occurrence of banned variable, or other reason to quit*)
+
+
 (*Flex argument: a term, its type, and the index that refers to it.*)
 type flarg = {t: term, T: typ, j: int};
 
 (*Form the arguments into records for deletion/sorting.*)
 fun flexargs ([], [], []) = [] : flarg list
   | flexargs (j :: js, t :: ts, T :: Ts) = {j = j, t = t, T = T} :: flexargs (js, ts, Ts)
-  | flexargs _ = raise Fail "flexargs";
-
-
-(*If an argument contains a banned Bound, then it should be deleted.
-  But if the only path is flexible, this is difficult; the code gives up!
-  In  %x y.?a(x) =?= %x y.?b(?c(y)) should we instantiate ?b or ?c *)
-exception CHANGE_FAIL;   (*flexible occurrence of banned variable*)
-
+  | flexargs _ = raise CHANGE_FAIL;
+(*We give up if we see a variable of function type not applied to a full list of 
+  arguments (remember, this code assumes that terms are fully eta-expanded).  This situation 
+  can occur if a type variable is instantiated with a function type.
+*)
 
 (*Check whether the 'banned' bound var indices occur rigidly in t*)
 fun rigid_bound (lev, banned) t =
@@ -516,7 +519,7 @@
     val (Var (v, T), ts) = strip_comb t;
     val (Ts, U) = strip_type env T
     and js = length ts - 1  downto 0;
-    val args = sort (make_ord arg_less) (List.foldr (change_arg banned) [] (flexargs (js, ts, Ts)))
+    val args = sort (make_ord arg_less) (List.foldr (change_arg banned) [] (flexargs (js, ts, Ts))) 
     val ts' = map #t args;
   in
     if decreasing (length Ts) args then (env, (list_comb (Var (v, T), ts')))