1300
|
1 |
(* Title: HOL/MiniML/Maybe.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Dieter Nazareth and Tobias Nipkow
|
|
4 |
Copyright 1995 TU Muenchen
|
|
5 |
|
|
6 |
Universal error monad.
|
|
7 |
*)
|
|
8 |
|
|
9 |
Maybe = List +
|
|
10 |
|
|
11 |
datatype 'a maybe = Ok 'a | Fail
|
|
12 |
|
1376
|
13 |
consts bind :: ['a maybe, 'a => 'b maybe] => 'b maybe (infixl 60)
|
1300
|
14 |
|
|
15 |
defs
|
|
16 |
bind_def "m bind f == case m of Ok r => f r | Fail => Fail"
|
|
17 |
|
1400
|
18 |
syntax "@bind" :: [pttrns,'a maybe,'b] => 'c ("(_ := _;//_)" 0)
|
|
19 |
translations "P := E; F" == "E bind (%P.F)"
|
|
20 |
|
1300
|
21 |
end
|