src/HOL/Imperative_HOL/Heap_Monad.thy
author haftmann
Thu, 26 Aug 2010 12:19:49 +0200
changeset 39006 f9837065b5e8
parent 39004 f9cd27cbe8a4
child 39194 e55deaa22fff
permissions -rw-r--r--
prevent line breaks after Scala symbolic operators
     1 (*  Title:      HOL/Imperative_HOL/Heap_Monad.thy
     2     Author:     John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
     3 *)
     4 
     5 header {* A monad with a polymorphic heap and primitive reasoning infrastructure *}
     6 
     7 theory Heap_Monad
     8 imports Heap Monad_Syntax Code_Natural
     9 begin
    10 
    11 subsection {* The monad *}
    12 
    13 subsubsection {* Monad construction *}
    14 
    15 text {* Monadic heap actions either produce values
    16   and transform the heap, or fail *}
    17 datatype 'a Heap = Heap "heap \<Rightarrow> ('a \<times> heap) option"
    18 
    19 primrec execute :: "'a Heap \<Rightarrow> heap \<Rightarrow> ('a \<times> heap) option" where
    20   [code del]: "execute (Heap f) = f"
    21 
    22 lemma Heap_cases [case_names succeed fail]:
    23   fixes f and h
    24   assumes succeed: "\<And>x h'. execute f h = Some (x, h') \<Longrightarrow> P"
    25   assumes fail: "execute f h = None \<Longrightarrow> P"
    26   shows P
    27   using assms by (cases "execute f h") auto
    28 
    29 lemma Heap_execute [simp]:
    30   "Heap (execute f) = f" by (cases f) simp_all
    31 
    32 lemma Heap_eqI:
    33   "(\<And>h. execute f h = execute g h) \<Longrightarrow> f = g"
    34     by (cases f, cases g) (auto simp: expand_fun_eq)
    35 
    36 ML {* structure Execute_Simps = Named_Thms(
    37   val name = "execute_simps"
    38   val description = "simplification rules for execute"
    39 ) *}
    40 
    41 setup Execute_Simps.setup
    42 
    43 lemma execute_Let [execute_simps]:
    44   "execute (let x = t in f x) = (let x = t in execute (f x))"
    45   by (simp add: Let_def)
    46 
    47 
    48 subsubsection {* Specialised lifters *}
    49 
    50 definition tap :: "(heap \<Rightarrow> 'a) \<Rightarrow> 'a Heap" where
    51   [code del]: "tap f = Heap (\<lambda>h. Some (f h, h))"
    52 
    53 lemma execute_tap [execute_simps]:
    54   "execute (tap f) h = Some (f h, h)"
    55   by (simp add: tap_def)
    56 
    57 definition heap :: "(heap \<Rightarrow> 'a \<times> heap) \<Rightarrow> 'a Heap" where
    58   [code del]: "heap f = Heap (Some \<circ> f)"
    59 
    60 lemma execute_heap [execute_simps]:
    61   "execute (heap f) = Some \<circ> f"
    62   by (simp add: heap_def)
    63 
    64 definition guard :: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> 'a \<times> heap) \<Rightarrow> 'a Heap" where
    65   [code del]: "guard P f = Heap (\<lambda>h. if P h then Some (f h) else None)"
    66 
    67 lemma execute_guard [execute_simps]:
    68   "\<not> P h \<Longrightarrow> execute (guard P f) h = None"
    69   "P h \<Longrightarrow> execute (guard P f) h = Some (f h)"
    70   by (simp_all add: guard_def)
    71 
    72 
    73 subsubsection {* Predicate classifying successful computations *}
    74 
    75 definition success :: "'a Heap \<Rightarrow> heap \<Rightarrow> bool" where
    76   "success f h \<longleftrightarrow> execute f h \<noteq> None"
    77 
    78 lemma successI:
    79   "execute f h \<noteq> None \<Longrightarrow> success f h"
    80   by (simp add: success_def)
    81 
    82 lemma successE:
    83   assumes "success f h"
    84   obtains r h' where "r = fst (the (execute c h))"
    85     and "h' = snd (the (execute c h))"
    86     and "execute f h \<noteq> None"
    87   using assms by (simp add: success_def)
    88 
    89 ML {* structure Success_Intros = Named_Thms(
    90   val name = "success_intros"
    91   val description = "introduction rules for success"
    92 ) *}
    93 
    94 setup Success_Intros.setup
    95 
    96 lemma success_tapI [success_intros]:
    97   "success (tap f) h"
    98   by (rule successI) (simp add: execute_simps)
    99 
   100 lemma success_heapI [success_intros]:
   101   "success (heap f) h"
   102   by (rule successI) (simp add: execute_simps)
   103 
   104 lemma success_guardI [success_intros]:
   105   "P h \<Longrightarrow> success (guard P f) h"
   106   by (rule successI) (simp add: execute_guard)
   107 
   108 lemma success_LetI [success_intros]:
   109   "x = t \<Longrightarrow> success (f x) h \<Longrightarrow> success (let x = t in f x) h"
   110   by (simp add: Let_def)
   111 
   112 lemma success_ifI:
   113   "(c \<Longrightarrow> success t h) \<Longrightarrow> (\<not> c \<Longrightarrow> success e h) \<Longrightarrow>
   114     success (if c then t else e) h"
   115   by (simp add: success_def)
   116 
   117 
   118 subsubsection {* Predicate for a simple relational calculus *}
   119 
   120 text {*
   121   The @{text crel} predicate states that when a computation @{text c}
   122   runs with the heap @{text h} will result in return value @{text r}
   123   and a heap @{text "h'"}, i.e.~no exception occurs.
   124 *}  
   125 
   126 definition crel :: "'a Heap \<Rightarrow> heap \<Rightarrow> heap \<Rightarrow> 'a \<Rightarrow> bool" where
   127   crel_def: "crel c h h' r \<longleftrightarrow> execute c h = Some (r, h')"
   128 
   129 lemma crelI:
   130   "execute c h = Some (r, h') \<Longrightarrow> crel c h h' r"
   131   by (simp add: crel_def)
   132 
   133 lemma crelE:
   134   assumes "crel c h h' r"
   135   obtains "r = fst (the (execute c h))"
   136     and "h' = snd (the (execute c h))"
   137     and "success c h"
   138 proof (rule that)
   139   from assms have *: "execute c h = Some (r, h')" by (simp add: crel_def)
   140   then show "success c h" by (simp add: success_def)
   141   from * have "fst (the (execute c h)) = r" and "snd (the (execute c h)) = h'"
   142     by simp_all
   143   then show "r = fst (the (execute c h))"
   144     and "h' = snd (the (execute c h))" by simp_all
   145 qed
   146 
   147 lemma crel_success:
   148   "crel c h h' r \<Longrightarrow> success c h"
   149   by (simp add: crel_def success_def)
   150 
   151 lemma success_crelE:
   152   assumes "success c h"
   153   obtains r h' where "crel c h h' r"
   154   using assms by (auto simp add: crel_def success_def)
   155 
   156 lemma crel_deterministic:
   157   assumes "crel f h h' a"
   158     and "crel f h h'' b"
   159   shows "a = b" and "h' = h''"
   160   using assms unfolding crel_def by auto
   161 
   162 ML {* structure Crel_Intros = Named_Thms(
   163   val name = "crel_intros"
   164   val description = "introduction rules for crel"
   165 ) *}
   166 
   167 ML {* structure Crel_Elims = Named_Thms(
   168   val name = "crel_elims"
   169   val description = "elimination rules for crel"
   170 ) *}
   171 
   172 setup "Crel_Intros.setup #> Crel_Elims.setup"
   173 
   174 lemma crel_LetI [crel_intros]:
   175   assumes "x = t" "crel (f x) h h' r"
   176   shows "crel (let x = t in f x) h h' r"
   177   using assms by simp
   178 
   179 lemma crel_LetE [crel_elims]:
   180   assumes "crel (let x = t in f x) h h' r"
   181   obtains "crel (f t) h h' r"
   182   using assms by simp
   183 
   184 lemma crel_ifI:
   185   assumes "c \<Longrightarrow> crel t h h' r"
   186     and "\<not> c \<Longrightarrow> crel e h h' r"
   187   shows "crel (if c then t else e) h h' r"
   188   by (cases c) (simp_all add: assms)
   189 
   190 lemma crel_ifE:
   191   assumes "crel (if c then t else e) h h' r"
   192   obtains "c" "crel t h h' r"
   193     | "\<not> c" "crel e h h' r"
   194   using assms by (cases c) simp_all
   195 
   196 lemma crel_tapI [crel_intros]:
   197   assumes "h' = h" "r = f h"
   198   shows "crel (tap f) h h' r"
   199   by (rule crelI) (simp add: assms execute_simps)
   200 
   201 lemma crel_tapE [crel_elims]:
   202   assumes "crel (tap f) h h' r"
   203   obtains "h' = h" and "r = f h"
   204   using assms by (rule crelE) (auto simp add: execute_simps)
   205 
   206 lemma crel_heapI [crel_intros]:
   207   assumes "h' = snd (f h)" "r = fst (f h)"
   208   shows "crel (heap f) h h' r"
   209   by (rule crelI) (simp add: assms execute_simps)
   210 
   211 lemma crel_heapE [crel_elims]:
   212   assumes "crel (heap f) h h' r"
   213   obtains "h' = snd (f h)" and "r = fst (f h)"
   214   using assms by (rule crelE) (simp add: execute_simps)
   215 
   216 lemma crel_guardI [crel_intros]:
   217   assumes "P h" "h' = snd (f h)" "r = fst (f h)"
   218   shows "crel (guard P f) h h' r"
   219   by (rule crelI) (simp add: assms execute_simps)
   220 
   221 lemma crel_guardE [crel_elims]:
   222   assumes "crel (guard P f) h h' r"
   223   obtains "h' = snd (f h)" "r = fst (f h)" "P h"
   224   using assms by (rule crelE)
   225     (auto simp add: execute_simps elim!: successE, cases "P h", auto simp add: execute_simps)
   226 
   227 
   228 subsubsection {* Monad combinators *}
   229 
   230 definition return :: "'a \<Rightarrow> 'a Heap" where
   231   [code del]: "return x = heap (Pair x)"
   232 
   233 lemma execute_return [execute_simps]:
   234   "execute (return x) = Some \<circ> Pair x"
   235   by (simp add: return_def execute_simps)
   236 
   237 lemma success_returnI [success_intros]:
   238   "success (return x) h"
   239   by (rule successI) (simp add: execute_simps)
   240 
   241 lemma crel_returnI [crel_intros]:
   242   "h = h' \<Longrightarrow> crel (return x) h h' x"
   243   by (rule crelI) (simp add: execute_simps)
   244 
   245 lemma crel_returnE [crel_elims]:
   246   assumes "crel (return x) h h' r"
   247   obtains "r = x" "h' = h"
   248   using assms by (rule crelE) (simp add: execute_simps)
   249 
   250 definition raise :: "string \<Rightarrow> 'a Heap" where -- {* the string is just decoration *}
   251   [code del]: "raise s = Heap (\<lambda>_. None)"
   252 
   253 lemma execute_raise [execute_simps]:
   254   "execute (raise s) = (\<lambda>_. None)"
   255   by (simp add: raise_def)
   256 
   257 lemma crel_raiseE [crel_elims]:
   258   assumes "crel (raise x) h h' r"
   259   obtains "False"
   260   using assms by (rule crelE) (simp add: success_def execute_simps)
   261 
   262 definition bind :: "'a Heap \<Rightarrow> ('a \<Rightarrow> 'b Heap) \<Rightarrow> 'b Heap" where
   263   [code del]: "bind f g = Heap (\<lambda>h. case execute f h of
   264                   Some (x, h') \<Rightarrow> execute (g x) h'
   265                 | None \<Rightarrow> None)"
   266 
   267 setup {*
   268   Adhoc_Overloading.add_variant 
   269     @{const_name Monad_Syntax.bind} @{const_name Heap_Monad.bind}
   270 *}
   271 
   272 lemma execute_bind [execute_simps]:
   273   "execute f h = Some (x, h') \<Longrightarrow> execute (f \<guillemotright>= g) h = execute (g x) h'"
   274   "execute f h = None \<Longrightarrow> execute (f \<guillemotright>= g) h = None"
   275   by (simp_all add: bind_def)
   276 
   277 lemma execute_bind_case:
   278   "execute (f \<guillemotright>= g) h = (case (execute f h) of
   279     Some (x, h') \<Rightarrow> execute (g x) h' | None \<Rightarrow> None)"
   280   by (simp add: bind_def)
   281 
   282 lemma execute_bind_success:
   283   "success f h \<Longrightarrow> execute (f \<guillemotright>= g) h = execute (g (fst (the (execute f h)))) (snd (the (execute f h)))"
   284   by (cases f h rule: Heap_cases) (auto elim!: successE simp add: bind_def)
   285 
   286 lemma success_bind_executeI:
   287   "execute f h = Some (x, h') \<Longrightarrow> success (g x) h' \<Longrightarrow> success (f \<guillemotright>= g) h"
   288   by (auto intro!: successI elim!: successE simp add: bind_def)
   289 
   290 lemma success_bind_crelI [success_intros]:
   291   "crel f h h' x \<Longrightarrow> success (g x) h' \<Longrightarrow> success (f \<guillemotright>= g) h"
   292   by (auto simp add: crel_def success_def bind_def)
   293 
   294 lemma crel_bindI [crel_intros]:
   295   assumes "crel f h h' r" "crel (g r) h' h'' r'"
   296   shows "crel (f \<guillemotright>= g) h h'' r'"
   297   using assms
   298   apply (auto intro!: crelI elim!: crelE successE)
   299   apply (subst execute_bind, simp_all)
   300   done
   301 
   302 lemma crel_bindE [crel_elims]:
   303   assumes "crel (f \<guillemotright>= g) h h'' r'"
   304   obtains h' r where "crel f h h' r" "crel (g r) h' h'' r'"
   305   using assms by (auto simp add: crel_def bind_def split: option.split_asm)
   306 
   307 lemma execute_bind_eq_SomeI:
   308   assumes "execute f h = Some (x, h')"
   309     and "execute (g x) h' = Some (y, h'')"
   310   shows "execute (f \<guillemotright>= g) h = Some (y, h'')"
   311   using assms by (simp add: bind_def)
   312 
   313 lemma return_bind [simp]: "return x \<guillemotright>= f = f x"
   314   by (rule Heap_eqI) (simp add: execute_bind execute_simps)
   315 
   316 lemma bind_return [simp]: "f \<guillemotright>= return = f"
   317   by (rule Heap_eqI) (simp add: bind_def execute_simps split: option.splits)
   318 
   319 lemma bind_bind [simp]: "(f \<guillemotright>= g) \<guillemotright>= k = (f :: 'a Heap) \<guillemotright>= (\<lambda>x. g x \<guillemotright>= k)"
   320   by (rule Heap_eqI) (simp add: bind_def execute_simps split: option.splits)
   321 
   322 lemma raise_bind [simp]: "raise e \<guillemotright>= f = raise e"
   323   by (rule Heap_eqI) (simp add: execute_simps)
   324 
   325 
   326 subsection {* Generic combinators *}
   327 
   328 subsubsection {* Assertions *}
   329 
   330 definition assert :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a Heap" where
   331   "assert P x = (if P x then return x else raise ''assert'')"
   332 
   333 lemma execute_assert [execute_simps]:
   334   "P x \<Longrightarrow> execute (assert P x) h = Some (x, h)"
   335   "\<not> P x \<Longrightarrow> execute (assert P x) h = None"
   336   by (simp_all add: assert_def execute_simps)
   337 
   338 lemma success_assertI [success_intros]:
   339   "P x \<Longrightarrow> success (assert P x) h"
   340   by (rule successI) (simp add: execute_assert)
   341 
   342 lemma crel_assertI [crel_intros]:
   343   "P x \<Longrightarrow> h' = h \<Longrightarrow> r = x \<Longrightarrow> crel (assert P x) h h' r"
   344   by (rule crelI) (simp add: execute_assert)
   345  
   346 lemma crel_assertE [crel_elims]:
   347   assumes "crel (assert P x) h h' r"
   348   obtains "P x" "r = x" "h' = h"
   349   using assms by (rule crelE) (cases "P x", simp_all add: execute_assert success_def)
   350 
   351 lemma assert_cong [fundef_cong]:
   352   assumes "P = P'"
   353   assumes "\<And>x. P' x \<Longrightarrow> f x = f' x"
   354   shows "(assert P x >>= f) = (assert P' x >>= f')"
   355   by (rule Heap_eqI) (insert assms, simp add: assert_def)
   356 
   357 
   358 subsubsection {* Plain lifting *}
   359 
   360 definition lift :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b Heap" where
   361   "lift f = return o f"
   362 
   363 lemma lift_collapse [simp]:
   364   "lift f x = return (f x)"
   365   by (simp add: lift_def)
   366 
   367 lemma bind_lift:
   368   "(f \<guillemotright>= lift g) = (f \<guillemotright>= (\<lambda>x. return (g x)))"
   369   by (simp add: lift_def comp_def)
   370 
   371 
   372 subsubsection {* Iteration -- warning: this is rarely useful! *}
   373 
   374 primrec fold_map :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b list Heap" where
   375   "fold_map f [] = return []"
   376 | "fold_map f (x # xs) = do {
   377      y \<leftarrow> f x;
   378      ys \<leftarrow> fold_map f xs;
   379      return (y # ys)
   380    }"
   381 
   382 lemma fold_map_append:
   383   "fold_map f (xs @ ys) = fold_map f xs \<guillemotright>= (\<lambda>xs. fold_map f ys \<guillemotright>= (\<lambda>ys. return (xs @ ys)))"
   384   by (induct xs) simp_all
   385 
   386 lemma execute_fold_map_unchanged_heap [execute_simps]:
   387   assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<exists>y. execute (f x) h = Some (y, h)"
   388   shows "execute (fold_map f xs) h =
   389     Some (List.map (\<lambda>x. fst (the (execute (f x) h))) xs, h)"
   390 using assms proof (induct xs)
   391   case Nil show ?case by (simp add: execute_simps)
   392 next
   393   case (Cons x xs)
   394   from Cons.prems obtain y
   395     where y: "execute (f x) h = Some (y, h)" by auto
   396   moreover from Cons.prems Cons.hyps have "execute (fold_map f xs) h =
   397     Some (map (\<lambda>x. fst (the (execute (f x) h))) xs, h)" by auto
   398   ultimately show ?case by (simp, simp only: execute_bind(1), simp add: execute_simps)
   399 qed
   400 
   401 subsection {* Code generator setup *}
   402 
   403 subsubsection {* Logical intermediate layer *}
   404 
   405 primrec raise' :: "String.literal \<Rightarrow> 'a Heap" where
   406   [code del, code_post]: "raise' (STR s) = raise s"
   407 
   408 lemma raise_raise' [code_inline]:
   409   "raise s = raise' (STR s)"
   410   by simp
   411 
   412 code_datatype raise' -- {* avoid @{const "Heap"} formally *}
   413 
   414 
   415 subsubsection {* SML and OCaml *}
   416 
   417 code_type Heap (SML "unit/ ->/ _")
   418 code_const bind (SML "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())")
   419 code_const return (SML "!(fn/ ()/ =>/ _)")
   420 code_const Heap_Monad.raise' (SML "!(raise/ Fail/ _)")
   421 
   422 code_type Heap (OCaml "unit/ ->/ _")
   423 code_const bind (OCaml "!(fun/ f'_/ ()/ ->/ f'_/ (_/ ())/ ())")
   424 code_const return (OCaml "!(fun/ ()/ ->/ _)")
   425 code_const Heap_Monad.raise' (OCaml "failwith")
   426 
   427 
   428 subsubsection {* Haskell *}
   429 
   430 text {* Adaption layer *}
   431 
   432 code_include Haskell "Heap"
   433 {*import qualified Control.Monad;
   434 import qualified Control.Monad.ST;
   435 import qualified Data.STRef;
   436 import qualified Data.Array.ST;
   437 
   438 import Natural;
   439 
   440 type RealWorld = Control.Monad.ST.RealWorld;
   441 type ST s a = Control.Monad.ST.ST s a;
   442 type STRef s a = Data.STRef.STRef s a;
   443 type STArray s a = Data.Array.ST.STArray s Natural a;
   444 
   445 newSTRef = Data.STRef.newSTRef;
   446 readSTRef = Data.STRef.readSTRef;
   447 writeSTRef = Data.STRef.writeSTRef;
   448 
   449 newArray :: Natural -> a -> ST s (STArray s a);
   450 newArray k = Data.Array.ST.newArray (0, k);
   451 
   452 newListArray :: [a] -> ST s (STArray s a);
   453 newListArray xs = Data.Array.ST.newListArray (0, (fromInteger . toInteger . length) xs) xs;
   454 
   455 newFunArray :: Natural -> (Natural -> a) -> ST s (STArray s a);
   456 newFunArray k f = Data.Array.ST.newListArray (0, k) (map f [0..k-1]);
   457 
   458 lengthArray :: STArray s a -> ST s Natural;
   459 lengthArray a = Control.Monad.liftM snd (Data.Array.ST.getBounds a);
   460 
   461 readArray :: STArray s a -> Natural -> ST s a;
   462 readArray = Data.Array.ST.readArray;
   463 
   464 writeArray :: STArray s a -> Natural -> a -> ST s ();
   465 writeArray = Data.Array.ST.writeArray;*}
   466 
   467 code_reserved Haskell Heap
   468 
   469 text {* Monad *}
   470 
   471 code_type Heap (Haskell "Heap.ST/ Heap.RealWorld/ _")
   472 code_monad bind Haskell
   473 code_const return (Haskell "return")
   474 code_const Heap_Monad.raise' (Haskell "error")
   475 
   476 
   477 subsubsection {* Scala *}
   478 
   479 code_include Scala "Heap"
   480 {*import collection.mutable.ArraySeq
   481 
   482 def bind[A, B](f: Unit => A, g: A => Unit => B): Unit => B = (_: Unit) => g (f ()) ()
   483 
   484 class Ref[A](x: A) {
   485   var value = x
   486 }
   487 
   488 object Ref {
   489   def apply[A](x: A): Ref[A] =
   490     new Ref[A](x)
   491   def lookup[A](r: Ref[A]): A =
   492     r.value
   493   def update[A](r: Ref[A], x: A): Unit =
   494     { r.value = x }
   495 }
   496 
   497 object Array {
   498   def alloc[A](n: Natural.Nat)(x: A): ArraySeq[A] =
   499     ArraySeq.fill(n.as_Int)(x)
   500   def make[A](n: Natural.Nat)(f: Natural.Nat => A): ArraySeq[A] =
   501     ArraySeq.tabulate(n.as_Int)((k: Int) => f(Natural.Nat(k)))
   502   def len[A](a: ArraySeq[A]): Natural.Nat =
   503     Natural.Nat(a.length)
   504   def nth[A](a: ArraySeq[A], n: Natural.Nat): A =
   505     a(n.as_Int)
   506   def upd[A](a: ArraySeq[A], n: Natural.Nat, x: A): Unit =
   507     a.update(n.as_Int, x)
   508   def freeze[A](a: ArraySeq[A]): List[A] =
   509     a.toList
   510 }*}
   511 
   512 code_reserved Scala bind Ref Array
   513 
   514 code_type Heap (Scala "Unit/ =>/ _")
   515 code_const bind (Scala "Heap.bind")
   516 code_const return (Scala "('_: Unit)/ =>/ _")
   517 code_const Heap_Monad.raise' (Scala "!error((_))")
   518 
   519 
   520 subsubsection {* Target variants with less units *}
   521 
   522 setup {*
   523 
   524 let
   525 
   526 open Code_Thingol;
   527 
   528 fun imp_program naming =
   529 
   530   let
   531     fun is_const c = case lookup_const naming c
   532      of SOME c' => (fn c'' => c' = c'')
   533       | NONE => K false;
   534     val is_bind = is_const @{const_name bind};
   535     val is_return = is_const @{const_name return};
   536     val dummy_name = "";
   537     val dummy_case_term = IVar NONE;
   538     (*assumption: dummy values are not relevant for serialization*)
   539     val (unitt, unitT) = case lookup_const naming @{const_name Unity}
   540      of SOME unit' => (IConst (unit', (([], []), [])), the (lookup_tyco naming @{type_name unit}) `%% [])
   541       | NONE => error ("Must include " ^ @{const_name Unity} ^ " in generated constants.");
   542     fun dest_abs ((v, ty) `|=> t, _) = ((v, ty), t)
   543       | dest_abs (t, ty) =
   544           let
   545             val vs = fold_varnames cons t [];
   546             val v = Name.variant vs "x";
   547             val ty' = (hd o fst o unfold_fun) ty;
   548           in ((SOME v, ty'), t `$ IVar (SOME v)) end;
   549     fun force (t as IConst (c, _) `$ t') = if is_return c
   550           then t' else t `$ unitt
   551       | force t = t `$ unitt;
   552     fun tr_bind'' [(t1, _), (t2, ty2)] =
   553       let
   554         val ((v, ty), t) = dest_abs (t2, ty2);
   555       in ICase (((force t1, ty), [(IVar v, tr_bind' t)]), dummy_case_term) end
   556     and tr_bind' t = case unfold_app t
   557      of (IConst (c, (_, ty1 :: ty2 :: _)), [x1, x2]) => if is_bind c
   558           then tr_bind'' [(x1, ty1), (x2, ty2)]
   559           else force t
   560       | _ => force t;
   561     fun imp_monad_bind'' ts = (SOME dummy_name, unitT) `|=> ICase (((IVar (SOME dummy_name), unitT),
   562       [(unitt, tr_bind'' ts)]), dummy_case_term)
   563     fun imp_monad_bind' (const as (c, (_, tys))) ts = if is_bind c then case (ts, tys)
   564        of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' [(t1, ty1), (t2, ty2)]
   565         | ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' [(t1, ty1), (t2, ty2)] `$ t3
   566         | (ts, _) => imp_monad_bind (eta_expand 2 (const, ts))
   567       else IConst const `$$ map imp_monad_bind ts
   568     and imp_monad_bind (IConst const) = imp_monad_bind' const []
   569       | imp_monad_bind (t as IVar _) = t
   570       | imp_monad_bind (t as _ `$ _) = (case unfold_app t
   571          of (IConst const, ts) => imp_monad_bind' const ts
   572           | (t, ts) => imp_monad_bind t `$$ map imp_monad_bind ts)
   573       | imp_monad_bind (v_ty `|=> t) = v_ty `|=> imp_monad_bind t
   574       | imp_monad_bind (ICase (((t, ty), pats), t0)) = ICase
   575           (((imp_monad_bind t, ty),
   576             (map o pairself) imp_monad_bind pats),
   577               imp_monad_bind t0);
   578 
   579   in (Graph.map_nodes o map_terms_stmt) imp_monad_bind end;
   580 
   581 in
   582 
   583 Code_Target.extend_target ("SML_imp", ("SML", imp_program))
   584 #> Code_Target.extend_target ("OCaml_imp", ("OCaml", imp_program))
   585 #> Code_Target.extend_target ("Scala_imp", ("Scala", imp_program))
   586 
   587 end
   588 
   589 *}
   590 
   591 
   592 hide_const (open) Heap heap guard raise' fold_map
   593 
   594 end