| author | wenzelm | 
| Tue, 07 Aug 2001 21:27:00 +0200 | |
| changeset 11472 | d08d4e17a5f6 | 
| parent 9001 | 93af64f54bf2 | 
| permissions | -rw-r--r-- | 
(* Title: Option.thy ID: $Id$ Author: Tobias Nipkow Copyright 1994 TU Muenchen Datatype 'a option *) Option = Datatype + datatype 'a option = None | Some 'a consts the :: "'a option => 'a" o2s :: "'a option => 'a set" primrec "the (Some x) = x" primrec "o2s None = {}" "o2s (Some x) = {x}" constdefs option_map :: "('a => 'b) => ('a option => 'b option)" "option_map == %f y. case y of None => None | Some x => Some (f x)" end