src/HOL/Imperative_HOL/Array.thy
author haftmann
Thu, 08 Jan 2009 17:10:41 +0100
changeset 29396 ebcd69a00872
parent 28562 src/HOL/Library/Array.thy@4e74209f113e
child 29730 86cac1fab613
permissions -rw-r--r--
split of Imperative_HOL theories from HOL-Library
haftmann@26170
     1
(*  Title:      HOL/Library/Array.thy
haftmann@26170
     2
    ID:         $Id$
haftmann@26170
     3
    Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
haftmann@26170
     4
*)
haftmann@26170
     5
haftmann@26170
     6
header {* Monadic arrays *}
haftmann@26170
     7
haftmann@26170
     8
theory Array
haftmann@26182
     9
imports Heap_Monad Code_Index
haftmann@26170
    10
begin
haftmann@26170
    11
haftmann@26170
    12
subsection {* Primitives *}
haftmann@26170
    13
haftmann@26170
    14
definition
haftmann@26170
    15
  new :: "nat \<Rightarrow> 'a\<Colon>heap \<Rightarrow> 'a array Heap" where
haftmann@26170
    16
  [code del]: "new n x = Heap_Monad.heap (Heap.array n x)"
haftmann@26170
    17
haftmann@26170
    18
definition
haftmann@26170
    19
  of_list :: "'a\<Colon>heap list \<Rightarrow> 'a array Heap" where
haftmann@26170
    20
  [code del]: "of_list xs = Heap_Monad.heap (Heap.array_of_list xs)"
haftmann@26170
    21
haftmann@26170
    22
definition
haftmann@26170
    23
  length :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where
haftmann@26170
    24
  [code del]: "length arr = Heap_Monad.heap (\<lambda>h. (Heap.length arr h, h))"
haftmann@26170
    25
haftmann@26170
    26
definition
haftmann@26170
    27
  nth :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a Heap"
haftmann@26170
    28
where
haftmann@26170
    29
  [code del]: "nth a i = (do len \<leftarrow> length a;
haftmann@26170
    30
                 (if i < len
haftmann@26170
    31
                     then Heap_Monad.heap (\<lambda>h. (get_array a h ! i, h))
haftmann@26170
    32
                     else raise (''array lookup: index out of range''))
haftmann@26170
    33
              done)"
haftmann@26170
    34
haftmann@26170
    35
-- {* FIXME adjustion for List theory *}
haftmann@26170
    36
no_syntax
haftmann@26170
    37
  nth  :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!" 100)
haftmann@26170
    38
haftmann@26170
    39
abbreviation
haftmann@26170
    40
  nth_list :: "'a list \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!" 100)
haftmann@26170
    41
where
haftmann@26170
    42
  "nth_list \<equiv> List.nth"
haftmann@26170
    43
haftmann@26170
    44
definition
haftmann@26170
    45
  upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a\<Colon>heap array Heap"
haftmann@26170
    46
where
haftmann@26170
    47
  [code del]: "upd i x a = (do len \<leftarrow> length a;
haftmann@26170
    48
                      (if i < len
haftmann@26719
    49
                           then Heap_Monad.heap (\<lambda>h. (a, Heap.upd a i x h))
haftmann@26719
    50
                           else raise (''array update: index out of range''))
haftmann@26170
    51
                   done)" 
haftmann@26170
    52
haftmann@26170
    53
lemma upd_return:
haftmann@26170
    54
  "upd i x a \<guillemotright> return a = upd i x a"
haftmann@26719
    55
proof (rule Heap_eqI)
haftmann@26719
    56
  fix h
haftmann@26719
    57
  obtain len h' where "Heap_Monad.execute (Array.length a) h = (len, h')"
haftmann@26719
    58
    by (cases "Heap_Monad.execute (Array.length a) h")
haftmann@26719
    59
  then show "Heap_Monad.execute (upd i x a \<guillemotright> return a) h = Heap_Monad.execute (upd i x a) h"
haftmann@28145
    60
    by (auto simp add: upd_def bindM_def split: sum.split)
haftmann@26719
    61
qed
haftmann@26170
    62
haftmann@26170
    63
haftmann@26170
    64
subsection {* Derivates *}
haftmann@26170
    65
haftmann@26170
    66
definition
haftmann@26170
    67
  map_entry :: "nat \<Rightarrow> ('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
haftmann@26170
    68
