tuned array theory
authorhaftmann
Fri, 09 Jul 2010 09:48:52 +0200
changeset 37752 d0a384c84d69
parent 37751 89e16802b6cc
child 37753 3ac6867279f0
tuned array theory
src/HOL/Imperative_HOL/Array.thy
--- a/src/HOL/Imperative_HOL/Array.thy	Fri Jul 09 08:11:10 2010 +0200
+++ b/src/HOL/Imperative_HOL/Array.thy	Fri Jul 09 09:48:52 2010 +0200
@@ -8,42 +8,81 @@
 imports Heap_Monad
 begin
 
-subsection {* Primitive layer *}
+subsection {* Primitives *}
 
-definition 
+definition (*FIXME present :: "heap \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> bool" where*)
   array_present :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> bool" where
   "array_present a h \<longleftrightarrow> addr_of_array a < lim h"
 
-definition
+definition (*FIXME get :: "heap \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a list" where*)
   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
+definition (*FIXME set*)
   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
+definition (*FIXME alloc*)
+  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
+definition (*FIXME length :: "heap \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> nat" where*)
+  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
+definition (*FIXME update*)
+  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 *}
+definition (*FIXME noteq*)
+  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"
+
+
+subsection {* Monad operations *}
+
+definition new :: "nat \<Rightarrow> 'a\<Colon>heap \<Rightarrow> 'a array Heap" where
+  [code del]: "new n x = Heap_Monad.heap (array (replicate n x))"
+
+definition of_list :: "'a\<Colon>heap list \<Rightarrow> 'a array Heap" where
+  [code del]: "of_list xs = Heap_Monad.heap (array xs)"
+
+definition make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a\<Colon>heap) \<Rightarrow> 'a array Heap" where
+  [code del]: "make n f = Heap_Monad.heap (array (map f [0 ..< n]))"
+
+definition len :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where
+  [code del]: "len a = Heap_Monad.heap (\<lambda>h. (length a h, h))"
+
+definition nth :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a Heap" where
+  [code del]: "nth a i = Heap_Monad.guard (\<lambda>h. i < length a h)
+    (\<lambda>h. (get_array a h ! i, h))"
+
+definition upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a\<Colon>heap array Heap" where
+  [code del]: "upd i x a = Heap_Monad.guard (\<lambda>h. i < length a h)
+    (\<lambda>h. (a, change a i x h))"
+
+definition map_entry :: "nat \<Rightarrow> ('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap" where
+  [code del]: "map_entry i f a = Heap_Monad.guard (\<lambda>h. i < length a h)
+    (\<lambda>h. (a, change a i (f (get_array a h ! i)) h))"
+
+definition swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a Heap" where
+  [code del]: "swap i x a = Heap_Monad.guard (\<lambda>h. i < length a h)
+    (\<lambda>h. (get_array a h ! i, change a i x h))"
+
+definition freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap" where
+  [code del]: "freeze a = Heap_Monad.heap (\<lambda>h. (get_array a h, h))"
+
+
+subsection {* Properties *}
 
 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
@@ -114,97 +153,65 @@
   "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 *}
+lemma execute_new [simp]:
+  "Heap_Monad.execute (new n x) h = Some (array (replicate n x) h)"
+  by (simp add: new_def)
 
-definition
-  new :: "nat \<Rightarrow> 'a\<Colon>heap \<Rightarrow> 'a array Heap" where
-  [code del]: "new n x = Heap_Monad.heap (Array.array (replicate n x))"
+lemma execute_of_list [simp]:
+  "Heap_Monad.execute (of_list xs) h = Some (array xs h)"
+  by (simp add: of_list_def)
 
-definition
-  of_list :: "'a\<Colon>heap list \<Rightarrow> 'a array Heap" where
-  [code del]: "of_list xs = Heap_Monad.heap (Array.array xs)"
+lemma execute_make [simp]:
+  "Heap_Monad.execute (make n f) h = Some (array (map f [0 ..< n]) h)"
+  by (simp add: make_def)
 
