diff -r f4da791d4850 -r d960cc4a6afc doc-src/TutorialI/Misc/Option2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/TutorialI/Misc/Option2.thy Fri Dec 01 13:47:37 2000 +0100 @@ -0,0 +1,33 @@ +(*<*) +theory Option2 = Main: +hide const None Some +hide type option +(*>*) + +text{*\indexbold{*option}\indexbold{*None}\indexbold{*Some} +Our final datatype is very simple but still eminently useful: +*} + +datatype 'a option = None | Some 'a; + +text{*\noindent +Frequently one needs to add a distiguished element to some existing type. +For example, type @{text"t option"} can model the result of a computation that +may either terminate with an error (represented by @{term None}) or return +some value @{term v} (represented by @{term"Some v"}). +Similarly, @{typ nat} extended with $\infty$ can be modeled by type +@{typ"nat option"}. In both cases one could define a new datatype with +customized constructors like @{term Error} and @{term Infinity}, +but it is often simpler to use @{text option}. For an application see +\S\ref{sec:Trie}. +*} +(*<*) +(* +constdefs + infplus :: "nat option \ nat option \ nat option" +"infplus x y \ case x of None \ None + | Some m \ (case y of None \ None | Some n \ Some(m+n))" + +*) +end +(*>*)