src/HOL/Imperative_HOL/Heap_Monad.thy
author haftmann
Thu, 08 Jan 2009 17:10:41 +0100
changeset 29396 ebcd69a00872
parent 28742 src/HOL/Library/Heap_Monad.thy@07073b1087dd
child 29730 86cac1fab613
permissions -rw-r--r--
split of Imperative_HOL theories from HOL-Library
haftmann@26170
     1
(*  Title:      HOL/Library/Heap_Monad.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 {* A monad with a polymorphic heap *}
haftmann@26170
     7
haftmann@26170
     8
theory Heap_Monad
haftmann@26170
     9
imports Heap
haftmann@26170
    10
begin
haftmann@26170
    11
haftmann@26170
    12
subsection {* The monad *}
haftmann@26170
    13
haftmann@26170
    14
subsubsection {* Monad combinators *}
haftmann@26170
    15
haftmann@26170
    16
datatype exception = Exn
haftmann@26170
    17
haftmann@26170
    18
text {* Monadic heap actions either produce values
haftmann@26170
    19
  and transform the heap, or fail *}
haftmann@26170
    20
datatype 'a Heap = Heap "heap \<Rightarrow> ('a + exception) \<times> heap"
haftmann@26170
    21
haftmann@26170
    22
primrec
haftmann@26170
    23
  execute :: "'a Heap \<Rightarrow> heap \<Rightarrow> ('a + exception) \<times> heap" where
haftmann@26170
    24
  "execute (Heap f) = f"
haftmann@26170
    25
lemmas [code del] = execute.simps
haftmann@26170
    26
haftmann@26170
    27
lemma Heap_execute [simp]:
haftmann@26170
    28
  "Heap (execute f) = f" by (cases f) simp_all
haftmann@26170
    29
haftmann@26170
    30
lemma Heap_eqI:
haftmann@26170
    31
  "(\<And>h. execute f h = execute g h) \<Longrightarrow> f = g"
haftmann@26170
    32
    by (cases f, cases g) (auto simp: expand_fun_eq)
haftmann@26170
    33
haftmann@26170
    34
lemma Heap_eqI':
haftmann@26170
    35
  "(\<And>h. (\<lambda>x. execute (f x) h) = (\<lambda>y. execute (g y) h)) \<Longrightarrow> f = g"
haftmann@26170
    36
    by (auto simp: expand_fun_eq intro: Heap_eqI)
haftmann@26170
    37
haftmann@26170
    38
lemma Heap_strip: "(\<And>f. PROP P f) \<equiv> (\<And>g. PROP P (Heap g))"
haftmann@26170
    39
proof
haftmann@26170
    40
  fix g :: "heap \<Rightarrow> ('a + exception) \<times> heap" 
haftmann@26170
    41
  assume "\<And>f. PROP P f"
haftmann@26170
    42
  then show "PROP P (Heap g)" .
haftmann@26170
    43
next
haftmann@26170
    44
  fix f :: "'a Heap" 
haftmann@26170
    45
  assume assm: "\<And>g. PROP P (Heap g)"
haftmann@26170
    46
  then have "PROP P (Heap (execute f))" .
haftmann@26170
    47
  then show "PROP P f" by simp
haftmann@26170
    48
qed
haftmann@26170
    49
haftmann@26170
    50
definition
haftmann@26170
    51
  heap :: "(heap \<Rightarrow> 'a \<times> heap) \<Rightarrow> 'a Heap" where
haftmann@26170
    52
  [code del]: "heap f = Heap (\<lambda>h. apfst Inl (f h))"
haftmann@26170
    53
haftmann@26170
    54
lemma execute_heap [simp]:
haftmann@26170
    55
  "execute (heap f) h = apfst Inl (f h)"
haftmann@26170
    56
  by (simp add: heap_def)
haftmann@26170
    57
haftmann@26170
    58
definition
haftmann@26170
    59
  bindM :: "'a Heap \<Rightarrow> ('a \<Rightarrow> 'b Heap) \<Rightarrow> 'b Heap" (infixl ">>=" 54) where
haftmann@26170
    60
  [code del]: "f >>= g = Heap (\<lambda>h. case execute f h of
