| author | blanchet | 
| Wed, 23 Jun 2010 09:40:06 +0200 | |
| changeset 37511 | 26afa11a1fb2 | 
| parent 36176 | 3fe7e97ccca8 | 
| child 37709 | 70fafefbcc98 | 
| permissions | -rw-r--r-- | 
| 31870 | 1 | (* Title: HOL/Imperative_HOL/Array.thy | 
| 26170 | 2 | Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen | 
| 3 | *) | |
| 4 | ||
| 5 | header {* Monadic arrays *}
 | |
| 6 | ||
| 7 | theory Array | |
| 31203 
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
 haftmann parents: 
29822diff
changeset | 8 | imports Heap_Monad | 
| 26170 | 9 | begin | 
| 10 | ||
| 11 | subsection {* Primitives *}
 | |
| 12 | ||
| 13 | definition | |
| 14 | new :: "nat \<Rightarrow> 'a\<Colon>heap \<Rightarrow> 'a array Heap" where | |
| 15 | [code del]: "new n x = Heap_Monad.heap (Heap.array n x)" | |
| 16 | ||
| 17 | definition | |
| 18 | of_list :: "'a\<Colon>heap list \<Rightarrow> 'a array Heap" where | |
| 19 | [code del]: "of_list xs = Heap_Monad.heap (Heap.array_of_list xs)" | |
| 20 | ||
| 21 | definition | |
| 22 | length :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where | |
| 23 | [code del]: "length arr = Heap_Monad.heap (\<lambda>h. (Heap.length arr h, h))" | |
| 24 | ||
| 25 | definition | |
| 26 | nth :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a Heap" | |
| 27 | where | |
| 28 | [code del]: "nth a i = (do len \<leftarrow> length a; | |
| 29 | (if i < len | |
| 30 | then Heap_Monad.heap (\<lambda>h. (get_array a h ! i, h)) | |
| 31 |                      else raise (''array lookup: index out of range''))
 | |
| 32 | done)" | |
| 33 | ||
| 34 | definition | |
| 35 | upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a\<Colon>heap array Heap" | |
| 36 | where | |
| 37 | [code del]: "upd i x a = (do len \<leftarrow> length a; | |
| 38 | (if i < len | |
| 26719 | 39 | then Heap_Monad.heap (\<lambda>h. (a, Heap.upd a i x h)) | 
| 40 |                            else raise (''array update: index out of range''))
 | |
| 26170 | 41 | done)" | 
| 42 | ||
| 43 | lemma upd_return: | |
| 44 | "upd i x a \<guillemotright> return a = upd i x a" | |
| 26719 | 45 | proof (rule Heap_eqI) | 
| 46 | fix h | |
| 47 | obtain len h' where "Heap_Monad.execute (Array.length a) h = (len, h')" | |
| 48 | by (cases "Heap_Monad.execute (Array.length a) h") | |
| 49 | then show "Heap_Monad.execute (upd i x a \<guillemotright> return a) h = Heap_Monad.execute (upd i x a) h" | |
| 28145 | 50 | by (auto simp add: upd_def bindM_def split: sum.split) | 
| 26719 | 51 | qed | 
| 26170 | 52 | |
| 53 | ||
| 54 | subsection {* Derivates *}
 | |
| 55 | ||
| 56 | definition | |
| 57 |   map_entry :: "nat \<Rightarrow> ('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
 | |
| 58 | where | |
| 59 | "map_entry i f a = (do | |
| 60 | x \<leftarrow> nth a i; | |
| 61 | upd i (f x) a | |
| 62 | done)" | |
| 63 | ||
| 64 | definition | |
| 65 | swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a Heap" | |
| 66 | where | |
| 67 | "swap i x a = (do | |
| 68 | y \<leftarrow> nth a i; | |
| 69 | upd i x a; | |
| 27596 | 70 | return y | 
| 26170 | 71 | done)" | 
| 72 | ||
| 73 | definition | |
| 74 | make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a\<Colon>heap) \<Rightarrow> 'a array Heap" | |
| 75 | where | |
| 76 | "make n f = of_list (map f [0 ..< n])" | |
| 77 | ||
| 78 | definition | |
| 79 | freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap" | |
| 80 | where | |
| 81 | "freeze a = (do | |
| 82 | n \<leftarrow> length a; | |
| 83 | mapM (nth a) [0..<n] | |
| 84 | done)" | |
| 85 | ||
| 27656 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: 
27596diff
changeset | 86 | definition | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: 
27596diff
changeset | 87 |    map :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
 | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: 
