author | oheimb |
Thu, 12 Mar 1998 13:17:13 +0100 | |
changeset 4743 | b3bfcbd9fb93 |
parent 4192 | c38ab5af38b5 |
child 4752 | 1c334bd00038 |
permissions | -rw-r--r-- |
(* Title: Option.thy ID: $Id$ Author: Tobias Nipkow Copyright 1994 TU Muenchen Datatype 'a option *) Option = Arith + datatype 'a option = None | Some 'a constdefs the :: "'a option => 'a" "the == %y. case y of None => arbitrary | Some x => x" option_map :: "('a => 'b) => ('a option => 'b option)" "option_map == %f y. case y of None => None | Some x => Some (f x)" end