haftmann@26170
    61
                  (Inl x, h') \<Rightarrow> execute (g x) h'
haftmann@26170
    62
                | r \<Rightarrow> r)"
haftmann@26170
    63
haftmann@26170
    64
notation
haftmann@26170
    65
  bindM (infixl "\<guillemotright>=" 54)
haftmann@26170
    66
haftmann@26170
    67
abbreviation
haftmann@26170
    68
  chainM :: "'a Heap \<Rightarrow> 'b Heap \<Rightarrow> 'b Heap"  (infixl ">>" 54) where
haftmann@26170
    69
  "f >> g \<equiv> f >>= (\<lambda>_. g)"
haftmann@26170
    70
haftmann@26170
    71
notation
haftmann@26170
    72
  chainM (infixl "\<guillemotright>" 54)
haftmann@26170
    73
haftmann@26170
    74
definition
haftmann@26170
    75
  return :: "'a \<Rightarrow> 'a Heap" where
haftmann@26170
    76
  [code del]: "return x = heap (Pair x)"
haftmann@26170
    77
haftmann@26170
    78
lemma execute_return [simp]:
haftmann@26170
    79
  "execute (return x) h = apfst Inl (x, h)"
haftmann@26170
    80
  by (simp add: return_def)
haftmann@26170
    81
haftmann@26170
    82
definition
haftmann@26170
    83
  raise :: "string \<Rightarrow> 'a Heap" where -- {* the string is just decoration *}
haftmann@26170
    84
  [code del]: "raise s = Heap (Pair (Inr Exn))"
haftmann@26170
    85
haftmann@26170
    86
notation (latex output)
haftmann@26170
    87
  "raise" ("\<^raw:{\textsf{raise}}>")
haftmann@26170
    88
haftmann@26170
    89
lemma execute_raise [simp]:
haftmann@26170
    90
  "execute (raise s) h = (Inr Exn, h)"
haftmann@26170
    91
  by (simp add: raise_def)
haftmann@26170
    92
haftmann@26170
    93
haftmann@26170
    94
subsubsection {* do-syntax *}
haftmann@26170
    95
haftmann@26170
    96
text {*
haftmann@26170
    97
  We provide a convenient do-notation for monadic expressions
haftmann@26170
    98
  well-known from Haskell.  @{const Let} is printed
haftmann@26170
    99
  specially in do-expressions.
haftmann@26170
   100
*}
haftmann@26170
   101
haftmann@26170
   102
nonterminals do_expr
haftmann@26170
   103
haftmann@26170
   104
syntax
haftmann@26170
   105
  "_do" :: "do_expr \<Rightarrow> 'a"
haftmann@26170
   106
    ("(do (_)//done)" [12] 100)
haftmann@26170
   107
  "_bindM" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
haftmann@26170
   108
    ("_ <- _;//_" [1000, 13, 12] 12)
haftmann@26170
   109
  "_chainM" :: "'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
haftmann@26170
   110
    ("_;//_" [13, 12] 12)
haftmann@26170
   111
  "_let" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
haftmann@26170
   112
    ("let _ = _;//_" [1000, 13, 12] 12)
haftmann@26170
   113
  "_nil" :: "'a \<Rightarrow> do_expr"
haftmann@26170
   114
    ("_" [12] 12)
haftmann@26170
   115
haftmann@26170
   116
syntax (xsymbols)
haftmann@26170
   117
  "_bindM" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
haftmann@26170
   118
    ("_ \<leftarrow> _;//_" [1000, 13, 12] 12)
haftmann@26170
   119
syntax (latex output)
haftmann@26170
   120
  "_do" :: "do_expr \<Rightarrow> 'a"
haftmann@26170
   121
    ("(\<^raw:{\textsf{do}}> (_))" [12] 100)
haftmann@26170
   122
  "_let" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
haftmann@26170
   123
    ("\<^raw:\textsf{let}> _ = _;//_" [1000, 13, 12] 12)
haftmann@26170
   124
notation (latex output)
haftmann@26170
   125
  "return" ("\<^raw:{\textsf{return}}>")
haftmann@26170
   126
haftmann@26170
   127
translations
haftmann@28145
   128
  "_do f" => "f"