27596diff
changeset | 88 | where | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: 
27596diff
changeset | 89 | "map f a = (do | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: 
27596diff
changeset | 90 | n \<leftarrow> length a; | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: 
27596diff
changeset | 91 | mapM (\<lambda>n. map_entry n f a) [0..<n]; | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: 
27596diff
changeset | 92 | return a | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: 
27596diff
changeset | 93 | done)" | 
| 
d4f6e64ee7cc
added verification framework for the HeapMonad and quicksort as example for this framework
 bulwahn parents: 
27596diff
changeset | 94 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35846diff
changeset | 95 | hide_const (open) new map -- {* avoid clashed with some popular names *}
 | 
| 26170 | 96 | |
| 97 | ||
| 98 | subsection {* Properties *}
 | |
| 99 | ||
| 28562 | 100 | lemma array_make [code]: | 
| 26170 | 101 | "Array.new n x = make n (\<lambda>_. x)" | 
| 102 | by (induct n) (simp_all add: make_def new_def Heap_Monad.heap_def | |
| 103 | monad_simp array_of_list_replicate [symmetric] | |
| 104 | map_replicate_trivial replicate_append_same | |
| 105 | of_list_def) | |
| 106 | ||
| 28562 | 107 | lemma array_of_list_make [code]: | 
| 26170 | 108 | "of_list xs = make (List.length xs) (\<lambda>n. xs ! n)" | 
| 109 | unfolding make_def map_nth .. | |
| 110 | ||
| 26182 | 111 | |
| 112 | subsection {* Code generator setup *}
 | |
| 113 | ||
| 114 | subsubsection {* Logical intermediate layer *}
 | |
