src/HOL/MiniML/Maybe.thy
author nipkow
Fri, 08 Dec 1995 19:48:15 +0100
changeset 1400 5d909faf0e04
parent 1376 92f83b9d17e1
child 1557 fe30812f5b5e
permissions -rw-r--r--
Introduced Monad syntax Pat := Val; Cont
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     1
(* Title:     HOL/MiniML/Maybe.thy
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     2
   ID:        $Id$
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     3
   Author:    Dieter Nazareth and Tobias Nipkow
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     4
   Copyright  1995 TU Muenchen
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     5
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     6
Universal error monad.
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     7
*)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     8
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     9
Maybe = List +
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    10
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    11
datatype 'a maybe =  Ok 'a | Fail
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    12
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1300
diff changeset
    13
consts bind :: ['a maybe, 'a => 'b maybe] => 'b maybe (infixl 60)
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    14
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    15
defs
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    16
  bind_def "m bind f == case m of Ok r => f r | Fail => Fail"
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    17
1400
5d909faf0e04 Introduced Monad syntax Pat := Val; Cont
nipkow
parents: 1376
diff changeset
    18
syntax "@bind" :: [pttrns,'a maybe,'b] => 'c ("(_ := _;//_)" 0)
5d909faf0e04 Introduced Monad syntax Pat := Val; Cont
nipkow
parents: 1376
diff changeset
    19
translations "P := E; F" == "E bind (%P.F)"
5d909faf0e04 Introduced Monad syntax Pat := Val; Cont
nipkow
parents: 1376
diff changeset
    20
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    21
end