src/HOL/Imperative_HOL/Heap_Monad.thy
author wenzelm
Thu, 11 Feb 2010 22:19:58 +0100
changeset 35116 1a0c129bb2e0
parent 34047 1a82e2e29d67
child 35533 6ef9525a5727
permissions -rw-r--r--
modernized translations;
formal markup of @{syntax_const} and @{const_syntax};
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
wenzelm@35116
   150
          Const (@{syntax_const "_bindM"}, dummyT) $ v $ f $ unfold_monad g'
haftmann@26170
   151
        else
wenzelm@35116
   152
          Const (@{syntax_const "_chainM"}, dummyT) $ f $ unfold_monad g'
haftmann@26170
   153
        end
haftmann@26170
   154
    | unfold_monad (Const (@{const_syntax chainM}, _) $ f $ g) =
wenzelm@35116
   155
        Const (@{syntax_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;
wenzelm@35116
   159
        in Const (@{syntax_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
wenzelm@35116
   167
    (Const (@{syntax_const "_do"}, dummyT) $
wenzelm@35116
   168
      unfold_monad (Const (@{const_syntax bindM}, dummyT) $ f $ g), ts);
wenzelm@35116
   169
  fun Let_monad_tr' (f :: (g as Abs (_, _, g')) :: ts) =
wenzelm@35116
   170
    if contains_bindM g' then list_comb
wenzelm@35116
   171
      (Const (@{syntax_const "_do"}, dummyT) $
wenzelm@35116
   172
        unfold_monad (Const (@{const_syntax Let}, dummyT) $ f $ g), ts)
haftmann@28145
   173
    else raise Match;
wenzelm@35116
   174
in
wenzelm@35116
   175
 [(@{const_syntax bindM}, bindM_monad_tr'),
wenzelm@35116
   176
  (@{const_syntax Let}, Let_monad_tr')]
wenzelm@35116
   177
end;
haftmann@26170
   178
*}
haftmann@26170
   179
haftmann@26170
   180
haftmann@26170
   181
subsection {* Monad properties *}
haftmann@26170
   182
haftmann@26170
   183
subsubsection {* Monad laws *}
haftmann@26170
   184
haftmann@26170
   185
lemma return_bind: "return x \<guillemotright>= f = f x"
haftmann@26170
   186
  by (simp add: bindM_def return_def)
haftmann@26170
   187
haftmann@26170
   188
lemma bind_return: "f \<guillemotright>= return = f"
haftmann@26170
   189
proof (rule Heap_eqI)
haftmann@26170
   190
  fix h
haftmann@26170
   191
  show "execute (f \<guillemotright>= return) h = execute f h"
haftmann@26170
   192
    by (auto simp add: bindM_def return_def split: sum.splits prod.splits)
haftmann@26170
   193
qed
haftmann@26170
   194
haftmann@26170
   195
lemma bind_bind: "(f \<guillemotright>= g) \<guillemotright>= h = f \<guillemotright>= (\<lambda>x. g x \<guillemotright>= h)"
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 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
   199
  by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits)
haftmann@26170
   200
haftmann@26170
   201
lemma raise_bind: "raise e \<guillemotright>= f = raise e"
haftmann@26170
   202
  by (simp add: raise_def bindM_def)
haftmann@26170
   203
haftmann@26170
   204
haftmann@26170
   205
lemmas monad_simp = return_bind bind_return bind_bind raise_bind
haftmann@26170
   206
haftmann@26170
   207
haftmann@26170
   208
subsection {* Generic combinators *}
haftmann@26170
   209
haftmann@26170
   210
definition
haftmann@26170
   211
  liftM :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b Heap"
haftmann@26170
   212
where
haftmann@26170
   213
  "liftM f = return o f"
haftmann@26170
   214
haftmann@26170
   215
definition
haftmann@26170
   216
  compM :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> ('b \<Rightarrow> 'c Heap) \<Rightarrow> 'a \<Rightarrow> 'c Heap" (infixl ">>==" 54)
haftmann@26170
   217
where
haftmann@26170
   218
  "(f >>== g) = (\<lambda>x. f x \<guillemotright>= g)"
haftmann@26170
   219
haftmann@26170
   220
notation
haftmann@26170
   221
  compM (infixl "\<guillemotright>==" 54)
haftmann@26170
   222
haftmann@26170
   223
lemma liftM_collapse: "liftM f x = return (f x)"
haftmann@26170
   224
  by (simp add: liftM_def)
haftmann@26170
   225
haftmann@26170
   226
lemma liftM_compM: "liftM f \<guillemotright>== g = g o f"
haftmann@26170
   227
  by (auto intro: Heap_eqI' simp add: expand_fun_eq liftM_def compM_def bindM_def)
haftmann@26170
   228
haftmann@26170
   229
lemma compM_return: "f \<guillemotright>== return = f"
haftmann@26170
   230
  by (simp add: compM_def monad_simp)
haftmann@26170
   231
haftmann@26170
   232
lemma compM_compM: "(f \<guillemotright>== g) \<guillemotright>== h = f \<guillemotright>== (g \<guillemotright>== h)"
haftmann@26170
   233
  by (simp add: compM_def monad_simp)
haftmann@26170
   234
haftmann@26170
   235
lemma liftM_bind:
haftmann@26170
   236
  "(\<lambda>x. liftM f x \<guillemotright>= liftM g) = liftM (\<lambda>x. g (f x))"
haftmann@26170
   237
  by (rule Heap_eqI') (simp add: monad_simp liftM_def bindM_def)
haftmann@26170
   238
haftmann@26170
   239
lemma liftM_comp:
haftmann@26170
   240
  "liftM f o g = liftM (f o g)"
haftmann@26170
   241
  by (rule Heap_eqI') (simp add: liftM_def)
haftmann@26170
   242
haftmann@26170
   243
lemmas monad_simp' = monad_simp liftM_compM compM_return
haftmann@26170
   244
  compM_compM liftM_bind liftM_comp
haftmann@26170
   245
haftmann@26170
   246
primrec 
haftmann@26170
   247
  mapM :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b list Heap"
haftmann@26170
   248
where
haftmann@26170
   249
  "mapM f [] = return []"
haftmann@26170
   250
  | "mapM f (x#xs) = do y \<leftarrow> f x;
haftmann@26170
   251
                        ys \<leftarrow> mapM f xs;
haftmann@26170
   252
                        return (y # ys)
haftmann@26170
   253
                     done"
haftmann@26170
   254
haftmann@26170
   255
primrec
haftmann@26170
   256
  foldM :: "('a \<Rightarrow> 'b \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b Heap"
haftmann@26170
   257
where
haftmann@26170
   258
  "foldM f [] s = return s"
haftmann@26170
   259
  | "foldM f (x#xs) s = f x s \<guillemotright>= foldM f xs"
haftmann@26170
   260
haftmann@28742
   261
definition
haftmann@28742
   262
  assert :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a Heap"
haftmann@28742
   263
where
haftmann@28742
   264
  "assert P x = (if P x then return x else raise (''assert''))"
haftmann@28742
   265
haftmann@28742
   266
lemma assert_cong [fundef_cong]:
haftmann@28742
   267
  assumes "P = P'"
haftmann@28742
   268
  assumes "\<And>x. P' x \<Longrightarrow> f x = f' x"
haftmann@28742
   269
  shows "(assert P x >>= f) = (assert P' x >>= f')"
haftmann@28742
   270
  using assms by (auto simp add: assert_def return_bind raise_bind)
haftmann@28742
   271
bulwahn@34047
   272
subsubsection {* A monadic combinator for simple recursive functions *}
bulwahn@34047
   273
 
bulwahn@34047
   274
function (default "\<lambda>(f,g,x,h). (Inr Exn, undefined)") 
bulwahn@34047
   275
  mrec 
bulwahn@34047
   276
where
bulwahn@34047
   277
  "mrec f g x h = 
bulwahn@34047
   278
   (case Heap_Monad.execute (f x) h of
bulwahn@34047
   279
     (Inl (Inl r), h') \<Rightarrow> (Inl r, h')
bulwahn@34047
   280
   | (Inl (Inr s), h') \<Rightarrow> 
bulwahn@34047
   281
          (case mrec f g s h' of
bulwahn@34047
   282
             (Inl z, h'') \<Rightarrow> Heap_Monad.execute (g x s z) h''
bulwahn@34047
   283
           | (Inr e, h'') \<Rightarrow> (Inr e, h''))
bulwahn@34047
   284
   | (Inr e, h') \<Rightarrow> (Inr e, h')
bulwahn@34047
   285
   )"
bulwahn@34047
   286
by auto
bulwahn@34047
   287
bulwahn@34047
   288
lemma graph_implies_dom:
bulwahn@34047
   289
	"mrec_graph x y \<Longrightarrow> mrec_dom x"
bulwahn@34047
   290
apply (induct rule:mrec_graph.induct) 
bulwahn@34047
   291
apply (rule accpI)
bulwahn@34047
   292
apply (erule mrec_rel.cases)
bulwahn@34047
   293
by simp
bulwahn@34047
   294
bulwahn@34047
   295
lemma f_default: "\<not> mrec_dom (f, g, x, h) \<Longrightarrow> mrec f g x h = (Inr Exn, undefined)"
bulwahn@34047
   296
	unfolding mrec_def 
bulwahn@34047
   297
  by (rule fundef_default_value[OF mrec_sumC_def graph_implies_dom, of _ _ "(f, g, x, h)", simplified])
bulwahn@34047
   298
bulwahn@34047
   299
lemma f_di_reverse: 
bulwahn@34047
   300
  assumes "\<not> mrec_dom (f, g, x, h)"
bulwahn@34047
   301
  shows "
bulwahn@34047
   302
   (case Heap_Monad.execute (f x) h of
bulwahn@34047
   303
     (Inl (Inl r), h') \<Rightarrow> mrecalse
bulwahn@34047
   304
   | (Inl (Inr s), h') \<Rightarrow> \<not> mrec_dom (f, g, s, h')
bulwahn@34047
   305
   | (Inr e, h') \<Rightarrow> mrecalse
bulwahn@34047
   306
   )" 
bulwahn@34047
   307
using assms
bulwahn@34047
   308
by (auto split:prod.splits sum.splits)
bulwahn@34047
   309
 (erule notE, rule accpI, elim mrec_rel.cases, simp)+
bulwahn@34047
   310
bulwahn@34047
   311
bulwahn@34047
   312
lemma mrec_rule:
bulwahn@34047
   313
  "mrec f g x h = 
bulwahn@34047
   314
   (case Heap_Monad.execute (f x) h of
bulwahn@34047
   315
     (Inl (Inl r), h') \<Rightarrow> (Inl r, h')
bulwahn@34047
   316
   | (Inl (Inr s), h') \<Rightarrow> 
bulwahn@34047
   317
          (case mrec f g s h' of
bulwahn@34047
   318
             (Inl z, h'') \<Rightarrow> Heap_Monad.execute (g x s z) h''
bulwahn@34047
   319
           | (Inr e, h'') \<Rightarrow> (Inr e, h''))
bulwahn@34047
   320
   | (Inr e, h') \<Rightarrow> (Inr e, h')
bulwahn@34047
   321
   )"
bulwahn@34047
   322
apply (cases "mrec_dom (f,g,x,h)", simp)
bulwahn@34047
   323
apply (frule f_default)
bulwahn@34047
   324
apply (frule f_di_reverse, simp)
bulwahn@34047
   325
by (auto split: sum.split prod.split simp: f_default)
bulwahn@34047
   326
bulwahn@34047
   327
bulwahn@34047
   328
definition
bulwahn@34047
   329
  "MREC f g x = Heap (mrec f g x)"
bulwahn@34047
   330
bulwahn@34047
   331
lemma MREC_rule:
bulwahn@34047
   332
  "MREC f g x = 
bulwahn@34047
   333
  (do y \<leftarrow> f x;
bulwahn@34047
   334
                (case y of 
bulwahn@34047
   335
                Inl r \<Rightarrow> return r
bulwahn@34047
   336
              | Inr s \<Rightarrow> 
bulwahn@34047
   337
                do z \<leftarrow> MREC f g s ;
bulwahn@34047
   338
                   g x s z
bulwahn@34047
   339
                done) done)"
bulwahn@34047
   340
  unfolding MREC_def
bulwahn@34047
   341
  unfolding bindM_def return_def
bulwahn@34047
   342
  apply simp
bulwahn@34047
   343
  apply (rule ext)
bulwahn@34047
   344
  apply (unfold mrec_rule[of f g x])
bulwahn@34047
   345
  by (auto split:prod.splits sum.splits)
bulwahn@34047
   346
haftmann@26170
   347
hide (open) const heap execute
haftmann@26170
   348
haftmann@26182
   349
haftmann@26182
   350
subsection {* Code generator setup *}
haftmann@26182
   351
haftmann@26182
   352
subsubsection {* Logical intermediate layer *}
haftmann@26182
   353
haftmann@26182
   354
definition
haftmann@31205
   355
  Fail :: "String.literal \<Rightarrow> exception"
haftmann@26182
   356
where
haftmann@28562
   357
  [code del]: "Fail s = Exn"
haftmann@26182
   358
haftmann@26182
   359
definition
haftmann@26182
   360
  raise_exc :: "exception \<Rightarrow> 'a Heap"
haftmann@26182
   361
where
haftmann@28562
   362
  [code del]: "raise_exc e = raise []"
haftmann@26182
   363
haftmann@32061
   364
lemma raise_raise_exc [code, code_unfold]:
haftmann@26182
   365
  "raise s = raise_exc (Fail (STR s))"
haftmann@26182
   366
  unfolding Fail_def raise_exc_def raise_def ..
haftmann@26182
   367
haftmann@26182
   368
hide (open) const Fail raise_exc
haftmann@26182
   369
haftmann@26182
   370
haftmann@27707
   371
subsubsection {* SML and OCaml *}
haftmann@26182
   372
haftmann@26752
   373
code_type Heap (SML "unit/ ->/ _")
haftmann@26182
   374
code_const Heap (SML "raise/ (Fail/ \"bare Heap\")")
haftmann@27826
   375
code_const "op \<guillemotright>=" (SML "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())")
haftmann@27707
   376
code_const return (SML "!(fn/ ()/ =>/ _)")
haftmann@26182
   377
code_const "Heap_Monad.Fail" (SML "Fail")
haftmann@27707
   378
code_const "Heap_Monad.raise_exc" (SML "!(fn/ ()/ =>/ raise/ _)")
haftmann@26182
   379
haftmann@26182
   380
code_type Heap (OCaml "_")
haftmann@26182
   381
code_const Heap (OCaml "failwith/ \"bare Heap\"")
haftmann@27826
   382
code_const "op \<guillemotright>=" (OCaml "!(fun/ f'_/ ()/ ->/ f'_/ (_/ ())/ ())")
haftmann@27707
   383
code_const return (OCaml "!(fun/ ()/ ->/ _)")
haftmann@26182
   384
code_const "Heap_Monad.Fail" (OCaml "Failure")
haftmann@27707
   385
code_const "Heap_Monad.raise_exc" (OCaml "!(fun/ ()/ ->/ raise/ _)")
haftmann@27707
   386
haftmann@31870
   387
setup {*
haftmann@27707
   388
haftmann@31870
   389
let
haftmann@27707
   390
haftmann@31870
   391
open Code_Thingol;
haftmann@27707
   392
haftmann@31870
   393
fun imp_program naming =
haftmann@31870
   394
haftmann@31870
   395
  let
haftmann@31870
   396
    fun is_const c = case lookup_const naming c
haftmann@31870
   397
     of SOME c' => (fn c'' => c' = c'')
haftmann@31870
   398
      | NONE => K false;
haftmann@31870
   399
    val is_bindM = is_const @{const_name bindM};
haftmann@31870
   400
    val is_return = is_const @{const_name return};
haftmann@31893
   401
    val dummy_name = "";
haftmann@31870
   402
    val dummy_type = ITyVar dummy_name;
haftmann@31893
   403
    val dummy_case_term = IVar NONE;
haftmann@31870
   404
    (*assumption: dummy values are not relevant for serialization*)
haftmann@31870
   405
    val unitt = case lookup_const naming @{const_name Unity}
haftmann@31870
   406
     of SOME unit' => IConst (unit', (([], []), []))
haftmann@31870
   407
      | NONE => error ("Must include " ^ @{const_name Unity} ^ " in generated constants.");
haftmann@31870
   408
    fun dest_abs ((v, ty) `|=> t, _) = ((v, ty), t)
haftmann@31870
   409
      | dest_abs (t, ty) =
haftmann@31870
   410
          let
haftmann@31870
   411
            val vs = fold_varnames cons t [];
haftmann@31870
   412
            val v = Name.variant vs "x";
haftmann@31870
   413
            val ty' = (hd o fst o unfold_fun) ty;
haftmann@31893
   414
          in ((SOME v, ty'), t `$ IVar (SOME v)) end;
haftmann@31870
   415
    fun force (t as IConst (c, _) `$ t') = if is_return c
haftmann@31870
   416
          then t' else t `$ unitt
haftmann@31870
   417
      | force t = t `$ unitt;
haftmann@31870
   418
    fun tr_bind' [(t1, _), (t2, ty2)] =
haftmann@31870
   419
      let
haftmann@31870
   420
        val ((v, ty), t) = dest_abs (t2, ty2);
haftmann@31870
   421
      in ICase (((force t1, ty), [(IVar v, tr_bind'' t)]), dummy_case_term) end
haftmann@31870
   422
    and tr_bind'' t = case unfold_app t
haftmann@31870
   423
         of (IConst (c, (_, ty1 :: ty2 :: _)), [x1, x2]) => if is_bindM c
haftmann@31870
   424
              then tr_bind' [(x1, ty1), (x2, ty2)]
haftmann@31870
   425
              else force t
haftmann@31870
   426
          | _ => force t;
haftmann@31893
   427
    fun imp_monad_bind'' ts = (SOME dummy_name, dummy_type) `|=> ICase (((IVar (SOME dummy_name), dummy_type),
haftmann@31870
   428
      [(unitt, tr_bind' ts)]), dummy_case_term)
haftmann@31870
   429
    and imp_monad_bind' (const as (c, (_, tys))) ts = if is_bindM c then case (ts, tys)
haftmann@31870
   430
       of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' [(t1, ty1), (t2, ty2)]
haftmann@31870
   431
        | ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' [(t1, ty1), (t2, ty2)] `$ t3
haftmann@31870
   432
        | (ts, _) => imp_monad_bind (eta_expand 2 (const, ts))
haftmann@31870
   433
      else IConst const `$$ map imp_monad_bind ts
haftmann@31870
   434
    and imp_monad_bind (IConst const) = imp_monad_bind' const []
haftmann@31870
   435
      | imp_monad_bind (t as IVar _) = t
haftmann@31870
   436
      | imp_monad_bind (t as _ `$ _) = (case unfold_app t
haftmann@31870
   437
         of (IConst const, ts) => imp_monad_bind' const ts
haftmann@31870
   438
          | (t, ts) => imp_monad_bind t `$$ map imp_monad_bind ts)
haftmann@31870
   439
      | imp_monad_bind (v_ty `|=> t) = v_ty `|=> imp_monad_bind t
haftmann@31870
   440
      | imp_monad_bind (ICase (((t, ty), pats), t0)) = ICase
haftmann@31870
   441
          (((imp_monad_bind t, ty),
haftmann@31870
   442
            (map o pairself) imp_monad_bind pats),
haftmann@31870
   443
              imp_monad_bind t0);
haftmann@31870
   444
haftmann@31870
   445
  in (Graph.map_nodes o map_terms_stmt) imp_monad_bind end;
haftmann@27707
   446
haftmann@27707
   447
in
haftmann@27707
   448
haftmann@31870
   449
Code_Target.extend_target ("SML_imp", ("SML", imp_program))
haftmann@31870
   450
#> Code_Target.extend_target ("OCaml_imp", ("OCaml", imp_program))
haftmann@27707
   451
haftmann@27707
   452
end
haftmann@31870
   453
haftmann@27707
   454
*}
haftmann@27707
   455
haftmann@26182
   456
code_reserved OCaml Failure raise
haftmann@26182
   457
haftmann@26182
   458
haftmann@26182
   459
subsubsection {* Haskell *}
haftmann@26182
   460
haftmann@26182
   461
text {* Adaption layer *}
haftmann@26182
   462
haftmann@29730
   463
code_include Haskell "Heap"
haftmann@26182
   464
{*import qualified Control.Monad;
haftmann@26182
   465
import qualified Control.Monad.ST;
haftmann@26182
   466
import qualified Data.STRef;
haftmann@26182
   467
import qualified Data.Array.ST;
haftmann@26182
   468
haftmann@27695
   469
type RealWorld = Control.Monad.ST.RealWorld;
haftmann@26182
   470
type ST s a = Control.Monad.ST.ST s a;
haftmann@26182
   471
type STRef s a = Data.STRef.STRef s a;
haftmann@27673
   472
type STArray s a = Data.Array.ST.STArray s Int a;
haftmann@26182
   473
haftmann@26182
   474
newSTRef = Data.STRef.newSTRef;
haftmann@26182
   475
readSTRef = Data.STRef.readSTRef;
haftmann@26182
   476
writeSTRef = Data.STRef.writeSTRef;
haftmann@26182
   477
haftmann@27673
   478
newArray :: (Int, Int) -> a -> ST s (STArray s a);
haftmann@26182
   479
newArray = Data.Array.ST.newArray;
haftmann@26182
   480
haftmann@27673
   481
newListArray :: (Int, Int) -> [a] -> ST s (STArray s a);
haftmann@26182
   482
newListArray = Data.Array.ST.newListArray;
haftmann@26182
   483
haftmann@27673
   484
lengthArray :: STArray s a -> ST s Int;
haftmann@27673
   485
lengthArray a = Control.Monad.liftM snd (Data.Array.ST.getBounds a);
haftmann@26182
   486
haftmann@27673
   487
readArray :: STArray s a -> Int -> ST s a;
haftmann@26182
   488
readArray = Data.Array.ST.readArray;
haftmann@26182
   489
haftmann@27673
   490
writeArray :: STArray s a -> Int -> a -> ST s ();
haftmann@26182
   491
writeArray = Data.Array.ST.writeArray;*}
haftmann@26182
   492
haftmann@29730
   493
code_reserved Haskell Heap
haftmann@26182
   494
haftmann@26182
   495
text {* Monad *}
haftmann@26182
   496
haftmann@29730
   497
code_type Heap (Haskell "Heap.ST/ Heap.RealWorld/ _")
haftmann@27695
   498
code_const Heap (Haskell "error/ \"bare Heap\"")
haftmann@28145
   499
code_monad "op \<guillemotright>=" Haskell
haftmann@26182
   500
code_const return (Haskell "return")
haftmann@26182
   501
code_const "Heap_Monad.Fail" (Haskell "_")
haftmann@26182
   502
code_const "Heap_Monad.raise_exc" (Haskell "error")
haftmann@26182
   503
haftmann@26170
   504
end