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