haftmann@26170
   129
  "_bindM x f g" => "f \<guillemotright>= (\<lambda>x. g)"
haftmann@26170
   130
  "_chainM f g" => "f \<guillemotright> g"
haftmann@26170
   131
  "_let x t f" => "CONST Let t (\<lambda>x. f)"
haftmann@26170
   132
  "_nil f" => "f"
haftmann@26170
   133
haftmann@26170
   134
print_translation {*
haftmann@26170
   135
let
haftmann@26170
   136
  fun dest_abs_eta (Abs (abs as (_, ty, _))) =
haftmann@26170
   137
        let
haftmann@26170
   138
          val (v, t) = Syntax.variant_abs abs;
haftmann@28145
   139
        in (Free (v, ty), t) end
haftmann@26170
   140
    | dest_abs_eta t =
haftmann@26170
   141
        let
haftmann@26170
   142
          val (v, t) = Syntax.variant_abs ("", dummyT, t $ Bound 0);
haftmann@28145
   143
        in (Free (v, dummyT), t) end;
haftmann@26170
   144
  fun unfold_monad (Const (@{const_syntax bindM}, _) $ f $ g) =
haftmann@26170
   145
        let
haftmann@28145
   146
          val (v, g') = dest_abs_eta g;
haftmann@28145
   147
          val vs = fold_aterms (fn Free (v, _) => insert (op =) v | _ => I) v [];
haftmann@26170
   148
          val v_used = fold_aterms
haftmann@28145
   149
            (fn Free (w, _) => (fn s => s orelse member (op =) vs w) | _ => I) g' false;
haftmann@26170
   150
        in if v_used then
haftmann@28145
   151
          Const ("_bindM", dummyT) $ v $ f $ unfold_monad g'
haftmann@26170
   152
        else
haftmann@26170
   153
          Const ("_chainM", dummyT) $ f $ unfold_monad g'
haftmann@26170
   154
        end
haftmann@26170
   155
    | unfold_monad (Const (@{const_syntax chainM}, _) $ f $ g) =
haftmann@26170
   156
        Const ("_chainM", dummyT) $ f $ unfold_monad g
haftmann@26170
   157
    | unfold_monad (Const (@{const_syntax Let}, _) $ f $ g) =
haftmann@26170
   158
        let
haftmann@28145
   159
          val (v, g') = dest_abs_eta g;
haftmann@28145
   160
        in Const ("_let", dummyT) $ v $ f $ unfold_monad g' end
haftmann@26170
   161
    | unfold_monad (Const (@{const_syntax Pair}, _) $ f) =
haftmann@28145
   162
        Const (@{const_syntax return}, dummyT) $ f
haftmann@26170
   163
    | unfold_monad f = f;
haftmann@28145
   164
  fun contains_bindM (Const (@{const_syntax bindM}, _) $ _ $ _) = true
haftmann@28145
   165
    | contains_bindM (Const (@{const_syntax Let}, _) $ _ $ Abs (_, _, t)) =
haftmann@28145
   166
        contains_bindM t;
haftmann@28145
   167
  fun bindM_monad_tr' (f::g::ts) = list_comb
haftmann@28145
   168
    (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax bindM}, dummyT) $ f $ g), ts);
haftmann@28145
   169
  fun Let_monad_tr' (f :: (g as Abs (_, _, g')) :: ts) = if contains_bindM g' then list_comb
haftmann@28145
   170
      (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax Let}, dummyT) $ f $ g), ts)
haftmann@28145
   171
    else raise Match;
haftmann@28145
   172
