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