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