where
haftmann@26170
    69
  "map_entry i f a = (do
haftmann@26170
    70
     x \<leftarrow> nth a i;
haftmann@26170
    71
     upd i (f x) a
haftmann@26170
    72
   done)"
haftmann@26170
    73
haftmann@26170
    74
definition
haftmann@26170
    75
  swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a Heap"
haftmann@26170
    76
where
haftmann@26170
    77
  "swap i x a = (do
haftmann@26170
    78
     y \<leftarrow> nth a i;
haftmann@26170
    79
     upd i x a;
haftmann@27596
    80
     return y
haftmann@26170
    81
   done)"
haftmann@26170
    82
haftmann@26170
    83
definition
haftmann@26170
    84
  make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a\<Colon>heap) \<Rightarrow> 'a array Heap"
haftmann@26170
    85
where
haftmann@26170
    86
  "make n f = of_list (map f [0 ..< n])"
haftmann@26170
    87
haftmann@26170
    88
definition
haftmann@26170
    89
  freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap"
haftmann@26170
    90
where
haftmann@26170
    91
  "freeze a = (do
haftmann@26170
    92
     n \<leftarrow> length a;
haftmann@26170
    93
     mapM (nth a) [0..<n]
haftmann@26170
    94
   done)"
haftmann@26170
    95
bulwahn@27656
    96
definition
bulwahn@27656
    97
   map :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap"
bulwahn@27656
    98
where
bulwahn@27656
    99
  "map f a = (do
bulwahn@27656
   100
     n \<leftarrow> length a;
bulwahn@27656
   101
     mapM (\<lambda>n. map_entry n f a) [0..<n];
bulwahn@27656
   102
     return a
bulwahn@27656
   103
   done)"
bulwahn@27656
   104
bulwahn@27656
   105
hide (open) const new map -- {* avoid clashed with some popular names *}
haftmann@26170
   106
haftmann@26170
   107
haftmann@26170
   108
subsection {* Properties *}
haftmann@26170
   109
haftmann@28562
   110
lemma array_make [code]:
haftmann@26170
   111
  "Array.new n x = make n (\<lambda>_. x)"
haftmann@26170
   112
  by (induct n) (simp_all add: make_def new_def Heap_Monad.heap_def
haftmann@26170
   113
    monad_simp array_of_list_replicate [symmetric]
haftmann@26170
   114
    map_replicate_trivial replicate_append_same
haftmann@26170
   115
    of_list_def)
haftmann@26170
   116
haftmann@28562
   117
lemma array_of_list_make [code]:
haftmann@26170
   118
  "of_list xs = make (List.length xs) (\<lambda>n. xs ! n)"
haftmann@26170
   119
  unfolding make_def map_nth ..
haftmann@26170
   120
haftmann@26182
   121
haftmann@26182
   122
subsection {* Code generator setup *}
haftmann@26182
   123
haftmann@26182
   124
subsubsection {* Logical intermediate layer *}
haftmann@26182
   125
haftmann@26182
   126
definition new' where
haftmann@26182
   127
  [code del]: "new' = Array.new o nat_of_index"
haftmann@26182
   128
hide (open) const new'
haftmann@28562
   129
lemma [code]:
haftmann@26182
   130
  "Array.new = Array.new' o index_of_nat"
