src/HOL/Library/Monad_Syntax.thy
changeset 45990 b7b905b23b2a
parent 45964 7b3a18670a9f
child 46143 c932c80d3eae
equal deleted inserted replaced
45989:b39256df5f8a 45990:b7b905b23b2a
     3 *)
     3 *)
     4 
     4 
     5 header {* Monad notation for arbitrary types *}
     5 header {* Monad notation for arbitrary types *}
     6 
     6 
     7 theory Monad_Syntax
     7 theory Monad_Syntax
     8 imports Main "~~/src/Tools/Adhoc_Overloading" "~~/src/HOL/Library/More_List"
     8 imports Main "~~/src/Tools/Adhoc_Overloading"
     9 begin
     9 begin
    10 
    10 
    11 text {*
    11 text {*
    12   We provide a convenient do-notation for monadic expressions
    12   We provide a convenient do-notation for monadic expressions
    13   well-known from Haskell.  @{const Let} is printed
    13   well-known from Haskell.  @{const Let} is printed