--- a/src/HOL/Bali/Basis.thy Thu Feb 21 20:08:09 2002 +0100
+++ b/src/HOL/Bali/Basis.thy Thu Feb 21 20:09:19 2002 +0100
@@ -256,7 +256,7 @@
(* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *)
translations
- "option"<= (type) "Option.option"
+ "option"<= (type) "Datatype.option"
"list" <= (type) "List.list"
"sum3" <= (type) "Basis.sum3"
--- a/src/HOL/Bali/Eval.thy Thu Feb 21 20:08:09 2002 +0100
+++ b/src/HOL/Bali/Eval.thy Thu Feb 21 20:09:19 2002 +0100
@@ -779,7 +779,7 @@
ML {*
local
- fun is_None (Const ("Option.option.None",_)) = true
+ fun is_None (Const ("Datatype.option.None",_)) = true
| is_None _ = false
fun pred (t as (_ $ (Const ("Pair",_) $
(Const ("Pair", _) $ x $ _) $ _ ) $ _)) = is_None x
@@ -806,7 +806,7 @@
ML {*
local
- fun is_Some (Const ("Pair",_) $ (Const ("Option.option.Some",_) $ _)$ _) =true
+ fun is_Some (Const ("Pair",_) $ (Const ("Datatype.option.Some",_) $ _)$ _) =true
| is_Some _ = false
fun pred (_ $ (Const ("Pair",_) $
_ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
--- a/src/HOL/Bali/Evaln.thy Thu Feb 21 20:08:09 2002 +0100
+++ b/src/HOL/Bali/Evaln.thy Thu Feb 21 20:09:19 2002 +0100
@@ -323,7 +323,7 @@
ML {*
local
- fun is_Some (Const ("Pair",_) $ (Const ("Option.option.Some",_) $ _)$ _) =true
+ fun is_Some (Const ("Pair",_) $ (Const ("Datatype.option.Some",_) $ _)$ _) =true
| is_Some _ = false
fun pred (_ $ (Const ("Pair",_) $
_ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
--- a/src/HOL/Map.thy Thu Feb 21 20:08:09 2002 +0100
+++ b/src/HOL/Map.thy Thu Feb 21 20:09:19 2002 +0100
@@ -6,7 +6,7 @@
The datatype of `maps' (written ~=>); strongly resembles maps in VDM.
*)
-Map = List + Option +
+Map = List +
types ('a,'b) "~=>" = 'a => 'b option (infixr 0)
--- a/src/HOL/MiniML/Maybe.thy Thu Feb 21 20:08:09 2002 +0100
+++ b/src/HOL/MiniML/Maybe.thy Thu Feb 21 20:09:19 2002 +0100
@@ -6,7 +6,7 @@
Universal error monad.
*)
-Maybe = Option + List +
+Maybe = Main +
constdefs
option_bind :: ['a option, 'a => 'b option] => 'b option
--- a/src/HOL/PreList.thy Thu Feb 21 20:08:09 2002 +0100
+++ b/src/HOL/PreList.thy Thu Feb 21 20:09:19 2002 +0100
@@ -8,8 +8,7 @@
*)
theory PreList =
- Option + Wellfounded_Relations + NatSimprocs + Recdef + Record +
- Relation_Power:
+ Wellfounded_Relations + NatSimprocs + Recdef + Relation_Power:
(*belongs to theory Divides*)
declare dvdI [intro?] dvdE [elim?] dvd_trans [trans]
--- a/src/HOL/UNITY/Simple/Channel.thy Thu Feb 21 20:08:09 2002 +0100
+++ b/src/HOL/UNITY/Simple/Channel.thy Thu Feb 21 20:09:19 2002 +0100
@@ -8,7 +8,7 @@
From Misra, "A Logic for Concurrent Programming" (1994), section 13.3
*)
-Channel = WFair + Option +
+Channel = WFair +
types state = nat set
--- a/src/HOLCF/IOA/meta_theory/Automata.thy Thu Feb 21 20:08:09 2002 +0100
+++ b/src/HOLCF/IOA/meta_theory/Automata.thy Thu Feb 21 20:09:19 2002 +0100
@@ -7,7 +7,7 @@
*)
-Automata = Option + Asig + Inductive +
+Automata = Asig +
default type