in [
haftmann@28145
   173
  (@{const_syntax bindM}, bindM_monad_tr'),
haftmann@28145
   174
  (@{const_syntax Let}, Let_monad_tr')
haftmann@28145
   175
] end;
haftmann@26170
   176
*}
haftmann@26170
   177
haftmann@26170
   178
haftmann@26170
   179
subsection {* Monad properties *}
haftmann@26170
   180
haftmann@26170
   181
subsubsection {* Monad laws *}
haftmann@26170
   182
haftmann@26170
   183
lemma return_bind: "return x \<guillemotright>= f = f x"
haftmann@26170
   184
  by (simp add: bindM_def return_def)
haftmann@26170
   185
haftmann@26170
   186
lemma bind_return: "f \<guillemotright>= return = f"
haftmann@26170
   187
proof (rule Heap_eqI)
haftmann@26170
   188
  fix h
haftmann@26170
   189
  show "execute (f \<guillemotright>= return) h = execute f h"
haftmann@26170
   190
    by (auto simp add: bindM_def return_def split: sum.splits prod.splits)
haftmann@26170
   191
qed
haftmann@26170
   192
haftmann@26170
   193
lemma bind_bind: "(f \<guillemotright>= g) \<guillemotright>= h = f \<guillemotright>= (\<lambda>x. g x \<guillemotright>= h)"
haftmann@26170
   194
  by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits)
haftmann@26170
   195
haftmann@26170
   196
lemma bind_bind': "f \<guillemotright>= (\<lambda>x. g x \<guillemotright>= h x) = f \<guillemotright>= (\<lambda>x. g x \<guillemotright>= (\<lambda>y. return (x, y))) \<guillemotright>= (\<lambda>(x, y). h x y)"
haftmann@26170
   197
  by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits)
haftmann@26170
   198
haftmann@26170
   199
lemma raise_bind: "raise e \<guillemotright>= f = raise e"
haftmann@26170
   200
  by (simp add: raise_def bindM_def)
haftmann@26170
   201
haftmann@26170
   202
haftmann@26170
   203
lemmas monad_simp = return_bind bind_return bind_bind raise_bind
haftmann@26170
   204
haftmann@26170
   205
haftmann@26170
   206
subsection {* Generic combinators *}
haftmann@26170
   207
haftmann@26170
   208
definition
haftmann@26170
   209
  liftM :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b Heap"
haftmann@26170
   210
where
haftmann@26170
   211
  "liftM f = return o f"
haftmann@26170
   212
haftmann@26170
   213
definition
haftmann@26170
   214
  compM :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> ('b \<Rightarrow> 'c Heap) \<Rightarrow> 'a \<Rightarrow> 'c Heap" (infixl ">>==" 54)
haftmann@26170
   215
where
haftmann@26170
   216
  "(f >>== g) = (\<lambda>x. f x \<guillemotright>= g)"
haftmann@26170
   217
haftmann@26170
   218
notation
haftmann@26170
   219
  compM (infixl "\<guillemotright>==" 54)
haftmann@26170
   220
haftmann@26170
   221
lemma liftM_collapse: "liftM f x = return (f x)"
haftmann@26170
   222
  by (simp add: liftM_def)
haftmann@26170
   223
haftmann@26170
   224