haftmann@26182
   131
  by (simp add: new'_def o_def)
haftmann@26182
   132
haftmann@26182
   133
definition of_list' where
haftmann@26182
   134
  [code del]: "of_list' i xs = Array.of_list (take (nat_of_index i) xs)"
haftmann@26182
   135
hide (open) const of_list'
haftmann@28562
   136
lemma [code]:
haftmann@26182
   137
  "Array.of_list xs = Array.of_list' (index_of_nat (List.length xs)) xs"
haftmann@26182
   138
  by (simp add: of_list'_def)
haftmann@26182
   139
haftmann@26182
   140
definition make' where
haftmann@26182
   141
  [code del]: "make' i f = Array.make (nat_of_index i) (f o index_of_nat)"
haftmann@26182
   142
hide (open) const make'
haftmann@28562
   143
lemma [code]:
haftmann@26182
   144
  "Array.make n f = Array.make' (index_of_nat n) (f o nat_of_index)"
haftmann@26182
   145
  by (simp add: make'_def o_def)
haftmann@26182
   146
haftmann@26182
   147
definition length' where
haftmann@26182
   148
  [code del]: "length' = Array.length \<guillemotright>== liftM index_of_nat"
haftmann@26182
   149
hide (open) const length'
haftmann@28562
   150
lemma [code]:
haftmann@26182
   151
  "Array.length = Array.length' \<guillemotright>== liftM nat_of_index"
haftmann@26182
   152
  by (simp add: length'_def monad_simp',
haftmann@26182
   153
    simp add: liftM_def comp_def monad_simp,
haftmann@26182
   154
    simp add: monad_simp')
haftmann@26182
   155
haftmann@26182
   156
definition nth' where
haftmann@26182
   157
  [code del]: "nth' a = Array.nth a o nat_of_index"
haftmann@26182
   158
hide (open) const nth'
haftmann@28562
   159
lemma [code]:
haftmann@26182
   160
  "Array.nth a n = Array.nth' a (index_of_nat n)"
haftmann@26182
   161
  by (simp add: nth'_def)
haftmann@26182
   162
haftmann@26182
   163
definition upd' where
haftmann@26182
   164
  [code del]: "upd' a i x = Array.upd (nat_of_index i) x a \<guillemotright> return ()"
haftmann@26182
   165
hide (open) const upd'
haftmann@28562
   166
lemma [code]:
haftmann@26182
   167
  "Array.upd i x a = Array.upd' a (index_of_nat i) x \<guillemotright> return a"
haftmann@26743
   168
  by (simp add: upd'_def monad_simp upd_return)
haftmann@26182
   169
haftmann@26182
   170
haftmann@26182
   171
subsubsection {* SML *}
haftmann@26182
   172
haftmann@26182
   173
code_type array (SML "_/ array")
haftmann@26182
   174
code_const Array (SML "raise/ (Fail/ \"bare Array\")")
haftmann@26752
   175
code_const Array.new' (SML "(fn/ ()/ =>/ Array.array/ ((_),/ (_)))")
haftmann@26752
   176
code_const Array.of_list (SML "(fn/ ()/ =>/ Array.fromList/ _)")
haftmann@26752
   177
code_const Array.make' (SML "(fn/ ()/ =>/ Array.tabulate/ ((_),/ (_)))")
haftmann@26752
   178
code_const Array.length' (SML "(fn/ ()/ =>/ Array.length/ _)")
haftmann@26752
   179
code_const Array.nth' (SML "(fn/ ()/ =>/ Array.sub/ ((_),/ (_)))")
haftmann@26752
   180
code_const Array.upd' (SML "(fn/ ()/ =>/ Array.update/ ((_),/ (_),/ (_)))")
haftmann@26182
   181
haftmann@26182
   182
code_reserved SML Array
haftmann@26182
   183
haftmann@26182
   184
haftmann@26182
   185
subsubsection {* OCaml *}
haftmann@26182
   186
haftmann@26182
   187
code_type array (OCaml "_/ array")
haftmann@26182
   188
code_const Array (OCaml "failwith/ \"bare Array\"")
haftmann@27673
   189
code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ _/ _)")
haftmann@27673
   190
code_const Array.of_list (OCaml "(fun/ ()/ ->/ Array.of'_list/ _)")
haftmann@27673
   191
code_const Array.make' (OCaml "(fun/ ()/ ->/ Array.init/ _/ _)")
haftmann@27673
   192
code_const Array.length' (OCaml "(fun/ ()/ ->/ Array.length/ _)")
haftmann@27673
   193
code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ _)")
haftmann@27673
   194
code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ _/ _)")
haftmann@26182
   195
haftmann@26182
   196
code_reserved OCaml Array
haftmann@26182
   197
haftmann@26182
   198
haftmann@26182
   199
subsubsection {* Haskell *}
haftmann@26182
   200
haftmann@27695
   201
code_type array (Haskell "STArray/ RealWorld/ _")
haftmann@26182
   202
code_const Array (Haskell "error/ \"bare Array\"")
haftmann@26182
   203
code_const Array.new' (Haskell "newArray/ (0,/ _)")
haftmann@26182
   204
code_const Array.of_list' (Haskell "newListArray/ (0,/ _)")
haftmann@27673
   205
code_const Array.length' (Haskell "lengthArray")
haftmann@26182
   206
code_const Array.nth' (Haskell "readArray")
haftmann@26182
   207
code_const Array.upd' (Haskell "writeArray")
haftmann@26182
   208
haftmann@26170
   209
end