src/HOL/Library/State_Monad.thy
author wenzelm
Wed, 31 Dec 2008 18:53:16 +0100
changeset 29270 0eade173f77e
parent 28145 af3923ed4786
child 29799 7c7f759c438e
permissions -rw-r--r--
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/State_Monad.thy
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
     2
    ID:         $Id$
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
     4
*)
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
     5
27487
c8a6ce181805 absolute imports of HOL/*.thy theories
haftmann
parents: 27368
diff changeset
     6
header {* Combinator syntax for generic, open state monads (single threaded monads) *}
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
     7
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
     8
theory State_Monad
27487
c8a6ce181805 absolute imports of HOL/*.thy theories
haftmann
parents: 27368
diff changeset
     9
imports Plain "~~/src/HOL/List"
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    10
begin
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    11
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    12
subsection {* Motivation *}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    13
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    14
text {*
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    15
  The logic HOL has no notion of constructor classes, so
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    16
  it is not possible to model monads the Haskell way
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    17
  in full genericity in Isabelle/HOL.
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    18
  
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    19
  However, this theory provides substantial support for
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    20
  a very common class of monads: \emph{state monads}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    21
  (or \emph{single-threaded monads}, since a state
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    22
  is transformed single-threaded).
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    23
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    24
  To enter from the Haskell world,
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    25
  \url{http://www.engr.mun.ca/~theo/Misc/haskell_and_monads.htm}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    26
  makes a good motivating start.  Here we just sketch briefly
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    27
  how those monads enter the game of Isabelle/HOL.
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    28
*}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    29
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    30
subsection {* State transformations and combinators *}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    31
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    32
text {*
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    33
  We classify functions operating on states into two categories:
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    34
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    35
  \begin{description}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    36
    \item[transformations]
26266
haftmann
parents: 26260
diff changeset
    37
      with type signature @{text "\<sigma> \<Rightarrow> \<sigma>'"},
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    38
      transforming a state.
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    39
    \item[``yielding'' transformations]
26266
haftmann
parents: 26260
diff changeset
    40
      with type signature @{text "\<sigma> \<Rightarrow> \<alpha> \<times> \<sigma>'"},
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    41
      ``yielding'' a side result while transforming a state.
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    42
    \item[queries]
26266
haftmann
parents: 26260
diff changeset
    43
      with type signature @{text "\<sigma> \<Rightarrow> \<alpha>"},
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    44
      computing a result dependent on a state.
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    45
  \end{description}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    46
26266
haftmann
parents: 26260
diff changeset
    47
  By convention we write @{text "\<sigma>"} for types representing states
haftmann
parents: 26260
diff changeset
    48
  and @{text "\<alpha>"}, @{text "\<beta>"}, @{text "\<gamma>"}, @{text "\<dots>"}
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    49
  for types representing side results.  Type changes due
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    50
  to transformations are not excluded in our scenario.
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    51
26266
haftmann
parents: 26260
diff changeset
    52
  We aim to assert that values of any state type @{text "\<sigma>"}
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    53
  are used in a single-threaded way: after application
26266
haftmann
parents: 26260
diff changeset
    54
  of a transformation on a value of type @{text "\<sigma>"}, the
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    55
  former value should not be used again.  To achieve this,
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    56
  we use a set of monad combinators:
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    57
*}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    58
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
    59
notation fcomp (infixl "o>" 60)
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
    60
notation (xsymbols) fcomp (infixl "o>" 60)
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
    61
notation scomp (infixl "o->" 60)
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
    62
notation (xsymbols) scomp (infixl "o\<rightarrow>" 60)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21283
diff changeset
    63
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26266
diff changeset
    64
abbreviation (input)
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26266
diff changeset
    65
  "return \<equiv> Pair"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21283
diff changeset
    66
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    67
text {*
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    68
  Given two transformations @{term f} and @{term g}, they
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
    69
  may be directly composed using the @{term "op o>"} combinator,
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
    70
  forming a forward composition: @{prop "(f o> g) s = f (g s)"}.
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    71
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    72
  After any yielding transformation, we bind the side result
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    73
  immediately using a lambda abstraction.  This 
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
    74
  is the purpose of the @{term "op o\<rightarrow>"} combinator:
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
    75
  @{prop "(f o\<rightarrow> (\<lambda>x. g)) s = (let (x, s') = f s in g s')"}.
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    76
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    77
  For queries, the existing @{term "Let"} is appropriate.
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    78
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    79
  Naturally, a computation may yield a side result by pairing
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    80
  it to the state from the left;  we introduce the
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    81
  suggestive abbreviation @{term return} for this purpose.
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    82
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    83
  The most crucial distinction to Haskell is that we do
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    84
  not need to introduce distinguished type constructors
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    85
  for different kinds of state.  This has two consequences:
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    86
  \begin{itemize}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    87
    \item The monad model does not state anything about
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    88
       the kind of state; the model for the state is
26260
haftmann
parents: 26141
diff changeset
    89
       completely orthogonal and may be
haftmann
parents: 26141
diff changeset
    90
       specified completely independently.
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    91
    \item There is no distinguished type constructor
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    92
       encapsulating away the state transformation, i.e.~transformations
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    93
       may be applied directly without using any lifting
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    94
       or providing and dropping units (``open monad'').
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    95
    \item The type of states may change due to a transformation.
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    96
  \end{itemize}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    97
*}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    98
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
    99
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   100
subsection {* Monad laws *}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   101
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   102
text {*
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   103
  The common monadic laws hold and may also be used
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   104
  as normalization rules for monadic expressions:
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   105
*}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   106
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   107
lemmas monad_simp = Pair_scomp scomp_Pair id_fcomp fcomp_id
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   108
  scomp_scomp scomp_fcomp fcomp_scomp fcomp_assoc
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   109
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   110
text {*
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   111
  Evaluation of monadic expressions by force:
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   112
*}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   113
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   114
lemmas monad_collapse = monad_simp fcomp_apply scomp_apply split_beta
26260
haftmann
parents: 26141
diff changeset
   115
haftmann
parents: 26141
diff changeset
   116
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   117
subsection {* Syntax *}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   118
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   119
text {*
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   120
  We provide a convenient do-notation for monadic expressions
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   121
  well-known from Haskell.  @{const Let} is printed
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   122
  specially in do-expressions.
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   123
*}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   124
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   125
nonterminals do_expr
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   126
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   127
syntax
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   128
  "_do" :: "do_expr \<Rightarrow> 'a"
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   129
    ("do _ done" [12] 12)
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26266
diff changeset
   130
  "_scomp" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   131
    ("_ <- _;// _" [1000, 13, 12] 12)
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   132
  "_fcomp" :: "'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   133
    ("_;// _" [13, 12] 12)
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   134
  "_let" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   135
    ("let _ = _;// _" [1000, 13, 12] 12)
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   136
  "_done" :: "'a \<Rightarrow> do_expr"
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   137
    ("_" [12] 12)
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   138
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   139
syntax (xsymbols)
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26266
diff changeset
   140
  "_scomp" :: "pttrn \<Rightarrow> 'a \<Rightarrow> do_expr \<Rightarrow> do_expr"
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   141
    ("_ \<leftarrow> _;// _" [1000, 13, 12] 12)
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   142
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   143
translations
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   144
  "_do f" => "f"
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   145
  "_scomp x f g" => "f o\<rightarrow> (\<lambda>x. g)"
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   146
  "_fcomp f g" => "f o> g"
24195
haftmann
parents: 22664
diff changeset
   147
  "_let x t f" => "CONST Let t (\<lambda>x. f)"
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   148
  "_done f" => "f"
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   149
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   150
print_translation {*
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   151
let
24253
3d7f74fd9fd9 fixed syntax
haftmann
parents: 24224
diff changeset
   152
  fun dest_abs_eta (Abs (abs as (_, ty, _))) =
3d7f74fd9fd9 fixed syntax
haftmann
parents: 24224
diff changeset
   153
        let
3d7f74fd9fd9 fixed syntax
haftmann
parents: 24224
diff changeset
   154
          val (v, t) = Syntax.variant_abs abs;
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   155
        in (Free (v, ty), t) end
24253
3d7f74fd9fd9 fixed syntax
haftmann
parents: 24224
diff changeset
   156
    | dest_abs_eta t =
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   157
        let
24253
3d7f74fd9fd9 fixed syntax
haftmann
parents: 24224
diff changeset
   158
          val (v, t) = Syntax.variant_abs ("", dummyT, t $ Bound 0);
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   159
        in (Free (v, dummyT), t) end;
26588
d83271bfaba5 removed syntax from monad combinators; renamed mbind to scomp
haftmann
parents: 26266
diff changeset
   160
  fun unfold_monad (Const (@{const_syntax scomp}, _) $ f $ g) =
24253
3d7f74fd9fd9 fixed syntax
haftmann
parents: 24224
diff changeset
   161
        let
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   162
          val (v, g') = dest_abs_eta g;
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   163
        in Const ("_scomp", dummyT) $ v $ f $ unfold_monad g' end
24195
haftmann
parents: 22664
diff changeset
   164
    | unfold_monad (Const (@{const_syntax fcomp}, _) $ f $ g) =
haftmann
parents: 22664
diff changeset
   165
        Const ("_fcomp", dummyT) $ f $ unfold_monad g
24253
3d7f74fd9fd9 fixed syntax
haftmann
parents: 24224
diff changeset
   166
    | unfold_monad (Const (@{const_syntax Let}, _) $ f $ g) =
24195
haftmann
parents: 22664
diff changeset
   167
        let
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   168
          val (v, g') = dest_abs_eta g;
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   169
        in Const ("_let", dummyT) $ v $ f $ unfold_monad g' end
24195
haftmann
parents: 22664
diff changeset
   170
    | unfold_monad (Const (@{const_syntax Pair}, _) $ f) =
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   171
        Const ("return", dummyT) $ f
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   172
    | unfold_monad f = f;
28145
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   173
  fun contains_scomp (Const (@{const_syntax scomp}, _) $ _ $ _) = true
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   174
    | contains_scomp (Const (@{const_syntax fcomp}, _) $ _ $ t) =
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   175
        contains_scomp t
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   176
    | contains_scomp (Const (@{const_syntax Let}, _) $ _ $ Abs (_, _, t)) =
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   177
        contains_scomp t;
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   178
  fun scomp_monad_tr' (f::g::ts) = list_comb
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   179
    (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax scomp}, dummyT) $ f $ g), ts);
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   180
  fun fcomp_monad_tr' (f::g::ts) = if contains_scomp g then list_comb
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   181
      (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax fcomp}, dummyT) $ f $ g), ts)
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   182
    else raise Match;
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   183
  fun Let_monad_tr' (f :: (g as Abs (_, _, g')) :: ts) = if contains_scomp g' then list_comb
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   184
      (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax Let}, dummyT) $ f $ g), ts)
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   185
    else raise Match;
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   186
in [
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   187
  (@{const_syntax scomp}, scomp_monad_tr'),
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   188
  (@{const_syntax fcomp}, fcomp_monad_tr'),
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   189
  (@{const_syntax Let}, Let_monad_tr')
af3923ed4786 dropped "run" marker in monad syntax
haftmann
parents: 27487
diff changeset
   190
] end;
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   191
*}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   192
21418
4bc2882f80af added combinators and lemmas
haftmann
parents: 21404
diff changeset
   193
text {*
24195
haftmann
parents: 22664
diff changeset
   194
  For an example, see HOL/ex/Random.thy.
21192
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   195
*}
5fe5cd5fede7 added state monad to HOL library
haftmann
parents:
diff changeset
   196
22664
e965391e2864 do translation: CONST;
wenzelm
parents: 22492
diff changeset
   197
end