| 115 | ||
| 116 | definition new' where | |
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 117 | [code del]: "new' = Array.new o Code_Numeral.nat_of" | 
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35846diff
changeset | 118 | hide_const (open) new' | 
| 28562 | 119 | lemma [code]: | 
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 120 | "Array.new = Array.new' o Code_Numeral.of_nat" | 
| 26182 | 121 | by (simp add: new'_def o_def) | 
| 122 | ||
| 123 | definition of_list' where | |
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 124 | [code del]: "of_list' i xs = Array.of_list (take (Code_Numeral.nat_of i) xs)" | 
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35846diff
changeset | 125 | hide_const (open) of_list' | 
| 28562 | 126 | lemma [code]: | 
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 127 | "Array.of_list xs = Array.of_list' (Code_Numeral.of_nat (List.length xs)) xs" | 
| 26182 | 128 | by (simp add: of_list'_def) | 
| 129 | ||
| 130 | definition make' where | |
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 131 | [code del]: "make' i f = Array.make (Code_Numeral.nat_of i) (f o Code_Numeral.of_nat)" | 
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35846diff
changeset | 132 | hide_const (open) make' | 
| 28562 | 133 | lemma [code]: | 
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 134 | "Array.make n f = Array.make' (Code_Numeral.of_nat n) (f o Code_Numeral.nat_of)" | 
| 26182 | 135 | by (simp add: make'_def o_def) | 
| 136 | ||
| 137 | definition length' where | |
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 138 | [code del]: "length' = Array.length \<guillemotright>== liftM Code_Numeral.of_nat" | 
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35846diff
changeset | 139 | hide_const (open) length' | 
| 28562 | 140 | lemma [code]: | 
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 141 | "Array.length = Array.length' \<guillemotright>== liftM Code_Numeral.nat_of" | 
| 26182 | 142 | by (simp add: length'_def monad_simp', | 
| 143 | simp add: liftM_def comp_def monad_simp, | |
| 144 | simp add: monad_simp') | |
| 145 | ||
| 146 | definition nth' where | |
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 147 | [code del]: "nth' a = Array.nth a o Code_Numeral.nat_of" | 
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35846diff
changeset | 148 | hide_const (open) nth' | 
| 28562 | 149 | lemma [code]: | 
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 150 | "Array.nth a n = Array.nth' a (Code_Numeral.of_nat n)" | 
| 26182 | 151 | by (simp add: nth'_def) | 
| 152 | ||
| 153 | definition upd' where | |
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 154 | [code del]: "upd' a i x = Array.upd (Code_Numeral.nat_of i) x a \<guillemotright> return ()" | 
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35846diff
changeset | 155 | hide_const (open) upd' | 
| 28562 | 156 | lemma [code]: | 
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31203diff
changeset | 157 | "Array.upd i x a = Array.upd' a (Code_Numeral.of_nat i) x \<guillemotright> return a" | 
| 26743 | 158 | by (simp add: upd'_def monad_simp upd_return) | 
| 26182 | 159 | |
| 160 | ||
| 161 | subsubsection {* SML *}
 | |
| 162 | ||
| 163 | code_type array (SML "_/ array") | |
| 164 | code_const Array (SML "raise/ (Fail/ \"bare Array\")") | |
| 26752 | 165 | code_const Array.new' (SML "(fn/ ()/ =>/ Array.array/ ((_),/ (_)))") | 
| 35846 | 166 | code_const Array.of_list' (SML "(fn/ ()/ =>/ Array.fromList/ _)") | 
| 26752 | 167 | code_const Array.make' (SML "(fn/ ()/ =>/ Array.tabulate/ ((_),/ (_)))") | 
| 168 | code_const Array.length' (SML "(fn/ ()/ =>/ Array.length/ _)") | |
| 169 | code_const Array.nth' (SML "(fn/ ()/ =>/ Array.sub/ ((_),/ (_)))") | |
| 170 | code_const Array.upd' (SML "(fn/ ()/ =>/ Array.update/ ((_),/ (_),/ (_)))") | |
| 26182 | 171 | |
| 172 | code_reserved SML Array | |
| 173 | ||
| 174 | ||
| 175 | subsubsection {* OCaml *}
 | |
| 176 | ||
| 177 | code_type array (OCaml "_/ array") | |
| 178 | code_const Array (OCaml "failwith/ \"bare Array\"") | |
| 32580 | 179 | code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ (Big'_int.int'_of'_big'_int/ _)/ _)") | 
| 35846 | 180 | code_const Array.of_list' (OCaml "(fun/ ()/ ->/ Array.of'_list/ _)") | 
| 32580 | 181 | code_const Array.length' (OCaml "(fun/ ()/ ->/ Big'_int.big'_int'_of'_int/ (Array.length/ _))") | 
| 182 | code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ (Big'_int.int'_of'_big'_int/ _))") | |
| 183 | code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ (Big'_int.int'_of'_big'_int/ _)/ _)") | |
| 26182 | 184 | |
| 185 | code_reserved OCaml Array | |
| 186 | ||
| 187 | ||
| 188 | subsubsection {* Haskell *}
 | |
| 189 | ||
| 29793 | 190 | code_type array (Haskell "Heap.STArray/ Heap.RealWorld/ _") | 
| 26182 | 191 | code_const Array (Haskell "error/ \"bare Array\"") | 
| 29793 | 192 | code_const Array.new' (Haskell "Heap.newArray/ (0,/ _)") | 
| 193 | code_const Array.of_list' (Haskell "Heap.newListArray/ (0,/ _)") | |
| 194 | code_const Array.length' (Haskell "Heap.lengthArray") | |
| 195 | code_const Array.nth' (Haskell "Heap.readArray") | |
| 196 | code_const Array.upd' (Haskell "Heap.writeArray") | |
| 26182 | 197 | |
| 26170 | 198 | end |