lemma liftM_compM: "liftM f \<guillemotright>== g = g o f"
haftmann@26170
   225
  by (auto intro: Heap_eqI' simp add: expand_fun_eq liftM_def compM_def bindM_def)
haftmann@26170
   226
haftmann@26170
   227
lemma compM_return: "f \<guillemotright>== return = f"
haftmann@26170
   228
  by (simp add: compM_def monad_simp)
haftmann@26170
   229
haftmann@26170
   230
lemma compM_compM: "(f \<guillemotright>== g) \<guillemotright>== h = f \<guillemotright>== (g \<guillemotright>== h)"
haftmann@26170
   231
  by (simp add: compM_def monad_simp)
haftmann@26170
   232
haftmann@26170
   233
lemma liftM_bind:
haftmann@26170
   234
  "(\<lambda>x. liftM f x \<guillemotright>= liftM g) = liftM (\<lambda>x. g (f x))"
haftmann@26170
   235
  by (rule Heap_eqI') (simp add: monad_simp liftM_def bindM_def)
haftmann@26170
   236
haftmann@26170
   237
lemma liftM_comp:
haftmann@26170
   238
  "liftM f o g = liftM (f o g)"
haftmann@26170
   239
  by (rule Heap_eqI') (simp add: liftM_def)
haftmann@26170
   240
haftmann@26170
   241
lemmas monad_simp' = monad_simp liftM_compM compM_return
haftmann@26170
   242
  compM_compM liftM_bind liftM_comp
haftmann@26170
   243
haftmann@26170
   244
primrec 
haftmann@26170
   245
  mapM :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b list Heap"
haftmann@26170
   246
where
haftmann@26170
   247
  "mapM f [] = return []"
haftmann@26170
   248
  | "mapM f (x#xs) = do y \<leftarrow> f x;
haftmann@26170
   249
                        ys \<leftarrow> mapM f xs;
haftmann@26170
   250
                        return (y # ys)
haftmann@26170
   251
                     done"
haftmann@26170
   252
haftmann@26170
   253
primrec
haftmann@26170
   254
  foldM :: "('a \<Rightarrow> 'b \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b Heap"
haftmann@26170
   255
where
haftmann@26170
   256
  "foldM f [] s = return s"
haftmann@26170
   257
  | "foldM f (x#xs) s = f x s \<guillemotright>= foldM f xs"
haftmann@26170
   258
haftmann@28742
   259
definition
haftmann@28742
   260
  assert :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a Heap"
haftmann@28742
   261
where
haftmann@28742
   262
  "assert P x = (if P x then return x else raise (''assert''))"
haftmann@28742
   263
haftmann@28742
   264
lemma assert_cong [fundef_cong]:
haftmann@28742
   265
  assumes "P = P'"
haftmann@28742
   266
  assumes "\<And>x. P' x \<Longrightarrow> f x = f' x"
haftmann@28742
   267
  shows "(assert P x >>= f) = (assert P' x >>= f')"
haftmann@28742
   268
  using assms by (auto simp add: assert_def return_bind raise_bind)
haftmann@28742
   269
haftmann@26170
   270
hide (open) const heap execute
haftmann@26170
   271
haftmann@26182
   272
haftmann@26182
   273
subsection {* Code generator setup *}
haftmann@26182
   274
haftmann@26182
   275
subsubsection {* Logical intermediate layer *}
haftmann@26182
   276
haftmann@26182
   277
definition
haftmann@26182
   278
  Fail :: "message_string \<Rightarrow> exception"
haftmann@26182
   279
where
haftmann@28562
   280
  [code del]: "Fail s = Exn"
haftmann@26182
   281
haftmann@26182
   282
definition
haftmann@26182
   283
  raise_exc :: "exception \<Rightarrow> 'a Heap"
haftmann@26182
   284
where
haftmann@28562
   285
  [code del]: "raise_exc e = raise []"
haftmann@26182
   286
haftmann@28562
   287
lemma raise_raise_exc [code, code inline]:
haftmann@26182
   288
  "raise s = raise_exc (Fail (STR s))"
haftmann@26182
   289
  unfolding Fail_def raise_exc_def raise_def ..
haftmann@26182
   290
haftmann@26182
   291
hide (open) const Fail raise_exc
haftmann@26182
   292
haftmann@26182
   293
haftmann@27707
   294
subsubsection {* SML and OCaml *}
haftmann@26182
   295
haftmann@26752
   296
code_type Heap (SML "unit/ ->/ _")
haftmann@26182
   297
code_const Heap (SML "raise/ (Fail/ \"bare Heap\")")
haftmann@27826
   298
code_const "op \<guillemotright>=" (SML "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())")
haftmann@27707
   299
code_const return (SML "!(fn/ ()/ =>/ _)")
haftmann@26182
   300
code_const "Heap_Monad.Fail" (SML "Fail")
haftmann@27707
   301
code_const "Heap_Monad.raise_exc" (SML "!(fn/ ()/ =>/ raise/ _)")
haftmann@26182
   302
haftmann@26182
   303
code_type Heap (OCaml "_")
haftmann@26182
   304
code_const Heap (OCaml "failwith/ \"bare Heap\"")
haftmann@27826
   305
code_const "op \<guillemotright>=" (OCaml "!(fun/ f'_/ ()/ ->/ f'_/ (_/ ())/ ())")
haftmann@27707
   306
code_const return (OCaml "!(fun/ ()/ ->/ _)")
haftmann@26182
   307
code_const "Heap_Monad.Fail" (OCaml "Failure")
haftmann@27707
   308
code_const "Heap_Monad.raise_exc" (OCaml "!(fun/ ()/ ->/ raise/ _)")
haftmann@27707
   309
haftmann@28663
   310
setup {* let
haftmann@28663
   311
  open Code_Thingol;
haftmann@27707
   312
haftmann@28663
   313
  fun lookup naming = the o Code_Thingol.lookup_const naming;
haftmann@27707
   314
haftmann@28663
   315
  fun imp_monad_bind'' bind' return' unit' ts =
haftmann@28663
   316
    let
haftmann@28663
   317
      val dummy_name = "";
haftmann@28663
   318
      val dummy_type = ITyVar dummy_name;
haftmann@28663
   319
      val dummy_case_term = IVar dummy_name;
haftmann@28663
   320
      (*assumption: dummy values are not relevant for serialization*)
haftmann@28663
   321
      val unitt = IConst (unit', ([], []));
haftmann@28663
   322
      fun dest_abs ((v, ty) `|-> t, _) = ((v, ty), t)
haftmann@28663
   323
        | dest_abs (t, ty) =
haftmann@28663
   324
            let
haftmann@28663
   325
              val vs = Code_Thingol.fold_varnames cons t [];
haftmann@28663
   326
              val v = Name.variant vs "x";
haftmann@28663
   327
              val ty' = (hd o fst o Code_Thingol.unfold_fun) ty;
haftmann@28663
   328
            in ((v, ty'), t `$ IVar v) end;
haftmann@28663
   329
      fun force (t as IConst (c, _) `$ t') = if c = return'
haftmann@28663
   330
            then t' else t `$ unitt
haftmann@28663
   331
        | force t = t `$ unitt;
haftmann@28663
   332
      fun tr_bind' [(t1, _), (t2, ty2)] =
haftmann@28663
   333
        let
haftmann@28663
   334
          val ((v, ty), t) = dest_abs (t2, ty2);
haftmann@28663
   335
        in ICase (((force t1, ty), [(IVar v, tr_bind'' t)]), dummy_case_term) end
haftmann@28663
   336
      and tr_bind'' t = case Code_Thingol.unfold_app t
haftmann@28663
   337
           of (IConst (c, (_, ty1 :: ty2 :: _)), [x1, x2]) => if c = bind'
haftmann@28663
   338
                then tr_bind' [(x1, ty1), (x2, ty2)]
haftmann@28663
   339
                else force t
haftmann@28663
   340
            | _ => force t;
haftmann@28663
   341
    in (dummy_name, dummy_type) `|-> ICase (((IVar dummy_name, dummy_type),
haftmann@28663
   342
      [(unitt, tr_bind' ts)]), dummy_case_term) end
haftmann@28663
   343
  and imp_monad_bind' bind' return' unit' (const as (c, (_, tys))) ts = if c = bind' then case (ts, tys)
haftmann@28663
   344
     of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)]
haftmann@28663
   345
      | ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)] `$ t3
haftmann@28663
   346
      | (ts, _) => imp_monad_bind bind' return' unit' (eta_expand 2 (const, ts))
haftmann@28663
   347
    else IConst const `$$ map (imp_monad_bind bind' return' unit') ts
haftmann@28663
   348
  and imp_monad_bind bind' return' unit' (IConst const) = imp_monad_bind' bind' return' unit' const []
haftmann@28663
   349
    | imp_monad_bind bind' return' unit' (t as IVar _) = t
haftmann@28663
   350
    | imp_monad_bind bind' return' unit' (t as _ `$ _) = (case unfold_app t
haftmann@28663
   351
       of (IConst const, ts) => imp_monad_bind' bind' return' unit' const ts
haftmann@28663
   352
        | (t, ts) => imp_monad_bind bind' return' unit' t `$$ map (imp_monad_bind bind' return' unit') ts)
haftmann@28663
   353
    | imp_monad_bind bind' return' unit' (v_ty `|-> t) = v_ty `|-> imp_monad_bind bind' return' unit' t
haftmann@28663
   354
    | imp_monad_bind bind' return' unit' (ICase (((t, ty), pats), t0)) = ICase
haftmann@28663
   355
        (((imp_monad_bind bind' return' unit' t, ty), (map o pairself) (imp_monad_bind bind' return' unit') pats), imp_monad_bind bind' return' unit' t0);
haftmann@27707
   356
haftmann@28663
   357
   fun imp_program naming = (Graph.map_nodes o map_terms_stmt)
haftmann@28663
   358
     (imp_monad_bind (lookup naming @{const_name bindM})
haftmann@28663
   359
       (lookup naming @{const_name return})
haftmann@28663
   360
       (lookup naming @{const_name Unity}));
haftmann@27707
   361
haftmann@27707
   362
in
haftmann@27707
   363
haftmann@28663
   364
  Code_Target.extend_target ("SML_imp", ("SML", imp_program))
haftmann@28663
   365
  #> Code_Target.extend_target ("OCaml_imp", ("OCaml", imp_program))
haftmann@27707
   366
haftmann@27707
   367
end
haftmann@27707
   368
*}
haftmann@27707
   369
haftmann@26182
   370
haftmann@26182
   371
code_reserved OCaml Failure raise
haftmann@26182
   372
haftmann@26182
   373
haftmann@26182
   374
subsubsection {* Haskell *}
haftmann@26182
   375
haftmann@26182
   376
text {* Adaption layer *}
haftmann@26182
   377
haftmann@26182
   378
code_include Haskell "STMonad"
haftmann@26182
   379
{*import qualified Control.Monad;
haftmann@26182
   380
import qualified Control.Monad.ST;
haftmann@26182
   381
import qualified Data.STRef;
haftmann@26182
   382
import qualified Data.Array.ST;
haftmann@26182
   383
haftmann@27695
   384
type RealWorld = Control.Monad.ST.RealWorld;
haftmann@26182
   385
type ST s a = Control.Monad.ST.ST s a;
haftmann@26182
   386
type STRef s a = Data.STRef.STRef s a;
haftmann@27673
   387
type STArray s a = Data.Array.ST.STArray s Int a;
haftmann@26182
   388
haftmann@26182
   389
runST :: (forall s. ST s a) -> a;
haftmann@26182
   390
runST s = Control.Monad.ST.runST s;
haftmann@26182
   391
haftmann@26182
   392
newSTRef = Data.STRef.newSTRef;
haftmann@26182
   393
readSTRef = Data.STRef.readSTRef;
haftmann@26182
   394
writeSTRef = Data.STRef.writeSTRef;
haftmann@26182
   395
haftmann@27673
   396
newArray :: (Int, Int) -> a -> ST s (STArray s a);
haftmann@26182
   397
newArray = Data.Array.ST.newArray;
haftmann@26182
   398
haftmann@27673
   399
newListArray :: (Int, Int) -> [a] -> ST s (STArray s a);
haftmann@26182
   400
newListArray = Data.Array.ST.newListArray;
haftmann@26182
   401
haftmann@27673
   402
lengthArray :: STArray s a -> ST s Int;
haftmann@27673
   403
lengthArray a = Control.Monad.liftM snd (Data.Array.ST.getBounds a);
haftmann@26182
   404
haftmann@27673
   405
readArray :: STArray s a -> Int -> ST s a;
haftmann@26182
   406
readArray = Data.Array.ST.readArray;
haftmann@26182
   407
haftmann@27673
   408
writeArray :: STArray s a -> Int -> a -> ST s ();
haftmann@26182
   409
writeArray = Data.Array.ST.writeArray;*}
haftmann@26182
   410
haftmann@27695
   411
code_reserved Haskell RealWorld ST STRef Array
haftmann@26182
   412
  runST
haftmann@26182
   413
  newSTRef reasSTRef writeSTRef
haftmann@27673
   414
  newArray newListArray lengthArray readArray writeArray
haftmann@26182
   415
haftmann@26182
   416
text {* Monad *}
haftmann@26182
   417
haftmann@27695
   418
code_type Heap (Haskell "ST/ RealWorld/ _")
haftmann@27695
   419
code_const Heap (Haskell "error/ \"bare Heap\"")
haftmann@28145
   420
code_monad "op \<guillemotright>=" Haskell
haftmann@26182
   421
code_const return (Haskell "return")
haftmann@26182
   422
code_const "Heap_Monad.Fail" (Haskell "_")
haftmann@26182
   423
code_const "Heap_Monad.raise_exc" (Haskell "error")
haftmann@26182
   424
haftmann@26170
   425
end