-definition
-  len :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where
-  [code del]: "len arr = Heap_Monad.heap (\<lambda>h. (Array.length arr h, h))"
+lemma execute_len [simp]:
+  "Heap_Monad.execute (len a) h = Some (length a h, h)"
+  by (simp add: len_def)
+
+lemma execute_nth [simp]:
+  "i < length a h \<Longrightarrow>
+    Heap_Monad.execute (nth a i) h = Some (get_array a h ! i, h)"
+  "i \<ge> length a h \<Longrightarrow> Heap_Monad.execute (nth a i) h = None"
+  by (simp_all add: nth_def)
 
-definition
-  nth :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a Heap"
-where
-  [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'')
-              done)"
+lemma execute_upd [simp]:
+  "i < length a h \<Longrightarrow>
+    Heap_Monad.execute (upd i x a) h = Some (a, change a i x h)"
+  "i \<ge> length a h \<Longrightarrow> Heap_Monad.execute (nth a i) h = None"
+  by (simp_all add: upd_def)
 
-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> len a;
-                      (if i < len
-                           then Heap_Monad.heap (\<lambda>h. (a, change a i x h))
-                           else raise ''array update: index out of range'')
-                   done)" 
+lemma execute_map_entry [simp]:
+  "i < length a h \<Longrightarrow>
+   Heap_Monad.execute (map_entry i f a) h =
+      Some (a, change a i (f (get_array a h ! i)) h)"
+  "i \<ge> length a h \<Longrightarrow> Heap_Monad.execute (nth a i) h = None"
+  by (simp_all add: map_entry_def)
+
+lemma execute_swap [simp]:
+  "i < length a h \<Longrightarrow>
+   Heap_Monad.execute (swap i x a) h =
+      Some (get_array a h ! i, change a i x h)"
+  "i \<ge> length a h \<Longrightarrow> Heap_Monad.execute (nth a i) h = None"
+  by (simp_all add: swap_def)
+
+lemma execute_freeze [simp]:
+  "Heap_Monad.execute (freeze a) h = Some (get_array a h, h)"
+  by (simp add: freeze_def)
 
 lemma upd_return:
   "upd i x a \<guillemotright> return a = upd i x a"
-  by (rule Heap_eqI) (simp add: upd_def bindM_def split: option.split) 
-
-
-subsection {* Derivates *}
-
-definition
-  map_entry :: "nat \<Rightarrow> ('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
-where
-  "map_entry i f a = (do
-     x \<leftarrow> nth a i;
-     upd i (f x) a
-   done)"
+  by (rule Heap_eqI) (simp add: bindM_def guard_def upd_def)
 
-definition
-  swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a Heap"
-where
-  "swap i x a = (do
-     y \<leftarrow> nth a i;
-     upd i x a;
-     return y
-   done)"
-
-definition
-  make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a\<Colon>heap) \<Rightarrow> 'a array Heap"
-where
-  "make n f = of_list (map f [0 ..< n])"
+lemma array_make:
+  "new n x = make n (\<lambda>_. x)"
+  by (rule Heap_eqI) (simp add: map_replicate_trivial)
 
-definition
-  freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap"
-where
-  "freeze a = (do
-     n \<leftarrow> len a;
-     mapM (nth a) [0..<n]
-   done)"
+lemma array_of_list_make:
+  "of_list xs = make (List.length xs) (\<lambda>n. xs ! n)"
+  by (rule Heap_eqI) (simp add: map_nth)
 
-definition
-   map :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
-where
-  "map f a = (do
-     n \<leftarrow> len a;
-     mapM (\<lambda>n. map_entry n f a) [0..<n];
-     return a
-   done)"
-
-
-
-subsection {* Properties *}
-
-lemma array_make [code]:
-  "Array.new n x = make n (\<lambda>_. x)"
-  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)"
-  by (rule Heap_eqI) (simp add: make_def map_nth)
+hide_const (open) new map
 
 
 subsection {* Code generator setup *}
@@ -213,48 +220,95 @@
 
 definition new' where
   [code del]: "new' = Array.new o Code_Numeral.nat_of"
-hide_const (open) new'
+
 lemma [code]:
-  "Array.new = Array.new' o Code_Numeral.of_nat"
+  "Array.new = new' o Code_Numeral.of_nat"
   by (simp add: new'_def o_def)
 
 definition of_list' where
   [code del]: "of_list' i xs = Array.of_list (take (Code_Numeral.nat_of i) xs)"
-hide_const (open) of_list'
+
 lemma [code]:
-  "Array.of_list xs = Array.of_list' (Code_Numeral.of_nat (List.length xs)) xs"
+  "Array.of_list xs = of_list' (Code_Numeral.of_nat (List.length xs)) xs"
   by (simp add: of_list'_def)
 
 definition make' where
   [code del]: "make' i f = Array.make (Code_Numeral.nat_of i) (f o Code_Numeral.of_nat)"
-hide_const (open) make'
+
 lemma [code]:
-  "Array.make n f = Array.make' (Code_Numeral.of_nat n) (f o Code_Numeral.nat_of)"
+  "Array.make n f = make' (Code_Numeral.of_nat n) (f o Code_Numeral.nat_of)"
   by (simp add: make'_def o_def)
 
 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.len a = Array.len' a \<guillemotright>= (\<lambda>i. return (Code_Numeral.nat_of i))"
+  "Array.len a = 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"
-hide_const (open) nth'
+
 lemma [code]:
-  "Array.nth a n = Array.nth' a (Code_Numeral.of_nat n)"
+  "Array.nth a n = nth' a (Code_Numeral.of_nat n)"
   by (simp add: nth'_def)
 
 definition upd' where
   [code del]: "upd' a i x = Array.upd (Code_Numeral.nat_of i) x a \<guillemotright> return ()"
-hide_const (open) upd'
+
 lemma [code]:
-  "Array.upd i x a = Array.upd' a (Code_Numeral.of_nat i) x \<guillemotright> return a"
+  "Array.upd i x a = upd' a (Code_Numeral.of_nat i) x \<guillemotright> return a"
   by (simp add: upd'_def upd_return)
 
+lemma [code]:
+  "map_entry i f a = (do
+     x \<leftarrow> nth a i;
+     upd i (f x) a
+   done)"
+  by (rule Heap_eqI) (simp add: bindM_def guard_def map_entry_def)
 
-subsubsection {* SML *}
+lemma [code]:
+  "swap i x a = (do
+     y \<leftarrow> nth a i;
+     upd i x a;
+     return y
+   done)"
+  by (rule Heap_eqI) (simp add: bindM_def guard_def swap_def)
+
+lemma [code]:
+  "freeze a = (do
+     n \<leftarrow> len a;
+     mapM (\<lambda>i. nth a i) [0..<n]
+   done)"
+proof (rule Heap_eqI)
+  fix h
+  have *: "List.map
+     (\<lambda>x. fst (the (if x < length a h
+                    then Some (get_array a h ! x, h) else None)))
+     [0..<length a h] =
+       List.map (List.nth (get_array a h)) [0..<length a h]"
+    by simp
+  have "Heap_Monad.execute (mapM (Array.nth a) [0..<length a h]) h =
+    Some (get_array a h, h)"
+    apply (subst execute_mapM_unchanged_heap)
+    apply (simp_all add: nth_def guard_def *)
+    apply (simp add: length_def map_nth)
+    done
+  then have "Heap_Monad.execute (do
+      n \<leftarrow> len a;
+      mapM (Array.nth a) [0..<n]
+    done) h = Some (get_array a h, h)"
+    by (auto intro: execute_eq_SomeI)
+  then show "Heap_Monad.execute (freeze a) h = Heap_Monad.execute (do
+      n \<leftarrow> len a;
+      mapM (Array.nth a) [0..<n]
+    done) h" by simp
+qed
+
+hide_const (open) new' of_list' make' len' nth' upd'
+
+
+text {* SML *}
 
 code_type array (SML "_/ array")
 code_const Array (SML "raise/ (Fail/ \"bare Array\")")
@@ -268,7 +322,7 @@
 code_reserved SML Array
 
 
-subsubsection {* OCaml *}
+text {* OCaml *}
 
 code_type array (OCaml "_/ array")
 code_const Array (OCaml "failwith/ \"bare Array\"")
@@ -281,7 +335,7 @@
 code_reserved OCaml Array
 
 
-subsubsection {* Haskell *}
+text {* Haskell *}
 
 code_type array (Haskell "Heap.STArray/ Heap.RealWorld/ _")
 code_const Array (Haskell "error/ \"bare Array\"")
@@ -291,6 +345,4 @@
 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