Deleted Library.option type.
--- a/Admin/Benchmarks/HOL-datatype/ROOT.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/Admin/Benchmarks/HOL-datatype/ROOT.ML Sun Feb 13 17:15:14 2005 +0100
@@ -1,7 +1,7 @@
(* Title: Admin/Benchmarks/HOL-datatype/ROOT.ML
ID: $Id$
-Some rather large datatype examples (from John Harrison).
+SOME rather large datatype examples (from John Harrison).
*)
val tests = ["Brackin", "Instructions", "SML", "Verilog"];
--- a/NEWS Fri Feb 11 18:51:00 2005 +0100
+++ b/NEWS Sun Feb 13 17:15:14 2005 +0100
@@ -6,6 +6,11 @@
*** General ***
+* The type Library.option is no more, along with the exception
+ Library.OPTION: Isabelle now uses the standard option type. The
+ functions the, is_some, is_none, etc. are still in Library, but
+ the constructors are now SOME and NONE instead of Some and None.
+
* Document preparation: new antiquotations @{lhs thm} and @{rhs thm}
printing the lhs/rhs of definitions, equations, inequations etc.
--- a/TFL/casesplit.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/TFL/casesplit.ML Sun Feb 13 17:15:14 2005 +0100
@@ -72,11 +72,11 @@
(* finding a free var to split *)
val find_term_split :
- Term.term * Term.term -> (string * Term.typ) Library.option
+ Term.term * Term.term -> (string * Term.typ) option
val find_thm_split :
- Thm.thm -> int -> Thm.thm -> (string * Term.typ) Library.option
+ Thm.thm -> int -> Thm.thm -> (string * Term.typ) option
val find_thms_split :
- Thm.thm list -> int -> Thm.thm -> (string * Term.typ) Library.option
+ Thm.thm list -> int -> Thm.thm -> (string * Term.typ) option
(* try to recursively split conjectured thm to given list of thms *)
val splitto : Thm.thm list -> Thm.thm -> Thm.thm
@@ -129,8 +129,8 @@
| TVar((s,i),_) => raise ERROR_MESSAGE
("Free variable: " ^ s)
val dt = case (Symtab.lookup (dtypestab,ty_str))
- of Some dt => dt
- | None => raise ERROR_MESSAGE ("Not a Datatype: " ^ ty_str)
+ of SOME dt => dt
+ | NONE => raise ERROR_MESSAGE ("Not a Datatype: " ^ ty_str)
in
cases_thm_of_induct_thm (#induction dt)
end;
@@ -206,11 +206,11 @@
fun getter x =
let val (n,ty) = Term.dest_Free x in
(if vstr = n orelse vstr = Syntax.dest_skolem n
- then Some (n,ty) else None )
- handle LIST _ => None
+ then SOME (n,ty) else NONE )
+ handle LIST _ => NONE
end;
val (n,ty) = case Library.get_first getter freets
- of Some (n, ty) => (n, ty)
+ of SOME (n, ty) => (n, ty)
| _ => raise ERROR_MESSAGE ("no such variable " ^ vstr);
val sgn = Thm.sign_of_thm th;
@@ -247,22 +247,22 @@
subterm, or constant in another, ie assume that one term is a plit of
another, then gives back the free variable that has been split. *)
exception find_split_exp of string
-fun find_term_split (Free v, _ $ _) = Some v
- | find_term_split (Free v, Const _) = Some v
- | find_term_split (Free v, Abs _) = Some v (* do we really want this case? *)
- | find_term_split (Free v, Var _) = None (* keep searching *)
+fun find_term_split (Free v, _ $ _) = SOME v
+ | find_term_split (Free v, Const _) = SOME v
+ | find_term_split (Free v, Abs _) = SOME v (* do we really want this case? *)
+ | find_term_split (Free v, Var _) = NONE (* keep searching *)
| find_term_split (a $ b, a2 $ b2) =
(case find_term_split (a, a2) of
- None => find_term_split (b,b2)
+ NONE => find_term_split (b,b2)
| vopt => vopt)
| find_term_split (Abs(_,ty,t1), Abs(_,ty2,t2)) =
find_term_split (t1, t2)
| find_term_split (Const (x,ty), Const(x2,ty2)) =
- if x = x2 then None else (* keep searching *)
+ if x = x2 then NONE else (* keep searching *)
raise find_split_exp (* stop now *)
"Terms are not identical upto a free varaible! (Consts)"
| find_term_split (Bound i, Bound j) =
- if i = j then None else (* keep searching *)
+ if i = j then NONE else (* keep searching *)
raise find_split_exp (* stop now *)
"Terms are not identical upto a free varaible! (Bound)"
| find_term_split (a, b) =
@@ -274,7 +274,7 @@
splitth. *)
fun find_thm_split splitth i genth =
find_term_split (Logic.get_goal (Thm.prop_of genth) i,
- Thm.concl_of splitth) handle find_split_exp _ => None;
+ Thm.concl_of splitth) handle find_split_exp _ => NONE;
(* as above but searches "splitths" for a theorem that suggest a case split *)
fun find_thms_split splitths i genth =
@@ -304,12 +304,12 @@
fun split th =
(case find_thms_split splitths 1 th of
- None =>
+ NONE =>
(writeln "th:";
Display.print_thm th; writeln "split ths:";
Display.print_thms splitths; writeln "\n--";
raise ERROR_MESSAGE "splitto: cannot find variable to split on")
- | Some v =>
+ | SOME v =>
let
val gt = Data.dest_Trueprop (nth_elem(0, Thm.prems_of th));
val split_thm = mk_casesplit_goal_thm sgn v gt;
@@ -322,8 +322,8 @@
(* note: multiple unifiers! we only take the first element,
probably fine -- there is probably only one anyway. *)
(case Library.get_first (Seq.pull o solve_by_splitth th) splitths of
- None => split th
- | Some (solved_th, more) => solved_th)
+ NONE => split th
+ | SOME (solved_th, more) => solved_th)
in
recsplitf genth
end;
@@ -340,7 +340,7 @@
the well-foundness conditions have been solved. *)
local
fun get_related_thms i =
- mapfilter ((fn (r,x) => if x = i then Some r else None));
+ mapfilter ((fn (r,x) => if x = i then SOME r else NONE));
fun solve_eq (th, [], i) =
raise ERROR_MESSAGE "derive_init_eqs: missing rules"
--- a/TFL/dcterm.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/TFL/dcterm.ML Sun Feb 13 17:15:14 2005 +0100
@@ -71,7 +71,7 @@
(*---------------------------------------------------------------------------
- * Some simple constructor functions.
+ * SOME simple constructor functions.
*---------------------------------------------------------------------------*)
val mk_hol_const = Thm.cterm_of (Theory.sign_of HOL.thy) o Const;
@@ -124,7 +124,7 @@
in (dest_monop expected M, N) handle U.ERR _ => err () end;
fun dest_binder expected tm =
- dest_abs None (dest_monop expected tm)
+ dest_abs NONE (dest_monop expected tm)
handle U.ERR _ => raise ERR "dest_binder" ("Not a(n) " ^ quote expected);
@@ -175,7 +175,7 @@
val strip_comb = strip (Library.swap o dest_comb) (* Goes to the "left" *)
val strip_imp = rev2swap o strip dest_imp
-val strip_abs = rev2swap o strip (dest_abs None)
+val strip_abs = rev2swap o strip (dest_abs NONE)
val strip_forall = rev2swap o strip dest_forall
val strip_exists = rev2swap o strip dest_exists
--- a/TFL/post.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/TFL/post.ML Sun Feb 13 17:15:14 2005 +0100
@@ -207,7 +207,7 @@
-- Lucas Dixon, Aug 2004 *)
local
fun get_related_thms i =
- mapfilter ((fn (r,x) => if x = i then Some r else None));
+ mapfilter ((fn (r,x) => if x = i then SOME r else NONE));
fun solve_eq (th, [], i) =
raise ERROR_MESSAGE "derive_init_eqs: missing rules"
--- a/TFL/rules.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/TFL/rules.ML Sun Feb 13 17:15:14 2005 +0100
@@ -430,7 +430,7 @@
fun SUBS thl =
rewrite_rule (map (fn th => th RS eq_reflection handle THM _ => th) thl);
-val rew_conv = MetaSimplifier.rewrite_cterm (true, false, false) (K (K None));
+val rew_conv = MetaSimplifier.rewrite_cterm (true, false, false) (K (K NONE));
fun simpl_conv ss thl ctm =
rew_conv (ss addsimps thl) ctm RS meta_eq_to_obj_eq;
@@ -753,8 +753,8 @@
else q_eliminate (thm,imp,sign))
(* Assume that the leading constant is ==, *)
| _ => thm (* if it is not a ==> *)
- in Some(eliminate (rename thm)) end
- handle U.ERR _ => None (* FIXME handle THM as well?? *)
+ in SOME(eliminate (rename thm)) end
+ handle U.ERR _ => NONE (* FIXME handle THM as well?? *)
fun restrict_prover ss thm =
let val dummy = say "restrict_prover:"
@@ -797,8 +797,8 @@
(LIST_CONJ rcontext)
end
val th'' = th' RS thm
- in Some (th'')
- end handle U.ERR _ => None (* FIXME handle THM as well?? *)
+ in SOME (th'')
+ end handle U.ERR _ => NONE (* FIXME handle THM as well?? *)
in
(if (is_cong thm) then cong_prover else restrict_prover) ss thm
end
--- a/TFL/tfl.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/TFL/tfl.ML Sun Feb 13 17:15:14 2005 +0100
@@ -260,8 +260,8 @@
let val pty as Type (ty_name,_) = type_of p
in
case (ty_info ty_name)
- of None => mk_case_fail("Not a known datatype: "^ty_name)
- | Some{case_const,constructors} =>
+ of NONE => mk_case_fail("Not a known datatype: "^ty_name)
+ | SOME{case_const,constructors} =>
let
val case_const_name = #1(dest_Const case_const)
val nrows = List.concat (map (expand constructors pty) rows)
@@ -371,7 +371,7 @@
(*For Isabelle, the lhs of a definition must be a constant.*)
fun mk_const_def sign (Name, Ty, rhs) =
- Sign.infer_types (Sign.pp sign) sign (K None) (K None) [] false
+ Sign.infer_types (Sign.pp sign) sign (K NONE) (K NONE) [] false
([Const("==",dummyT) $ Const(Name,Ty) $ rhs], propT)
|> #1;
@@ -623,7 +623,7 @@
val vlist = #2(S.strip_comb (S.rhs body))
val plist = ListPair.zip (vlist, xlist)
val args = map (fn qv => the (gen_assoc (op aconv) (plist, qv))) qvars
- handle Library.OPTION => sys_error
+ handle Option => sys_error
"TFL fault [alpha_ex_unroll]: no correspondence"
fun build ex [] = []
| build (_$rex) (v::rst) =
@@ -666,8 +666,8 @@
let val Type (ty_name,_) = type_of p
in
case (ty_info ty_name)
- of None => fail("Not a known datatype: "^ty_name)
- | Some{constructors,nchotomy} =>
+ of NONE => fail("Not a known datatype: "^ty_name)
+ | SOME{constructors,nchotomy} =>
let val thm' = R.ISPEC (tych u) nchotomy
val disjuncts = S.strip_disj (concl thm')
val subproblems = divide(constructors, rows)
--- a/TFL/thry.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/TFL/thry.ML Sun Feb 13 17:15:14 2005 +0100
@@ -58,14 +58,14 @@
fun match_info thy tname =
case (DatatypePackage.case_const_of thy tname, DatatypePackage.constrs_of thy tname) of
- (Some case_const, Some constructors) =>
- Some {case_const = case_const, constructors = constructors}
- | _ => None;
+ (SOME case_const, SOME constructors) =>
+ SOME {case_const = case_const, constructors = constructors}
+ | _ => NONE;
fun induct_info thy tname = case get_info thy tname of
- None => None
- | Some {nchotomy, ...} =>
- Some {nchotomy = nchotomy,
+ NONE => NONE
+ | SOME {nchotomy, ...} =>
+ SOME {nchotomy = nchotomy,
constructors = the (DatatypePackage.constrs_of thy tname)};
fun extract_info thy =
--- a/TFL/usyntax.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/TFL/usyntax.ML Sun Feb 13 17:15:14 2005 +0100
@@ -385,11 +385,11 @@
(* Search a term for a sub-term satisfying the predicate p. *)
fun find_term p =
let fun find tm =
- if (p tm) then Some tm
+ if (p tm) then SOME tm
else case tm of
Abs(_,_,body) => find body
- | (t$u) => (case find t of None => find u | some => some)
- | _ => None
+ | (t$u) => (case find t of NONE => find u | some => some)
+ | _ => NONE
in find
end;
--- a/src/CCL/CCL.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/CCL/CCL.ML Sun Feb 13 17:15:14 2005 +0100
@@ -88,7 +88,7 @@
| arg_str n a s = arg_str (n-1) a ("," ^ a ^ (chr((ord "a")+n-1)) ^ s);
val sg = sign_of thy;
val T = case Sign.const_type sg (Sign.intern_const (sign_of thy) sy) of
- None => error(sy^" not declared") | Some(T) => T;
+ NONE => error(sy^" not declared") | SOME(T) => T;
val arity = length (fst (strip_type T));
in sy ^ (arg_str arity name "") end;
--- a/src/CTT/ex/elim.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/CTT/ex/elim.ML Sun Feb 13 17:15:14 2005 +0100
@@ -3,7 +3,7 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1991 University of Cambridge
-Some examples taken from P. Martin-L\"of, Intuitionistic type theory
+SOME examples taken from P. Martin-L\"of, Intuitionistic type theory
(Bibliopolis, 1984).
by (safe_tac prems 1);
--- a/src/Cube/ex/ex.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Cube/ex/ex.ML Sun Feb 13 17:15:14 2005 +0100
@@ -201,7 +201,7 @@
by (EVERY[etac pi_elim 1, assume_tac 1, assume_tac 1]);
uresult();
-(* Some random examples *)
+(* SOME random examples *)
goal LP2.thy "A:* c:A f:A->A |- \
\ Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T";
--- a/src/FOL/IFOL_lemmas.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOL/IFOL_lemmas.ML Sun Feb 13 17:15:14 2005 +0100
@@ -204,7 +204,7 @@
by (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ;
qed "ex1I";
-(*Sometimes easier to use: the premises have no shared variables. Safe!*)
+(*SOMEtimes easier to use: the premises have no shared variables. Safe!*)
val [ex,eq] = Goal
"[| EX x. P(x); !!x y. [| P(x); P(y) |] ==> x=y |] ==> EX! x. P(x)";
by (rtac (ex RS exE) 1);
--- a/src/FOL/eqrule_FOL_data.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOL/eqrule_FOL_data.ML Sun Feb 13 17:15:14 2005 +0100
@@ -17,10 +17,10 @@
struct
fun mk_eq th = case concl_of th of
- Const("==",_)$_$_ => Some (th)
- | _$(Const("op <->",_)$_$_) => Some (th RS iff_reflection)
- | _$(Const("op =",_)$_$_) => Some (th RS eq_reflection)
- | _ => None;
+ Const("==",_)$_$_ => SOME (th)
+ | _$(Const("op <->",_)$_$_) => SOME (th RS iff_reflection)
+ | _$(Const("op =",_)$_$_) => SOME (th RS eq_reflection)
+ | _ => NONE;
val tranformation_pairs =
[("op -->", [mp]), ("op &", [conjunct1,conjunct2]),
@@ -37,8 +37,8 @@
(case Term.head_of p of
Const(a,_) =>
(case Library.assoc(pairs,a) of
- Some(rls) => flat (map atoms ([th] RL rls))
- | None => [th])
+ SOME(rls) => flat (map atoms ([th] RL rls))
+ | NONE => [th])
| _ => [th])
| _ => [th])
in atoms end;
--- a/src/FOL/ex/int.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOL/ex/int.ML Sun Feb 13 17:15:14 2005 +0100
@@ -296,7 +296,7 @@
writeln"Hard examples with quantifiers";
(*The ones that have not been proved are not known to be valid!
- Some will require quantifier duplication -- not currently available*)
+ SOME will require quantifier duplication -- not currently available*)
writeln"Problem ~~18";
Goal "~~(EX y. ALL x. P(y)-->P(x))";
--- a/src/FOL/ex/intro.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOL/ex/intro.ML Sun Feb 13 17:15:14 2005 +0100
@@ -14,7 +14,7 @@
context FOL.thy;
-(**** Some simple backward proofs ****)
+(**** SOME simple backward proofs ****)
Goal "P|P --> P";
by (rtac impI 1);
--- a/src/FOL/ex/quant.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOL/ex/quant.ML Sun Feb 13 17:15:14 2005 +0100
@@ -34,7 +34,7 @@
result();
-writeln"Some harder ones";
+writeln"SOME harder ones";
Goal "(EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))";
by tac;
@@ -104,7 +104,7 @@
result();
-writeln"Some slow ones";
+writeln"SOME slow ones";
(*Principia Mathematica *11.53 *)
--- a/src/FOL/simpdata.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOL/simpdata.ML Sun Feb 13 17:15:14 2005 +0100
@@ -116,8 +116,8 @@
(case head_of p of
Const(a,_) =>
(case assoc(pairs,a) of
- Some(rls) => flat (map atoms ([th] RL rls))
- | None => [th])
+ SOME(rls) => flat (map atoms ([th] RL rls))
+ | NONE => [th])
| _ => [th])
| _ => [th])
in atoms end;
@@ -239,12 +239,12 @@
structure Quantifier1 = Quantifier1Fun(
struct
(*abstract syntax*)
- fun dest_eq((c as Const("op =",_)) $ s $ t) = Some(c,s,t)
- | dest_eq _ = None;
- fun dest_conj((c as Const("op &",_)) $ s $ t) = Some(c,s,t)
- | dest_conj _ = None;
- fun dest_imp((c as Const("op -->",_)) $ s $ t) = Some(c,s,t)
- | dest_imp _ = None;
+ fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t)
+ | dest_eq _ = NONE;
+ fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t)
+ | dest_conj _ = NONE;
+ fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t)
+ | dest_imp _ = NONE;
val conj = FOLogic.conj
val imp = FOLogic.imp
(*rules*)
--- a/src/FOLP/ex/int.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOLP/ex/int.ML Sun Feb 13 17:15:14 2005 +0100
@@ -231,7 +231,7 @@
writeln"Hard examples with quantifiers";
(*The ones that have not been proved are not known to be valid!
- Some will require quantifier duplication -- not currently available*)
+ SOME will require quantifier duplication -- not currently available*)
writeln"Problem ~~18";
goal IFOLP.thy "?p : ~~(EX y. ALL x. P(y)-->P(x))";
--- a/src/FOLP/ex/intro.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOLP/ex/intro.ML Sun Feb 13 17:15:14 2005 +0100
@@ -12,7 +12,7 @@
*)
-(**** Some simple backward proofs ****)
+(**** SOME simple backward proofs ****)
goal FOLP.thy "?p:P|P --> P";
by (rtac impI 1);
--- a/src/FOLP/ex/quant.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOLP/ex/quant.ML Sun Feb 13 17:15:14 2005 +0100
@@ -34,7 +34,7 @@
result();
-writeln"Some harder ones";
+writeln"SOME harder ones";
Goal "?p : (EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))";
by tac;
@@ -104,7 +104,7 @@
result();
-writeln"Some slow ones";
+writeln"SOME slow ones";
(*Principia Mathematica *11.53 *)
--- a/src/FOLP/simp.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOLP/simp.ML Sun Feb 13 17:15:14 2005 +0100
@@ -136,8 +136,8 @@
(*Applies tactic and returns the first resulting state, FAILS if none!*)
fun one_result(tac,thm) = case Seq.pull(tac thm) of
- Some(thm',_) => thm'
- | None => raise THM("Simplifier: could not continue", 0, [thm]);
+ SOME(thm',_) => thm'
+ | NONE => raise THM("Simplifier: could not continue", 0, [thm]);
fun res1(thm,thms,i) = one_result(resolve_tac thms i,thm);
@@ -180,8 +180,8 @@
Abs (_,_,body) => add_vars(body,vars)
| r$s => (case head_of tm of
Const(c,T) => (case assoc(new_asms,c) of
- None => add_vars(r,add_vars(s,vars))
- | Some(al) => add_list(tm,al,vars))
+ NONE => add_vars(r,add_vars(s,vars))
+ | SOME(al) => add_list(tm,al,vars))
| _ => add_vars(r,add_vars(s,vars)))
| _ => vars
in add_vars end;
@@ -211,7 +211,7 @@
| Free _ => resolve_tac congs 1 ORELSE refl1_tac
| _ => refl1_tac)
val add_norm_tac = DEPTH_FIRST (has_fewer_prems nops) norm_step_tac
- val Some(thm'',_) = Seq.pull(add_norm_tac thm')
+ val SOME(thm'',_) = Seq.pull(add_norm_tac thm')
in thm'' end;
fun add_norm_tags congs =
@@ -423,14 +423,14 @@
if var_lhs(thm,i) then (ss,thm,anet,ats,cs) else
if lhs_is_NORM(thm,i) then (ss, res1(thm,trans_norms,i), anet,ats,cs)
else case Seq.pull(cong_tac i thm) of
- Some(thm',_) =>
+ SOME(thm',_) =>
let val ps = prems_of thm and ps' = prems_of thm';
val n = length(ps')-length(ps);
val a = length(strip_assums_hyp(nth_elem(i-1,ps)))
val l = map (fn p => length(strip_assums_hyp(p)))
(take(n,drop(i-1,ps')));
in (simp_refl(rev(l),a,REW::ss),thm',anet,ats,cs) end
- | None => (REW::ss,thm,anet,ats,cs);
+ | NONE => (REW::ss,thm,anet,ats,cs);
(*NB: the "Adding rewrites:" trace will look strange because assumptions
are represented by rules, generalized over their parameters*)
@@ -447,22 +447,22 @@
end;
fun rew(seq,thm,ss,anet,ats,cs, more) = case Seq.pull seq of
- Some(thm',seq') =>
+ SOME(thm',seq') =>
let val n = (nprems_of thm') - (nprems_of thm)
in pr_rew(i,n,thm,thm',more);
if n=0 then (SIMP_LHS::ss, thm', anet, ats, cs)
else ((replicate n PROVE) @ (POP_CS::SIMP_LHS::ss),
thm', anet, ats, (ss,thm,anet,ats,seq',more)::cs)
end
- | None => if more
+ | NONE => if more
then rew((lhs_net_tac anet i THEN assume_tac i) thm,
thm,ss,anet,ats,cs,false)
else (ss,thm,anet,ats,cs);
fun try_true(thm,ss,anet,ats,cs) =
case Seq.pull(auto_tac i thm) of
- Some(thm',_) => (ss,thm',anet,ats,cs)
- | None => let val (ss0,thm0,anet0,ats0,seq,more)::cs0 = cs
+ SOME(thm',_) => (ss,thm',anet,ats,cs)
+ | NONE => let val (ss0,thm0,anet0,ats0,seq,more)::cs0 = cs
in if !tracing
then (writeln"*** Failed to prove precondition. Normal form:";
pr_goal_concl i thm; writeln"")
@@ -472,8 +472,8 @@
fun if_exp(thm,ss,anet,ats,cs) =
case Seq.pull (IF1_TAC (cong_tac i) i thm) of
- Some(thm',_) => (SIMP_LHS::IF::ss,thm',anet,ats,cs)
- | None => (ss,thm,anet,ats,cs);
+ SOME(thm',_) => (SIMP_LHS::IF::ss,thm',anet,ats,cs)
+ | NONE => (ss,thm,anet,ats,cs);
fun step(s::ss, thm, anet, ats, cs) = case s of
MK_EQ => (ss, res1(thm,[red2],i), anet, ats, cs)
@@ -549,10 +549,10 @@
let val (Const(_,cT), va, vb) = dest_red(hd(prems_of thm));
val [P] = add_term_vars(concl_of thm,[]) \\ [va,vb]
val eqT::_ = binder_types cT
- in if Type.typ_instance tsig (T,eqT) then Some(thm,va,vb,P)
+ in if Type.typ_instance tsig (T,eqT) then SOME(thm,va,vb,P)
else find thms
end
- | find [] = None
+ | find [] = NONE
in find subst_thms end;
fun mk_cong sg (f,aTs,rT) (refl,eq) =
@@ -569,8 +569,8 @@
fun mk(c,T::Ts,i,yik) =
let val si = radixstring(26,"a",i)
in case find_subst tsig T of
- None => mk(c,Ts,i-1,eta_Var(("X"^si,0),T)::yik)
- | Some s => let val c' = c RSN (2,ri(s,i,si,T,yik))
+ NONE => mk(c,Ts,i-1,eta_Var(("X"^si,0),T)::yik)
+ | SOME s => let val c' = c RSN (2,ri(s,i,si,T,yik))
in mk(c',Ts,i-1,eta_Var(("Y"^si,0),T)::yik) end
end
| mk(c,[],_,_) = c;
@@ -582,17 +582,17 @@
fun find_refl(r::rs) =
let val (Const(eq,eqT),_,_) = dest_red(concl_of r)
in if Type.typ_instance tsig (rT, hd(binder_types eqT))
- then Some(r,(eq,body_type eqT)) else find_refl rs
+ then SOME(r,(eq,body_type eqT)) else find_refl rs
end
- | find_refl([]) = None;
+ | find_refl([]) = NONE;
in case find_refl refl_thms of
- None => [] | Some(refl) => [mk_cong sg (f,aTs,rT) refl]
+ NONE => [] | SOME(refl) => [mk_cong sg (f,aTs,rT) refl]
end;
fun mk_cong_thy thy f =
let val sg = sign_of thy;
val T = case Sign.const_type sg f of
- None => error(f^" not declared") | Some(T) => T;
+ NONE => error(f^" not declared") | SOME(T) => T;
val T' = incr_tvar 9 T;
in mk_cong_type sg (Const(f,T'),T') end;
@@ -602,9 +602,9 @@
let val sg = sign_of thy;
val S0 = Sign.defaultS sg;
fun readfT(f,s) =
- let val T = incr_tvar 9 (Sign.read_typ(sg,K(Some(S0))) s);
+ let val T = incr_tvar 9 (Sign.read_typ(sg,K(SOME(S0))) s);
val t = case Sign.const_type sg f of
- Some(_) => Const(f,T) | None => Free(f,T)
+ SOME(_) => Const(f,T) | NONE => Free(f,T)
in (t,T) end
in flat o map (mk_cong_type sg o readfT) end;
--- a/src/FOLP/simpdata.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/FOLP/simpdata.ML Sun Feb 13 17:15:14 2005 +0100
@@ -85,8 +85,8 @@
(case head_of p of
Const(a,_) =>
(case assoc(pairs,a) of
- Some(rls) => flat (map atoms ([th] RL rls))
- | None => [th])
+ SOME(rls) => flat (map atoms ([th] RL rls))
+ | NONE => [th])
| _ => [th])
| _ => [th])
in atoms end;
--- a/src/HOL/Algebra/abstract/Ring.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Algebra/abstract/Ring.thy Sun Feb 13 17:15:14 2005 +0100
@@ -170,8 +170,8 @@
|> mk_meta_eq;
val (t', u) = Logic.dest_equals (Thm.prop_of rew);
in if t' aconv u
- then None
- else Some rew
+ then NONE
+ else SOME rew
end;
in
val ring_simproc = Simplifier.simproc (sign_of (the_context ())) "ring" lhss proc;
--- a/src/HOL/Algebra/abstract/order.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Algebra/abstract/order.ML Sun Feb 13 17:15:14 2005 +0100
@@ -60,7 +60,7 @@
fun termless_ring (a, b) = (term_lpo ring_ord (a, b) = LESS);
-(* Some code useful for debugging
+(* SOME code useful for debugging
val intT = HOLogic.intT;
val a = Free ("a", intT);
--- a/src/HOL/Algebra/ringsimp.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Algebra/ringsimp.ML Sun Feb 13 17:15:14 2005 +0100
@@ -82,9 +82,9 @@
("Non-commutative rings in proof context: " ^ commas non_comm_rings ^
"\nCommutative rings in proof context: " ^ commas comm_rings);
val simps =
- flat (map (fn s => ProofContext.get_thms ctxt (s ^ ".ring_simprules", None))
+ flat (map (fn s => ProofContext.get_thms ctxt (s ^ ".ring_simprules", NONE))
non_comm_rings) @
- flat (map (fn s => ProofContext.get_thms ctxt (s ^ ".cring_simprules", None))
+ flat (map (fn s => ProofContext.get_thms ctxt (s ^ ".cring_simprules", NONE))
comm_rings);
in Method.SIMPLE_METHOD' HEADGOAL (asm_full_simp_tac (cring_ss addsimps simps))
end;
--- a/src/HOL/Bali/Basis.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Bali/Basis.thy Sun Feb 13 17:15:14 2005 +0100
@@ -20,7 +20,7 @@
ML {*
fun cond_simproc name pat pred thm = Simplifier.simproc (Thm.sign_of_thm thm) name [pat]
- (fn _ => fn _ => fn t => if pred t then None else Some (mk_meta_eq thm));
+ (fn _ => fn _ => fn t => if pred t then NONE else SOME (mk_meta_eq thm));
*}
declare split_if_asm [split] option.split [split] option.split_asm [split]
--- a/src/HOL/Extraction.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Extraction.thy Sun Feb 13 17:15:14 2005 +0100
@@ -15,19 +15,19 @@
ML_setup {*
fun realizes_set_proc (Const ("realizes", Type ("fun", [Type ("Null", []), _])) $ r $
(Const ("op :", _) $ x $ S)) = (case strip_comb S of
- (Var (ixn, U), ts) => Some (list_comb (Var (ixn, binder_types U @
+ (Var (ixn, U), ts) => SOME (list_comb (Var (ixn, binder_types U @
[HOLogic.dest_setT (body_type U)] ---> HOLogic.boolT), ts @ [x]))
- | (Free (s, U), ts) => Some (list_comb (Free (s, binder_types U @
+ | (Free (s, U), ts) => SOME (list_comb (Free (s, binder_types U @
[HOLogic.dest_setT (body_type U)] ---> HOLogic.boolT), ts @ [x]))
- | _ => None)
+ | _ => NONE)
| realizes_set_proc (Const ("realizes", Type ("fun", [T, _])) $ r $
(Const ("op :", _) $ x $ S)) = (case strip_comb S of
- (Var (ixn, U), ts) => Some (list_comb (Var (ixn, T :: binder_types U @
+ (Var (ixn, U), ts) => SOME (list_comb (Var (ixn, T :: binder_types U @
[HOLogic.dest_setT (body_type U)] ---> HOLogic.boolT), r :: ts @ [x]))
- | (Free (s, U), ts) => Some (list_comb (Free (s, T :: binder_types U @
+ | (Free (s, U), ts) => SOME (list_comb (Free (s, T :: binder_types U @
[HOLogic.dest_setT (body_type U)] ---> HOLogic.boolT), r :: ts @ [x]))
- | _ => None)
- | realizes_set_proc _ = None;
+ | _ => NONE)
+ | realizes_set_proc _ = NONE;
fun mk_realizes_set r rT s (setT as Type ("set", [elT])) =
Abs ("x", elT, Const ("realizes", rT --> HOLogic.boolT --> HOLogic.boolT) $
@@ -36,8 +36,8 @@
Context.>> (fn thy => thy |>
Extraction.add_types
- [("bool", ([], None)),
- ("set", ([realizes_set_proc], Some mk_realizes_set))] |>
+ [("bool", ([], NONE)),
+ ("set", ([realizes_set_proc], SOME mk_realizes_set))] |>
Extraction.set_preprocessor (fn sg =>
Proofterm.rewrite_proof_notypes
([], ("HOL/elim_cong", RewriteHOLProof.elim_cong) ::
--- a/src/HOL/Fun.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Fun.thy Sun Feb 13 17:15:14 2005 +0100
@@ -474,14 +474,14 @@
(* simplifies terms of the form f(...,x:=y,...,x:=z,...) to f(...,x:=z,...) *)
local
- fun gen_fun_upd None T _ _ = None
- | gen_fun_upd (Some f) T x y = Some (Const ("Fun.fun_upd",T) $ f $ x $ y)
+ fun gen_fun_upd NONE T _ _ = NONE
+ | gen_fun_upd (SOME f) T x y = SOME (Const ("Fun.fun_upd",T) $ f $ x $ y)
fun dest_fun_T1 (Type (_, T :: Ts)) = T
fun find_double (t as Const ("Fun.fun_upd",T) $ f $ x $ y) =
let
fun find (Const ("Fun.fun_upd",T) $ g $ v $ w) =
- if v aconv x then Some g else gen_fun_upd (find g) T v w
- | find t = None
+ if v aconv x then SOME g else gen_fun_upd (find g) T v w
+ | find t = NONE
in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end
val ss = simpset ()
@@ -491,8 +491,8 @@
Simplifier.simproc (Theory.sign_of (the_context ()))
"fun_upd2" ["f(v := w, x := y)"]
(fn sg => fn _ => fn t =>
- case find_double t of (T, None) => None
- | (T, Some rhs) => Some (Tactic.prove sg [] [] (Term.equals T $ t $ rhs) fun_upd_prover))
+ case find_double t of (T, NONE) => NONE
+ | (T, SOME rhs) => SOME (Tactic.prove sg [] [] (Term.equals T $ t $ rhs) fun_upd_prover))
end;
Addsimprocs[fun_upd2_simproc];
--- a/src/HOL/Hoare/hoare.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Hoare/hoare.ML Sun Feb 13 17:15:14 2005 +0100
@@ -89,7 +89,7 @@
(*****************************************************************************)
(** Simplifying: **)
-(** Some useful lemmata, lists and simplification tactics to control which **)
+(** SOME useful lemmata, lists and simplification tactics to control which **)
(** theorems are used to simplify at each moment, so that the original **)
(** input does not suffer any unexpected transformation **)
(*****************************************************************************)
--- a/src/HOL/Hoare/hoareAbort.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Hoare/hoareAbort.ML Sun Feb 13 17:15:14 2005 +0100
@@ -90,7 +90,7 @@
(*****************************************************************************)
(** Simplifying: **)
-(** Some useful lemmata, lists and simplification tactics to control which **)
+(** SOME useful lemmata, lists and simplification tactics to control which **)
(** theorems are used to simplify at each moment, so that the original **)
(** input does not suffer any unexpected transformation **)
(*****************************************************************************)
--- a/src/HOL/HoareParallel/OG_Syntax.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/HoareParallel/OG_Syntax.thy Sun Feb 13 17:15:14 2005 +0100
@@ -96,8 +96,8 @@
fun upd_tr' (x_upd, T) =
(case try (unsuffix RecordPackage.updateN) x_upd of
- Some x => (x, if T = dummyT then T else Term.domain_type T)
- | None => raise Match);
+ SOME x => (x, if T = dummyT then T else Term.domain_type T)
+ | NONE => raise Match);
fun update_name_tr' (Free x) = Free (upd_tr' x)
| update_name_tr' ((c as Const ("_free", _)) $ Free x) =
--- a/src/HOL/HoareParallel/RG_Syntax.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/HoareParallel/RG_Syntax.thy Sun Feb 13 17:15:14 2005 +0100
@@ -69,8 +69,8 @@
fun upd_tr' (x_upd, T) =
(case try (unsuffix RecordPackage.updateN) x_upd of
- Some x => (x, if T = dummyT then T else Term.domain_type T)
- | None => raise Match);
+ SOME x => (x, if T = dummyT then T else Term.domain_type T)
+ | NONE => raise Match);
fun update_name_tr' (Free x) = Free (upd_tr' x)
| update_name_tr' ((c as Const ("_free", _)) $ Free x) =
--- a/src/HOL/Hyperreal/NSA.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Hyperreal/NSA.thy Sun Feb 13 17:15:14 2005 +0100
@@ -606,10 +606,10 @@
0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
fun reorient_proc sg _ (_ $ t $ u) =
case u of
- Const("0", _) => None
- | Const("1", _) => None
- | Const("Numeral.number_of", _) $ _ => None
- | _ => Some (case t of
+ Const("0", _) => NONE
+ | Const("1", _) => NONE
+ | Const("Numeral.number_of", _) $ _ => NONE
+ | _ => SOME (case t of
Const("0", _) => meta_zero_approx_reorient
| Const("1", _) => meta_one_approx_reorient
| Const("Numeral.number_of", _) $ _ =>
--- a/src/HOL/Import/hol4rews.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Import/hol4rews.ML Sun Feb 13 17:15:14 2005 +0100
@@ -149,7 +149,7 @@
val merge : T * T -> T = StringPair.merge (K true)
fun print sg tab =
Pretty.writeln (Pretty.big_list "HOL4 mappings:"
- (StringPair.foldl (fn (xl,((bthy,bthm),isathm)) => (Pretty.str (bthy ^ "." ^ bthm ^ (case isathm of Some th => " --> " ^ th | None => "IGNORED"))::xl)) ([],tab)))
+ (StringPair.foldl (fn (xl,((bthy,bthm),isathm)) => (Pretty.str (bthy ^ "." ^ bthm ^ (case isathm of SOME th => " --> " ^ th | NONE => "IGNORED"))::xl)) ([],tab)))
end;
structure HOL4Maps = TheoryDataFun(HOL4MapsArgs);
@@ -277,7 +277,7 @@
let
val _ = message ("Ignoring " ^ bthy ^ "." ^ bthm)
val curmaps = HOL4Maps.get thy
- val newmaps = StringPair.update_new(((bthy,bthm),None),curmaps)
+ val newmaps = StringPair.update_new(((bthy,bthm),NONE),curmaps)
val upd_thy = HOL4Maps.put newmaps thy
in
upd_thy
@@ -312,8 +312,8 @@
val moves = HOL4Moves.get thy
fun F thmname =
case Symtab.lookup(moves,thmname) of
- Some name => F name
- | None => thmname
+ SOME name => F name
+ | NONE => thmname
in
F thmname
end
@@ -338,8 +338,8 @@
val moves = HOL4CMoves.get thy
fun F thmname =
case Symtab.lookup(moves,thmname) of
- Some name => F name
- | None => thmname
+ SOME name => F name
+ | NONE => thmname
in
F thmname
end
@@ -349,7 +349,7 @@
val isathm = follow_name isathm thy
val _ = message ("Adding theorem map: " ^ bthy ^ "." ^ bthm ^ " --> " ^ isathm)
val curmaps = HOL4Maps.get thy
- val newmaps = StringPair.update_new(((bthy,bthm),Some isathm),curmaps)
+ val newmaps = StringPair.update_new(((bthy,bthm),SOME isathm),curmaps)
val upd_thy = HOL4Maps.put newmaps thy
in
upd_thy
@@ -378,7 +378,7 @@
else thy
val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
val curmaps = HOL4ConstMaps.get thy
- val newmaps = StringPair.update_new(((bthy,bconst),(internal,isaconst,None)),curmaps)
+ val newmaps = StringPair.update_new(((bthy,bconst),(internal,isaconst,NONE)),curmaps)
val upd_thy = HOL4ConstMaps.put newmaps thy
in
upd_thy
@@ -404,8 +404,8 @@
fun get_hol4_const_mapping bthy bconst thy =
let
val bconst = case get_hol4_const_renaming bthy bconst thy of
- Some name => name
- | None => bconst
+ SOME name => name
+ | NONE => bconst
val maps = HOL4ConstMaps.get thy
in
StringPair.lookup(maps,(bthy,bconst))
@@ -420,7 +420,7 @@
else thy
val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
val curmaps = HOL4ConstMaps.get thy
- val newmaps = StringPair.update_new(((bthy,bconst),(internal,isaconst,Some typ)),curmaps)
+ val newmaps = StringPair.update_new(((bthy,bconst),(internal,isaconst,SOME typ)),curmaps)
val upd_thy = HOL4ConstMaps.put newmaps thy
in
upd_thy
@@ -549,7 +549,7 @@
fun should_ignore thyname thy thmname =
case get_hol4_mapping thyname thmname thy of
- Some None => true
+ SOME NONE => true
| _ => false
val trans_string =
@@ -577,8 +577,8 @@
StringPair.foldl (fn ((ign,map),((bthy,bthm),v)) =>
if bthy = thyname
then (case v of
- None => (bthm::ign,map)
- | Some w => (ign,(bthm,w)::map))
+ NONE => (bthm::ign,map)
+ | SOME w => (ign,(bthm,w)::map))
else (ign,map))
(([],[]),HOL4Maps.get thy)
val constmaps =
@@ -650,8 +650,8 @@
val _ = app (fn (hol,(internal,isa,opt_ty)) =>
(out ("\n " ^ (trans_string hol) ^ " > " ^ (trans_string (follow_cname isa thy)));
case opt_ty of
- Some ty => out (" :: \"" ^ (string_of_ctyp (ctyp_of sg ty)) ^ "\"")
- | None => ())) constmaps
+ SOME ty => out (" :: \"" ^ (string_of_ctyp (ctyp_of sg ty)) ^ "\"")
+ | NONE => ())) constmaps
val _ = if null constmaps
then ()
else out "\n\n"
@@ -721,8 +721,8 @@
fun F dname = (dname,add_defmap thyname name dname thy)
in
case StringPair.lookup(maps,(thyname,name)) of
- Some thmname => (thmname,thy)
- | None =>
+ SOME thmname => (thmname,thy)
+ | NONE =>
let
val used = HOL4UNames.get thy
val defname = def_name name
--- a/src/HOL/Import/import_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Import/import_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -14,10 +14,10 @@
struct
val name = "HOL4/import_data"
type T = ProofKernel.thm option array option
-val empty = None
+val empty = NONE
val copy = I
val prep_ext = I
-fun merge _ = None
+fun merge _ = NONE
fun print sg import_segment =
Pretty.writeln (Pretty.str ("Pretty printing of importer data not implemented"))
end
@@ -44,8 +44,8 @@
val prem = hd (prems_of th)
val _ = message ("Import trying to prove " ^ (string_of_cterm (cterm_of sg prem)))
val int_thms = case ImportData.get thy of
- None => fst (Replay.setup_int_thms thyname thy)
- | Some a => a
+ NONE => fst (Replay.setup_int_thms thyname thy)
+ | SOME a => a
val proof = snd (ProofKernel.import_proof thyname thmname thy) thy
val hol4thm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
val thm = snd (ProofKernel.to_isa_thm hol4thm)
@@ -56,7 +56,7 @@
val _ = message ("Import proved " ^ (string_of_thm thm))
in
case Shuffler.set_prop thy prem' [("",thm)] of
- Some (_,thm) =>
+ SOME (_,thm) =>
let
val _ = if prem' aconv (prop_of thm)
then message "import: Terms match up"
@@ -66,7 +66,7 @@
in
res
end
- | None => (message "import: set_prop didn't succeed"; no_tac th)
+ | NONE => (message "import: set_prop didn't succeed"; no_tac th)
end
val import_meth = Method.simple_args (Args.name -- Args.name)
--- a/src/HOL/Import/import_syntax.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Import/import_syntax.ML Sun Feb 13 17:15:14 2005 +0100
@@ -113,8 +113,8 @@
let
val thyname = get_import_thy thy
in
- foldl (fn (thy,((hol,isa),None)) => add_hol4_const_mapping thyname hol false isa thy
- | (thy,((hol,isa),Some ty)) => add_hol4_const_wt_mapping thyname hol false isa (typ_of (read_ctyp (sign_of thy) ty)) thy) (thy,constmaps)
+ foldl (fn (thy,((hol,isa),NONE)) => add_hol4_const_mapping thyname hol false isa thy
+ | (thy,((hol,isa),SOME ty)) => add_hol4_const_wt_mapping thyname hol false isa (typ_of (read_ctyp (sign_of thy) ty)) thy) (thy,constmaps)
end)
val const_moves = Scan.repeat1 (P.name --| P.$$$ ">" -- P.xname -- Scan.option (P.$$$ "::" |-- P.typ))
@@ -123,8 +123,8 @@
let
val thyname = get_import_thy thy
in
- foldl (fn (thy,((hol,isa),None)) => add_hol4_const_mapping thyname hol true isa thy
- | (thy,((hol,isa),Some ty)) => add_hol4_const_wt_mapping thyname hol true isa (typ_of (read_ctyp (sign_of thy) ty)) thy) (thy,constmaps)
+ foldl (fn (thy,((hol,isa),NONE)) => add_hol4_const_mapping thyname hol true isa thy
+ | (thy,((hol,isa),SOME ty)) => add_hol4_const_wt_mapping thyname hol true isa (typ_of (read_ctyp (sign_of thy) ty)) thy) (thy,constmaps)
end)
fun import_thy s =
@@ -160,16 +160,16 @@
then thy
else
case ImportData.get thy of
- None => ImportData.put (Some (fst (Replay.setup_int_thms thyname thy))) thy
- | Some _ => error "Import data not closed - forgotten an end_setup, mayhap?"
+ NONE => ImportData.put (SOME (fst (Replay.setup_int_thms thyname thy))) thy
+ | SOME _ => error "Import data not closed - forgotten an end_setup, mayhap?"
fun clear_int_thms thy =
if !SkipProof.quick_and_dirty
then thy
else
case ImportData.get thy of
- None => error "Import data already cleared - forgotten a setup_theory?"
- | Some _ => ImportData.put None thy
+ NONE => error "Import data already cleared - forgotten a setup_theory?"
+ | SOME _ => ImportData.put NONE thy
val setup_theory = P.name
>>
--- a/src/HOL/Import/proof_kernel.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Import/proof_kernel.ML Sun Feb 13 17:15:14 2005 +0100
@@ -225,11 +225,11 @@
end
fun disk_info_of (Proof(Info{disk_info,...},_)) = !disk_info
-fun mk_proof p = Proof(Info{disk_info = ref None},p)
+fun mk_proof p = Proof(Info{disk_info = ref NONE},p)
fun content_of (Proof(_,p)) = p
fun set_disk_info_of (Proof(Info{disk_info,...},_)) thyname thmname =
- disk_info := Some(thyname,thmname)
+ disk_info := SOME(thyname,thmname)
structure Lib =
struct
@@ -307,11 +307,11 @@
fun F [] toks = (c,toks)
| F (c::cs) toks =
case LazySeq.getItem toks of
- Some(c',toks') =>
+ SOME(c',toks') =>
if c = c'
then F cs toks'
else raise SyntaxError
- | None => raise SyntaxError
+ | NONE => raise SyntaxError
in
F (String.explode str)
end
@@ -326,12 +326,12 @@
in
fun scan_nonquote toks =
case LazySeq.getItem toks of
- Some (c,toks') =>
+ SOME (c,toks') =>
(case c of
#"\"" => raise SyntaxError
| #"&" => scan_entity toks'
| c => (c,toks'))
- | None => raise SyntaxError
+ | NONE => raise SyntaxError
end
val scan_string = $$ #"\"" |-- repeat scan_nonquote --| $$ #"\"" >>
@@ -383,15 +383,15 @@
fun intern_const_name thyname const thy =
case get_hol4_const_mapping thyname const thy of
- Some (_,cname,_) => cname
- | None => (case get_hol4_const_renaming thyname const thy of
- Some cname => Sign.intern_const (sign_of thy) (thyname ^ "." ^ cname)
- | None => Sign.intern_const (sign_of thy) (thyname ^ "." ^ const))
+ SOME (_,cname,_) => cname
+ | NONE => (case get_hol4_const_renaming thyname const thy of
+ SOME cname => Sign.intern_const (sign_of thy) (thyname ^ "." ^ cname)
+ | NONE => Sign.intern_const (sign_of thy) (thyname ^ "." ^ const))
fun intern_type_name thyname const thy =
case get_hol4_type_mapping thyname const thy of
- Some (_,cname) => cname
- | None => Sign.intern_const (sign_of thy) (thyname ^ "." ^ const)
+ SOME (_,cname) => cname
+ | NONE => Sign.intern_const (sign_of thy) (thyname ^ "." ^ const)
fun mk_vartype name = TFree(name,["HOL.type"])
fun mk_thy_type thy Thy Tyop Args = Type(intern_type_name Thy Tyop thy,Args)
@@ -406,8 +406,8 @@
local
fun get_type sg thyname name =
case Sign.const_type sg name of
- Some ty => ty
- | None => raise ERR "get_type" (name ^ ": No such constant")
+ SOME ty => ty
+ | NONE => raise ERR "get_type" (name ^ ": No such constant")
in
fun prim_mk_const thy Thy Name =
let
@@ -415,7 +415,7 @@
val cmaps = HOL4ConstMaps.get thy
in
case StringPair.lookup(cmaps,(Thy,Name)) of
- Some(_,_,Some ty) => Const(name,ty)
+ SOME(_,_,SOME ty) => Const(name,ty)
| _ => Const(name,get_type (sign_of thy) Thy name)
end
end
@@ -511,8 +511,8 @@
| NONE => raise ERR "get_term_from_index" "Bad index"
and get_term_from_xml thy thyname types terms =
let
- fun get_type [] = None
- | get_type [ty] = Some (TypeNet.get_type_from_xml thy thyname types ty)
+ fun get_type [] = NONE
+ | get_type [ty] = SOME (TypeNet.get_type_from_xml thy thyname types ty)
| get_type _ = raise ERR "get_term" "Bad type"
fun gtfx (Elem("tmv",[("n",name),("t",tyi)],[])) =
@@ -566,17 +566,17 @@
let
val import_segment =
case get_segment2 thyname thy of
- Some seg => seg
- | None => get_import_segment thy
+ SOME seg => seg
+ | NONE => get_import_segment thy
val defpath = [OS.Path.joinDirFile {dir=getenv "ISABELLE_HOME_USER",file="proofs"}]
val path = space_explode ":" (getenv "PROOF_DIRS") @ defpath
- fun find [] = None
+ fun find [] = NONE
| find (p::ps) =
(let
val dir = OS.Path.joinDirFile {dir = p,file=import_segment}
in
if OS.FileSys.isDir dir
- then Some dir
+ then SOME dir
else find ps
end) handle OS.SysErr _ => find ps
in
@@ -586,8 +586,8 @@
fun proof_file_name thyname thmname thy =
let
val path = case get_proof_dir thyname thy of
- Some p => p
- | None => error "Cannot find proof files"
+ SOME p => p
+ | NONE => error "Cannot find proof files"
val _ = OS.FileSys.mkDir path handle OS.SysErr _ => ()
in
OS.Path.joinDirFile {dir = path, file = OS.Path.joinBaseExt {base = thmname, ext = SOME "prf"}}
@@ -705,14 +705,14 @@
val p = x2p prf
val tml = map xml_to_term tms
in
- mk_proof (PGenAbs(p,None,tml))
+ mk_proof (PGenAbs(p,NONE,tml))
end
| x2p (Elem("pgenabs",[("i",is)],prf::tms)) =
let
val p = x2p prf
val tml = map xml_to_term tms
in
- mk_proof (PGenAbs(p,Some (index_to_term is),tml))
+ mk_proof (PGenAbs(p,SOME (index_to_term is),tml))
end
| x2p (Elem("pimpas",[],[prf1,prf2])) =
let
@@ -848,8 +848,8 @@
val terms = TermNet.input_terms thyname types xterms
in
(case rest of
- [] => None
- | [xtm] => Some (fn thy => TermNet.get_term_from_xml thy thyname types terms xtm)
+ [] => NONE
+ | [xtm] => SOME (fn thy => TermNet.get_term_from_xml thy thyname types terms xtm)
| _ => raise ERR "import_proof" "Bad argument list",
xml_to_proof thyname types terms prf)
end
@@ -861,10 +861,10 @@
val res = bicompose false (false,th,m) i st
in
case Seq.pull res of
- Some (th,rest) => (case Seq.pull rest of
- Some _ => raise ERR "uniq_compose" "Not unique!"
- | None => th)
- | None => raise ERR "uniq_compose" "No result"
+ SOME (th,rest) => (case Seq.pull rest of
+ SOME _ => raise ERR "uniq_compose" "Not unique!"
+ | NONE => th)
+ | NONE => raise ERR "uniq_compose" "No result"
end
val reflexivity_thm = thm "refl"
@@ -932,7 +932,7 @@
val cvty = ctyp_of_term cv
val th1 = implies_elim_all th
val th2 = beta_eta_thm (forall_intr cv th1)
- val th3 = th2 COMP (beta_eta_thm (Drule.instantiate' [Some cvty] [Some clc] gen_thm))
+ val th3 = th2 COMP (beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME clc] gen_thm))
val c = prop_of th3
val vname = fst(dest_Free v)
val (cold,cnew) = case c of
@@ -978,8 +978,8 @@
let
fun G (ty as TFree _) =
(case try (Lib.assoc ty) tyinst of
- Some ty' => ty'
- | None => ty)
+ SOME ty' => ty'
+ | NONE => ty)
| G (Type(tyname,tys)) = Type(tyname,map G tys)
| G (TVar _) = raise ERR "apply_tyinst_typ" "Scematic variable found"
in
@@ -1004,8 +1004,8 @@
fun F (tm as Bound _) = tm
| F (tm as Free _) =
(case try (Lib.assoc tm) tminst of
- Some tm' => tm'
- | None => tm)
+ SOME tm' => tm'
+ | NONE => tm)
| F (tm as Const _) = tm
| F (tm1 $ tm2) = (F tm1) $ (F tm2)
| F (Abs(vname,ty,body)) = Abs(vname,ty,F body)
@@ -1045,8 +1045,8 @@
if v mem vars
then (vars,rens,inst)
else (case try (Lib.assoc v) rens of
- Some vnew => (vars,rens,(v,vnew)::inst)
- | None => if exists (fn Free(vname',_) => vname = vname' | _ => raise ERR "disamb_helper" "Bad varlist") vars
+ SOME vnew => (vars,rens,(v,vnew)::inst)
+ | NONE => if exists (fn Free(vname',_) => vname = vname' | _ => raise ERR "disamb_helper" "Bad varlist") vars
then
let
val tmnames = map name_of_var varstm
@@ -1130,10 +1130,10 @@
foldr (fn((ren as (vold as Free(vname,_),vnew)),
(rens,inst,vars)) =>
(case Library.find_first (fn Free(v,_) => v = vname | _ => false) vars of
- Some v' => if v' = vold
+ SOME v' => if v' = vold
then (rens,(vnew,vold)::inst,vold::vars)
else (ren::rens,(vold,vnew)::inst,vnew::vars)
- | None => (rens,(vnew,vold)::inst,vold::vars))
+ | NONE => (rens,(vnew,vold)::inst,vold::vars))
| _ => raise ERR "norm_hthm" "Internal error")
(rens,([],[],vars))
val (dom,rng) = ListPair.unzip inst
@@ -1156,14 +1156,14 @@
fun get_hol4_thm thyname thmname thy =
case get_hol4_theorem thyname thmname thy of
- Some hth => Some (HOLThm hth)
- | None =>
+ SOME hth => SOME (HOLThm hth)
+ | NONE =>
let
val pending = HOL4Pending.get thy
in
case StringPair.lookup (pending,(thyname,thmname)) of
- Some hth => Some (HOLThm hth)
- | None => None
+ SOME hth => SOME (HOLThm hth)
+ | NONE => NONE
end
fun non_trivial_term_consts tm =
@@ -1199,10 +1199,10 @@
val (newstr,u) = apboth Substring.string (Substring.splitr (fn c => c = #"_") f)
in
if not (idx = "") andalso u = "_"
- then Some (newstr,valOf(Int.fromString idx))
- else None
+ then SOME (newstr,valOf(Int.fromString idx))
+ else NONE
end
- handle _ => None
+ handle _ => NONE
fun rewrite_hol4_term t thy =
let
@@ -1233,25 +1233,25 @@
fun mk_res th = HOLThm(rens_of info,equal_elim i2h_conc th)
in
case get_hol4_mapping thyname thmname thy of
- Some (Some thmname) =>
+ SOME (SOME thmname) =>
let
val _ = message ("Looking for " ^ thmname)
- val th1 = (Some (transform_error (PureThy.get_thm thy) (thmname, None))
+ val th1 = (SOME (transform_error (PureThy.get_thm thy) (thmname, NONE))
handle ERROR_MESSAGE _ =>
(case split_name thmname of
- Some (listname,idx) => (Some (nth_elem(idx-1,PureThy.get_thms thy (listname, None)))
- handle _ => None)
- | None => None))
+ SOME (listname,idx) => (SOME (nth_elem(idx-1,PureThy.get_thms thy (listname, NONE)))
+ handle _ => NONE)
+ | NONE => NONE))
in
case th1 of
- Some th2 =>
+ SOME th2 =>
(case Shuffler.set_prop thy isaconc [(thmname,th2)] of
- Some (_,th) => (message "YES";(thy, Some (mk_res th)))
- | None => (message "NO2";error "get_isabelle_thm" "Bad mapping"))
- | None => (message "NO1";error "get_isabelle_thm" "Bad mapping")
+ SOME (_,th) => (message "YES";(thy, SOME (mk_res th)))
+ | NONE => (message "NO2";error "get_isabelle_thm" "Bad mapping"))
+ | NONE => (message "NO1";error "get_isabelle_thm" "Bad mapping")
end
- | Some None => error ("Trying to access ignored theorem " ^ thmname)
- | None =>
+ | SOME NONE => error ("Trying to access ignored theorem " ^ thmname)
+ | NONE =>
let
val _ = (message "Looking for conclusion:";
if_debug prin isaconc)
@@ -1262,32 +1262,32 @@
val _ = message ((Int.toString (length pot_thms)) ^ " potential theorems")
in
case Shuffler.set_prop thy isaconc pot_thms of
- Some (isaname,th) =>
+ SOME (isaname,th) =>
let
val hth as HOLThm args = mk_res th
val thy' = thy |> add_hol4_theorem thyname thmname args
|> add_hol4_mapping thyname thmname isaname
in
- (thy',Some hth)
+ (thy',SOME hth)
end
- | None => (thy,None)
+ | NONE => (thy,NONE)
end
end
- handle _ => (thy,None)
+ handle _ => (thy,NONE)
fun get_thm thyname thmname thy =
case get_hol4_thm thyname thmname thy of
- Some hth => (thy,Some hth)
- | None => ((case fst (import_proof thyname thmname thy) of
- Some f => get_isabelle_thm thyname thmname (f thy) thy
- | None => (thy,None))
- handle e as IO.Io _ => (thy,None)
- | e as PK _ => (thy,None))
+ SOME hth => (thy,SOME hth)
+ | NONE => ((case fst (import_proof thyname thmname thy) of
+ SOME f => get_isabelle_thm thyname thmname (f thy) thy
+ | NONE => (thy,NONE))
+ handle e as IO.Io _ => (thy,NONE)
+ | e as PK _ => (thy,NONE))
fun rename_const thyname thy name =
case get_hol4_const_renaming thyname name thy of
- Some cname => cname
- | None => name
+ SOME cname => cname
+ | NONE => name
fun get_def thyname constname rhs thy =
let
@@ -1300,7 +1300,7 @@
fun get_axiom thyname axname thy =
case get_thm thyname axname thy of
- arg as (_,Some _) => arg
+ arg as (_,SOME _) => arg
| _ => raise ERR "get_axiom" ("Trying to retrieve axiom (" ^ axname ^ ")")
fun intern_store_thm gen_output thyname thmname hth thy =
@@ -1325,7 +1325,7 @@
let
val cty = Thm.ctyp_of_term ctm
in
- Drule.instantiate' [Some cty] [Some ctm] reflexivity_thm
+ Drule.instantiate' [SOME cty] [SOME ctm] reflexivity_thm
end
fun REFL tm thy =
@@ -1363,8 +1363,8 @@
val tys_after = add_term_tvars (prop_of th1,[])
val tyinst = map (fn (bef,(i,_)) =>
(case try (Lib.assoc (TFree bef)) lambda of
- Some ty => (i,ctyp_of sg ty)
- | None => (i,ctyp_of sg (TFree bef))
+ SOME ty => (i,ctyp_of sg ty)
+ | NONE => (i,ctyp_of sg (TFree bef))
))
(zip tys_before tys_after)
val res = Drule.instantiate (tyinst,[]) th1
@@ -1433,9 +1433,9 @@
val fty = type_of f
val (fd,fr) = dom_rng fty
val comb_thm' = Drule.instantiate'
- [Some (ctyp_of sg fd),Some (ctyp_of sg fr)]
- [Some (cterm_of sg f),Some (cterm_of sg g),
- Some (cterm_of sg x),Some (cterm_of sg y)] comb_thm
+ [SOME (ctyp_of sg fd),SOME (ctyp_of sg fr)]
+ [SOME (cterm_of sg f),SOME (cterm_of sg g),
+ SOME (cterm_of sg x),SOME (cterm_of sg y)] comb_thm
in
[th1,th2] MRS comb_thm'
end
@@ -1495,7 +1495,7 @@
val (info',tm') = disamb_term_from info tm
val sg = sign_of thy
val ct = Thm.cterm_of sg tm'
- val disj1_thm' = Drule.instantiate' [] [None,Some ct] disj1_thm
+ val disj1_thm' = Drule.instantiate' [] [NONE,SOME ct] disj1_thm
val res = HOLThm(rens_of info',th RS disj1_thm')
val _ = message "RESULT:"
val _ = if_debug pth res
@@ -1512,7 +1512,7 @@
val (info',tm') = disamb_term_from info tm
val sg = sign_of thy
val ct = Thm.cterm_of sg tm'
- val disj2_thm' = Drule.instantiate' [] [None,Some ct] disj2_thm
+ val disj2_thm' = Drule.instantiate' [] [NONE,SOME ct] disj2_thm
val res = HOLThm(rens_of info',th RS disj2_thm')
val _ = message "RESULT:"
val _ = if_debug pth res
@@ -1611,7 +1611,7 @@
(Const("Ex",_) $ a) => a
| _ => raise ERR "EXISTS" "Argument not existential"
val ca = cterm_of sg a
- val exists_thm' = beta_eta_thm (Drule.instantiate' [Some cty] [Some ca,Some cwit] exists_thm)
+ val exists_thm' = beta_eta_thm (Drule.instantiate' [SOME cty] [SOME ca,SOME cwit] exists_thm)
val th1 = beta_eta_thm th
val th2 = implies_elim_all th1
val th3 = th2 COMP exists_thm'
@@ -1645,7 +1645,7 @@
_ $ (Const("Ex",_) $ a) => a
| _ => raise ERR "CHOOSE" "Conclusion not existential"
val ca = cterm_of (sign_of_thm th1) a
- val choose_thm' = beta_eta_thm (Drule.instantiate' [Some cvty] [Some ca,Some cc] choose_thm)
+ val choose_thm' = beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME ca,SOME cc] choose_thm)
val th21 = rearrange sg (HOLogic.mk_Trueprop (a $ v')) th2
val th22 = strip ((nprems_of th21)-1) (cprems_of th21) th21
val th23 = beta_eta_thm (forall_intr cv th22)
@@ -1683,7 +1683,7 @@
val sg = sign_of thy
val ctm = Thm.cterm_of sg tm'
val cty = Thm.ctyp_of_term ctm
- val spec' = Drule.instantiate' [Some cty] [None,Some ctm] spec_thm
+ val spec' = Drule.instantiate' [SOME cty] [NONE,SOME ctm] spec_thm
val th = th RS spec'
val res = HOLThm(rens_of info',th)
val _ = message "RESULT:"
@@ -1732,7 +1732,7 @@
val sg = sign_of thy
val ct = cterm_of sg tm'
val th1 = rearrange sg (HOLogic.mk_Trueprop (Const("Not",boolT-->boolT) $ tm')) th
- val ccontr_thm' = Drule.instantiate' [] [Some ct] ccontr_thm
+ val ccontr_thm' = Drule.instantiate' [] [SOME ct] ccontr_thm
val res1 = uniq_compose ((nprems_of th1) - 1) th1 1 ccontr_thm'
val res = HOLThm(rens_of info',res1)
val _ = message "RESULT:"
@@ -1749,7 +1749,7 @@
_ $ (Const("op =",_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
| _ => raise ERR "mk_ABS" "Bad conclusion"
val (fd,fr) = dom_rng (type_of f)
- val abs_thm' = Drule.instantiate' [Some (ctyp_of sg fd), Some (ctyp_of sg fr)] [Some (cterm_of sg f), Some (cterm_of sg g)] abs_thm
+ val abs_thm' = Drule.instantiate' [SOME (ctyp_of sg fd), SOME (ctyp_of sg fr)] [SOME (cterm_of sg f), SOME (cterm_of sg g)] abs_thm
val th2 = forall_intr cv th1
val th3 = th2 COMP abs_thm'
val res = implies_intr_hyps th3
@@ -1776,8 +1776,8 @@
let
val _ = message "GEN_ABS:"
val _ = case copt of
- Some c => if_debug prin c
- | None => ()
+ SOME c => if_debug prin c
+ | NONE => ()
val _ = if_debug (app prin) vlist
val _ = if_debug pth hth
val (info,th) = disamb_thm hth
@@ -1785,7 +1785,7 @@
val sg = sign_of thy
val th1 =
case copt of
- Some (c as Const(cname,cty)) =>
+ SOME (c as Const(cname,cty)) =>
let
fun inst_type ty1 ty2 (TVar _) = raise ERR "GEN_ABS" "Type variable found!"
| inst_type ty1 ty2 (ty as TFree _) = if ty1 = ty
@@ -1804,8 +1804,8 @@
mk_COMB (mk_REFL cc) (mk_ABS v th sg) sg
end) (vlist',th)
end
- | Some _ => raise ERR "GEN_ABS" "Bad constant"
- | None =>
+ | SOME _ => raise ERR "GEN_ABS" "Bad constant"
+ | NONE =>
foldr (fn (v,th) => mk_ABS v th sg) (vlist',th)
val res = HOLThm(rens_of info',th1)
val _ = message "RESULT:"
@@ -1824,7 +1824,7 @@
_ $ (Const("op -->",_) $ a $ Const("False",_)) => a
| _ => raise ERR "NOT_INTRO" "Conclusion of bad form"
val ca = cterm_of sg a
- val th2 = equal_elim (Drule.instantiate' [] [Some ca] not_intro_thm) th1
+ val th2 = equal_elim (Drule.instantiate' [] [SOME ca] not_intro_thm) th1
val res = HOLThm(rens,implies_intr_hyps th2)
val _ = message "RESULT:"
val _ = if_debug pth res
@@ -1842,7 +1842,7 @@
_ $ (Const("Not",_) $ a) => a
| _ => raise ERR "NOT_ELIM" "Conclusion of bad form"
val ca = cterm_of sg a
- val th2 = equal_elim (Drule.instantiate' [] [Some ca] not_elim_thm) th1
+ val th2 = equal_elim (Drule.instantiate' [] [SOME ca] not_elim_thm) th1
val res = HOLThm(rens,implies_intr_hyps th2)
val _ = message "RESULT:"
val _ = if_debug pth res
@@ -1905,8 +1905,8 @@
thy''
val hth = case Shuffler.set_prop thy22 (HOLogic.mk_Trueprop tm24) [("",th)] of
- Some (_,res) => HOLThm(rens_of linfo,res)
- | None => raise ERR "new_definition" "Bad conclusion"
+ SOME (_,res) => HOLThm(rens_of linfo,res)
+ | NONE => raise ERR "new_definition" "Bad conclusion"
val fullname = Sign.full_name sg thmname
val thy22' = case opt_get_output_thy thy22 of
"" => add_hol4_mapping thyname thmname fullname thy22
@@ -1973,7 +1973,7 @@
val thy1 = foldr (fn(name,thy)=>
snd (get_defname thyname name thy)) (names,thy1)
fun new_name name = fst (get_defname thyname name thy1)
- val (thy',res) = SpecificationPackage.add_specification_i None
+ val (thy',res) = SpecificationPackage.add_specification_i NONE
(map (fn name => (new_name name,name,false)) names)
(thy1,th)
val res' = Drule.freeze_all res
@@ -2041,7 +2041,7 @@
val tnames = map fst tfrees
val tsyn = mk_syn thy tycname
val typ = (tycname,tnames,tsyn)
- val (thy',typedef_info) = TypedefPackage.add_typedef_i false (Some thmname) typ c None (rtac th2 1) thy
+ val (thy',typedef_info) = TypedefPackage.add_typedef_i false (SOME thmname) typ c NONE (rtac th2 1) thy
val th3 = (#type_definition typedef_info) RS typedef_hol2hol4
@@ -2093,9 +2093,9 @@
val sg = sign_of thy
val tT = type_of t
val light_nonempty' =
- Drule.instantiate' [Some (ctyp_of sg tT)]
- [Some (cterm_of sg P),
- Some (cterm_of sg t)] light_nonempty
+ Drule.instantiate' [SOME (ctyp_of sg tT)]
+ [SOME (cterm_of sg P),
+ SOME (cterm_of sg t)] light_nonempty
val th2 = beta_eta_thm (td_th RS (beta_eta_thm light_nonempty'))
val c = case concl_of th2 of
_ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
@@ -2104,7 +2104,7 @@
val tnames = map fst tfrees
val tsyn = mk_syn thy tycname
val typ = (tycname,tnames,tsyn)
- val (thy',typedef_info) = TypedefPackage.add_typedef_i false (Some thmname) typ c (Some(rep_name,abs_name)) (rtac th2 1) thy
+ val (thy',typedef_info) = TypedefPackage.add_typedef_i false (SOME thmname) typ c (SOME(rep_name,abs_name)) (rtac th2 1) thy
val th3 = (#type_definition typedef_info) RS typedef_hol2hollight
--- a/src/HOL/Import/replay.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Import/replay.ML Sun Feb 13 17:15:14 2005 +0100
@@ -57,8 +57,8 @@
| rp (PHyp tm) thy = P.ASSUME tm thy
| rp (PDef(seg,name,rhs)) thy =
(case P.get_def seg name rhs thy of
- (thy',Some res) => (thy',res)
- | (thy',None) =>
+ (thy',SOME res) => (thy',res)
+ | (thy',NONE) =>
if seg = thyname
then P.new_definition seg name rhs thy'
else raise ERR "replay_proof" "Too late for term definition")
@@ -208,25 +208,25 @@
in
case pc of
PDisk => (case disk_info_of p of
- Some(thyname',thmname) =>
+ SOME(thyname',thmname) =>
(case Int.fromString thmname of
SOME i =>
if thyname' = thyname
then
(case Array.sub(int_thms,i-1) of
- None =>
+ NONE =>
let
val (thy',th) = rp' (snd (import_proof thyname' thmname thy) thy) thy
- val _ = Array.update(int_thms,i-1,Some th)
+ val _ = Array.update(int_thms,i-1,SOME th)
in
(thy',th)
end
- | Some th => (thy,th))
+ | SOME th => (thy,th))
else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")
| NONE =>
(case P.get_thm thyname' thmname thy of
- (thy',Some res) => (thy',res)
- | (thy',None) =>
+ (thy',SOME res) => (thy',res)
+ | (thy',NONE) =>
if thyname' = thyname
then
let
@@ -242,11 +242,11 @@
| _ => P.store_thm thyname' thmname th thy'
end
else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")))
- | None => raise ERR "rp'.PDisk" "Not enough information")
+ | NONE => raise ERR "rp'.PDisk" "Not enough information")
| PAxm(name,c) =>
(case P.get_axiom thyname name thy of
- (thy',Some res) => (thy',res)
- | (thy',None) => P.new_axiom name c thy')
+ (thy',SOME res) => (thy',res)
+ | (thy',NONE) => P.new_axiom name c thy')
| PTmSpec(seg,names,prf') =>
let
val (thy',th) = rp' prf' thy
@@ -275,8 +275,8 @@
let
val fname =
case P.get_proof_dir thyname thy of
- Some p => OS.Path.joinDirFile {dir=p,file=OS.Path.joinBaseExt{base = "facts",ext=SOME "lst"}}
- | None => error "Cannot find proof files"
+ SOME p => OS.Path.joinDirFile {dir=p,file=OS.Path.joinBaseExt{base = "facts",ext=SOME "lst"}}
+ | NONE => error "Cannot find proof files"
val is = TextIO.openIn fname
val (num_int_thms,facts) =
let
@@ -290,7 +290,7 @@
get_facts []
end
val _ = TextIO.closeIn is
- val int_thms = Array.array(num_int_thms,None:thm option)
+ val int_thms = Array.array(num_int_thms,NONE:thm option)
in
(int_thms,facts)
end
--- a/src/HOL/Import/shuffler.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Import/shuffler.ML Sun Feb 13 17:15:14 2005 +0100
@@ -231,7 +231,7 @@
val res = (find_bound 0 body;2) handle RESULT i => i
in
case res of
- 0 => Some (rew_th sg (x,xT) (y,yT) body)
+ 0 => SOME (rew_th sg (x,xT) (y,yT) body)
| 1 => if string_ord(y,x) = LESS
then
let
@@ -239,12 +239,12 @@
val t_th = reflexive (cterm_of sg t)
val newt_th = reflexive (cterm_of sg newt)
in
- Some (transitive t_th newt_th)
+ SOME (transitive t_th newt_th)
end
- else None
+ else NONE
| _ => error "norm_term (quant_rewrite) internal error"
end
- | quant_rewrite _ _ _ = (warning "quant_rewrite: Unknown lhs"; None)
+ | quant_rewrite _ _ _ = (warning "quant_rewrite: Unknown lhs"; NONE)
fun freeze_thaw_term t =
let
@@ -304,7 +304,7 @@
val lhs = #1 (Logic.dest_equals (prop_of (final init)))
in
if not (lhs aconv origt)
- then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)";
+ then (writeln "SOMEthing is utterly wrong: (orig,lhs,frozen type,t,tinst)";
writeln (string_of_cterm (cterm_of sg origt));
writeln (string_of_cterm (cterm_of sg lhs));
writeln (string_of_cterm (cterm_of sg typet));
@@ -338,15 +338,15 @@
val res = final rew_th
val lhs = (#1 (Logic.dest_equals (prop_of res)))
in
- Some res
+ SOME res
end
- else None)
+ else NONE)
handle e => (writeln "eta_contract:";print_exn e))
- | _ => (error ("Bad eta_contract argument" ^ (string_of_cterm (cterm_of sg t))); None)
+ | _ => (error ("Bad eta_contract argument" ^ (string_of_cterm (cterm_of sg t))); NONE)
end
fun beta_fun sg assume t =
- Some (beta_conversion true (cterm_of sg t))
+ SOME (beta_conversion true (cterm_of sg t))
fun eta_expand sg assumes origt =
let
@@ -359,7 +359,7 @@
val lhs = #1 (Logic.dest_equals (prop_of (final init)))
in
if not (lhs aconv origt)
- then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)";
+ then (writeln "SOMEthing is utterly wrong: (orig,lhs,frozen type,t,tinst)";
writeln (string_of_cterm (cterm_of sg origt));
writeln (string_of_cterm (cterm_of sg lhs));
writeln (string_of_cterm (cterm_of sg typet));
@@ -396,11 +396,11 @@
val res = equal_intr th1 (th2' |> implies_intr concl)
val res' = final res
in
- Some res'
+ SOME res'
end
- | _ => None)
- else None
- | _ => (error ("Bad eta_expand argument" ^ (string_of_cterm (cterm_of sg t))); None)
+ | _ => NONE)
+ else NONE
+ | _ => (error ("Bad eta_expand argument" ^ (string_of_cterm (cterm_of sg t))); NONE)
end
handle e => (writeln "eta_expand internal error";print_exn e)
@@ -519,19 +519,19 @@
end
(* make_equal sg t u tries to construct the theorem t == u under the
-signature sg. If it succeeds, Some (t == u) is returned, otherwise
-None is returned. *)
+signature sg. If it succeeds, SOME (t == u) is returned, otherwise
+NONE is returned. *)
fun make_equal sg t u =
let
val t_is_t' = norm_term sg t
val u_is_u' = norm_term sg u
val th = transitive t_is_t' (symmetric u_is_u')
- val _ = message ("make_equal: Some " ^ (string_of_thm th))
+ val _ = message ("make_equal: SOME " ^ (string_of_thm th))
in
- Some th
+ SOME th
end
- handle e as THM _ => (message "make_equal: None";None)
+ handle e as THM _ => (message "make_equal: NONE";NONE)
fun match_consts ignore t (* th *) =
let
@@ -564,7 +564,7 @@
(* set_prop t thms tries to make a theorem with the proposition t from
one of the theorems thms, by shuffling the propositions around. If it
-succeeds, Some theorem is returned, otherwise None. *)
+succeeds, SOME theorem is returned, otherwise NONE. *)
fun set_prop thy t =
let
@@ -576,34 +576,34 @@
val rhs = snd (dest_equals (cprop_of rew_th))
val shuffles = ShuffleData.get thy
- fun process [] = None
+ fun process [] = NONE
| process ((name,th)::thms) =
let
val norm_th = varifyT (norm_thm thy (close_thm (transfer_sg sg th)))
val triv_th = trivial rhs
val _ = message ("Shuffler.set_prop: Gluing together " ^ (string_of_thm norm_th) ^ " and " ^ (string_of_thm triv_th))
val mod_th = case Seq.pull (bicompose true (false,norm_th,0) 1 triv_th) of
- Some(th,_) => Some th
- | None => None
+ SOME(th,_) => SOME th
+ | NONE => NONE
in
case mod_th of
- Some mod_th =>
+ SOME mod_th =>
let
val closed_th = equal_elim (symmetric rew_th) mod_th
in
message ("Shuffler.set_prop succeeded by " ^ name);
- Some (name,forall_elim_list (map (cterm_of sg) vars) closed_th)
+ SOME (name,forall_elim_list (map (cterm_of sg) vars) closed_th)
end
- | None => process thms
+ | NONE => process thms
end
handle e as THM _ => process thms
in
fn thms =>
case process thms of
- res as Some (name,th) => if (prop_of th) aconv t
+ res as SOME (name,th) => if (prop_of th) aconv t
then res
else error "Internal error in set_prop"
- | None => None
+ | NONE => NONE
end
handle e => (writeln "set_prop internal error"; print_exn e)
@@ -624,8 +624,8 @@
val set = set_prop thy t
fun process_tac thms st =
case set thms of
- Some (_,th) => Seq.of_list (compose (th,i,st))
- | None => Seq.empty
+ SOME (_,th) => Seq.of_list (compose (th,i,st))
+ | NONE => Seq.empty
in
(process_tac thms APPEND (if search
then process_tac (find_potential thy t)
--- a/src/HOL/Integ/NatBin.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Integ/NatBin.thy Sun Feb 13 17:15:14 2005 +0100
@@ -859,14 +859,14 @@
ML {*
fun number_of_codegen thy gr s b (Const ("Numeral.number_of",
Type ("fun", [_, Type ("IntDef.int", [])])) $ bin) =
- (Some (gr, Pretty.str (string_of_int (HOLogic.dest_binum bin)))
- handle TERM _ => None)
+ (SOME (gr, Pretty.str (string_of_int (HOLogic.dest_binum bin)))
+ handle TERM _ => NONE)
| number_of_codegen thy gr s b (Const ("Numeral.number_of",
Type ("fun", [_, Type ("nat", [])])) $ bin) =
- Some (Codegen.invoke_codegen thy s b (gr,
+ SOME (Codegen.invoke_codegen thy s b (gr,
Const ("IntDef.nat", HOLogic.intT --> HOLogic.natT) $
(Const ("Numeral.number_of", HOLogic.binT --> HOLogic.intT) $ bin)))
- | number_of_codegen _ _ _ _ _ = None;
+ | number_of_codegen _ _ _ _ _ = NONE;
*}
setup {* [Codegen.add_codegen "number_of_codegen" number_of_codegen] *}
--- a/src/HOL/Integ/cooper_dec.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Integ/cooper_dec.ML Sun Feb 13 17:15:14 2005 +0100
@@ -92,7 +92,7 @@
fun dest_numeral (Const("0",Type ("IntDef.int", []))) = 0
|dest_numeral (Const("1",Type ("IntDef.int", []))) = 1
|dest_numeral (Const ("Numeral.number_of",_) $ n)= HOLogic.dest_binum n;
-(*Some terms often used for pattern matching*)
+(*SOME terms often used for pattern matching*)
val zero = mk_numeral 0;
val one = mk_numeral 1;
@@ -450,7 +450,7 @@
[("op =",op=), ("op <",op<), ("op >",op>), ("op <=",op<=) , ("op >=",op>=),
("Divides.op dvd",fn (x,y) =>((y mod x) = 0))];
-fun applyoperation (Some f) (a,b) = f (a, b)
+fun applyoperation (SOME f) (a,b) = f (a, b)
|applyoperation _ (_, _) = false;
(*Evaluation of constant atomic formulas*)
@@ -464,12 +464,12 @@
handle _ => at)
else
case assoc (operations,p) of
- Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)
+ SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)
handle _ => at)
| _ => at)
|Const("Not",_)$(Const (p,_) $ s $ t) =>(
case assoc (operations,p) of
- Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then
+ SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then
HOLogic.false_const else HOLogic.true_const)
handle _ => at)
| _ => at)
@@ -480,12 +480,12 @@
fun evalc_atom at = case at of
(Const (p,_) $ s $ t) =>
( case assoc (operations,p) of
- Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)
+ SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)
handle _ => at)
| _ => at)
|Const("Not",_)$(Const (p,_) $ s $ t) =>(
case assoc (operations,p) of
- Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then
+ SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then
HOLogic.false_const else HOLogic.true_const)
handle _ => at)
| _ => at)
@@ -711,14 +711,14 @@
fun inf_w mi d vars x p =
let val f = if mi then minusinf else plusinf in
case (simpl (minusinf x p)) of
- Const("True",_) => (Some (mk_numeral 1), HOLogic.true_const)
- |Const("False",_) => (None,HOLogic.false_const)
+ Const("True",_) => (SOME (mk_numeral 1), HOLogic.true_const)
+ |Const("False",_) => (NONE,HOLogic.false_const)
|F =>
let
fun h n =
case ((simpl o evalc) (linrep vars x (mk_numeral n) F)) of
- Const("True",_) => (Some (mk_numeral n),HOLogic.true_const)
- |F' => if n=1 then (None,F')
+ Const("True",_) => (SOME (mk_numeral n),HOLogic.true_const)
+ |F' => if n=1 then (NONE,F')
else let val (rw,rf) = h (n-1) in
(rw,HOLogic.mk_disj(F',rf))
end
@@ -729,9 +729,9 @@
fun set_w d b st vars x p = let
fun h ns = case ns of
- [] => (None,HOLogic.false_const)
+ [] => (NONE,HOLogic.false_const)
|n::nl => ( case ((simpl o evalc) (linrep vars x n p)) of
- Const("True",_) => (Some n,HOLogic.true_const)
+ Const("True",_) => (SOME n,HOLogic.true_const)
|F' => let val (rw,rf) = h nl
in (rw,HOLogic.mk_disj(F',rf))
end)
@@ -741,10 +741,10 @@
end;
fun withness d b st vars x p = case (inf_w b d vars x p) of
- (Some n,_) => (Some n,HOLogic.true_const)
- |(None,Pinf) => (case (set_w d b st vars x p) of
- (Some n,_) => (Some n,HOLogic.true_const)
- |(_,Pst) => (None,HOLogic.mk_disj(Pinf,Pst)));
+ (SOME n,_) => (SOME n,HOLogic.true_const)
+ |(NONE,Pinf) => (case (set_w d b st vars x p) of
+ (SOME n,_) => (SOME n,HOLogic.true_const)
+ |(_,Pst) => (NONE,HOLogic.mk_disj(Pinf,Pst)));
--- a/src/HOL/Integ/cooper_proof.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Integ/cooper_proof.ML Sun Feb 13 17:15:14 2005 +0100
@@ -194,8 +194,8 @@
fun simple_prove_goal_cterm2 G tacs =
let
- fun check None = error "prove_goal: tactic failed"
- | check (Some (thm, _)) = (case nprems_of thm of
+ fun check NONE = error "prove_goal: tactic failed"
+ | check (SOME (thm, _)) = (case nprems_of thm of
0 => thm
| i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
in check (Seq.pull (EVERY tacs (trivial G))) end;
@@ -320,7 +320,7 @@
val cx = cterm_of sg y
val pre = prove_elementar sg "lf"
(HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral n)))
- val th1 = (pre RS (instantiate' [] [Some ck,Some cc, Some cx, Some ct] (ac_pi_eq)))
+ val th1 = (pre RS (instantiate' [] [SOME ck,SOME cc, SOME cx, SOME ct] (ac_pi_eq)))
in ([], fn [] => [th1,(prove_elementar sg "sa" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
end
@@ -344,32 +344,32 @@
"op <" =>
let val pre = prove_elementar sg "lf"
(HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral k)))
- val th1 = (pre RS (instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct] (ac_lt_eq)))
+ val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_lt_eq)))
in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
end
|"op =" =>
let val pre = prove_elementar sg "lf"
(HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
- val th1 = (pre RS(instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct] (ac_eq_eq)))
+ val th1 = (pre RS(instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_eq_eq)))
in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
end
|"Divides.op dvd" =>
let val pre = prove_elementar sg "lf"
(HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
- val th1 = (pre RS (instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct]) (ac_dvd_eq))
+ val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct]) (ac_dvd_eq))
in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
end
end
- else ([], fn [] => instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl)
+ else ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl)
|( Const ("Not", _) $ p) => ([p], fn [th] => th RS qe_Not)
|( Const ("op &",_) $ p $ q) => ([p,q], fn [th1,th2] => [th1,th2] MRS qe_conjI)
|( Const ("op |",_) $ p $ q) =>([p,q], fn [th1,th2] => [th1,th2] MRS qe_disjI)
- |_ => ([], fn [] => instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl);
+ |_ => ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl);
fun proof_of_adjustcoeffeq sg x l = thm_of sg (decomp_adjustcoeffeq sg x l);
@@ -380,43 +380,43 @@
(*==================================================*)
fun rho_for_modd_minf x dlcm sg fm1 =
let
- (*Some certified Terms*)
+ (*SOME certified Terms*)
val ctrue = cterm_of sg HOLogic.true_const
val cfalse = cterm_of sg HOLogic.false_const
val fm = norm_zero_one fm1
in case fm1 of
(Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) =>
- if (x=y) andalso (c1= zero) andalso (c2= one) then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+ if (x=y) andalso (c1= zero) andalso (c2= one) then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one)
- then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+ then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if (y=x) andalso (c1 = zero) then
- if (pm1 = one) then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_minf)) else
- (instantiate' [Some cboolT] [Some ctrue] (fm_modd_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+ if (pm1 = one) then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_minf)) else
+ (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cz = cterm_of sg (norm_zero_one z)
val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
- in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_minf)))
+ in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_minf)))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
|(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
c $ y ) $ z))) =>
if y=x then let val cz = cterm_of sg (norm_zero_one z)
val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
- in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_minf)))
+ in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_minf)))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
- |_ => instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf)
+ |_ => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)
end;
(*=========================================================================*)
(*=========================================================================*)
@@ -426,36 +426,36 @@
in case fm1 of
(Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) =>
if (x=y) andalso (c1=zero) andalso (c2=one)
- then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (neq_eq_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (neq_eq_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one))
- then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (eq_eq_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (eq_eq_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if (y=x) andalso (c1 =zero) then
- if pm1 = one then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (le_eq_minf)) else
- (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (len_eq_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ if pm1 = one then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (le_eq_minf)) else
+ (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (len_eq_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cd = cterm_of sg (norm_zero_one d)
val cz = cterm_of sg (norm_zero_one z)
- in(instantiate' [] [Some cd, Some cz] (not_dvd_eq_minf))
+ in(instantiate' [] [SOME cd, SOME cz] (not_dvd_eq_minf))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
|(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cd = cterm_of sg (norm_zero_one d)
val cz = cterm_of sg (norm_zero_one z)
- in(instantiate' [] [Some cd, Some cz ] (dvd_eq_minf))
+ in(instantiate' [] [SOME cd, SOME cz ] (dvd_eq_minf))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
- |_ => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ |_ => (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
end;
(*=====================================================*)
@@ -476,7 +476,7 @@
(* -------------------------------------------------------------*)
fun rho_for_modd_pinf x dlcm sg fm1 =
let
- (*Some certified Terms*)
+ (*SOME certified Terms*)
val ctrue = cterm_of sg HOLogic.true_const
val cfalse = cterm_of sg HOLogic.false_const
@@ -484,37 +484,37 @@
in case fm1 of
(Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) =>
if ((x=y) andalso (c1= zero) andalso (c2= one))
- then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+ then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
if ((is_arith_rel fm) andalso (x = y) andalso (c1 = zero) andalso (c2 = one))
- then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+ then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if ((y=x) andalso (c1 = zero)) then
if (pm1 = one)
- then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_pinf))
- else (instantiate' [Some cboolT] [Some cfalse] (fm_modd_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+ then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cz = cterm_of sg (norm_zero_one z)
val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
- in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_pinf)))
+ in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_pinf)))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
|(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
c $ y ) $ z))) =>
if y=x then let val cz = cterm_of sg (norm_zero_one z)
val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
- in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_pinf)))
+ in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_pinf)))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
- |_ => instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf)
+ |_ => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)
end;
(* -------------------------------------------------------------*)
(* Finding rho for pinf_eq *)
@@ -525,36 +525,36 @@
in case fm1 of
(Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) =>
if (x=y) andalso (c1=zero) andalso (c2=one)
- then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (neq_eq_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (neq_eq_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one))
- then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (eq_eq_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (eq_eq_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if (y=x) andalso (c1 =zero) then
- if pm1 = one then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (le_eq_pinf)) else
- (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (len_eq_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ if pm1 = one then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (le_eq_pinf)) else
+ (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (len_eq_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cd = cterm_of sg (norm_zero_one d)
val cz = cterm_of sg (norm_zero_one z)
- in(instantiate' [] [Some cd, Some cz] (not_dvd_eq_pinf))
+ in(instantiate' [] [SOME cd, SOME cz] (not_dvd_eq_pinf))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
|(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cd = cterm_of sg (norm_zero_one d)
val cz = cterm_of sg (norm_zero_one z)
- in(instantiate' [] [Some cd, Some cz ] (dvd_eq_pinf))
+ in(instantiate' [] [SOME cd, SOME cz ] (dvd_eq_pinf))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
- |_ => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ |_ => (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
end;
@@ -593,7 +593,7 @@
val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
val cdlcm = cterm_of sg dlcm
val cB = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one bs))
- in instantiate' [] [Some cdlcm,Some cB, Some cp] (bst_thm)
+ in instantiate' [] [SOME cdlcm,SOME cB, SOME cp] (bst_thm)
end;
fun asetproof_of sg (x as Free(xn,xT)) fm ast dlcm =
@@ -601,7 +601,7 @@
val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
val cdlcm = cterm_of sg dlcm
val cA = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one ast))
- in instantiate' [] [Some cdlcm,Some cA, Some cp] (ast_thm)
+ in instantiate' [] [SOME cdlcm,SOME cA, SOME cp] (ast_thm)
end;
*)
@@ -623,9 +623,9 @@
then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B)
val th2 = prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)))
val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma]([th3,th1,th2] MRS (not_bst_p_ne)))
+ in (instantiate' [] [SOME cfma]([th3,th1,th2] MRS (not_bst_p_ne)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
if (is_arith_rel at) andalso (x=y)
@@ -633,40 +633,40 @@
in let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ bst_z $ B)
val th2 = prove_elementar sg "ss" (HOLogic.mk_eq (bst_z,Const("op -",T) $ (Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma] ([th3,th1,th2] MRS (not_bst_p_eq)))
+ in (instantiate' [] [SOME cfma] ([th3,th1,th2] MRS (not_bst_p_eq)))
end
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if (y=x) andalso (c1 =zero) then
if pm1 = one then
let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B)
val th2 = prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)))
- in (instantiate' [] [Some cfma, Some cdlcm]([th1,th2] MRS (not_bst_p_gt)))
+ in (instantiate' [] [SOME cfma, SOME cdlcm]([th1,th2] MRS (not_bst_p_gt)))
end
else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma, Some cB,Some (cterm_of sg (norm_zero_one z))] (th1 RS (not_bst_p_lt)))
+ in (instantiate' [] [SOME cfma, SOME cB,SOME (cterm_of sg (norm_zero_one z))] (th1 RS (not_bst_p_lt)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then
let val cz = cterm_of sg (norm_zero_one z)
val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
- in (instantiate' [] [Some cfma, Some cB,Some cz] (th1 RS (not_bst_p_ndvd)))
+ in (instantiate' [] [SOME cfma, SOME cB,SOME cz] (th1 RS (not_bst_p_ndvd)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
|(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then
let val cz = cterm_of sg (norm_zero_one z)
val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
- in (instantiate' [] [Some cfma,Some cB,Some cz] (th1 RS (not_bst_p_dvd)))
+ in (instantiate' [] [SOME cfma,SOME cB,SOME cz] (th1 RS (not_bst_p_dvd)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
- |_ => (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ |_ => (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
end;
@@ -714,9 +714,9 @@
then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ A)
val th2 = prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)))
val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma]([th3,th1,th2] MRS (not_ast_p_ne)))
+ in (instantiate' [] [SOME cfma]([th3,th1,th2] MRS (not_ast_p_ne)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
if (is_arith_rel at) andalso (x=y)
@@ -724,39 +724,39 @@
val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ ast_z $ A)
val th2 = prove_elementar sg "ss" (HOLogic.mk_eq (ast_z,Const("op +",T) $ (Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma] ([th3,th1,th2] MRS (not_ast_p_eq)))
+ in (instantiate' [] [SOME cfma] ([th3,th1,th2] MRS (not_ast_p_eq)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if (y=x) andalso (c1 =zero) then
if pm1 = (mk_numeral ~1) then
let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one z) $ A)
val th2 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm))
- in (instantiate' [] [Some cfma]([th2,th1] MRS (not_ast_p_lt)))
+ in (instantiate' [] [SOME cfma]([th2,th1] MRS (not_ast_p_lt)))
end
else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma, Some cA,Some (cterm_of sg (norm_zero_one z))] (th1 RS (not_ast_p_gt)))
+ in (instantiate' [] [SOME cfma, SOME cA,SOME (cterm_of sg (norm_zero_one z))] (th1 RS (not_ast_p_gt)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then
let val cz = cterm_of sg (norm_zero_one z)
val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
- in (instantiate' [] [Some cfma, Some cA,Some cz] (th1 RS (not_ast_p_ndvd)))
+ in (instantiate' [] [SOME cfma, SOME cA,SOME cz] (th1 RS (not_ast_p_ndvd)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
|(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then
let val cz = cterm_of sg (norm_zero_one z)
val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
- in (instantiate' [] [Some cfma,Some cA,Some cz] (th1 RS (not_ast_p_dvd)))
+ in (instantiate' [] [SOME cfma,SOME cA,SOME cz] (th1 RS (not_ast_p_dvd)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
- |_ => (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ |_ => (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
end;
@@ -788,21 +788,21 @@
fun rho_for_evalc sg at = case at of
(Const (p,_) $ s $ t) =>(
case assoc (operations,p) of
- Some f =>
+ SOME f =>
((if (f ((dest_numeral s),(dest_numeral t)))
then prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const))
else prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const)))
- handle _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl)
- | _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl )
+ handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl)
+ | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl )
|Const("Not",_)$(Const (p,_) $ s $ t) =>(
case assoc (operations,p) of
- Some f =>
+ SOME f =>
((if (f ((dest_numeral s),(dest_numeral t)))
then prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const))
else prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const)))
- handle _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl)
- | _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl )
- | _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl;
+ handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl)
+ | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl )
+ | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl;
(*=========================================================*)
@@ -828,7 +828,7 @@
|(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI)
|(Const("Not",_)$p) => ([p],fn [th] => th RS qe_Not)
|(Const("Divides.op dvd",_)$d$r) =>
- if is_numeral d then ([], fn [] => (prove_elementar sg "lf" (HOLogic.mk_eq (r, lint vars r))) RS (instantiate' [] [None , None, Some (cterm_of sg d)](linearize_dvd)))
+ if is_numeral d then ([], fn [] => (prove_elementar sg "lf" (HOLogic.mk_eq (r, lint vars r))) RS (instantiate' [] [NONE , NONE, SOME (cterm_of sg d)](linearize_dvd)))
else (warning "Nonlinear Term --- Non numeral leftside at dvd";
raise COOPER)
|_ => ([], fn [] => prove_elementar sg "lf" (HOLogic.mk_eq (t, linform vars t)));
@@ -897,7 +897,7 @@
val ss = presburger_ss addsimps
[simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
(* uth : EX x.P(l*x) = EX x. l dvd x & P x*)
- val uth = instantiate' [] [Some (cterm_of sg P) , Some (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
+ val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
(* e_ac_thm : Ex x. efm = EX x. fm*)
val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
(* A and B set of the formula*)
@@ -961,7 +961,7 @@
val ss = presburger_ss addsimps
[simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
(* uth : EX x.P(l*x) = EX x. l dvd x & P x*)
- val uth = instantiate' [] [Some (cterm_of sg P) , Some (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
+ val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
(* e_ac_thm : Ex x. efm = EX x. fm*)
val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
(* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*)
@@ -972,7 +972,7 @@
val (w,rs) = cooper_w [] cfm
val exp_cp_thm = case w of
(* FIXME - e_ac_thm just tipped to test syntactical correctness of the program!!!!*)
- Some n => e_ac_thm (* Prove cfm (n) and use exI and then Eq_TrueI*)
+ SOME n => e_ac_thm (* Prove cfm (n) and use exI and then Eq_TrueI*)
|_ => let
(* A and B set of the formula*)
val A = aset x cfm
--- a/src/HOL/Integ/int_arith1.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Integ/int_arith1.ML Sun Feb 13 17:15:14 2005 +0100
@@ -93,10 +93,10 @@
structure Bin_Simprocs =
struct
fun prove_conv tacs sg (hyps: thm list) xs (t, u) =
- if t aconv u then None
+ if t aconv u then NONE
else
let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u))
- in Some (Tactic.prove sg xs [] eq (K (EVERY tacs))) end
+ in SOME (Tactic.prove sg xs [] eq (K (EVERY tacs))) end
fun prove_conv_nohyps tacs sg = prove_conv tacs sg [];
fun prove_conv_nohyps_novars tacs sg = prove_conv tacs sg [] [];
@@ -119,10 +119,10 @@
0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
fun reorient_proc sg _ (_ $ t $ u) =
case u of
- Const("0", _) => None
- | Const("1", _) => None
- | Const("Numeral.number_of", _) $ _ => None
- | _ => Some (case t of
+ Const("0", _) => NONE
+ | Const("1", _) => NONE
+ | Const("Numeral.number_of", _) $ _ => NONE
+ | _ => SOME (case t of
Const("0", _) => meta_zero_reorient
| Const("1", _) => meta_one_reorient
| Const("Numeral.number_of", _) $ _ => meta_number_of_reorient)
@@ -305,8 +305,8 @@
[mult_assoc, minus_mult_left, minus_mult_eq_1_to_2];
(*Apply the given rewrite (if present) just once*)
-fun trans_tac None = all_tac
- | trans_tac (Some th) = ALLGOALS (rtac (th RS trans));
+fun trans_tac NONE = all_tac
+ | trans_tac (SOME th) = ALLGOALS (rtac (th RS trans));
fun simplify_meta_eq rules =
simplify (HOL_basic_ss addeqcongs[eq_cong2] addsimps rules)
@@ -538,7 +538,7 @@
Addsimprocs [fast_int_arith_simproc]
-(* Some test data
+(* SOME test data
Goal "!!a::int. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
by (fast_arith_tac 1);
Goal "!!a::int. [| a < b; c < d |] ==> a-d+ 2 <= b+(-c)";
--- a/src/HOL/Integ/presburger.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Integ/presburger.ML Sun Feb 13 17:15:14 2005 +0100
@@ -73,7 +73,7 @@
fun term_typed_consts t = add_term_typed_consts(t,[]);
-(* Some Types*)
+(* SOME Types*)
val bT = HOLogic.boolT;
val iT = HOLogic.intT;
val binT = HOLogic.binT;
@@ -246,7 +246,7 @@
val g' = if a then abstract_pres sg g else g
(* Transform the term*)
val (t,np,nh) = prepare_for_presburger sg q g'
- (* Some simpsets for dealing with mod div abs and nat*)
+ (* SOME simpsets for dealing with mod div abs and nat*)
val simpset0 = HOL_basic_ss
addsimps [mod_div_equality', Suc_plus1]
addsplits [split_zdiv, split_zmod, split_div', split_min, split_max]
@@ -313,7 +313,7 @@
"decision procedure for Presburger arithmetic"),
ArithTheoryData.map (fn {splits, inj_consts, discrete, presburger} =>
{splits = splits, inj_consts = inj_consts, discrete = discrete,
- presburger = Some (presburger_tac true false)})];
+ presburger = SOME (presburger_tac true false)})];
end;
--- a/src/HOL/Integ/qelim.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Integ/qelim.ML Sun Feb 13 17:15:14 2005 +0100
@@ -51,7 +51,7 @@
end
|(Const("All",_)$Abs(xn,xT,p)) => ([(HOLogic.exists_const xT)$Abs(xn,xT,HOLogic.Not $ p)], fn [th] => th RS qe_ALL)
- | _ => ([],fn [] => instantiate' [Some (ctyp_of sg (type_of P))] [Some (cterm_of sg P)] refl);
+ | _ => ([],fn [] => instantiate' [SOME (ctyp_of sg (type_of P))] [SOME (cterm_of sg P)] refl);
fun tproof_of_mlift_qelim sg isat afnp nfnp qfnp p =
--- a/src/HOL/IsaMakefile Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/IsaMakefile Sun Feb 13 17:15:14 2005 +0100
@@ -82,7 +82,7 @@
$(SRC)/TFL/usyntax.ML $(SRC)/TFL/utils.ML \
$(SRC)/Provers/eqsubst.ML\
eqrule_HOL_data.ML\
- Datatype.thy Datatype_Universe.thy \
+ Datatype.ML Datatype.thy Datatype_Universe.thy \
Divides.thy Equiv_Relations.thy Extraction.thy Finite_Set.ML Finite_Set.thy \
Fun.thy Gfp.thy Hilbert_Choice.thy HOL.ML \
HOL.thy Inductive.thy Infinite_Set.thy Integ/Numeral.thy \
--- a/src/HOL/Isar_examples/Hoare.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Isar_examples/Hoare.thy Sun Feb 13 17:15:14 2005 +0100
@@ -262,8 +262,8 @@
fun upd_tr' (x_upd, T) =
(case try (unsuffix RecordPackage.updateN) x_upd of
- Some x => (x, if T = dummyT then T else Term.domain_type T)
- | None => raise Match);
+ SOME x => (x, if T = dummyT then T else Term.domain_type T)
+ | NONE => raise Match);
fun update_name_tr' (Free x) = Free (upd_tr' x)
| update_name_tr' ((c as Const ("_free", _)) $ Free x) =
--- a/src/HOL/Library/EfficientNat.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Library/EfficientNat.thy Sun Feb 13 17:15:14 2005 +0100
@@ -117,8 +117,8 @@
val eqs1 = flat (map
(fn th => map (pair th) (find_vars (lhs_of th))) thms);
val eqs2 = map (fn x as (_, ((_, ctzero), _)) => (x, mapfilter (fn th =>
- Some (th, Thm.cterm_match (lhs_of th, ctzero))
- handle Pattern.MATCH => None) thms)) eqs1;
+ SOME (th, Thm.cterm_match (lhs_of th, ctzero))
+ handle Pattern.MATCH => NONE) thms)) eqs1;
fun distinct_vars vs = forall is_Var vs andalso null (duplicates vs);
val eqs2' = map (apsnd (filter (fn (_, (_, s)) =>
distinct_vars (map (term_of o snd) s)))) eqs2;
@@ -129,21 +129,21 @@
Thm.implies_elim (Thm.implies_elim
(Drule.fconv_rule (Thm.beta_conversion true)
(Drule.instantiate'
- [Some (ctyp_of_term ct)] [Some (Thm.cabs cv ct),
- Some (rhs_of th''), Some (Thm.cabs cv' (rhs_of th)), Some cv']
+ [SOME (ctyp_of_term ct)] [SOME (Thm.cabs cv ct),
+ SOME (rhs_of th''), SOME (Thm.cabs cv' (rhs_of th)), SOME cv']
Suc_if_eq')) th'') (Thm.forall_intr cv' th)
in
mapfilter (fn thm =>
- if thm = th then Some th'''
- else if b andalso thm = th' then None
- else Some thm) thms
+ if thm = th then SOME th'''
+ else if b andalso thm = th' then NONE
+ else SOME thm) thms
end
in
case Library.find_first (not o null o snd) eqs2' of
- None => (case Library.find_first (not o null o snd) eqs2 of
- None => thms
- | Some x => remove_suc thy (mk_thms false x))
- | Some x => remove_suc thy (mk_thms true x)
+ NONE => (case Library.find_first (not o null o snd) eqs2 of
+ NONE => thms
+ | SOME x => remove_suc thy (mk_thms false x))
+ | SOME x => remove_suc thy (mk_thms true x)
end;
fun eqn_suc_preproc thy ths =
@@ -161,25 +161,25 @@
val Suc_clause' = Thm.transfer thy Suc_clause;
val vname = variant (map fst
(foldl add_term_varnames ([], map prop_of thms))) "x";
- fun find_var (t as Const ("Suc", _) $ (v as Var _)) = Some (t, v)
- | find_var (t $ u) = (case find_var t of None => find_var u | x => x)
- | find_var _ = None;
+ fun find_var (t as Const ("Suc", _) $ (v as Var _)) = SOME (t, v)
+ | find_var (t $ u) = (case find_var t of NONE => find_var u | x => x)
+ | find_var _ = NONE;
fun find_thm th =
let val th' = ObjectLogic.atomize_thm th
in apsome (pair (th, th')) (find_var (prop_of th')) end
in
case get_first find_thm thms of
- None => thms
- | Some ((th, th'), (Sucv, v)) =>
+ NONE => thms
+ | SOME ((th, th'), (Sucv, v)) =>
let
val cert = cterm_of (sign_of_thm th);
val th'' = ObjectLogic.rulify (Thm.implies_elim
(Drule.fconv_rule (Thm.beta_conversion true)
(Drule.instantiate' []
- [Some (cert (lambda v (Abs ("x", HOLogic.natT,
+ [SOME (cert (lambda v (Abs ("x", HOLogic.natT,
abstract_over (Sucv,
HOLogic.dest_Trueprop (prop_of th')))))),
- Some (cert v)] Suc_clause'))
+ SOME (cert v)] Suc_clause'))
(Thm.forall_intr (cert v) th'))
in
remove_suc_clause thy (map (fn th''' =>
--- a/src/HOL/Library/word_setup.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Library/word_setup.ML Sun Feb 13 17:15:14 2005 +0100
@@ -24,17 +24,17 @@
fun add_word thy =
let
(*lcp** val fast1_th = PureThy.get_thm thy "Word.fast_nat_to_bv_def"**)
- val fast2_th = PureThy.get_thm thy ("Word.fast_bv_to_nat_def", None)
+ val fast2_th = PureThy.get_thm thy ("Word.fast_bv_to_nat_def", NONE)
(*lcp** fun f sg thms (Const("Word.nat_to_bv",_) $ (Const("Numeral.number_of",_) $ t)) =
if num_is_usable t
- then Some (Drule.cterm_instantiate [(cterm_of sg (Var(("w",0),Type("Numeral.bin",[]))),cterm_of sg t)] fast1_th)
- else None
- | f _ _ _ = None **)
+ then SOME (Drule.cterm_instantiate [(cterm_of sg (Var(("w",0),Type("Numeral.bin",[]))),cterm_of sg t)] fast1_th)
+ else NONE
+ | f _ _ _ = NONE **)
fun g sg thms (Const("Word.bv_to_nat",_) $ (t as (Const("List.list.Cons",_) $ _ $ _))) =
if vec_is_usable t
- then Some (Drule.cterm_instantiate [(cterm_of sg (Var(("bs",0),Type("List.list",[Type("Word.bit",[])]))),cterm_of sg t)] fast2_th)
- else None
- | g _ _ _ = None
+ then SOME (Drule.cterm_instantiate [(cterm_of sg (Var(("bs",0),Type("List.list",[Type("Word.bit",[])]))),cterm_of sg t)] fast2_th)
+ else NONE
+ | g _ _ _ = NONE
(*lcp** val simproc1 = Simplifier.simproc (sign_of thy) "nat_to_bv" ["Word.nat_to_bv (number_of ?w)"] f ***)
val simproc2 = Simplifier.simproc (sign_of thy) "bv_to_nat" ["Word.bv_to_nat (?x # ?xs)"] g
in
--- a/src/HOL/List.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/List.thy Sun Feb 13 17:15:14 2005 +0100
@@ -430,12 +430,12 @@
val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr)
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2));
val thm = Tactic.prove sg [] [] eq (K (rearr_tac 1));
- in Some ((conv RS (thm RS trans)) RS eq_reflection) end;
+ in SOME ((conv RS (thm RS trans)) RS eq_reflection) end;
in
if list1 lastl andalso list1 lastr then rearr append1_eq_conv
else if lastl aconv lastr then rearr append_same_eq
- else None
+ else NONE
end;
in
@@ -2118,17 +2118,17 @@
fun list_codegen thy gr dep b t =
let val (gr', ps) = foldl_map (Codegen.invoke_codegen thy dep false)
(gr, HOLogic.dest_list t)
- in Some (gr', Pretty.list "[" "]" ps) end handle TERM _ => None;
+ in SOME (gr', Pretty.list "[" "]" ps) end handle TERM _ => NONE;
fun dest_nibble (Const (s, _)) = int_of_nibble (unprefix "List.nibble.Nibble" s)
| dest_nibble _ = raise Match;
fun char_codegen thy gr dep b (Const ("List.char.Char", _) $ c1 $ c2) =
(let val c = chr (dest_nibble c1 * 16 + dest_nibble c2)
- in if Symbol.is_printable c then Some (gr, Pretty.quote (Pretty.str c))
- else None
- end handle LIST _ => None | Match => None)
- | char_codegen thy gr dep b _ = None;
+ in if Symbol.is_printable c then SOME (gr, Pretty.quote (Pretty.str c))
+ else NONE
+ end handle LIST _ => NONE | Match => NONE)
+ | char_codegen thy gr dep b _ = NONE;
in
--- a/src/HOL/Main.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Main.thy Sun Feb 13 17:15:14 2005 +0100
@@ -42,18 +42,18 @@
val eq_codegen_setup = [Codegen.add_codegen "eq_codegen"
(fn thy => fn gr => fn dep => fn b => fn t =>
(case strip_comb t of
- (Const ("op =", Type (_, [Type ("fun", _), _])), _) => None
+ (Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE
| (Const ("op =", _), [t, u]) =>
let
val (gr', pt) = Codegen.invoke_codegen thy dep false (gr, t);
val (gr'', pu) = Codegen.invoke_codegen thy dep false (gr', u)
in
- Some (gr'', Codegen.parens
+ SOME (gr'', Codegen.parens
(Pretty.block [pt, Pretty.str " =", Pretty.brk 1, pu]))
end
- | (t as Const ("op =", _), ts) => Some (Codegen.invoke_codegen
+ | (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen
thy dep b (gr, Codegen.eta_expand t ts 2))
- | _ => None))];
+ | _ => NONE))];
fun gen_bool i = one_of [false, true];
--- a/src/HOL/Matrix/Cplex.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Matrix/Cplex.ML Sun Feb 13 17:15:14 2005 +0100
@@ -81,8 +81,8 @@
| rev_cmp cplexGeq = cplexLeq
| rev_cmp cplexEq = cplexEq
-fun the None = raise (Load_cplexFile "Some expected")
- | the (Some x) = x;
+fun the NONE = raise (Load_cplexFile "SOME expected")
+ | the (SOME x) = x;
fun modulo_signed is_something (cplexNeg u) = is_something u
| modulo_signed is_something u = is_something u
@@ -189,25 +189,25 @@
val a = to_upper x
in
if a = "BOUNDS" orelse a = "BOUND" then
- Some "BOUNDS"
+ SOME "BOUNDS"
else if a = "MINIMIZE" orelse a = "MINIMUM" orelse a = "MIN" then
- Some "MINIMIZE"
+ SOME "MINIMIZE"
else if a = "MAXIMIZE" orelse a = "MAXIMUM" orelse a = "MAX" then
- Some "MAXIMIZE"
+ SOME "MAXIMIZE"
else if a = "ST" orelse a = "S.T." orelse a = "ST." then
- Some "ST"
+ SOME "ST"
else if a = "FREE" orelse a = "END" then
- Some a
+ SOME a
else if a = "GENERAL" orelse a = "GENERALS" orelse a = "GEN" then
- Some "GENERAL"
+ SOME "GENERAL"
else if a = "INTEGER" orelse a = "INTEGERS" orelse a = "INT" then
- Some "INTEGER"
+ SOME "INTEGER"
else if a = "BINARY" orelse a = "BINARIES" orelse a = "BIN" then
- Some "BINARY"
+ SOME "BINARY"
else if a = "INF" orelse a = "INFINITY" then
- Some "INF"
+ SOME "INF"
else
- None
+ NONE
end
val TOKEN_ERROR = ~1
@@ -303,12 +303,12 @@
let val u = hd (!rest) in
(
rest := tl (!rest);
- Some u
+ SOME u
)
end
else
let val s = TextIO.inputLine f in
- if s = "" then None else
+ if s = "" then NONE else
let val t = tokenize s in
if (length t >= 2 andalso
snd(hd (tl t)) = ":")
@@ -326,12 +326,12 @@
fun readToken_helper2 () =
let val c = readToken_helper () in
- if c = None then None
+ if c = NONE then NONE
else if !ignore_NL andalso fst (the c) = TOKEN_NL then
readToken_helper2 ()
else if fst (the c) = TOKEN_SYMBOL
- andalso keyword (snd (the c)) <> None
- then Some (TOKEN_KEYWORD, the (keyword (snd (the c))))
+ andalso keyword (snd (the c)) <> NONE
+ then SOME (TOKEN_KEYWORD, the (keyword (snd (the c))))
else c
end
@@ -354,78 +354,78 @@
fun readTerm_Product only_num =
let val c = readToken () in
- if c = None then None
+ if c = NONE then NONE
else if fst (the c) = TOKEN_SYMBOL
then (
- if only_num then (pushToken (the c); None)
- else Some (cplexVar (snd (the c)))
+ if only_num then (pushToken (the c); NONE)
+ else SOME (cplexVar (snd (the c)))
)
else if only_num andalso is_value (the c) then
- Some (get_value (the c))
+ SOME (get_value (the c))
else if is_value (the c) then
let val t1 = get_value (the c)
val d = readToken ()
in
- if d = None then Some t1
+ if d = NONE then SOME t1
else if fst (the d) = TOKEN_SYMBOL then
- Some (cplexProd (t1, cplexVar (snd (the d))))
+ SOME (cplexProd (t1, cplexVar (snd (the d))))
else
- (pushToken (the d); Some t1)
+ (pushToken (the d); SOME t1)
end
- else (pushToken (the c); None)
+ else (pushToken (the c); NONE)
end
fun readTerm_Signed only_signed only_num =
let
val c = readToken ()
in
- if c = None then None
+ if c = NONE then NONE
else
let val d = the c in
if d = (TOKEN_DELIMITER, "+") then
readTerm_Product only_num
else if d = (TOKEN_DELIMITER, "-") then
- Some (cplexNeg (the (readTerm_Product
+ SOME (cplexNeg (the (readTerm_Product
only_num)))
else (pushToken d;
- if only_signed then None
+ if only_signed then NONE
else readTerm_Product only_num)
end
end
fun readTerm_Sum first_signed =
let val c = readTerm_Signed first_signed false in
- if c = None then [] else (the c)::(readTerm_Sum true)
+ if c = NONE then [] else (the c)::(readTerm_Sum true)
end
fun readTerm () =
let val c = readTerm_Sum false in
- if c = [] then None
- else if tl c = [] then Some (hd c)
- else Some (cplexSum c)
+ if c = [] then NONE
+ else if tl c = [] then SOME (hd c)
+ else SOME (cplexSum c)
end
fun readLabeledTerm () =
let val c = readToken () in
- if c = None then (None, None)
+ if c = NONE then (NONE, NONE)
else if fst (the c) = TOKEN_LABEL then
let val t = readTerm () in
- if t = None then
+ if t = NONE then
raise (Load_cplexFile ("term after label "^
(snd (the c))^
" expected"))
- else (Some (snd (the c)), t)
+ else (SOME (snd (the c)), t)
end
- else (pushToken (the c); (None, readTerm ()))
+ else (pushToken (the c); (NONE, readTerm ()))
end
fun readGoal () =
let
val g = readToken ()
in
- if g = Some (TOKEN_KEYWORD, "MAXIMIZE") then
+ if g = SOME (TOKEN_KEYWORD, "MAXIMIZE") then
cplexMaximize (the (snd (readLabeledTerm ())))
- else if g = Some (TOKEN_KEYWORD, "MINIMIZE") then
+ else if g = SOME (TOKEN_KEYWORD, "MINIMIZE") then
cplexMinimize (the (snd (readLabeledTerm ())))
else raise (Load_cplexFile "MAXIMIZE or MINIMIZE expected")
end
@@ -447,17 +447,17 @@
(str2cmp b,
(t1, t2))
in
- if snd t = None then None
+ if snd t = NONE then NONE
else
let val c = readToken () in
- if c = None orelse fst (the c) <> TOKEN_CMP
+ if c = NONE orelse fst (the c) <> TOKEN_CMP
then raise (Load_cplexFile "TOKEN_CMP expected")
else
let val n = readTerm_Signed false true in
- if n = None then
+ if n = NONE then
raise (Load_cplexFile "num expected")
else
- Some (fst t,
+ SOME (fst t,
make_constraint (snd (the c))
(the (snd t))
(the n))
@@ -469,10 +469,10 @@
let
fun readbody () =
let val t = readConstraint () in
- if t = None then []
+ if t = NONE then []
else if (is_normed_Constr (snd (the t))) then
(the t)::(readbody ())
- else if (fst (the t) <> None) then
+ else if (fst (the t) <> NONE) then
raise (Load_cplexFile
("constraint '"^(the (fst (the t)))^
"'is not normed"))
@@ -481,7 +481,7 @@
"constraint is not normed")
end
in
- if readToken () = Some (TOKEN_KEYWORD, "ST")
+ if readToken () = SOME (TOKEN_KEYWORD, "ST")
then
readbody ()
else
@@ -490,15 +490,15 @@
fun readCmp () =
let val c = readToken () in
- if c = None then None
+ if c = NONE then NONE
else if fst (the c) = TOKEN_CMP then
- Some (str2cmp (snd (the c)))
- else (pushToken (the c); None)
+ SOME (str2cmp (snd (the c)))
+ else (pushToken (the c); NONE)
end
fun skip_NL () =
let val c = readToken () in
- if c <> None andalso fst (the c) = TOKEN_NL then
+ if c <> NONE andalso fst (the c) = TOKEN_NL then
skip_NL ()
else
(pushToken (the c); ())
@@ -515,17 +515,17 @@
val _ = skip_NL ()
val t1 = readTerm ()
in
- if t1 = None then None
+ if t1 = NONE then NONE
else
let
val c1 = readCmp ()
in
- if c1 = None then
+ if c1 = NONE then
let
val c = readToken ()
in
- if c = Some (TOKEN_KEYWORD, "FREE") then
- Some (
+ if c = SOME (TOKEN_KEYWORD, "FREE") then
+ SOME (
cplexBounds (cplexNeg cplexInf,
cplexLeq,
the t1,
@@ -538,16 +538,16 @@
let
val t2 = readTerm ()
in
- if t2 = None then
+ if t2 = NONE then
raise (Load_cplexFile "term expected")
else
let val c2 = readCmp () in
- if c2 = None then
- Some (make_bounds (the c1)
+ if c2 = NONE then
+ SOME (make_bounds (the c1)
(the t1)
(the t2))
else
- Some (
+ SOME (
cplexBounds (the t1,
the c1,
the t2,
@@ -565,7 +565,7 @@
let
val b = readBound ()
in
- if b = None then []
+ if b = NONE then []
else if (is_normed_Bounds (the b)) then
(the b)::(readbody())
else (
@@ -574,13 +574,13 @@
(makestring (the b)))))
end
in
- if readToken () = Some (TOKEN_KEYWORD, "BOUNDS") then
+ if readToken () = SOME (TOKEN_KEYWORD, "BOUNDS") then
readbody ()
else raise (Load_cplexFile "BOUNDS expected")
end
fun readEnd () =
- if readToken () = Some (TOKEN_KEYWORD, "END") then ()
+ if readToken () = SOME (TOKEN_KEYWORD, "END") then ()
else raise (Load_cplexFile "END expected")
val result_Goal = readGoal ()
@@ -648,7 +648,7 @@
fun write_constraints [] = ()
| write_constraints (c::cs) =
- (if (fst c <> None)
+ (if (fst c <> NONE)
then
(write (the (fst c)); write ": ")
else
@@ -767,7 +767,7 @@
val zero = cplexNum "0"
fun make_pos_constr v =
- (None, cplexConstr (cplexGeq, ((cplexVar v), zero)))
+ (NONE, cplexConstr (cplexGeq, ((cplexVar v), zero)))
fun make_free_bound v =
cplexBounds (cplexNeg cplexInf, cplexLeq,
@@ -777,7 +777,7 @@
val pos_constrs = rev(Symtab.foldl
(fn (l, (k,v)) => (make_pos_constr k) :: l)
([], positive_vars))
- val bound_constrs = map (fn x => (None, x))
+ val bound_constrs = map (fn x => (NONE, x))
(flat (map bound2constr bounds))
val constraints' = constraints @ pos_constrs @ bound_constrs
val bounds' = rev(Symtab.foldl (fn (l, (v,_)) =>
@@ -859,17 +859,17 @@
let val u = hd (!rest) in
(
rest := tl (!rest);
- Some u
+ SOME u
)
end
else
let val s = TextIO.inputLine f in
- if s = "" then None else (rest := tokenize s; readToken_helper())
+ if s = "" then NONE else (rest := tokenize s; readToken_helper())
end
- fun is_tt tok ty = (tok <> None andalso (fst (the tok)) = ty)
+ fun is_tt tok ty = (tok <> NONE andalso (fst (the tok)) = ty)
- fun pushToken a = if a = None then () else (rest := ((the a)::(!rest)))
+ fun pushToken a = if a = NONE then () else (rest := ((the a)::(!rest)))
fun readToken () =
let val t = readToken_helper () in
@@ -878,14 +878,14 @@
else if is_tt t TOKEN_NL then
let val t2 = readToken_helper () in
if is_tt t2 TOKEN_SIGN then
- (pushToken (Some (TOKEN_SEPARATOR, "-")); t)
+ (pushToken (SOME (TOKEN_SEPARATOR, "-")); t)
else
(pushToken t2; t)
end
else if is_tt t TOKEN_SIGN then
let val t2 = readToken_helper () in
if is_tt t2 TOKEN_NUM then
- (Some (TOKEN_NUM, (snd (the t))^(snd (the t2))))
+ (SOME (TOKEN_NUM, (snd (the t))^(snd (the t2))))
else
(pushToken t2; t)
end
@@ -897,7 +897,7 @@
let
val t = readToken ()
in
- if is_tt t TOKEN_NL orelse t = None
+ if is_tt t TOKEN_NL orelse t = NONE
then P
else readRestOfLine P
end
@@ -940,26 +940,26 @@
val k = readToken ()
in
if is_tt k TOKEN_NUM then
- readRestOfLine (Some (snd (the t2), snd (the k)))
+ readRestOfLine (SOME (snd (the t2), snd (the k)))
else
raise (Load_cplexResult "number expected")
end
else
- (pushToken t2; pushToken t1; None)
+ (pushToken t2; pushToken t1; NONE)
end
fun load_values () =
let val v = load_value () in
- if v = None then [] else (the v)::(load_values ())
+ if v = NONE then [] else (the v)::(load_values ())
end
val header = readHeader ()
val result =
case assoc (header, "STATUS") of
- Some "INFEASIBLE" => Infeasible
- | Some "UNBOUNDED" => Unbounded
- | Some "OPTIMAL" => Optimal (the (assoc (header, "OBJECTIVE")),
+ SOME "INFEASIBLE" => Infeasible
+ | SOME "UNBOUNDED" => Unbounded
+ | SOME "OPTIMAL" => Optimal (the (assoc (header, "OBJECTIVE")),
(skip_until_sep ();
skip_until_sep ();
load_values ()))
--- a/src/HOL/Matrix/CplexMatrixConverter.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Matrix/CplexMatrixConverter.ML Sun Feb 13 17:15:14 2005 +0100
@@ -59,10 +59,10 @@
val (varcount, i2s_tab, s2i_tab) = build_naming 0 Inttab.empty Symtab.empty bounds
- fun i2s i = case Inttab.lookup (i2s_tab, i) of None => raise (Converter "index not found")
- | Some n => n
- fun s2i s = case Symtab.lookup (s2i_tab, s) of None => raise (Converter ("name not found: "^s))
- | Some i => i
+ fun i2s i = case Inttab.lookup (i2s_tab, i) of NONE => raise (Converter "index not found")
+ | SOME n => n
+ fun s2i s = case Symtab.lookup (s2i_tab, s) of NONE => raise (Converter ("name not found: "^s))
+ | SOME i => i
fun num2str positive (cplexNeg t) = num2str (not positive) t
| num2str positive (cplexNum num) = if positive then num else "-"^num
| num2str _ _ = raise (Converter "term is not a (possibly signed) number")
@@ -82,8 +82,8 @@
| constrs2Ab j A b ((_, cplexConstr (cplexGeq, (t1,t2)))::cs) =
constrs2Ab (j+1) (set_vector A j (sum2vec (neg_term t1))) (set_elem b j (num2str true (neg_term t2))) cs
| constrs2Ab j A b ((_, cplexConstr (cplexEq, (t1,t2)))::cs) =
- constrs2Ab j A b ((None, cplexConstr (cplexLeq, (t1,t2)))::
- (None, cplexConstr (cplexGeq, (t1, t2)))::cs)
+ constrs2Ab j A b ((NONE, cplexConstr (cplexLeq, (t1,t2)))::
+ (NONE, cplexConstr (cplexGeq, (t1, t2)))::cs)
| constrs2Ab _ _ _ _ = raise (Converter "no strict constraints allowed")
val (A, b) = constrs2Ab 0 empty_matrix empty_vector constrs
--- a/src/HOL/Matrix/FloatSparseMatrixBuilder.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Matrix/FloatSparseMatrixBuilder.ML Sun Feb 13 17:15:14 2005 +0100
@@ -226,8 +226,8 @@
let
fun upd m j i x =
case Inttab.lookup (m, j) of
- Some v => Inttab.update ((j, Inttab.update ((i, x), v)), m)
- | None => Inttab.update ((j, Inttab.update ((i, x), Inttab.empty)), m)
+ SOME v => Inttab.update ((j, Inttab.update ((i, x), v)), m)
+ | NONE => Inttab.update ((j, Inttab.update ((i, x), Inttab.empty)), m)
fun updv j (m, (i, s)) = upd m i j s
@@ -274,11 +274,11 @@
fun mk_constr index vector c =
let
- val s = case Inttab.lookup (c, index) of Some s => s | None => "0"
+ val s = case Inttab.lookup (c, index) of SOME s => s | NONE => "0"
val (p, s) = split_numstr s
val num = if p then cplex.cplexNum s else cplex.cplexNeg (cplex.cplexNum s)
in
- (None, cplex.cplexConstr (cplex.cplexLeq, (vec2sum vector, num)))
+ (NONE, cplex.cplexConstr (cplex.cplexLeq, (vec2sum vector, num)))
end
fun delete index c = Inttab.delete index c handle Inttab.UNDEF _ => c
@@ -327,11 +327,11 @@
fun mk_constr index vector c =
let
- val s = case Inttab.lookup (c, index) of Some s => s | None => "0"
+ val s = case Inttab.lookup (c, index) of SOME s => s | NONE => "0"
val (p, s) = split_numstr s
val num = if p then cplex.cplexNum s else cplex.cplexNeg (cplex.cplexNum s)
in
- (None, cplex.cplexConstr (cplex.cplexEq, (vec2sum vector, num)))
+ (NONE, cplex.cplexConstr (cplex.cplexEq, (vec2sum vector, num)))
end
fun delete index c = Inttab.delete index c handle Inttab.UNDEF _ => c
@@ -361,12 +361,12 @@
fun cut_matrix vfilter vsize m =
let
fun app (m, (i, v)) =
- if (Inttab.lookup (vfilter, i) = None) then
+ if (Inttab.lookup (vfilter, i) = NONE) then
m
else
case vsize of
- None => Inttab.update ((i,v), m)
- | Some s => Inttab.update((i, cut_vector s v),m)
+ NONE => Inttab.update ((i,v), m)
+ | SOME s => Inttab.update((i, cut_vector s v),m)
in
Inttab.foldl app (empty_matrix, m)
end
@@ -376,10 +376,10 @@
fun v_only_elem v =
case Inttab.min_key v of
- None => None
- | Some vmin => (case Inttab.max_key v of
- None => Some vmin
- | Some vmax => if vmin = vmax then Some vmin else None)
+ NONE => NONE
+ | SOME vmin => (case Inttab.max_key v of
+ NONE => SOME vmin
+ | SOME vmax => if vmin = vmax then SOME vmin else NONE)
fun v_fold f a v = Inttab.foldl f (a,v)
--- a/src/HOL/Matrix/codegen_prep.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Matrix/codegen_prep.ML Sun Feb 13 17:15:14 2005 +0100
@@ -74,8 +74,8 @@
val key = f x
in
case assoc (g, key) of
- None => (key, [x])::g
- | Some l => overwrite (g, (key, x::l))
+ NONE => (key, [x])::g
+ | SOME l => overwrite (g, (key, x::l))
end
fun prepare_thms ths =
--- a/src/HOL/Matrix/eq_codegen.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Matrix/eq_codegen.ML Sun Feb 13 17:15:14 2005 +0100
@@ -12,8 +12,8 @@
val (_, eq) = Thm.dest_comb (cprop_of thm);
val (ct, rhs) = Thm.dest_comb eq;
val (_, lhs) = Thm.dest_comb ct
- in Thm.implies_elim (Drule.instantiate' [Some (ctyp_of_term lhs)]
- [Some lhs, Some rhs] eq_reflection) thm
+ in Thm.implies_elim (Drule.instantiate' [SOME (ctyp_of_term lhs)]
+ [SOME lhs, SOME rhs] eq_reflection) thm
end;
structure SimprocsCodegen =
@@ -51,7 +51,7 @@
decomp_term_code cn ((v1 :: v2 :: vs,
bs @ [(Free (s, T), v1)],
ps @ [mk_val [v1, v2] [Pretty.str "Thm.dest_abs", Pretty.brk 1,
- Pretty.str "None", Pretty.brk 1, Pretty.str v]]), (v2, t))
+ Pretty.str "NONE", Pretty.brk 1, Pretty.str v]]), (v2, t))
end
| t $ u =>
let
@@ -143,7 +143,7 @@
| mk_cterm_code b ty_bs ts xs (vals, Abs (s, T, t)) =
let
val u = Free (s, T);
- val Some s' = assoc (ts, u);
+ val SOME s' = assoc (ts, u);
val p = Pretty.str s';
val (vals', p') = mk_cterm_code true ty_bs ts (p :: xs)
(if null (typ_tvars T) then vals
@@ -153,13 +153,13 @@
end
| mk_cterm_code b ty_bs ts xs (vals, Bound i) = (vals, nth_elem (i, xs))
| mk_cterm_code b ty_bs ts xs (vals, t) = (case assoc (vals, t) of
- None =>
- let val Some s = assoc (ts, t)
+ NONE =>
+ let val SOME s = assoc (ts, t)
in (if is_Const t andalso not (null (term_tvars t)) then
vals @ [(t, (("", s), [mk_val [s] [inst_ty true ty_bs t s]]))]
else vals, Pretty.str s)
end
- | Some ((_, s), _) => (vals, Pretty.str s));
+ | SOME ((_, s), _) => (vals, Pretty.str s));
fun get_cases sg =
Symtab.foldl (fn (tab, (k, {case_rewrites, ...})) => Symtab.update_new
@@ -195,8 +195,8 @@
*)
fun lookup sg fs t = (case Library.find_first (is_instance sg t o fst) fs of
- None => (bla := (t ins !bla); None)
- | Some (_, x) => Some x);
+ NONE => (bla := (t ins !bla); NONE)
+ | SOME (_, x) => SOME x);
fun unint sg fs t = forall (is_none o lookup sg fs) (term_consts' t);
@@ -220,15 +220,15 @@
fun mk_simpl_code sg case_tab mkeq fs ts ty_bs thm_bs ((vs, vals), t) =
(case assoc (vals, t) of
- Some ((eq, ct), ps) => (* binding already generated *)
+ SOME ((eq, ct), ps) => (* binding already generated *)
if mkeq andalso eq="" then
let val eq' = variant vs "eq"
in ((eq' :: vs, overwrite (vals,
(t, ((eq', ct), ps @ [mk_refleq eq' ct])))), (eq', ct))
end
else ((vs, vals), (eq, ct))
- | None => (case assoc (ts, t) of
- Some v => (* atomic term *)
+ | NONE => (case assoc (ts, t) of
+ SOME v => (* atomic term *)
let val xs = if not (null (term_tvars t)) andalso is_Const t then
[mk_val [v] [inst_ty false ty_bs t v]] else []
in
@@ -240,10 +240,10 @@
else ((vs, if null xs then vals else vals @
[(t, (("", v), xs))]), ("", v))
end
- | None => (* complex term *)
+ | NONE => (* complex term *)
let val (f as Const (cname, _), us) = strip_comb t
in case Symtab.lookup (case_tab, cname) of
- Some cases => (* case expression *)
+ SOME cases => (* case expression *)
let
val (us', u) = split_last us;
val b = unint sg fs u;
@@ -329,7 +329,7 @@
end
end
- | None =>
+ | NONE =>
let
val b = forall (unint sg fs) us;
val (q, eqs) = foldl_map
@@ -345,18 +345,18 @@
(Pretty.str eqf) (map (Pretty.str o fst) eqs)
in
case (lookup sg fs f, b) of
- (None, true) => (* completely uninterpreted *)
+ (NONE, true) => (* completely uninterpreted *)
if mkeq then ((ct :: eq :: vs', vals' @
[(t, ((eq, ct), [ctv, mk_refleq eq ct]))]), (eq, ct))
else ((ct :: vs', vals' @ [(t, (("", ct), [ctv]))]), ("", ct))
- | (None, false) => (* function uninterpreted *)
+ | (NONE, false) => (* function uninterpreted *)
((eq :: ct :: vs', vals' @
[(t, ((eq, ct), [ctv, mk_val [eq] [combp false]]))]), (eq, ct))
- | (Some (s, _, _), true) => (* arguments uninterpreted *)
+ | (SOME (s, _, _), true) => (* arguments uninterpreted *)
((eq :: ct :: vs', vals' @
[(t, ((eq, ct), [mk_val [ct, eq] (separate (Pretty.brk 1)
(Pretty.str s :: map (Pretty.str o snd) eqs))]))]), (eq, ct))
- | (Some (s, _, _), false) => (* function and arguments interpreted *)
+ | (SOME (s, _, _), false) => (* function and arguments interpreted *)
let val eq' = variant (eq :: ct :: vs') "eq"
in ((eq' :: eq :: ct :: vs', vals' @ [(t, ((eq', ct),
[mk_val [ct, eq] (separate (Pretty.brk 1)
@@ -373,8 +373,8 @@
fun mk_funs_code sg case_tab fs fs' =
let
val case_thms = mapfilter (fn s => (case Symtab.lookup (case_tab, s) of
- None => None
- | Some thms => Some (unsuffix "_case" (Sign.base_name s) ^ ".cases",
+ NONE => NONE
+ | SOME thms => SOME (unsuffix "_case" (Sign.base_name s) ^ ".cases",
map (fn i => Sign.base_name s ^ "_" ^ string_of_int i)
(1 upto length thms) ~~ thms)))
(foldr add_term_consts (map (prop_of o snd)
@@ -423,7 +423,7 @@
val default = if d then
let
- val Some s = assoc (thm_bs, f);
+ val SOME s = assoc (thm_bs, f);
val ct = variant vs' "ct"
in [Pretty.brk 1, Pretty.str "handle", Pretty.brk 1,
Pretty.str "Match =>", Pretty.brk 1, mk_let "let" 2
--- a/src/HOL/Matrix/fspmlp.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Matrix/fspmlp.ML Sun Feb 13 17:15:14 2005 +0100
@@ -47,12 +47,12 @@
let
val x =
case VarGraph.lookup (g, dest_key) of
- None => (None, Inttab.update ((row_index, (row_bound, [])), Inttab.empty))
- | Some (sure_bound, f) =>
+ NONE => (NONE, Inttab.update ((row_index, (row_bound, [])), Inttab.empty))
+ | SOME (sure_bound, f) =>
(sure_bound,
case Inttab.lookup (f, row_index) of
- None => Inttab.update ((row_index, (row_bound, [])), f)
- | Some _ => raise (Internal "add_row_bound"))
+ NONE => Inttab.update ((row_index, (row_bound, [])), f)
+ | SOME _ => raise (Internal "add_row_bound"))
in
VarGraph.update ((dest_key, x), g)
end
@@ -61,10 +61,10 @@
let
val x =
case VarGraph.lookup (g, key) of
- None => (Some bound, Inttab.empty)
- | Some (None, f) => (Some bound, f)
- | Some (Some old_bound, f) =>
- (Some ((case btype of
+ NONE => (SOME bound, Inttab.empty)
+ | SOME (NONE, f) => (SOME bound, f)
+ | SOME (SOME old_bound, f) =>
+ (SOME ((case btype of
UPPER => FloatArith.min
| LOWER => FloatArith.max)
old_bound bound), f)
@@ -74,32 +74,32 @@
fun get_sure_bound g key =
case VarGraph.lookup (g, key) of
- None => None
- | Some (sure_bound, _) => sure_bound
+ NONE => NONE
+ | SOME (sure_bound, _) => sure_bound
(*fun get_row_bound g key row_index =
case VarGraph.lookup (g, key) of
- None => None
- | Some (sure_bound, f) =>
+ NONE => NONE
+ | SOME (sure_bound, f) =>
(case Inttab.lookup (f, row_index) of
- None => None
- | Some (row_bound, _) => (sure_bound, row_bound))*)
+ NONE => NONE
+ | SOME (row_bound, _) => (sure_bound, row_bound))*)
fun add_edge g src_key dest_key row_index coeff =
case VarGraph.lookup (g, dest_key) of
- None => raise (Internal "add_edge: dest_key not found")
- | Some (sure_bound, f) =>
+ NONE => raise (Internal "add_edge: dest_key not found")
+ | SOME (sure_bound, f) =>
(case Inttab.lookup (f, row_index) of
- None => raise (Internal "add_edge: row_index not found")
- | Some (row_bound, sources) =>
+ NONE => raise (Internal "add_edge: row_index not found")
+ | SOME (row_bound, sources) =>
VarGraph.update ((dest_key, (sure_bound, Inttab.update ((row_index, (row_bound, (coeff, src_key) :: sources)), f))), g))
fun split_graph g =
let
fun split ((r1, r2), (key, (sure_bound, _))) =
case sure_bound of
- None => (r1, r2)
- | Some bound =>
+ NONE => (r1, r2)
+ | SOME bound =>
(case key of
(u, UPPER) => (r1, Inttab.update ((u, bound), r2))
| (u, LOWER) => (Inttab.update ((u, bound), r1), r2))
@@ -118,7 +118,7 @@
If safe is false, termination is not guaranteed, but on termination the sure bounds are optimal (relative to the algorithm) *)
fun propagate_sure_bounds safe names g =
let
- (* returns None if no new sure bound could be calculated, otherwise the new sure bound is returned *)
+ (* returns NONE if no new sure bound could be calculated, otherwise the new sure bound is returned *)
fun calc_sure_bound_from_sources g (key as (_, btype)) =
let
fun mult_upper x (lower, upper) =
@@ -139,40 +139,40 @@
let
fun add_src_bound (sum, (coeff, src_key)) =
case sum of
- None => None
- | Some x =>
+ NONE => NONE
+ | SOME x =>
(case get_sure_bound g src_key of
- None => None
- | Some src_sure_bound => Some (FloatArith.add x (mult_btype src_sure_bound coeff)))
+ NONE => NONE
+ | SOME src_sure_bound => SOME (FloatArith.add x (mult_btype src_sure_bound coeff)))
in
- case foldl add_src_bound (Some row_bound, sources) of
- None => sure_bound
- | new_sure_bound as (Some new_bound) =>
+ case foldl add_src_bound (SOME row_bound, sources) of
+ NONE => sure_bound
+ | new_sure_bound as (SOME new_bound) =>
(case sure_bound of
- None => new_sure_bound
- | Some old_bound =>
- Some (case btype of
+ NONE => new_sure_bound
+ | SOME old_bound =>
+ SOME (case btype of
UPPER => FloatArith.min old_bound new_bound
| LOWER => FloatArith.max old_bound new_bound))
end
in
case VarGraph.lookup (g, key) of
- None => None
- | Some (sure_bound, f) =>
+ NONE => NONE
+ | SOME (sure_bound, f) =>
let
val x = Inttab.foldl calc_sure_bound (sure_bound, f)
in
- if x = sure_bound then None else x
+ if x = sure_bound then NONE else x
end
end
fun propagate ((g, b), (key, _)) =
case calc_sure_bound_from_sources g key of
- None => (g,b)
- | Some bound => (update_sure_bound g key bound,
+ NONE => (g,b)
+ | SOME bound => (update_sure_bound g key bound,
if safe then
case get_sure_bound g key of
- None => true
+ NONE => true
| _ => b
else
true)
@@ -203,7 +203,7 @@
fun calcr (g, (row_index, a)) =
let
val b = FloatSparseMatrixBuilder.v_elem_at b row_index
- val (_, b2) = ExactFloatingPoint.approx_decstr_by_bin prec (case b of None => "0" | Some b => b)
+ val (_, b2) = ExactFloatingPoint.approx_decstr_by_bin prec (case b of NONE => "0" | SOME b => b)
val approx_a = FloatSparseMatrixBuilder.v_fold (fn (l, (i,s)) =>
(i, ExactFloatingPoint.approx_decstr_by_bin prec s)::l) [] a
@@ -265,8 +265,8 @@
let
val index = xlen-i
val r = abs_estimate (i-1) r1 r2
- val b1 = case Inttab.lookup (r1, index) of None => raise (Load ("x-value not bounded from below: "^(names index))) | Some x => x
- val b2 = case Inttab.lookup (r2, index) of None => raise (Load ("x-value not bounded from above: "^(names index))) | Some x => x
+ val b1 = case Inttab.lookup (r1, index) of NONE => raise (Load ("x-value not bounded from below: "^(names index))) | SOME x => x
+ val b2 = case Inttab.lookup (r2, index) of NONE => raise (Load ("x-value not bounded from above: "^(names index))) | SOME x => x
val abs_max = FloatArith.max (FloatArith.neg (FloatArith.negative_part b1)) (FloatArith.positive_part b2)
val vec = FloatSparseMatrixBuilder.cons_spvec (FloatSparseMatrixBuilder.mk_spvec_entry 0 abs_max) FloatSparseMatrixBuilder.empty_spvec
in
@@ -287,7 +287,7 @@
val (dualprog, indexof) = FloatSparseMatrixBuilder.dual_cplexProg c A b
val results = Cplex.solve dualprog
val (optimal,v) = CplexFloatSparseMatrixConverter.convert_results results indexof
- val A = FloatSparseMatrixBuilder.cut_matrix v None A
+ val A = FloatSparseMatrixBuilder.cut_matrix v NONE A
fun id x = x
val v = FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 v
val b = FloatSparseMatrixBuilder.transpose_matrix (FloatSparseMatrixBuilder.set_vector FloatSparseMatrixBuilder.empty_matrix 0 b)
--- a/src/HOL/MicroJava/J/JListExample.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/MicroJava/J/JListExample.thy Sun Feb 13 17:15:14 2005 +0100
@@ -109,7 +109,7 @@
ML {*
-val Library.Some ((_, (heap, locs)), _) = Seq.pull test;
+val SOME ((_, (heap, locs)), _) = Seq.pull test;
locs l1_name;
locs l2_name;
locs l3_name;
--- a/src/HOL/Modelcheck/EindhovenSyn.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Modelcheck/EindhovenSyn.ML Sun Feb 13 17:15:14 2005 +0100
@@ -26,7 +26,7 @@
val pair_eta_expand_proc =
Simplifier.simproc (Theory.sign_of (the_context ())) "pair_eta_expand" ["f::'a*'b=>'c"]
- (fn _ => fn _ => fn t => case t of Abs _ => Some (mk_meta_eq pair_eta_expand) | _ => None);
+ (fn _ => fn _ => fn t => case t of Abs _ => SOME (mk_meta_eq pair_eta_expand) | _ => NONE);
val Eindhoven_ss =
simpset() addsimprocs [pair_eta_expand_proc] addsimps [Let_def];
--- a/src/HOL/Modelcheck/MuckeSyn.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Modelcheck/MuckeSyn.ML Sun Feb 13 17:15:14 2005 +0100
@@ -1,22 +1,22 @@
(* search_mu t searches for Mu terms in term t. In the case of nested Mu's,
- it yields innermost one. If no Mu term is present, search_mu yields None
+ it yields innermost one. If no Mu term is present, search_mu yields NONE
*)
(* extended for treatment of nu (TH) *)
fun search_mu ((Const("MuCalculus.mu",tp)) $ t2) =
(case (search_mu t2) of
- Some t => Some t
- | None => Some ((Const("MuCalculus.mu",tp)) $ t2))
+ SOME t => SOME t
+ | NONE => SOME ((Const("MuCalculus.mu",tp)) $ t2))
| search_mu ((Const("MuCalculus.nu",tp)) $ t2) =
(case (search_mu t2) of
- Some t => Some t
- | None => Some ((Const("MuCalculus.nu",tp)) $ t2))
+ SOME t => SOME t
+ | NONE => SOME ((Const("MuCalculus.nu",tp)) $ t2))
| search_mu (t1 $ t2) =
(case (search_mu t1) of
- Some t => Some t
- | None => search_mu t2)
+ SOME t => SOME t
+ | NONE => search_mu t2)
| search_mu (Abs(_,_,t)) = search_mu t
- | search_mu _ = None;
+ | search_mu _ = NONE;
@@ -27,11 +27,11 @@
fun search_var s t =
case t of
t1 $ t2 => (case (search_var s t1) of
- Some tt => Some tt |
- None => search_var s t2) |
+ SOME tt => SOME tt |
+ NONE => search_var s t2) |
Abs(_,_,t) => search_var s t |
- Var((s1,_),_) => if s = s1 then Some t else None |
- _ => None;
+ Var((s1,_),_) => if s = s1 then SOME t else NONE |
+ _ => NONE;
(* function move_mus:
@@ -65,8 +65,8 @@
if t < 0 then 1 else t+1 end;
in
case redex of
- None => all_tac state |
- Some redexterm =>
+ NONE => all_tac state |
+ SOME redexterm =>
let val Credex = cterm_of sign redexterm;
val aiCterm =
cterm_of sig_move_thm (Logic.incr_indexes ([],idx) (term_of aCterm));
@@ -99,8 +99,8 @@
infix cc;
-fun None cc xl = xl
- | (Some x) cc xl = x::xl;
+fun NONE cc xl = xl
+ | (SOME x) cc xl = x::xl;
fun getargs ((x $ y) $ (Var ((z,_),_))) = getargs (x $ y) ^ " " ^z
| getargs (x $ (Var ((y,_),_))) = y;
@@ -126,18 +126,18 @@
in
(* extension with removing bold font (TH) *)
-fun mk_lam_def (_::_) _ _ = None
- | mk_lam_def [] ((Const("==",_) $ (Const _)) $ RHS) t = Some t
+fun mk_lam_def (_::_) _ _ = NONE
+ | mk_lam_def [] ((Const("==",_) $ (Const _)) $ RHS) t = SOME t
| mk_lam_def [] ((Const("==",_) $ LHS) $ RHS) t =
let val tsig = #sign (rep_thm t);
val fnam = Sign.string_of_term tsig (getfun LHS);
val rhs = Sign.string_of_term tsig (freeze_thaw RHS)
val gl = delete_bold_string (fnam ^" == % " ^ (getargs LHS) ^" . " ^ rhs);
in
- Some (prove_goal (theory_of_sign tsig) gl (fn prems =>
+ SOME (prove_goal (theory_of_sign tsig) gl (fn prems =>
[(REPEAT (rtac ext_rl 1)), (rtac t 1) ]))
end
-| mk_lam_def [] _ t= None;
+| mk_lam_def [] _ t= NONE;
fun mk_lam_defs ([]:thm list) = ([]: thm list)
| mk_lam_defs (t::l) =
@@ -157,8 +157,8 @@
Simplifier.simproc (Theory.sign_of (the_context ()))
"pair_eta_expand" ["f::'a*'b=>'c"]
(fn _ => fn _ => fn t =>
- case t of Abs _ => Some (mk_meta_eq pair_eta_expand)
- | _ => None);
+ case t of Abs _ => SOME (mk_meta_eq pair_eta_expand)
+ | _ => NONE);
val Mucke_ss = simpset() addsimprocs [pair_eta_expand_proc] addsimps [Let_def];
--- a/src/HOL/Modelcheck/mucke_oracle.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Modelcheck/mucke_oracle.ML Sun Feb 13 17:15:14 2005 +0100
@@ -981,8 +981,8 @@
val s = post_last_dot(fst(qtn a));
fun param_types ((Const("Trueprop",_)) $ (_ $ (Var(_,Type(_,t))))) = t |
param_types _ = error "malformed induct-theorem in preprocess_td";
- val pl = param_types (concl_of (get_thm (theory_of_sign sg) (s ^ ".induct", None)));
- val l = split_constrs sg (snd(qtn a)) pl (prems_of (get_thm (theory_of_sign sg) (s ^ ".induct", None)));
+ val pl = param_types (concl_of (get_thm (theory_of_sign sg) (s ^ ".induct", NONE)));
+ val l = split_constrs sg (snd(qtn a)) pl (prems_of (get_thm (theory_of_sign sg) (s ^ ".induct", NONE)));
val ntl = new_types l;
in
((a,l) :: (preprocess_td sg (ntl @ b) (a :: done)))
--- a/src/HOL/Orderings.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Orderings.thy Sun Feb 13 17:15:14 2005 +0100
@@ -290,21 +290,21 @@
let fun of_sort t = Sign.of_sort sign (type_of t, sort)
fun dec (Const ("Not", _) $ t) = (
case dec t of
- None => None
- | Some (t1, rel, t2) => Some (t1, "~" ^ rel, t2))
+ NONE => NONE
+ | SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2))
| dec (Const ("op =", _) $ t1 $ t2) =
if of_sort t1
- then Some (t1, "=", t2)
- else None
+ then SOME (t1, "=", t2)
+ else NONE
| dec (Const ("op <=", _) $ t1 $ t2) =
if of_sort t1
- then Some (t1, "<=", t2)
- else None
+ then SOME (t1, "<=", t2)
+ else NONE
| dec (Const ("op <", _) $ t1 $ t2) =
if of_sort t1
- then Some (t1, "<", t2)
- else None
- | dec _ = None
+ then SOME (t1, "<", t2)
+ else NONE
+ | dec _ = NONE
in dec t end;
structure Quasi_Tac = Quasi_Tac_Fun (
--- a/src/HOL/Product_Type.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Product_Type.thy Sun Feb 13 17:15:14 2005 +0100
@@ -34,7 +34,7 @@
val unit_eq_proc =
let val unit_meta_eq = mk_meta_eq (thm "unit_eq") in
Simplifier.simproc (Theory.sign_of (the_context ())) "unit_eq" ["x::unit"]
- (fn _ => fn _ => fn t => if HOLogic.is_unit t then None else Some unit_meta_eq)
+ (fn _ => fn _ => fn t => if HOLogic.is_unit t then NONE else SOME unit_meta_eq)
end;
Addsimprocs [unit_eq_proc];
@@ -409,9 +409,9 @@
| no_args k i (t $ u) = no_args k i t andalso no_args k i u
| no_args k i (Bound m) = m < k orelse m > k+i
| no_args _ _ _ = true;
- fun split_pat tp i (Abs (_,_,t)) = if tp 0 i t then Some (i,t) else None
+ fun split_pat tp i (Abs (_,_,t)) = if tp 0 i t then SOME (i,t) else NONE
| split_pat tp i (Const ("split", _) $ Abs (_, _, t)) = split_pat tp (i+1) t
- | split_pat tp i _ = None;
+ | split_pat tp i _ = NONE;
fun metaeq sg lhs rhs = mk_meta_eq (Tactic.prove sg [] []
(HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,rhs)))
(K (simp_tac (HOL_basic_ss addsimps [cond_split_eta]) 1)));
@@ -428,14 +428,14 @@
| subst arg k i t = t;
fun beta_proc sg _ (s as Const ("split", _) $ Abs (_, _, t) $ arg) =
(case split_pat beta_term_pat 1 t of
- Some (i,f) => Some (metaeq sg s (subst arg 0 i f))
- | None => None)
- | beta_proc _ _ _ = None;
+ SOME (i,f) => SOME (metaeq sg s (subst arg 0 i f))
+ | NONE => NONE)
+ | beta_proc _ _ _ = NONE;
fun eta_proc sg _ (s as Const ("split", _) $ Abs (_, _, t)) =
(case split_pat eta_term_pat 1 t of
- Some (_,ft) => Some (metaeq sg s (let val (f $ arg) = ft in f end))
- | None => None)
- | eta_proc _ _ _ = None;
+ SOME (_,ft) => SOME (metaeq sg s (let val (f $ arg) = ft in f end))
+ | NONE => NONE)
+ | eta_proc _ _ _ = NONE;
in
val split_beta_proc = Simplifier.simproc (Theory.sign_of (the_context ()))
"split_beta" ["split f z"] beta_proc;
@@ -810,13 +810,13 @@
val (gr2, pr) = Codegen.invoke_codegen thy dep false (gr1, r);
in (gr2, (pl, pr)) end
in case dest_let t of
- ([], _) => None
+ ([], _) => NONE
| (ps, u) =>
let
val (gr1, qs) = foldl_map mk_code (gr, ps);
val (gr2, pu) = Codegen.invoke_codegen thy dep false (gr1, u)
in
- Some (gr2, Pretty.blk (0, [Pretty.str "let ", Pretty.blk (0, flat
+ SOME (gr2, Pretty.blk (0, [Pretty.str "let ", Pretty.blk (0, flat
(separate [Pretty.str ";", Pretty.brk 1] (map (fn (pl, pr) =>
[Pretty.block [Pretty.str "val ", pl, Pretty.str " =",
Pretty.brk 1, pr]]) qs))),
@@ -824,7 +824,7 @@
Pretty.brk 1, Pretty.str "end"]))
end
end
- | let_codegen thy gr dep brack t = None;
+ | let_codegen thy gr dep brack t = NONE;
fun split_codegen thy gr dep brack (t as Const ("split", _) $ _) =
(case strip_abs 1 t of
@@ -833,11 +833,11 @@
val (gr1, q) = Codegen.invoke_codegen thy dep false (gr, p);
val (gr2, pu) = Codegen.invoke_codegen thy dep false (gr1, u)
in
- Some (gr2, Pretty.block [Pretty.str "(fn ", q, Pretty.str " =>",
+ SOME (gr2, Pretty.block [Pretty.str "(fn ", q, Pretty.str " =>",
Pretty.brk 1, pu, Pretty.str ")"])
end
- | _ => None)
- | split_codegen thy gr dep brack t = None;
+ | _ => NONE)
+ | split_codegen thy gr dep brack t = NONE;
in
--- a/src/HOL/Real/real_arith.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Real/real_arith.ML Sun Feb 13 17:15:14 2005 +0100
@@ -136,7 +136,7 @@
end;
-(* Some test data [omitting examples that assume the ordering to be discrete!]
+(* SOME test data [omitting examples that assume the ordering to be discrete!]
Goal "!!a::real. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
by (fast_arith_tac 1);
qed "";
--- a/src/HOL/TLA/TLA.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/TLA/TLA.ML Sun Feb 13 17:15:14 2005 +0100
@@ -818,7 +818,7 @@
by (asm_simp_tac (simpset() addsimps [split_box_conj,box_stp_actI]) 1);
qed "WF1";
-(* Sometimes easier to use; designed for action B rather than state predicate Q *)
+(* SOMEtimes easier to use; designed for action B rather than state predicate Q *)
val [prem1,prem2,prem3] = goalw thy [leadsto_def]
"[| |- N & $P --> $Enabled (<A>_v); \
\ |- N & <A>_v --> B; \
--- a/src/HOL/Tools/Presburger/cooper_dec.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/Presburger/cooper_dec.ML Sun Feb 13 17:15:14 2005 +0100
@@ -92,7 +92,7 @@
fun dest_numeral (Const("0",Type ("IntDef.int", []))) = 0
|dest_numeral (Const("1",Type ("IntDef.int", []))) = 1
|dest_numeral (Const ("Numeral.number_of",_) $ n)= HOLogic.dest_binum n;
-(*Some terms often used for pattern matching*)
+(*SOME terms often used for pattern matching*)
val zero = mk_numeral 0;
val one = mk_numeral 1;
@@ -450,7 +450,7 @@
[("op =",op=), ("op <",op<), ("op >",op>), ("op <=",op<=) , ("op >=",op>=),
("Divides.op dvd",fn (x,y) =>((y mod x) = 0))];
-fun applyoperation (Some f) (a,b) = f (a, b)
+fun applyoperation (SOME f) (a,b) = f (a, b)
|applyoperation _ (_, _) = false;
(*Evaluation of constant atomic formulas*)
@@ -464,12 +464,12 @@
handle _ => at)
else
case assoc (operations,p) of
- Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)
+ SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)
handle _ => at)
| _ => at)
|Const("Not",_)$(Const (p,_) $ s $ t) =>(
case assoc (operations,p) of
- Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then
+ SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then
HOLogic.false_const else HOLogic.true_const)
handle _ => at)
| _ => at)
@@ -480,12 +480,12 @@
fun evalc_atom at = case at of
(Const (p,_) $ s $ t) =>
( case assoc (operations,p) of
- Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)
+ SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)
handle _ => at)
| _ => at)
|Const("Not",_)$(Const (p,_) $ s $ t) =>(
case assoc (operations,p) of
- Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then
+ SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then
HOLogic.false_const else HOLogic.true_const)
handle _ => at)
| _ => at)
@@ -711,14 +711,14 @@
fun inf_w mi d vars x p =
let val f = if mi then minusinf else plusinf in
case (simpl (minusinf x p)) of
- Const("True",_) => (Some (mk_numeral 1), HOLogic.true_const)
- |Const("False",_) => (None,HOLogic.false_const)
+ Const("True",_) => (SOME (mk_numeral 1), HOLogic.true_const)
+ |Const("False",_) => (NONE,HOLogic.false_const)
|F =>
let
fun h n =
case ((simpl o evalc) (linrep vars x (mk_numeral n) F)) of
- Const("True",_) => (Some (mk_numeral n),HOLogic.true_const)
- |F' => if n=1 then (None,F')
+ Const("True",_) => (SOME (mk_numeral n),HOLogic.true_const)
+ |F' => if n=1 then (NONE,F')
else let val (rw,rf) = h (n-1) in
(rw,HOLogic.mk_disj(F',rf))
end
@@ -729,9 +729,9 @@
fun set_w d b st vars x p = let
fun h ns = case ns of
- [] => (None,HOLogic.false_const)
+ [] => (NONE,HOLogic.false_const)
|n::nl => ( case ((simpl o evalc) (linrep vars x n p)) of
- Const("True",_) => (Some n,HOLogic.true_const)
+ Const("True",_) => (SOME n,HOLogic.true_const)
|F' => let val (rw,rf) = h nl
in (rw,HOLogic.mk_disj(F',rf))
end)
@@ -741,10 +741,10 @@
end;
fun withness d b st vars x p = case (inf_w b d vars x p) of
- (Some n,_) => (Some n,HOLogic.true_const)
- |(None,Pinf) => (case (set_w d b st vars x p) of
- (Some n,_) => (Some n,HOLogic.true_const)
- |(_,Pst) => (None,HOLogic.mk_disj(Pinf,Pst)));
+ (SOME n,_) => (SOME n,HOLogic.true_const)
+ |(NONE,Pinf) => (case (set_w d b st vars x p) of
+ (SOME n,_) => (SOME n,HOLogic.true_const)
+ |(_,Pst) => (NONE,HOLogic.mk_disj(Pinf,Pst)));
--- a/src/HOL/Tools/Presburger/cooper_proof.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/Presburger/cooper_proof.ML Sun Feb 13 17:15:14 2005 +0100
@@ -194,8 +194,8 @@
fun simple_prove_goal_cterm2 G tacs =
let
- fun check None = error "prove_goal: tactic failed"
- | check (Some (thm, _)) = (case nprems_of thm of
+ fun check NONE = error "prove_goal: tactic failed"
+ | check (SOME (thm, _)) = (case nprems_of thm of
0 => thm
| i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
in check (Seq.pull (EVERY tacs (trivial G))) end;
@@ -320,7 +320,7 @@
val cx = cterm_of sg y
val pre = prove_elementar sg "lf"
(HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral n)))
- val th1 = (pre RS (instantiate' [] [Some ck,Some cc, Some cx, Some ct] (ac_pi_eq)))
+ val th1 = (pre RS (instantiate' [] [SOME ck,SOME cc, SOME cx, SOME ct] (ac_pi_eq)))
in ([], fn [] => [th1,(prove_elementar sg "sa" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
end
@@ -344,32 +344,32 @@
"op <" =>
let val pre = prove_elementar sg "lf"
(HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral k)))
- val th1 = (pre RS (instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct] (ac_lt_eq)))
+ val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_lt_eq)))
in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
end
|"op =" =>
let val pre = prove_elementar sg "lf"
(HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
- val th1 = (pre RS(instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct] (ac_eq_eq)))
+ val th1 = (pre RS(instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_eq_eq)))
in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
end
|"Divides.op dvd" =>
let val pre = prove_elementar sg "lf"
(HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
- val th1 = (pre RS (instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct]) (ac_dvd_eq))
+ val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct]) (ac_dvd_eq))
in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
end
end
- else ([], fn [] => instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl)
+ else ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl)
|( Const ("Not", _) $ p) => ([p], fn [th] => th RS qe_Not)
|( Const ("op &",_) $ p $ q) => ([p,q], fn [th1,th2] => [th1,th2] MRS qe_conjI)
|( Const ("op |",_) $ p $ q) =>([p,q], fn [th1,th2] => [th1,th2] MRS qe_disjI)
- |_ => ([], fn [] => instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl);
+ |_ => ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl);
fun proof_of_adjustcoeffeq sg x l = thm_of sg (decomp_adjustcoeffeq sg x l);
@@ -380,43 +380,43 @@
(*==================================================*)
fun rho_for_modd_minf x dlcm sg fm1 =
let
- (*Some certified Terms*)
+ (*SOME certified Terms*)
val ctrue = cterm_of sg HOLogic.true_const
val cfalse = cterm_of sg HOLogic.false_const
val fm = norm_zero_one fm1
in case fm1 of
(Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) =>
- if (x=y) andalso (c1= zero) andalso (c2= one) then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+ if (x=y) andalso (c1= zero) andalso (c2= one) then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one)
- then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+ then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if (y=x) andalso (c1 = zero) then
- if (pm1 = one) then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_minf)) else
- (instantiate' [Some cboolT] [Some ctrue] (fm_modd_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+ if (pm1 = one) then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_minf)) else
+ (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cz = cterm_of sg (norm_zero_one z)
val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
- in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_minf)))
+ in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_minf)))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
|(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
c $ y ) $ z))) =>
if y=x then let val cz = cterm_of sg (norm_zero_one z)
val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
- in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_minf)))
+ in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_minf)))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
- |_ => instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf)
+ |_ => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)
end;
(*=========================================================================*)
(*=========================================================================*)
@@ -426,36 +426,36 @@
in case fm1 of
(Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) =>
if (x=y) andalso (c1=zero) andalso (c2=one)
- then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (neq_eq_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (neq_eq_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one))
- then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (eq_eq_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (eq_eq_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if (y=x) andalso (c1 =zero) then
- if pm1 = one then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (le_eq_minf)) else
- (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (len_eq_minf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ if pm1 = one then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (le_eq_minf)) else
+ (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (len_eq_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cd = cterm_of sg (norm_zero_one d)
val cz = cterm_of sg (norm_zero_one z)
- in(instantiate' [] [Some cd, Some cz] (not_dvd_eq_minf))
+ in(instantiate' [] [SOME cd, SOME cz] (not_dvd_eq_minf))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
|(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cd = cterm_of sg (norm_zero_one d)
val cz = cterm_of sg (norm_zero_one z)
- in(instantiate' [] [Some cd, Some cz ] (dvd_eq_minf))
+ in(instantiate' [] [SOME cd, SOME cz ] (dvd_eq_minf))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
- |_ => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ |_ => (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
end;
(*=====================================================*)
@@ -476,7 +476,7 @@
(* -------------------------------------------------------------*)
fun rho_for_modd_pinf x dlcm sg fm1 =
let
- (*Some certified Terms*)
+ (*SOME certified Terms*)
val ctrue = cterm_of sg HOLogic.true_const
val cfalse = cterm_of sg HOLogic.false_const
@@ -484,37 +484,37 @@
in case fm1 of
(Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) =>
if ((x=y) andalso (c1= zero) andalso (c2= one))
- then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+ then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
if ((is_arith_rel fm) andalso (x = y) andalso (c1 = zero) andalso (c2 = one))
- then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+ then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if ((y=x) andalso (c1 = zero)) then
if (pm1 = one)
- then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_pinf))
- else (instantiate' [Some cboolT] [Some cfalse] (fm_modd_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+ then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cz = cterm_of sg (norm_zero_one z)
val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
- in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_pinf)))
+ in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_pinf)))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
|(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
c $ y ) $ z))) =>
if y=x then let val cz = cterm_of sg (norm_zero_one z)
val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
- in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_pinf)))
+ in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_pinf)))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
- |_ => instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf)
+ |_ => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)
end;
(* -------------------------------------------------------------*)
(* Finding rho for pinf_eq *)
@@ -525,36 +525,36 @@
in case fm1 of
(Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) =>
if (x=y) andalso (c1=zero) andalso (c2=one)
- then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (neq_eq_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (neq_eq_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one))
- then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (eq_eq_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (eq_eq_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if (y=x) andalso (c1 =zero) then
- if pm1 = one then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (le_eq_pinf)) else
- (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (len_eq_pinf))
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ if pm1 = one then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (le_eq_pinf)) else
+ (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (len_eq_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cd = cterm_of sg (norm_zero_one d)
val cz = cterm_of sg (norm_zero_one z)
- in(instantiate' [] [Some cd, Some cz] (not_dvd_eq_pinf))
+ in(instantiate' [] [SOME cd, SOME cz] (not_dvd_eq_pinf))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
|(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then let val cd = cterm_of sg (norm_zero_one d)
val cz = cterm_of sg (norm_zero_one z)
- in(instantiate' [] [Some cd, Some cz ] (dvd_eq_pinf))
+ in(instantiate' [] [SOME cd, SOME cz ] (dvd_eq_pinf))
end
- else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
- |_ => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ |_ => (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
end;
@@ -593,7 +593,7 @@
val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
val cdlcm = cterm_of sg dlcm
val cB = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one bs))
- in instantiate' [] [Some cdlcm,Some cB, Some cp] (bst_thm)
+ in instantiate' [] [SOME cdlcm,SOME cB, SOME cp] (bst_thm)
end;
fun asetproof_of sg (x as Free(xn,xT)) fm ast dlcm =
@@ -601,7 +601,7 @@
val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
val cdlcm = cterm_of sg dlcm
val cA = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one ast))
- in instantiate' [] [Some cdlcm,Some cA, Some cp] (ast_thm)
+ in instantiate' [] [SOME cdlcm,SOME cA, SOME cp] (ast_thm)
end;
*)
@@ -623,9 +623,9 @@
then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B)
val th2 = prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)))
val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma]([th3,th1,th2] MRS (not_bst_p_ne)))
+ in (instantiate' [] [SOME cfma]([th3,th1,th2] MRS (not_bst_p_ne)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
if (is_arith_rel at) andalso (x=y)
@@ -633,40 +633,40 @@
in let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ bst_z $ B)
val th2 = prove_elementar sg "ss" (HOLogic.mk_eq (bst_z,Const("op -",T) $ (Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma] ([th3,th1,th2] MRS (not_bst_p_eq)))
+ in (instantiate' [] [SOME cfma] ([th3,th1,th2] MRS (not_bst_p_eq)))
end
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if (y=x) andalso (c1 =zero) then
if pm1 = one then
let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B)
val th2 = prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)))
- in (instantiate' [] [Some cfma, Some cdlcm]([th1,th2] MRS (not_bst_p_gt)))
+ in (instantiate' [] [SOME cfma, SOME cdlcm]([th1,th2] MRS (not_bst_p_gt)))
end
else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma, Some cB,Some (cterm_of sg (norm_zero_one z))] (th1 RS (not_bst_p_lt)))
+ in (instantiate' [] [SOME cfma, SOME cB,SOME (cterm_of sg (norm_zero_one z))] (th1 RS (not_bst_p_lt)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then
let val cz = cterm_of sg (norm_zero_one z)
val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
- in (instantiate' [] [Some cfma, Some cB,Some cz] (th1 RS (not_bst_p_ndvd)))
+ in (instantiate' [] [SOME cfma, SOME cB,SOME cz] (th1 RS (not_bst_p_ndvd)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
|(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then
let val cz = cterm_of sg (norm_zero_one z)
val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
- in (instantiate' [] [Some cfma,Some cB,Some cz] (th1 RS (not_bst_p_dvd)))
+ in (instantiate' [] [SOME cfma,SOME cB,SOME cz] (th1 RS (not_bst_p_dvd)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
- |_ => (instantiate' [] [Some cfma, Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+ |_ => (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
end;
@@ -714,9 +714,9 @@
then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ A)
val th2 = prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)))
val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma]([th3,th1,th2] MRS (not_ast_p_ne)))
+ in (instantiate' [] [SOME cfma]([th3,th1,th2] MRS (not_ast_p_ne)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
|(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
if (is_arith_rel at) andalso (x=y)
@@ -724,39 +724,39 @@
val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ ast_z $ A)
val th2 = prove_elementar sg "ss" (HOLogic.mk_eq (ast_z,Const("op +",T) $ (Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma] ([th3,th1,th2] MRS (not_ast_p_eq)))
+ in (instantiate' [] [SOME cfma] ([th3,th1,th2] MRS (not_ast_p_eq)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
|(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
if (y=x) andalso (c1 =zero) then
if pm1 = (mk_numeral ~1) then
let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one z) $ A)
val th2 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm))
- in (instantiate' [] [Some cfma]([th2,th1] MRS (not_ast_p_lt)))
+ in (instantiate' [] [SOME cfma]([th2,th1] MRS (not_ast_p_lt)))
end
else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
- in (instantiate' [] [Some cfma, Some cA,Some (cterm_of sg (norm_zero_one z))] (th1 RS (not_ast_p_gt)))
+ in (instantiate' [] [SOME cfma, SOME cA,SOME (cterm_of sg (norm_zero_one z))] (th1 RS (not_ast_p_gt)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
|Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then
let val cz = cterm_of sg (norm_zero_one z)
val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
- in (instantiate' [] [Some cfma, Some cA,Some cz] (th1 RS (not_ast_p_ndvd)))
+ in (instantiate' [] [SOME cfma, SOME cA,SOME cz] (th1 RS (not_ast_p_ndvd)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
|(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) =>
if y=x then
let val cz = cterm_of sg (norm_zero_one z)
val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
- in (instantiate' [] [Some cfma,Some cA,Some cz] (th1 RS (not_ast_p_dvd)))
+ in (instantiate' [] [SOME cfma,SOME cA,SOME cz] (th1 RS (not_ast_p_dvd)))
end
- else (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
- |_ => (instantiate' [] [Some cfma, Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+ |_ => (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
end;
@@ -788,21 +788,21 @@
fun rho_for_evalc sg at = case at of
(Const (p,_) $ s $ t) =>(
case assoc (operations,p) of
- Some f =>
+ SOME f =>
((if (f ((dest_numeral s),(dest_numeral t)))
then prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const))
else prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const)))
- handle _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl)
- | _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl )
+ handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl)
+ | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl )
|Const("Not",_)$(Const (p,_) $ s $ t) =>(
case assoc (operations,p) of
- Some f =>
+ SOME f =>
((if (f ((dest_numeral s),(dest_numeral t)))
then prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const))
else prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const)))
- handle _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl)
- | _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl )
- | _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl;
+ handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl)
+ | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl )
+ | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl;
(*=========================================================*)
@@ -828,7 +828,7 @@
|(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI)
|(Const("Not",_)$p) => ([p],fn [th] => th RS qe_Not)
|(Const("Divides.op dvd",_)$d$r) =>
- if is_numeral d then ([], fn [] => (prove_elementar sg "lf" (HOLogic.mk_eq (r, lint vars r))) RS (instantiate' [] [None , None, Some (cterm_of sg d)](linearize_dvd)))
+ if is_numeral d then ([], fn [] => (prove_elementar sg "lf" (HOLogic.mk_eq (r, lint vars r))) RS (instantiate' [] [NONE , NONE, SOME (cterm_of sg d)](linearize_dvd)))
else (warning "Nonlinear Term --- Non numeral leftside at dvd";
raise COOPER)
|_ => ([], fn [] => prove_elementar sg "lf" (HOLogic.mk_eq (t, linform vars t)));
@@ -897,7 +897,7 @@
val ss = presburger_ss addsimps
[simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
(* uth : EX x.P(l*x) = EX x. l dvd x & P x*)
- val uth = instantiate' [] [Some (cterm_of sg P) , Some (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
+ val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
(* e_ac_thm : Ex x. efm = EX x. fm*)
val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
(* A and B set of the formula*)
@@ -961,7 +961,7 @@
val ss = presburger_ss addsimps
[simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
(* uth : EX x.P(l*x) = EX x. l dvd x & P x*)
- val uth = instantiate' [] [Some (cterm_of sg P) , Some (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
+ val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
(* e_ac_thm : Ex x. efm = EX x. fm*)
val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
(* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*)
@@ -972,7 +972,7 @@
val (w,rs) = cooper_w [] cfm
val exp_cp_thm = case w of
(* FIXME - e_ac_thm just tipped to test syntactical correctness of the program!!!!*)
- Some n => e_ac_thm (* Prove cfm (n) and use exI and then Eq_TrueI*)
+ SOME n => e_ac_thm (* Prove cfm (n) and use exI and then Eq_TrueI*)
|_ => let
(* A and B set of the formula*)
val A = aset x cfm
--- a/src/HOL/Tools/Presburger/presburger.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/Presburger/presburger.ML Sun Feb 13 17:15:14 2005 +0100
@@ -73,7 +73,7 @@
fun term_typed_consts t = add_term_typed_consts(t,[]);
-(* Some Types*)
+(* SOME Types*)
val bT = HOLogic.boolT;
val iT = HOLogic.intT;
val binT = HOLogic.binT;
@@ -246,7 +246,7 @@
val g' = if a then abstract_pres sg g else g
(* Transform the term*)
val (t,np,nh) = prepare_for_presburger sg q g'
- (* Some simpsets for dealing with mod div abs and nat*)
+ (* SOME simpsets for dealing with mod div abs and nat*)
val simpset0 = HOL_basic_ss
addsimps [mod_div_equality', Suc_plus1]
addsplits [split_zdiv, split_zmod, split_div', split_min, split_max]
@@ -313,7 +313,7 @@
"decision procedure for Presburger arithmetic"),
ArithTheoryData.map (fn {splits, inj_consts, discrete, presburger} =>
{splits = splits, inj_consts = inj_consts, discrete = discrete,
- presburger = Some (presburger_tac true false)})];
+ presburger = SOME (presburger_tac true false)})];
end;
--- a/src/HOL/Tools/Presburger/qelim.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/Presburger/qelim.ML Sun Feb 13 17:15:14 2005 +0100
@@ -51,7 +51,7 @@
end
|(Const("All",_)$Abs(xn,xT,p)) => ([(HOLogic.exists_const xT)$Abs(xn,xT,HOLogic.Not $ p)], fn [th] => th RS qe_ALL)
- | _ => ([],fn [] => instantiate' [Some (ctyp_of sg (type_of P))] [Some (cterm_of sg P)] refl);
+ | _ => ([],fn [] => instantiate' [SOME (ctyp_of sg (type_of P))] [SOME (cterm_of sg P)] refl);
fun tproof_of_mlift_qelim sg isat afnp nfnp qfnp p =
--- a/src/HOL/Tools/datatype_aux.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/datatype_aux.ML Sun Feb 13 17:15:14 2005 +0100
@@ -130,11 +130,11 @@
val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop
(Logic.strip_imp_concl (nth_elem (i - 1, prems_of st))));
val getP = if can HOLogic.dest_imp (hd ts) then
- (apfst Some) o HOLogic.dest_imp else pair None;
+ (apfst SOME) o HOLogic.dest_imp else pair NONE;
fun abstr (t1, t2) = (case t1 of
- None => let val [Free (s, T)] = add_term_frees (t2, [])
+ NONE => let val [Free (s, T)] = add_term_frees (t2, [])
in absfree (s, T, t2) end
- | Some (_ $ t' $ _) => Abs ("x", fastype_of t', abstract_over (t', t2)))
+ | SOME (_ $ t' $ _) => Abs ("x", fastype_of t', abstract_over (t', t2)))
val cert = cterm_of (Thm.sign_of_thm st);
val Ps = map (cert o head_of o snd o getP) ts;
val indrule' = cterm_instantiate (Ps ~~
@@ -197,8 +197,8 @@
fun subst_DtTFree _ substs (T as (DtTFree name)) =
(case assoc (substs, name) of
- None => T
- | Some U => U)
+ NONE => T
+ | SOME U => U)
| subst_DtTFree i substs (DtType (name, ts)) =
DtType (name, map (subst_DtTFree i substs) ts)
| subst_DtTFree i _ (DtRec j) = DtRec (i + j);
@@ -237,8 +237,8 @@
| dtyp_of_typ _ (TVar _) = error "Illegal schematic type variable(s)"
| dtyp_of_typ new_dts (Type (tname, Ts)) =
(case assoc (new_dts, tname) of
- None => DtType (tname, map (dtyp_of_typ new_dts) Ts)
- | Some vs => if map (try dest_TFree) Ts = map Some vs then
+ NONE => DtType (tname, map (dtyp_of_typ new_dts) Ts)
+ | SOME vs => if map (try dest_TFree) Ts = map SOME vs then
DtRec (find_index (curry op = tname o fst) new_dts)
else error ("Illegal occurence of recursive type " ^ tname));
@@ -300,9 +300,9 @@
fun get_dt_descr T i tname dts =
(case Symtab.lookup (dt_info, tname) of
- None => typ_error T (tname ^ " is not a datatype - can't use it in\
+ NONE => typ_error T (tname ^ " is not a datatype - can't use it in\
\ nested recursion")
- | (Some {index, descr, ...}) =>
+ | (SOME {index, descr, ...}) =>
let val (_, vars, _) = the (assoc (descr, index));
val subst = ((map dest_DtTFree vars) ~~ dts) handle LIST _ =>
typ_error T ("Type constructor " ^ tname ^ " used with wrong\
--- a/src/HOL/Tools/datatype_codegen.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/datatype_codegen.ML Sun Feb 13 17:15:14 2005 +0100
@@ -27,17 +27,17 @@
fun find_nonempty (descr: DatatypeAux.descr) is i =
let
val (_, _, constrs) = the (assoc (descr, i));
- fun arg_nonempty (_, DatatypeAux.DtRec i) = if i mem is then None
+ fun arg_nonempty (_, DatatypeAux.DtRec i) = if i mem is then NONE
else apsome (curry op + 1 o snd) (find_nonempty descr (i::is) i)
- | arg_nonempty _ = Some 0;
+ | arg_nonempty _ = SOME 0;
fun max xs = foldl
- (fn (None, _) => None
- | (Some i, Some j) => Some (Int.max (i, j))
- | (_, None) => None) (Some 0, xs);
+ (fn (NONE, _) => NONE
+ | (SOME i, SOME j) => SOME (Int.max (i, j))
+ | (_, NONE) => NONE) (SOME 0, xs);
val xs = sort (int_ord o pairself snd)
(mapfilter (fn (s, dts) => apsome (pair s)
(max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs)
- in case xs of [] => None | x :: _ => Some x end;
+ in case xs of [] => NONE | x :: _ => SOME x end;
fun add_dt_defs thy dep gr (descr: DatatypeAux.descr) =
let
@@ -109,7 +109,7 @@
val sorts = map (rpair []) tvs;
val (cs1, cs2) =
partition (exists DatatypeAux.is_rec_type o snd) cs;
- val Some (cname, _) = find_nonempty descr [i] i;
+ val SOME (cname, _) = find_nonempty descr [i] i;
fun mk_delay p = Pretty.block
[Pretty.str "fn () =>", Pretty.brk 1, p];
@@ -173,10 +173,10 @@
handle Graph.CYCLES _ => gr) handle Graph.UNDEF _ =>
let
val gr1 = Graph.add_edge (dname, dep)
- (Graph.new_node (dname, (None, "")) gr);
+ (Graph.new_node (dname, (NONE, "")) gr);
val (gr2, dtdef) = mk_dtdef gr1 "datatype " descr';
in
- Graph.map_node dname (K (None,
+ Graph.map_node dname (K (NONE,
Pretty.string_of (Pretty.blk (0, separate Pretty.fbrk dtdef @
[Pretty.str ";"])) ^ "\n\n" ^
(if "term_of" mem !mode then
@@ -259,33 +259,33 @@
s = case_name orelse
is_some (assoc (#3 (the (assoc (descr, index))), s)))
(Symtab.dest (DatatypePackage.get_datatypes thy)) of
- None => None
- | Some (tname, {index, descr, ...}) =>
- if is_some (get_assoc_code thy s T) then None else
- let val Some (_, _, constrs) = assoc (descr, index)
+ NONE => NONE
+ | SOME (tname, {index, descr, ...}) =>
+ if is_some (get_assoc_code thy s T) then NONE else
+ let val SOME (_, _, constrs) = assoc (descr, index)
in (case (assoc (constrs, s), strip_type T) of
- (None, _) => Some (pretty_case thy gr dep brack
+ (NONE, _) => SOME (pretty_case thy gr dep brack
(#3 (the (assoc (descr, index)))) c ts)
- | (Some args, (_, Type _)) => Some (pretty_constr thy
+ | (SOME args, (_, Type _)) => SOME (pretty_constr thy
(fst (invoke_tycodegen thy dep false (gr, snd (strip_type T))))
dep brack args c ts)
- | _ => None)
+ | _ => NONE)
end)
- | _ => None);
+ | _ => NONE);
fun datatype_tycodegen thy gr dep brack (Type (s, Ts)) =
(case Symtab.lookup (DatatypePackage.get_datatypes thy, s) of
- None => None
- | Some {descr, ...} =>
- if is_some (get_assoc_type thy s) then None else
+ NONE => NONE
+ | SOME {descr, ...} =>
+ if is_some (get_assoc_type thy s) then NONE else
let val (gr', ps) = foldl_map
(invoke_tycodegen thy dep false) (gr, Ts)
- in Some (add_dt_defs thy dep gr' descr,
+ in SOME (add_dt_defs thy dep gr' descr,
Pretty.block ((if null Ts then [] else
[mk_tuple ps, Pretty.str " "]) @
[Pretty.str (mk_type_id (sign_of thy) s)]))
end)
- | datatype_tycodegen _ _ _ _ _ = None;
+ | datatype_tycodegen _ _ _ _ _ = NONE;
val setup =
--- a/src/HOL/Tools/datatype_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/datatype_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -114,28 +114,28 @@
fun datatype_info_sg sg name = Symtab.lookup (get_datatypes_sg sg, name);
fun datatype_info_sg_err sg name = (case datatype_info_sg sg name of
- Some info => info
- | None => error ("Unknown datatype " ^ quote name));
+ SOME info => info
+ | NONE => error ("Unknown datatype " ^ quote name));
val datatype_info = datatype_info_sg o Theory.sign_of;
fun datatype_info_err thy name = (case datatype_info thy name of
- Some info => info
- | None => error ("Unknown datatype " ^ quote name));
+ SOME info => info
+ | NONE => error ("Unknown datatype " ^ quote name));
fun constrs_of_sg sg tname = (case datatype_info_sg sg tname of
- Some {index, descr, ...} =>
+ SOME {index, descr, ...} =>
let val (_, _, constrs) = the (assoc (descr, index))
- in Some (map (fn (cname, _) => Const (cname, the (Sign.const_type sg cname))) constrs)
+ in SOME (map (fn (cname, _) => Const (cname, the (Sign.const_type sg cname))) constrs)
end
- | _ => None);
+ | _ => NONE);
val constrs_of = constrs_of_sg o Theory.sign_of;
fun case_const_of thy tname = (case datatype_info thy tname of
- Some {case_name, ...} => Some (Const (case_name, the (Sign.const_type
+ SOME {case_name, ...} => SOME (Const (case_name, the (Sign.const_type
(Theory.sign_of thy) case_name)))
- | _ => None);
+ | _ => NONE);
val weak_case_congs_of_sg = map (#weak_case_cong o #2) o Symtab.dest o get_datatypes_sg;
val weak_case_congs_of = weak_case_congs_of_sg o Theory.sign_of;
@@ -144,8 +144,8 @@
let val frees = map dest_Free (term_frees Bi)
val params = rename_wrt_term Bi (Logic.strip_params Bi);
in case assoc (frees @ params, var) of
- None => error ("No such variable in subgoal: " ^ quote var)
- | Some(Type (tn, _)) => tn
+ NONE => error ("No such variable in subgoal: " ^ quote var)
+ | SOME(Type (tn, _)) => tn
| _ => error ("Cannot determine type of " ^ quote var)
end;
@@ -156,7 +156,7 @@
val params = Logic.strip_params Bi; (*params of subgoal i*)
val params = rev (rename_wrt_term Bi params); (*as they are printed*)
val (types, sorts) = types_sorts state;
- fun types' (a, ~1) = (case assoc (params, a) of None => types(a, ~1) | sm => sm)
+ fun types' (a, ~1) = (case assoc (params, a) of NONE => types(a, ~1) | sm => sm)
| types' ixn = types ixn;
val (ct, _) = read_def_cterm (sign, types', sorts) [] false (aterm, TypeInfer.logicT);
in case #T (rep_cterm ct) of
@@ -179,8 +179,8 @@
local
-fun prep_var (Var (ixn, _), Some x) = Some (ixn, x)
- | prep_var _ = None;
+fun prep_var (Var (ixn, _), SOME x) = SOME (ixn, x)
+ | prep_var _ = NONE;
fun prep_inst (concl, xs) = (*exception LIST*)
let val vs = InductAttrib.vars_of concl
@@ -194,8 +194,8 @@
val {sign, ...} = Thm.rep_thm state;
val (rule, rule_name) =
(case opt_rule of
- Some r => (r, "Induction rule")
- | None =>
+ SOME r => (r, "Induction rule")
+ | NONE =>
let val tn = find_tname (hd (mapfilter I (flat varss))) Bi
in (#induction (datatype_info_sg_err sign tn), "Induction rule for type " ^ tn) end);
@@ -206,11 +206,11 @@
fun induct_tac s =
gen_induct_tac Tactic.res_inst_tac'
- (map (Library.single o Some) (Syntax.read_idents s), None);
+ (map (Library.single o SOME) (Syntax.read_idents s), NONE);
fun induct_thm_tac th s =
gen_induct_tac Tactic.res_inst_tac'
- ([map Some (Syntax.read_idents s)], Some th);
+ ([map SOME (Syntax.read_idents s)], SOME th);
end;
@@ -223,9 +223,9 @@
(hd (Logic.strip_assums_hyp (hd (Thm.prems_of rule))));
in inst_tac [(ixn, t)] rule i state end;
-fun gen_case_tac inst_tac (t, Some rule) i state =
+fun gen_case_tac inst_tac (t, SOME rule) i state =
case_inst_tac inst_tac t rule i state
- | gen_case_tac inst_tac (t, None) i state =
+ | gen_case_tac inst_tac (t, NONE) i state =
let val tn = infer_tname state i t in
if tn = HOLogic.boolN then inst_tac [(("P", 0), t)] case_split_thm i state
else case_inst_tac inst_tac t
@@ -233,7 +233,7 @@
i state
end handle THM _ => Seq.empty;
-fun case_tac t = gen_case_tac Tactic.res_inst_tac' (t, None);
+fun case_tac t = gen_case_tac Tactic.res_inst_tac' (t, NONE);
@@ -349,34 +349,34 @@
((i', Type (tname1, _)), (j', Type (tname2, _))) =>
if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
(case (constrs_of_sg sg tname1) of
- Some constrs => let val cnames = map (fst o dest_Const) constrs
+ SOME constrs => let val cnames = map (fst o dest_Const) constrs
in if cname1 mem cnames andalso cname2 mem cnames then
let val eq_t = Logic.mk_equals (t, Const ("False", HOLogic.boolT));
val eq_ct = cterm_of sg eq_t;
val Datatype_thy = theory "Datatype";
val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
- map (get_thm Datatype_thy o rpair None)
+ map (get_thm Datatype_thy o rpair NONE)
["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"]
in (case (#distinct (datatype_info_sg_err sg tname1)) of
- QuickAndDirty => Some (Thm.invoke_oracle
+ QuickAndDirty => SOME (Thm.invoke_oracle
Datatype_thy distinctN (sg, ConstrDistinct eq_t))
- | FewConstrs thms => Some (Tactic.prove sg [] [] eq_t (K
+ | FewConstrs thms => SOME (Tactic.prove sg [] [] eq_t (K
(EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
atac 2, resolve_tac thms 1, etac FalseE 1])))
- | ManyConstrs (thm, ss) => Some (Tactic.prove sg [] [] eq_t (K
+ | ManyConstrs (thm, ss) => SOME (Tactic.prove sg [] [] eq_t (K
(EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
full_simp_tac ss 1,
REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
etac FalseE 1]))))
end
- else None
+ else NONE
end
- | None => None)
- else None
- | _ => None)
- | _ => None)
- | distinct_proc sg _ _ = None;
+ | NONE => NONE)
+ else NONE
+ | _ => NONE)
+ | _ => NONE)
+ | distinct_proc sg _ _ = NONE;
val distinct_simproc =
Simplifier.simproc (Theory.sign_of HOL.thy) distinctN ["s = t"] distinct_proc;
@@ -397,52 +397,52 @@
fun dest_case1 (Const ("_case1", _) $ t $ u) = (case strip_comb t of
(Const (s, _), ts) => (Sign.intern_const sg s, ts)
| (Free (s, _), ts) => (Sign.intern_const sg s, ts)
- | _ => case_error "Head is not a constructor" None [t, u], u)
+ | _ => case_error "Head is not a constructor" NONE [t, u], u)
| dest_case1 t = raise TERM ("dest_case1", [t]);
fun dest_case2 (Const ("_case2", _) $ t $ u) = t :: dest_case2 u
| dest_case2 t = [t];
val cases as ((cname, _), _) :: _ = map dest_case1 (dest_case2 u);
val tab = Symtab.dest (get_datatypes_sg sg);
val (cases', default) = (case split_last cases of
- (cases', (("dummy_pattern", []), t)) => (cases', Some t)
- | _ => (cases, None))
+ (cases', (("dummy_pattern", []), t)) => (cases', SOME t)
+ | _ => (cases, NONE))
fun abstr (Free (x, T), body) = Term.absfree (x, T, body)
| abstr (Const ("_constrain", _) $ Free (x, T) $ tT, body) =
Syntax.const Syntax.constrainAbsC $ Term.absfree (x, T, body) $ tT
| abstr (Const ("Pair", _) $ x $ y, body) =
Syntax.const "split" $ abstr (x, abstr (y, body))
- | abstr (t, _) = case_error "Illegal pattern" None [t];
+ | abstr (t, _) = case_error "Illegal pattern" NONE [t];
in case find_first (fn (_, {descr, index, ...}) =>
exists (equal cname o fst) (#3 (snd (nth_elem (index, descr))))) tab of
- None => case_error ("Not a datatype constructor: " ^ cname) None [u]
- | Some (tname, {descr, case_name, index, ...}) =>
+ NONE => case_error ("Not a datatype constructor: " ^ cname) NONE [u]
+ | SOME (tname, {descr, case_name, index, ...}) =>
let
val _ = if exists (equal "dummy_pattern" o fst o fst) cases' then
- case_error "Illegal occurrence of '_' dummy pattern" (Some tname) [u] else ();
+ case_error "Illegal occurrence of '_' dummy pattern" (SOME tname) [u] else ();
val (_, (_, dts, constrs)) = nth_elem (index, descr);
val sorts = map (rpair [] o dest_DtTFree) dts;
fun find_case (cases, (s, dt)) =
(case find_first (equal s o fst o fst) cases' of
- None => (case default of
- None => case_error ("No clause for constructor " ^ s) (Some tname) [u]
- | Some t => (cases, list_abs (map (rpair dummyT) (DatatypeProp.make_tnames
+ NONE => (case default of
+ NONE => case_error ("No clause for constructor " ^ s) (SOME tname) [u]
+ | SOME t => (cases, list_abs (map (rpair dummyT) (DatatypeProp.make_tnames
(map (typ_of_dtyp descr sorts) dt)), t)))
- | Some (c as ((_, vs), t)) =>
+ | SOME (c as ((_, vs), t)) =>
if length dt <> length vs then
case_error ("Wrong number of arguments for constructor " ^ s)
- (Some tname) vs
+ (SOME tname) vs
else (cases \ c, foldr abstr (vs, t)))
val (cases'', fs) = foldl_map find_case (cases', constrs)
in case (cases'', length constrs = length cases', default) of
- ([], true, Some _) =>
- case_error "Extra '_' dummy pattern" (Some tname) [u]
+ ([], true, SOME _) =>
+ case_error "Extra '_' dummy pattern" (SOME tname) [u]
| (_ :: _, _, _) =>
let val extra = distinct (map (fst o fst) cases'')
in case extra \\ map fst constrs of
[] => case_error ("More than one clause for constructor(s) " ^
- commas extra) (Some tname) [u]
+ commas extra) (SOME tname) [u]
| extra' => case_error ("Illegal constructor(s): " ^ commas extra')
- (Some tname) [u]
+ (SOME tname) [u]
end
| _ => list_comb (Syntax.const case_name, fs) $ t
end
@@ -470,8 +470,8 @@
(constrs ~~ fs);
fun count_cases (cs, (_, _, true)) = cs
| count_cases (cs, (cname, (_, body), false)) = (case assoc (cs, body) of
- None => (body, [cname]) :: cs
- | Some cnames => overwrite (cs, (body, cnames @ [cname])));
+ NONE => (body, [cname]) :: cs
+ | SOME cnames => overwrite (cs, (body, cnames @ [cname])));
val cases' = sort (int_ord o Library.swap o pairself (length o snd))
(foldl count_cases ([], cases));
fun mk_case1 (cname, (vs, body), _) = Syntax.const "_case1" $
@@ -660,7 +660,7 @@
Theory.parent_path |>>>
add_and_get_axiomss "inject" new_type_names
(DatatypeProp.make_injs descr sorts);
- val size_thms = if no_size then [] else get_thms thy3 ("size", None);
+ val size_thms = if no_size then [] else get_thms thy3 ("size", NONE);
val (thy4, distinct) = add_and_get_axiomss "distinct" new_type_names
(DatatypeProp.make_distincts new_type_names descr sorts thy3) thy3;
@@ -945,7 +945,7 @@
val dt_info = get_datatypes thy;
val (descr, _) = unfold_datatypes sign dts' sorts dt_info dts' i;
val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
- if err then error ("Nonemptiness check failed for datatype " ^ s)
+ if err then error ("NONEmptiness check failed for datatype " ^ s)
else raise exn;
val descr' = flat descr;
@@ -978,7 +978,7 @@
fun mk_datatype args =
let
- val names = map (fn ((((None, _), t), _), _) => t | ((((Some t, _), _), _), _) => t) args;
+ val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
val specs = map (fn ((((_, vs), t), mx), cons) =>
(vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
in #1 o add_datatype false names specs end;
--- a/src/HOL/Tools/datatype_prop.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/datatype_prop.ML Sun Feb 13 17:15:14 2005 +0100
@@ -45,8 +45,8 @@
let
fun index (x :: xs) tab =
(case assoc (tab, x) of
- None => if x mem xs then (x ^ "1") :: index xs ((x, 2) :: tab) else x :: index xs tab
- | Some i => (x ^ Library.string_of_int i) :: index xs ((x, i + 1) :: tab))
+ NONE => if x mem xs then (x ^ "1") :: index xs ((x, 2) :: tab) else x :: index xs tab
+ | SOME i => (x ^ Library.string_of_int i) :: index xs ((x, i + 1) :: tab))
| index [] _ = [];
in index names [] end;
--- a/src/HOL/Tools/datatype_realizer.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/datatype_realizer.ML Sun Feb 13 17:15:14 2005 +0100
@@ -24,7 +24,7 @@
fun forall_intr_prf (t, prf) =
let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
- in Abst (a, Some T, Proofterm.prf_abstract_over t prf) end;
+ in Abst (a, SOME T, Proofterm.prf_abstract_over t prf) end;
fun prf_of thm =
let val {sign, prop, der = (_, prf), ...} = rep_thm thm
@@ -108,9 +108,9 @@
(descr ~~ recTs ~~ rec_result_Ts ~~ rec_names);
val r = if null is then Extraction.nullt else
foldr1 HOLogic.mk_prod (mapfilter (fn (((((i, _), T), U), s), tname) =>
- if i mem is then Some
+ if i mem is then SOME
(list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
- else None) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
+ else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
(map (fn ((((i, _), T), U), tname) =>
make_pred i U T (mk_proj i is r) (Free (tname, T)))
@@ -149,10 +149,10 @@
(case head_of (strip_abs_body f) of
Free (s, T) =>
let val T' = Type.varifyT T
- in Abst (s, Some T', Proofterm.prf_abstract_over
- (Var ((s, 0), T')) (AbsP ("H", Some p, prf)))
+ in Abst (s, SOME T', Proofterm.prf_abstract_over
+ (Var ((s, 0), T')) (AbsP ("H", SOME p, prf)))
end
- | _ => AbsP ("H", Some p, prf)))
+ | _ => AbsP ("H", SOME p, prf)))
(rec_fns ~~ prems_of thm, Proofterm.proof_combP
(prf_of thm', map PBound (length prems - 1 downto 0))));
@@ -183,7 +183,7 @@
list_comb (r, free_ts)))))
end;
- val Some (_, _, constrs) = assoc (descr, index);
+ val SOME (_, _, constrs) = assoc (descr, index);
val T = nth_elem (index, get_rec_types descr sorts);
val (rs, prems) = split_list (map (make_casedist_prem T) constrs);
val r = Const (case_name, map fastype_of rs ---> T --> rT);
@@ -210,7 +210,7 @@
val P = Var (("P", 0), rT' --> HOLogic.boolT);
val prf = forall_intr_prf (y, forall_intr_prf (P,
foldr (fn ((p, r), prf) =>
- forall_intr_prf (Logic.varify r, AbsP ("H", Some (Logic.varify p),
+ forall_intr_prf (Logic.varify r, AbsP ("H", SOME (Logic.varify p),
prf))) (prems ~~ rs, Proofterm.proof_combP (prf_of thm',
map PBound (length prems - 1 downto 0)))));
val r' = Logic.varify (Abs ("y", Type.varifyT T,
--- a/src/HOL/Tools/datatype_rep_proofs.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/datatype_rep_proofs.ML Sun Feb 13 17:15:14 2005 +0100
@@ -57,7 +57,7 @@
val [In0_inject, In1_inject, Scons_inject, Leaf_inject,
In0_eq, In1_eq, In0_not_In1, In1_not_In0,
- Lim_inject, Suml_inject, Sumr_inject] = map (get_thm Datatype_thy o rpair None)
+ Lim_inject, Suml_inject, Sumr_inject] = map (get_thm Datatype_thy o rpair NONE)
["In0_inject", "In1_inject", "Scons_inject", "Leaf_inject",
"In0_eq", "In1_eq", "In0_not_In1", "In1_not_In0",
"Lim_inject", "Suml_inject", "Sumr_inject"];
@@ -181,7 +181,7 @@
val thy3 = add_path flat_names big_name (foldl (fn (thy, ((((name, mx), tvs), c), name')) =>
setmp TypedefPackage.quiet_mode true
- (TypedefPackage.add_typedef_i false (Some name') (name, tvs, mx) c None
+ (TypedefPackage.add_typedef_i false (SOME name') (name, tvs, mx) c NONE
(rtac exI 1 THEN
QUIET_BREADTH_FIRST (has_fewer_prems 1)
(resolve_tac rep_intrs 1))) thy |> #1)
@@ -261,9 +261,9 @@
(* get axioms from theory *)
val newT_iso_axms = map (fn s =>
- (get_thm thy4 ("Abs_" ^ s ^ "_inverse", None),
- get_thm thy4 ("Rep_" ^ s ^ "_inverse", None),
- get_thm thy4 ("Rep_" ^ s, None))) new_type_names;
+ (get_thm thy4 ("Abs_" ^ s ^ "_inverse", NONE),
+ get_thm thy4 ("Rep_" ^ s ^ "_inverse", NONE),
+ get_thm thy4 ("Rep_" ^ s, NONE))) new_type_names;
(*------------------------------------------------*)
(* prove additional theorems: *)
--- a/src/HOL/Tools/inductive_codegen.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/inductive_codegen.ML Sun Feb 13 17:15:14 2005 +0100
@@ -70,13 +70,13 @@
fun get_clauses thy s =
let val {intros, graph, ...} = CodegenData.get thy
in case Symtab.lookup (intros, s) of
- None => (case InductivePackage.get_inductive thy s of
- None => None
- | Some ({names, ...}, {intrs, ...}) => Some (names, preprocess thy intrs))
- | Some _ =>
- let val Some names = find_first
+ NONE => (case InductivePackage.get_inductive thy s of
+ NONE => NONE
+ | SOME ({names, ...}, {intrs, ...}) => SOME (names, preprocess thy intrs))
+ | SOME _ =>
+ let val SOME names = find_first
(fn xs => s mem xs) (Graph.strong_conn graph)
- in Some (names, preprocess thy
+ in SOME (names, preprocess thy
(flat (map (fn s => the (Symtab.lookup (intros, s))) names)))
end
end;
@@ -115,17 +115,17 @@
fun infer_factors sg extra_fs (fs, (optf, t)) =
let fun err s = error (s ^ "\n" ^ Sign.string_of_term sg t)
in (case (optf, strip_comb t) of
- (Some f, (Const (name, _), args)) =>
+ (SOME f, (Const (name, _), args)) =>
(case assoc (extra_fs, name) of
- None => overwrite (fs, (name, if_none
+ NONE => overwrite (fs, (name, if_none
(apsome (mg_factor f) (assoc (fs, name))) f))
- | Some (fs', f') => (mg_factor f (FFix f');
+ | SOME (fs', f') => (mg_factor f (FFix f');
foldl (infer_factors sg extra_fs)
(fs, map (apsome FFix) fs' ~~ args)))
- | (Some f, (Var ((name, _), _), [])) =>
+ | (SOME f, (Var ((name, _), _), [])) =>
overwrite (fs, (name, if_none
(apsome (mg_factor f) (assoc (fs, name))) f))
- | (None, _) => fs
+ | (NONE, _) => fs
| _ => err "Illegal term")
handle Factors => err "Product factor mismatch in"
end;
@@ -145,8 +145,8 @@
fun check t = (case strip_comb t of
(Var _, []) => true
| (Const (s, _), ts) => (case assoc (cnstrs, s) of
- None => false
- | Some i => length ts = i andalso forall check ts)
+ NONE => false
+ | SOME i => length ts = i andalso forall check ts)
| _ => false)
in check end;
@@ -158,9 +158,9 @@
(**** mode inference ****)
fun string_of_mode (iss, is) = space_implode " -> " (map
- (fn None => "X"
- | Some js => enclose "[" "]" (commas (map string_of_int js)))
- (iss @ [Some is]));
+ (fn NONE => "X"
+ | SOME js => enclose "[" "]" (commas (map string_of_int js)))
+ (iss @ [SOME is]));
fun print_modes modes = message ("Inferred modes:\n" ^
space_implode "\n" (map (fn (s, ms) => s ^ ": " ^ commas (map
@@ -198,8 +198,8 @@
let
fun mk_modes name args = flat
(map (fn (m as (iss, is)) => map (Mode o pair m) (cprods (map
- (fn (None, _) => [None]
- | (Some js, arg) => map Some
+ (fn (NONE, _) => [NONE]
+ | (SOME js, arg) => map SOME
(filter (fn Mode ((_, js'), _) => js=js') (modes_of modes arg)))
(iss ~~ args)))) (the (assoc (modes, name))))
@@ -230,18 +230,18 @@
forall is_eqT dupTs
end)
(modes_of modes t handle OPTION => [Mode (([], []), [])])
- | Sidecond t => if term_vs t subset vs then Some (Mode (([], []), []))
- else None) ps);
+ | Sidecond t => if term_vs t subset vs then SOME (Mode (([], []), []))
+ else NONE) ps);
fun check_mode_clause thy arg_vs modes (iss, is) (ts, ps) =
let
val modes' = modes @ mapfilter
- (fn (_, None) => None | (v, Some js) => Some (v, [([], js)]))
+ (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
(arg_vs ~~ iss);
- fun check_mode_prems vs [] = Some vs
+ fun check_mode_prems vs [] = SOME vs
| check_mode_prems vs ps = (case select_mode_prem thy modes' vs ps of
- None => None
- | Some (x, _) => check_mode_prems
+ NONE => NONE
+ | SOME (x, _) => check_mode_prems
(case x of Prem (us, _) => vs union terms_vs us | _ => vs)
(filter_out (equal x) ps));
val (in_ts, in_ts') = partition (is_constrt thy) (fst (get_args is 1 ts));
@@ -251,12 +251,12 @@
forall is_eqT (map snd (duplicates (flat (map term_vTs in_ts)))) andalso
forall (is_eqT o fastype_of) in_ts' andalso
(case check_mode_prems (arg_vs union in_vs) ps of
- None => false
- | Some vs => concl_vs subset vs)
+ NONE => false
+ | SOME vs => concl_vs subset vs)
end;
fun check_modes_pred thy arg_vs preds modes (p, ms) =
- let val Some rs = assoc (preds, p)
+ let val SOME rs = assoc (preds, p)
in (p, filter (fn m => case find_index
(not o check_mode_clause thy arg_vs modes m) rs of
~1 => true
@@ -271,8 +271,8 @@
fun infer_modes thy extra_modes factors arg_vs preds = fixp (fn modes =>
map (check_modes_pred thy arg_vs preds (modes @ extra_modes)) modes)
(map (fn (s, (fs, f)) => (s, cprod (cprods (map
- (fn None => [None]
- | Some f' => map Some (subsets 1 (length f' + 1))) fs),
+ (fn NONE => [NONE]
+ | SOME f' => map SOME (subsets 1 (length f' + 1))) fs),
subsets 1 (length f + 1)))) factors);
(**** code generation ****)
@@ -332,8 +332,8 @@
end;
fun mk_v ((names, vs), s) = (case assoc (vs, s) of
- None => ((names, (s, [s])::vs), s)
- | Some xs => let val s' = variant names s in
+ NONE => ((names, (s, [s])::vs), s)
+ | SOME xs => let val s' = variant names s in
((s'::names, overwrite (vs, (s, s'::xs))), s') end);
fun distinct_v (nvs, Var ((s, 0), T)) =
@@ -369,11 +369,11 @@
(mk_const_id (sign_of thy) s ::
map (space_implode "_" o map string_of_int) (mapfilter I iss @ [is]));
-fun compile_expr thy dep brack (gr, (None, t)) =
+fun compile_expr thy dep brack (gr, (NONE, t)) =
apsnd single (invoke_codegen thy dep brack (gr, t))
- | compile_expr _ _ _ (gr, (Some _, Var ((name, _), _))) =
+ | compile_expr _ _ _ (gr, (SOME _, Var ((name, _), _))) =
(gr, [Pretty.str name])
- | compile_expr thy dep brack (gr, (Some (Mode (mode, ms)), t)) =
+ | compile_expr thy dep brack (gr, (SOME (Mode (mode, ms)), t)) =
let
val (Const (name, _), args) = strip_comb t;
val (gr', ps) = foldl_map
@@ -387,7 +387,7 @@
fun compile_clause thy gr dep all_vs arg_vs modes (iss, is) (ts, ps) =
let
val modes' = modes @ mapfilter
- (fn (_, None) => None | (v, Some js) => Some (v, [([], js)]))
+ (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
(arg_vs ~~ iss);
fun check_constrt ((names, eqs), t) =
@@ -427,7 +427,7 @@
| compile_prems out_ts vs names gr ps =
let
val vs' = distinct (flat (vs :: map term_vs out_ts));
- val Some (p, mode as Some (Mode ((_, js), _))) =
+ val SOME (p, mode as SOME (Mode ((_, js), _))) =
select_mode_prem thy modes' vs' ps;
val ps' = filter_out (equal p) ps;
val ((names', eqs), out_ts') =
@@ -503,28 +503,28 @@
(string * (int list list option list * int list list)) list;
fun lookup_modes gr dep = apfst flat (apsnd flat (ListPair.unzip
- (map ((fn (Some (Modes x), _) => x | _ => ([], [])) o Graph.get_node gr)
+ (map ((fn (SOME (Modes x), _) => x | _ => ([], [])) o Graph.get_node gr)
(Graph.all_preds gr [dep]))));
fun print_factors factors = message ("Factors:\n" ^
space_implode "\n" (map (fn (s, (fs, f)) => s ^ ": " ^
space_implode " -> " (map
- (fn None => "X" | Some f' => string_of_factors [] f')
- (fs @ [Some f]))) factors));
+ (fn NONE => "X" | SOME f' => string_of_factors [] f')
+ (fs @ [SOME f]))) factors));
fun prep_intrs intrs = map (rename_term o #prop o rep_thm o standard) intrs;
fun constrain cs [] = []
| constrain cs ((s, xs) :: ys) = (s, case assoc (cs, s) of
- None => xs
- | Some xs' => xs inter xs') :: constrain cs ys;
+ NONE => xs
+ | SOME xs' => xs inter xs') :: constrain cs ys;
fun mk_extra_defs thy gr dep names ts =
foldl (fn (gr, name) =>
if name mem names then gr
else (case get_clauses thy name of
- None => gr
- | Some (names, intrs) =>
+ NONE => gr
+ | SOME (names, intrs) =>
mk_ind_def thy gr dep names [] [] (prep_intrs intrs)))
(gr, foldr add_term_consts (ts, []))
@@ -539,8 +539,8 @@
fun dest_prem factors (_ $ (p as (Const ("op :", _) $ t $ u))) =
(case assoc (factors, case head_of u of
Const (name, _) => name | Var ((name, _), _) => name) of
- None => Prem (full_split_prod t, u)
- | Some f => Prem (split_prod [] f t, u))
+ NONE => Prem (full_split_prod t, u)
+ | SOME f => Prem (split_prod [] f t, u))
| dest_prem factors (_ $ ((eq as Const ("op =", _)) $ t $ u)) =
Prem ([t, u], eq)
| dest_prem factors (_ $ t) = Sidecond t;
@@ -562,20 +562,20 @@
fun add_prod_factors extra_fs (fs, _ $ (Const ("op :", _) $ t $ u)) =
if check_set (head_of u)
then infer_factors (sign_of thy) extra_fs
- (fs, (Some (FVar (prod_factors [] t)), u))
+ (fs, (SOME (FVar (prod_factors [] t)), u))
else fs
| add_prod_factors _ (fs, _) = fs;
val gr' = mk_extra_defs thy
(Graph.add_edge (hd ids, dep)
- (Graph.new_node (hd ids, (None, "")) gr)) (hd ids) names intrs;
+ (Graph.new_node (hd ids, (NONE, "")) gr)) (hd ids) names intrs;
val (extra_modes, extra_factors) = lookup_modes gr' (hd ids);
val fs = constrain factorcs (map (apsnd dest_factors)
(foldl (add_prod_factors extra_factors) ([], flat (map (fn t =>
Logic.strip_imp_concl t :: Logic.strip_imp_prems t) intrs))));
val factors = mapfilter (fn (name, f) =>
- if name mem arg_vs then None
- else Some (name, (map (curry assoc fs) arg_vs, f))) fs;
+ if name mem arg_vs then NONE
+ else SOME (name, (map (curry assoc fs) arg_vs, f))) fs;
val clauses =
foldl (add_clause (fs @ map (apsnd snd) extra_factors)) ([], intrs);
val modes = constrain modecs
@@ -585,19 +585,19 @@
val (gr'', s) = compile_preds thy gr' (hd ids) (terms_vs intrs) arg_vs
(modes @ extra_modes) clauses;
in
- (Graph.map_node (hd ids) (K (Some (Modes (modes, factors)), s)) gr'')
+ (Graph.map_node (hd ids) (K (SOME (Modes (modes, factors)), s)) gr'')
end
end;
fun find_mode s u modes is = (case find_first (fn Mode ((_, js), _) => is=js)
(modes_of modes u handle OPTION => []) of
- None => error ("No such mode for " ^ s ^ ": " ^ string_of_mode ([], is))
+ NONE => error ("No such mode for " ^ s ^ ": " ^ string_of_mode ([], is))
| mode => mode);
fun mk_ind_call thy gr dep t u is_query = (case head_of u of
Const (s, T) => (case (get_clauses thy s, get_assoc_code thy s T) of
- (None, _) => None
- | (Some (names, intrs), None) =>
+ (NONE, _) => NONE
+ | (SOME (names, intrs), NONE) =>
let
fun mk_mode (((ts, mode), i), Const ("dummy_pattern", _)) =
((ts, mode), i+1)
@@ -615,16 +615,16 @@
(invoke_codegen thy dep false) (gr1, ts');
val (gr3, ps) = compile_expr thy dep false (gr2, (mode, u))
in
- Some (gr3, Pretty.block
+ SOME (gr3, Pretty.block
(ps @ [Pretty.brk 1, mk_tuple in_ps]))
end
- | _ => None)
- | _ => None);
+ | _ => NONE)
+ | _ => NONE);
fun list_of_indset thy gr dep brack u = (case head_of u of
Const (s, T) => (case (get_clauses thy s, get_assoc_code thy s T) of
- (None, _) => None
- | (Some (names, intrs), None) =>
+ (NONE, _) => NONE
+ | (SOME (names, intrs), NONE) =>
let
val gr1 = mk_extra_defs thy
(mk_ind_def thy gr dep names [] [] (prep_intrs intrs)) dep names [u];
@@ -632,7 +632,7 @@
val mode = find_mode s u modes [];
val (gr2, ps) = compile_expr thy dep false (gr1, (mode, u))
in
- Some (gr2, (if brack then parens else I)
+ SOME (gr2, (if brack then parens else I)
(Pretty.block ([Pretty.str "Seq.list_of", Pretty.brk 1,
Pretty.str "("] @
conv_ntuple' (snd (the (assoc (factors, s))))
@@ -640,8 +640,8 @@
(ps @ [Pretty.brk 1, Pretty.str "()"]) @
[Pretty.str ")"])))
end
- | _ => None)
- | _ => None);
+ | _ => NONE)
+ | _ => NONE);
fun clause_of_eqn eqn =
let
@@ -671,7 +671,7 @@
parens (Pretty.block [Pretty.str (modename thy pname ([], mode)),
Pretty.brk 1, mk_tuple vars])]) ^ ";\n\n";
val gr' = mk_ind_def thy (Graph.add_edge (id, dep)
- (Graph.new_node (id, (None, s)) gr)) id [pname]
+ (Graph.new_node (id, (NONE, s)) gr)) id [pname]
[(pname, [([], mode)])]
[(pname, map (fn i => replicate i 2) (0 upto arity-1))]
clauses;
@@ -683,21 +683,21 @@
fun inductive_codegen thy gr dep brack (Const ("op :", _) $ t $ u) =
((case mk_ind_call thy gr dep (Term.no_dummy_patterns t) u false of
- None => None
- | Some (gr', call_p) => Some (gr', (if brack then parens else I)
+ NONE => NONE
+ | SOME (gr', call_p) => SOME (gr', (if brack then parens else I)
(Pretty.block [Pretty.str "?! (", call_p, Pretty.str ")"])))
handle TERM _ => mk_ind_call thy gr dep t u true)
| inductive_codegen thy gr dep brack t = (case strip_comb t of
(Const (s, _), ts) => (case Symtab.lookup (#eqns (CodegenData.get thy), s) of
- None => list_of_indset thy gr dep brack t
- | Some eqns =>
+ NONE => list_of_indset thy gr dep brack t
+ | SOME eqns =>
let
val gr' = mk_fun thy s (preprocess thy eqns) dep gr
val (gr'', ps) = foldl_map (invoke_codegen thy dep true) (gr', ts);
- in Some (gr'', mk_app brack (Pretty.str (mk_const_id
+ in SOME (gr'', mk_app brack (Pretty.str (mk_const_id
(sign_of thy) s)) ps)
end)
- | _ => None);
+ | _ => NONE);
val setup =
[add_codegen "inductive" inductive_codegen,
--- a/src/HOL/Tools/inductive_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/inductive_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -116,8 +116,8 @@
fun the_inductive thy name =
(case get_inductive thy name of
- None => error ("Unknown (co)inductive set " ^ quote name)
- | Some info => info);
+ NONE => error ("Unknown (co)inductive set " ^ quote name)
+ | SOME info => info);
val the_mk_cases = (#mk_cases o #2) oo the_inductive;
@@ -387,20 +387,20 @@
fun subst (s as ((m as Const ("op :", T)) $ t $ u)) =
(case pred_of u of
- None => (m $ fst (subst t) $ fst (subst u), None)
- | Some P => (HOLogic.mk_binop inductive_conj_name (s, P $ t), Some (s, P $ t)))
+ NONE => (m $ fst (subst t) $ fst (subst u), NONE)
+ | SOME P => (HOLogic.mk_binop inductive_conj_name (s, P $ t), SOME (s, P $ t)))
| subst s =
(case pred_of s of
- Some P => (HOLogic.mk_binop "op Int"
+ SOME P => (HOLogic.mk_binop "op Int"
(s, HOLogic.Collect_const (HOLogic.dest_setT
- (fastype_of s)) $ P), None)
- | None => (case s of
- (t $ u) => (fst (subst t) $ fst (subst u), None)
- | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), None)
- | _ => (s, None)));
+ (fastype_of s)) $ P), NONE)
+ | NONE => (case s of
+ (t $ u) => (fst (subst t) $ fst (subst u), NONE)
+ | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE)
+ | _ => (s, NONE)));
fun mk_prem (s, prems) = (case subst s of
- (_, Some (t, u)) => t :: u :: prems
+ (_, SOME (t, u)) => t :: u :: prems
| (t, _) => t :: prems);
val Const ("op :", _) $ t $ u =
@@ -562,8 +562,8 @@
fun mk_elim rl = Drule.standard (Tactic.rule_by_tactic tac (prem RS rl));
in
(case get_first (try mk_elim) elims of
- Some r => r
- | None => error (Pretty.string_of (Pretty.block
+ SOME r => r
+ | NONE => error (Pretty.string_of (Pretty.block
[Pretty.str mk_cases_err, Pretty.fbrk, Display.pretty_cterm cprop])))
end;
@@ -617,11 +617,11 @@
val sign = Theory.sign_of thy;
val sum_case_rewrites =
- (if PureThy.get_name thy = "Datatype" then PureThy.get_thms thy ("sum.cases", None)
+ (if PureThy.get_name thy = "Datatype" then PureThy.get_thms thy ("sum.cases", NONE)
else
(case ThyInfo.lookup_theory "Datatype" of
- None => []
- | Some thy' => PureThy.get_thms thy' ("sum.cases", None))) |> map mk_meta_eq;
+ NONE => []
+ | SOME thy' => PureThy.get_thms thy' ("sum.cases", NONE))) |> map mk_meta_eq;
val (preds, ind_prems, mutual_ind_concl, factors) =
mk_indrule cs cTs params intr_ts;
@@ -817,8 +817,8 @@
fun try_term f msg sign t =
(case Library.try f t of
- Some x => x
- | None => error (msg ^ Sign.string_of_term sign t));
+ SOME x => x
+ | NONE => error (msg ^ Sign.string_of_term sign t));
fun add_inductive_i verbose declare_consts alt_name coind no_elim no_ind cs pre_intros monos thy =
let
--- a/src/HOL/Tools/inductive_realizer.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/inductive_realizer.ML Sun Feb 13 17:15:14 2005 +0100
@@ -23,7 +23,7 @@
fun forall_intr_prf (t, prf) =
let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
- in Abst (a, Some T, Proofterm.prf_abstract_over t prf) end;
+ in Abst (a, SOME T, Proofterm.prf_abstract_over t prf) end;
fun subsets [] = [[]]
| subsets (x::xs) =
@@ -75,7 +75,7 @@
fun gen_rvar vs (t as Var ((a, 0), T)) =
let val U = TVar (("'" ^ a, 0), HOLogic.typeS)
in case try HOLogic.dest_setT T of
- None => if body_type T <> HOLogic.boolT then t else
+ NONE => if body_type T <> HOLogic.boolT then t else
let
val Ts = binder_types T;
val i = length Ts;
@@ -86,7 +86,7 @@
list_abs (("r", U) :: xs, mk_rlz U $ Bound i $ u)
else list_abs (xs, mk_rlz Extraction.nullT $ Extraction.nullt $ u)
end
- | Some T' => if a mem vs then
+ | SOME T' => if a mem vs then
Abs ("r", U, Abs ("x", T', mk_rlz U $ Bound 1 $
(HOLogic.mk_mem (Bound 0, t))))
else Abs ("x", T', mk_rlz Extraction.nullT $ Extraction.nullt $
@@ -112,7 +112,7 @@
val rname = space_implode "_" (s ^ "R" :: vs);
fun mk_Tprem n v =
- let val Some T = assoc (rvs, v)
+ let val SOME T = assoc (rvs, v)
in (Const ("typeof", T --> Type ("Type", [])) $ Var ((v, 0), T),
Extraction.mk_typ (if n then Extraction.nullT
else TVar (("'" ^ v, 0), HOLogic.typeS)))
@@ -199,11 +199,11 @@
let
val concls = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct));
val premss = mapfilter (fn (s, rs) => if s mem rsets then
- Some (map (fn r => nth_elem (find_index_eq (prop_of r) (map prop_of intrs),
- prems_of raw_induct)) rs) else None) rss;
+ SOME (map (fn r => nth_elem (find_index_eq (prop_of r) (map prop_of intrs),
+ prems_of raw_induct)) rs) else NONE) rss;
val concls' = mapfilter (fn (s, _) => if s mem rsets then
find_first (fn concl => s mem term_consts concl) concls
- else None) rss;
+ else NONE) rss;
val fs = flat (snd (foldl_map (fn (intrs, (prems, dummy)) =>
let
val (intrs1, intrs2) = splitAt (length prems, intrs);
@@ -217,8 +217,8 @@
val Ts = map fastype_of fs;
val rlzs = mapfilter (fn (a, concl) =>
let val T = Extraction.etype_of thy vs [] concl
- in if T = Extraction.nullT then None
- else Some (list_comb (Const (a, Ts ---> T), fs))
+ in if T = Extraction.nullT then NONE
+ else SOME (list_comb (Const (a, Ts ---> T), fs))
end) (rec_names ~~ concls')
in if null rlzs then Extraction.nullt else
let
@@ -259,8 +259,8 @@
fun mk_prf _ [] prf = prf
| mk_prf rs ((prem, rprem) :: prems) prf =
if Extraction.etype_of thy vs [] prem = Extraction.nullT
- then AbsP ("H", Some rprem, mk_prf rs prems prf)
- else forall_intr_prf (Var (hd rs), AbsP ("H", Some rprem,
+ then AbsP ("H", SOME rprem, mk_prf rs prems prf)
+ else forall_intr_prf (Var (hd rs), AbsP ("H", SOME rprem,
mk_prf (tl rs) prems prf));
in (Thm.name_of_thm rule, (vs,
@@ -301,13 +301,13 @@
(s, replicate ar HOLogic.typeS, HOLogic.typeS)) tnames) |>
Extraction.add_typeof_eqns_i ty_eqs;
val dts = mapfilter (fn (s, rs) => if s mem rsets then
- Some (dt_of_intrs thy1' vs rs) else None) rss;
+ SOME (dt_of_intrs thy1' vs rs) else NONE) rss;
(** datatype representing computational content of inductive set **)
val (thy2, (dummies, dt_info)) = thy1 |>
- (if null dts then rpair ([], None) else
- apsnd (apsnd Some) o add_dummies (DatatypePackage.add_datatype_i false false
+ (if null dts then rpair ([], NONE) else
+ apsnd (apsnd SOME) o add_dummies (DatatypePackage.add_datatype_i false false
(map #2 dts)) (map (pair false) dts) []) |>>
Extraction.add_typeof_eqns_i ty_eqs |>>
Extraction.add_realizes_eqns_i rlz_eqs;
@@ -346,7 +346,7 @@
(** realizer for induction rule **)
val Ps = mapfilter (fn _ $ M $ P => if set_of M mem rsets then
- Some (fst (fst (dest_Var (head_of P)))) else None)
+ SOME (fst (fst (dest_Var (head_of P)))) else NONE)
(HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct)));
fun add_ind_realizer (thy, Ps) =
@@ -434,7 +434,7 @@
val elimps = mapfilter (fn (s, intrs) => if s mem rsets then
apsome (rpair intrs) (find_first (fn (thm, _) =>
s mem term_consts (hd (prems_of thm))) (elims ~~ #elims ind_info))
- else None) rss;
+ else NONE) rss;
val thy6 = foldl (fn (thy, p as (((((elim, _), _), _), _), _)) => thy |>
add_elim_realizer [] p |> add_elim_realizer [fst (fst (dest_Var
(HOLogic.dest_Trueprop (concl_of elim))))] p) (thy5,
@@ -446,8 +446,8 @@
let
val (_, {intrs, induct, raw_induct, elims, ...}) =
(case InductivePackage.get_inductive thy name of
- None => error ("Unknown inductive set " ^ quote name)
- | Some info => info);
+ NONE => error ("Unknown inductive set " ^ quote name)
+ | SOME info => info);
val _ $ (_ $ _ $ S) = concl_of (hd intrs);
val vss = sort (int_ord o pairself length)
(subsets (map fst (relevant_vars S)))
@@ -465,8 +465,8 @@
handle TERM _ => err () | LIST _ => err ();
in
(add_ind_realizers (hd sets) (case arg of
- None => sets | Some None => []
- | Some (Some sets') => sets \\ map (Sign.intern_const (sign_of thy)) sets')
+ NONE => sets | SOME NONE => []
+ | SOME (SOME sets') => sets \\ map (Sign.intern_const (sign_of thy)) sets')
thy, thm)
end;
--- a/src/HOL/Tools/meson.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/meson.ML Sun Feb 13 17:15:14 2005 +0100
@@ -44,8 +44,8 @@
(*Permits forward proof from rules that discharge assumptions*)
fun forward_res nf st =
case Seq.pull (ALLGOALS (METAHYPS (fn [prem] => rtac (nf prem) 1)) st)
- of Some(th,_) => th
- | None => raise THM("forward_res", 0, [st]);
+ of SOME(th,_) => th
+ | NONE => raise THM("forward_res", 0, [st]);
(*Are any of the constants in "bs" present in the term?*)
@@ -126,8 +126,8 @@
(REPEAT
(METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1)
st)
- of Some(th,_) => th
- | None => raise THM("forward_res2", 0, [st]);
+ of SOME(th,_) => th
+ | NONE => raise THM("forward_res2", 0, [st]);
(*Remove duplicates in P|Q by assuming ~P in Q
rls (initially []) accumulates assumptions of the form P==>False*)
--- a/src/HOL/Tools/primrec_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/primrec_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -73,9 +73,9 @@
check_vars "extra variables on rhs: "
(map dest_Free (term_frees rhs) \\ lfrees);
case assoc (rec_fns, fname) of
- None =>
+ NONE =>
(fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns
- | Some (_, rpos', eqns) =>
+ | SOME (_, rpos', eqns) =>
if is_some (assoc (eqns, cname)) then
raise RecError "constructor already occurred as pattern"
else if rpos <> rpos' then
@@ -113,11 +113,11 @@
val (x, xs) = strip_comb x'
in
(case assoc (subs, x) of
- None =>
+ NONE =>
let
val (fs', ts') = foldl_map (subst subs) (fs, ts)
in (fs', list_comb (f, ts')) end
- | Some (i', y) =>
+ | SOME (i', y) =>
let
val (fs', ts') = foldl_map (subst subs) (fs, xs @ ls @ rs);
val fs'' = process_fun sign descr rec_eqns ((i', fname'), fs')
@@ -135,10 +135,10 @@
fun trans eqns ((cname, cargs), (fnames', fnss', fns)) =
(case assoc (eqns, cname) of
- None => (warning ("No equation for constructor " ^ quote cname ^
+ NONE => (warning ("No equation for constructor " ^ quote cname ^
"\nin definition of function " ^ quote fname);
(fnames', fnss', (Const ("arbitrary", dummyT))::fns))
- | Some (ls, cargs', rs, rhs, eq) =>
+ | SOME (ls, cargs', rs, rhs, eq) =>
let
val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
val rargs = map fst recs;
@@ -155,7 +155,7 @@
end)
in (case assoc (fnames, i) of
- None =>
+ NONE =>
if exists (equal fname o snd) fnames then
raise RecError ("inconsistent functions for datatype " ^ quote tname)
else
@@ -166,7 +166,7 @@
in
(fnames', (i, (fname, #1 (snd (hd eqns)), fns))::fnss')
end
- | Some fname' =>
+ | SOME fname' =>
if fname = fname' then (fnames, fnss)
else raise RecError ("inconsistent functions for datatype " ^ quote tname))
end;
@@ -176,7 +176,7 @@
fun get_fns fns (((i, (tname, _, constrs)), rec_name), (fs, defs)) =
case assoc (fns, i) of
- None =>
+ NONE =>
let
val dummy_fns = map (fn (_, cargs) => Const ("arbitrary",
replicate ((length cargs) + (length (filter is_rec_type cargs)))
@@ -185,7 +185,7 @@
in
(dummy_fns @ fs, defs)
end
- | Some (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname)::defs);
+ | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname)::defs);
(* make definition *)
@@ -206,8 +206,8 @@
fun find_dts (dt_info : datatype_info Symtab.table) _ [] = []
| find_dts dt_info tnames' (tname::tnames) =
(case Symtab.lookup (dt_info, tname) of
- None => primrec_err (quote tname ^ " is not a datatype")
- | Some dt =>
+ NONE => primrec_err (quote tname ^ " is not a datatype")
+ | SOME dt =>
if tnames' subset (map (#1 o snd) (#descr dt)) then
(tname, dt)::(find_dts dt_info tnames' tnames)
else find_dts dt_info tnames' tnames);
--- a/src/HOL/Tools/recdef_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/recdef_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -210,8 +210,8 @@
val ctxt0 = ProofContext.init thy;
val ctxt =
(case opt_src of
- None => ctxt0
- | Some src => Method.only_sectioned_args recdef_modifiers I src ctxt0);
+ NONE => ctxt0
+ | SOME src => Method.only_sectioned_args recdef_modifiers I src ctxt0);
val {simps, congs, wfs} = get_local_hints ctxt;
val cs = local_claset_of ctxt;
val ss = local_simpset_of ctxt addsimps simps;
@@ -310,8 +310,8 @@
val atts = map (prep_att thy) raw_atts;
val tcs =
(case get_recdef thy name of
- None => error ("No recdef definition of constant: " ^ quote name)
- | Some {tcs, ...} => tcs);
+ NONE => error ("No recdef definition of constant: " ^ quote name)
+ | SOME {tcs, ...} => tcs);
val i = if_none opt_i 1;
val tc = Library.nth_elem (i - 1, tcs) handle Library.LIST _ =>
error ("No termination condition #" ^ string_of_int i ^
--- a/src/HOL/Tools/recfun_codegen.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/recfun_codegen.ML Sun Feb 13 17:15:14 2005 +0100
@@ -53,8 +53,8 @@
val tab = CodegenData.get thy;
val (s, _) = const_of (prop_of thm);
in case Symtab.lookup (tab, s) of
- None => p
- | Some thms => (CodegenData.put (Symtab.update ((s,
+ NONE => p
+ | SOME thms => (CodegenData.put (Symtab.update ((s,
gen_rem eq_thm (thms, thm)), tab)) thy, thm)
end handle TERM _ => (warn thm; p);
@@ -68,14 +68,14 @@
fun get_equations thy (s, T) =
(case Symtab.lookup (CodegenData.get thy, s) of
- None => []
- | Some thms => preprocess thy (del_redundant thy []
+ NONE => []
+ | SOME thms => preprocess thy (del_redundant thy []
(filter (fn thm => is_instance thy T
(snd (const_of (prop_of thm)))) thms)));
fun mk_poly_id thy (s, T) = mk_const_id (sign_of thy) s ^
(case get_defn thy s T of
- Some (_, Some i) => "_def" ^ string_of_int i
+ SOME (_, SOME i) => "_def" ^ string_of_int i
| _ => "");
exception EQN of string * typ * string;
@@ -104,19 +104,19 @@
end;
fun put_code fundef = Graph.map_node dname
- (K (Some (EQN ("", dummyT, dname)), Pretty.string_of (Pretty.blk (0,
+ (K (SOME (EQN ("", dummyT, dname)), Pretty.string_of (Pretty.blk (0,
separate Pretty.fbrk fundef @ [Pretty.str ";"])) ^ "\n\n"));
in
(case try (Graph.get_node gr) dname of
- None =>
+ NONE =>
let
val gr1 = Graph.add_edge (dname, dep)
- (Graph.new_node (dname, (Some (EQN (s, T, "")), "")) gr);
+ (Graph.new_node (dname, (SOME (EQN (s, T, "")), "")) gr);
val (gr2, fundef) = mk_fundef "" "fun " gr1 eqs';
val xs = cycle gr2 ([], dname);
val cs = map (fn x => case Graph.get_node gr2 x of
- (Some (EQN (s, T, _)), _) => (s, T)
+ (SOME (EQN (s, T, _)), _) => (s, T)
| _ => error ("RecfunCodegen: illegal cyclic dependencies:\n" ^
implode (separate ", " xs))) xs
in (case xs of
@@ -129,12 +129,12 @@
val (gr3, fundef') = mk_fundef "" "fun "
(Graph.add_edge (dname, dep)
(foldr (uncurry Graph.new_node) (map (fn k =>
- (k, (Some (EQN ("", dummyT, dname)), ""))) xs,
+ (k, (SOME (EQN ("", dummyT, dname)), ""))) xs,
Graph.del_nodes xs gr2))) eqs''
in put_code fundef' gr3 end
else gr2)
end
- | Some (Some (EQN (_, _, s)), _) =>
+ | SOME (SOME (EQN (_, _, s)), _) =>
if s = "" then
if dname = dep then gr else Graph.add_edge (dname, dep) gr
else if s = dep then gr else Graph.add_edge (s, dep) gr)
@@ -142,15 +142,15 @@
fun recfun_codegen thy gr dep brack t = (case strip_comb t of
(Const (p as (s, T)), ts) => (case (get_equations thy p, get_assoc_code thy s T) of
- ([], _) => None
- | (_, Some _) => None
- | (eqns, None) =>
+ ([], _) => NONE
+ | (_, SOME _) => NONE
+ | (eqns, NONE) =>
let val (gr', ps) = foldl_map (invoke_codegen thy dep true) (gr, ts)
in
- Some (add_rec_funs thy gr' dep (map prop_of eqns),
+ SOME (add_rec_funs thy gr' dep (map prop_of eqns),
mk_app brack (Pretty.str (mk_poly_id thy p)) ps)
end)
- | _ => None);
+ | _ => NONE);
val setup =
--- a/src/HOL/Tools/reconstruction.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/reconstruction.ML Sun Feb 13 17:15:14 2005 +0100
@@ -149,11 +149,11 @@
case Thm.name_of_thm th of
"" => ResAxioms.meta_cnf_axiom th (*no name, so can't cache*)
| s => case Symtab.lookup (!cc,s) of
- None =>
+ NONE =>
let val cls = ResAxioms.meta_cnf_axiom th
in cc := Symtab.update ((s, (th,cls)), !cc); cls
end
- | Some(th',cls) =>
+ | SOME(th',cls) =>
if eq_thm(th,th') then cls
else (*New theorem stored under the same name? Possible??*)
let val cls = ResAxioms.meta_cnf_axiom th
--- a/src/HOL/Tools/record_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/record_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -82,8 +82,8 @@
(*** utilities ***)
fun but_last xs = fst (split_last xs);
-fun list None = []
- | list (Some x) = [x]
+fun list NONE = []
+ | list (SOME x) = [x]
(* messages *)
@@ -159,12 +159,12 @@
fun dest_recT (typ as Type (c_ext_type, Ts as (T::_))) =
(case try (unsuffix ext_typeN) c_ext_type of
- None => raise TYPE ("RecordPackage.dest_recT", [typ], [])
- | Some c => ((c, Ts), last_elem Ts))
+ NONE => raise TYPE ("RecordPackage.dest_recT", [typ], [])
+ | SOME c => ((c, Ts), last_elem Ts))
| dest_recT typ = raise TYPE ("RecordPackage.dest_recT", [typ], []);
fun is_recT T =
- (case try dest_recT T of None => false | Some _ => true);
+ (case try dest_recT T of NONE => false | SOME _ => true);
fun dest_recTs T =
let val ((c, Ts), U) = dest_recT T
@@ -174,9 +174,9 @@
fun last_extT T =
let val ((c, Ts), U) = dest_recT T
in (case last_extT U of
- None => Some (c,Ts)
- | Some l => Some l)
- end handle TYPE _ => None
+ NONE => SOME (c,Ts)
+ | SOME l => SOME l)
+ end handle TYPE _ => NONE
fun rec_id i T =
let val rTs = dest_recTs T
@@ -282,8 +282,8 @@
let
val prt_typ = Sign.pretty_typ sg;
- fun pretty_parent None = []
- | pretty_parent (Some (Ts, name)) =
+ fun pretty_parent NONE = []
+ | pretty_parent (SOME (Ts, name)) =
[Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]];
fun pretty_field (c, T) = Pretty.block
@@ -319,8 +319,8 @@
fun get_selectors sg name = Symtab.lookup (#selectors (get_sel_upd sg), name);
fun is_selector sg name =
case get_selectors sg (Sign.intern_const sg name) of
- None => false
- | Some _ => true
+ NONE => false
+ | SOME _ => true
fun get_updates sg name = Symtab.lookup (#updates (get_sel_upd sg), name);
@@ -399,8 +399,8 @@
(* extension of a record name *)
fun get_extension sg name =
case Symtab.lookup (#records (RecordsData.get_sg sg),name) of
- Some s => Some (#extension s)
- | None => None;
+ SOME s => SOME (#extension s)
+ | NONE => NONE;
(* access 'extfields' *)
@@ -461,18 +461,18 @@
(* parent records *)
-fun add_parents thy None parents = parents
- | add_parents thy (Some (types, name)) parents =
+fun add_parents thy NONE parents = parents
+ | add_parents thy (SOME (types, name)) parents =
let
val sign = Theory.sign_of thy;
fun err msg = error (msg ^ " parent record " ^ quote name);
val {args, parent, fields, extension, induct} =
- (case get_record thy name of Some info => info | None => err "Unknown");
+ (case get_record thy name of SOME info => info | NONE => err "Unknown");
val _ = if length types <> length args then err "Bad number of arguments for" else ();
fun bad_inst ((x, S), T) =
- if Sign.of_sort sign (T, S) then None else Some x
+ if Sign.of_sort sign (T, S) then NONE else SOME x
val bads = mapfilter bad_inst (args ~~ types);
val inst = map fst args ~~ types;
@@ -537,16 +537,16 @@
fun mk_ext (fargs as (name,arg)::_) =
(case get_fieldext sg (Sign.intern_const sg name) of
- Some (ext,_) => (case get_extfields sg ext of
- Some flds
+ SOME (ext,_) => (case get_extfields sg ext of
+ SOME flds
=> let val (args,rest) =
splitargs (map fst (but_last flds)) fargs;
val more' = mk_ext rest;
in list_comb (Syntax.const (suffix sfx ext),args@[more'])
end
- | None => raise TERM(msg ^ "no fields defined for "
+ | NONE => raise TERM(msg ^ "no fields defined for "
^ ext,[t]))
- | None => raise TERM (msg ^ name ^" is no proper field",[t]))
+ | NONE => raise TERM (msg ^ name ^" is no proper field",[t]))
| mk_ext [] = more
in mk_ext fieldargs end;
@@ -565,8 +565,8 @@
| splitargs _ _ = ([],[]);
fun get_sort xs n = (case assoc (xs,n) of
- Some s => s
- | None => Sign.defaultS sg);
+ SOME s => s
+ | NONE => Sign.defaultS sg);
fun to_type t = Sign.intern_typ sg
(Syntax.typ_of_term (get_sort (Syntax.raw_term_sorts t)) I t);
@@ -575,9 +575,9 @@
fun mk_ext (fargs as (name,arg)::_) =
(case get_fieldext sg (Sign.intern_const sg name) of
- Some (ext,alphas) =>
+ SOME (ext,alphas) =>
(case get_extfields sg ext of
- Some flds
+ SOME flds
=> (let
val flds' = but_last flds;
val types = map snd flds';
@@ -593,8 +593,8 @@
in list_comb (Syntax.const (suffix sfx ext),alphas'@[more'])
end handle TUNIFY => raise
TERM (msg ^ "type is no proper record (extension)", [t]))
- | None => raise TERM (msg ^ "no fields defined for " ^ ext,[t]))
- | None => raise TERM (msg ^ name ^" is no proper field",[t]))
+ | NONE => raise TERM (msg ^ "no fields defined for " ^ ext,[t]))
+ | NONE => raise TERM (msg ^ name ^" is no proper field",[t]))
| mk_ext [] = more
in mk_ext fieldargs end;
@@ -642,9 +642,9 @@
fun gen_field_upds_tr' mark sfx (tm as Const (name_field, _) $ t $ u) =
(case try (unsuffix sfx) name_field of
- Some name =>
+ SOME name =>
apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_field_upds_tr' mark sfx u)
- | None => ([], tm))
+ | NONE => ([], tm))
| gen_field_upds_tr' _ _ tm = ([], tm);
fun record_update_tr' tm =
@@ -663,17 +663,17 @@
(case strip_comb t of
(Const (ext,_),args)
=> (case try (unsuffix extN) (Sign.intern_const sg ext) of
- Some ext'
+ SOME ext'
=> (case get_extfields sg ext' of
- Some flds
+ SOME flds
=> (let
val (f::fs) = but_last (map fst flds);
val flds' = Sign.extern sg Sign.constK f::map NameSpace.base fs;
val (args',more) = split_last args;
in (flds'~~args')@field_lst more end
handle LIST _ => [("",t)])
- | None => [("",t)])
- | None => [("",t)])
+ | NONE => [("",t)])
+ | NONE => [("",t)])
| _ => [("",t)])
val (flds,(_,more)) = split_last (field_lst t);
@@ -705,8 +705,8 @@
(* type from tm so that we can continue on the type level rather then the term level.*)
fun get_sort xs n = (case assoc (xs,n) of
- Some s => s
- | None => Sign.defaultS sg);
+ SOME s => s
+ | NONE => Sign.defaultS sg);
(* WORKAROUND:
* If a record type occurs in an error message of type inference there
@@ -734,7 +734,7 @@
in if !print_record_type_abbr
then (case last_extT T of
- Some (name,_)
+ SOME (name,_)
=> if name = lastExt
then
(let
@@ -752,8 +752,8 @@
fun record_type_tr' sep mark record record_scheme sg t =
let
fun get_sort xs n = (case assoc (xs,n) of
- Some s => s
- | None => Sign.defaultS sg);
+ SOME s => s
+ | NONE => Sign.defaultS sg);
val T = Sign.intern_typ sg (Syntax.typ_of_term (get_sort (Syntax.raw_term_sorts t)) I t)
@@ -766,11 +766,11 @@
(case T of
Type (ext,args)
=> (case try (unsuffix ext_typeN) ext of
- Some ext'
+ SOME ext'
=> (case get_extfields sg ext' of
- Some flds
+ SOME flds
=> (case get_fieldext sg (fst (hd flds)) of
- Some (_,alphas)
+ SOME (_,alphas)
=> (let
val (f::fs) = but_last flds;
val flds' = apfst (Sign.extern sg Sign.constK) f
@@ -783,9 +783,9 @@
in flds''@field_lst more end
handle TUNIFY => [("",T)]
| LIST _=> [("",T)])
- | None => [("",T)])
- | None => [("",T)])
- | None => [("",T)])
+ | NONE => [("",T)])
+ | NONE => [("",T)])
+ | NONE => [("",T)])
| _ => [("",T)])
val (flds,(_,moreT)) = split_last (field_lst T);
@@ -843,10 +843,10 @@
foldl (fn (thms,(n,_)) => (list (Symtab.lookup (extsplit,n)))@thms)
([],dest_recTs T);
val thms = (case get_splits sg (rec_id (~1) T) of
- Some (all_thm,_,_,_) =>
+ SOME (all_thm,_,_,_) =>
all_thm::(case extsplits of [thm] => [] | _ => extsplits)
(* [thm] is the same as all_thm *)
- | None => extsplits)
+ | NONE => extsplits)
in (quick_and_dirty_prove true sg [] prop (fn _ => (simp_tac (simpset addsimps thms) 1)))
end;
@@ -873,53 +873,53 @@
Simplifier.simproc (Theory.sign_of HOL.thy) "record_simp" ["s (u k r)"]
(fn sg => fn _ => fn t =>
(case t of (sel as Const (s, Type (_,[domS,rangeS])))$((upd as Const (u, _)) $ k $ r)=>
- (case get_selectors sg s of Some () =>
- (case get_updates sg u of Some u_name =>
+ (case get_selectors sg s of SOME () =>
+ (case get_updates sg u of SOME u_name =>
let
fun mk_abs_var x t = (x, fastype_of t);
val {sel_upd={updates,...},extfields,...} = RecordsData.get_sg sg;
fun mk_eq_terms ((upd as Const (u,Type(_,[updT,_]))) $ k $ r) =
(case (Symtab.lookup (updates,u)) of
- None => None
- | Some u_name
+ NONE => NONE
+ | SOME u_name
=> if u_name = s
then let
val rv = mk_abs_var "r" r
val rb = Bound 0
val kv = mk_abs_var "k" k
val kb = Bound 1
- in Some (upd$kb$rb,kb,[kv,rv],true) end
+ in SOME (upd$kb$rb,kb,[kv,rv],true) end
else if u_name mem (map fst (get_fields extfields rangeS))
orelse s mem (map fst (get_fields extfields updT))
- then None
+ then NONE
else (case mk_eq_terms r of
- Some (trm,trm',vars,update_s)
+ SOME (trm,trm',vars,update_s)
=> let
val kv = mk_abs_var "k" k
val kb = Bound (length vars)
- in Some (upd$kb$trm,trm',kv::vars,update_s) end
- | None
+ in SOME (upd$kb$trm,trm',kv::vars,update_s) end
+ | NONE
=> let
val rv = mk_abs_var "r" r
val rb = Bound 0
val kv = mk_abs_var "k" k
val kb = Bound 1
- in Some (upd$kb$rb,rb,[kv,rv],false) end))
- | mk_eq_terms r = None
+ in SOME (upd$kb$rb,rb,[kv,rv],false) end))
+ | mk_eq_terms r = NONE
in
(case mk_eq_terms (upd$k$r) of
- Some (trm,trm',vars,update_s)
+ SOME (trm,trm',vars,update_s)
=> if update_s
- then Some (prove_split_simp sg domS
+ then SOME (prove_split_simp sg domS
(list_all(vars,(Logic.mk_equals (sel$trm,trm')))))
- else Some (prove_split_simp sg domS
+ else SOME (prove_split_simp sg domS
(list_all(vars,(Logic.mk_equals (sel$trm,sel$trm')))))
- | None => None)
+ | NONE => NONE)
end
- | None => None)
- | None => None)
- | _ => None));
+ | NONE => NONE)
+ | NONE => NONE)
+ | _ => NONE));
(* record_upd_simproc *)
(* simplify multiple updates:
@@ -953,9 +953,9 @@
fun is_upd_same (sprout,skeleton) u ((sel as Const (s,_))$r) =
if (unsuffix updateN u) = s andalso (seed s sprout) = r
- then Some (sel,seed s skeleton)
- else None
- | is_upd_same _ _ _ = None
+ then SOME (sel,seed s skeleton)
+ else NONE
+ | is_upd_same _ _ _ = NONE
fun init_seed r = ((r,Bound 0), [mk_abs_var "r" r]);
@@ -995,22 +995,22 @@
(case rest (Symtab.update ((u,()),already)) r of
Init ((sprout,skel),vars) =>
(case is_upd_same (sprout,skel) u k of
- Some (sel,skel') =>
+ SOME (sel,skel') =>
let
val (sprout',vars') = grow u uT k vars (sprout,skel);
in Inter(upd$(sel$skel')$skel,skel,vars',sprout') end
- | None =>
+ | NONE =>
let
val kv = mk_abs_var (sel_name u) k;
val kb = Bound (length vars);
in Init ((upd$k$sprout,upd$kb$skel),kv::vars) end)
| Inter (trm,trm',vars,sprout) =>
(case is_upd_same sprout u k of
- Some (sel,skel) =>
+ SOME (sel,skel) =>
let
val (sprout',vars') = grow u uT k vars sprout
in Inter(upd$(sel$skel)$trm,trm',vars',sprout') end
- | None =>
+ | NONE =>
let
val kv = mk_abs_var (sel_name u) k
val kb = Bound (length vars)
@@ -1022,11 +1022,11 @@
in (case mk_updterm updates Symtab.empty t of
Inter (trm,trm',vars,_)
- => Some (prove_split_simp sg T
+ => SOME (prove_split_simp sg T
(list_all(vars,(Logic.mk_equals (trm,trm')))))
- | _ => None)
+ | _ => NONE)
end
- | _ => None));
+ | _ => NONE));
end
(* record_eq_simproc *)
@@ -1047,11 +1047,11 @@
(fn sg => fn _ => fn t =>
(case t of Const ("op =", Type (_, [T, _])) $ _ $ _ =>
(case rec_id (~1) T of
- "" => None
+ "" => NONE
| name => (case get_equalities sg name of
- None => None
- | Some thm => Some (thm RS Eq_TrueI)))
- | _ => None));
+ NONE => NONE
+ | SOME thm => SOME (thm RS Eq_TrueI)))
+ | _ => NONE));
(* record_split_simproc *)
(* splits quantified occurrences of records, for which P holds. P can peek on the
@@ -1066,22 +1066,22 @@
(case t of (Const (quantifier, Type (_, [Type (_, [T, _]), _])))$trm =>
if quantifier = "All" orelse quantifier = "all" orelse quantifier = "Ex"
then (case rec_id (~1) T of
- "" => None
+ "" => NONE
| name
=> let val split = P t
in if split <> 0 then
(case get_splits sg (rec_id split T) of
- None => None
- | Some (all_thm, All_thm, Ex_thm,_)
- => Some (case quantifier of
+ NONE => NONE
+ | SOME (all_thm, All_thm, Ex_thm,_)
+ => SOME (case quantifier of
"all" => all_thm
| "All" => All_thm RS HOL.eq_reflection
| "Ex" => Ex_thm RS HOL.eq_reflection
| _ => error "record_split_simproc"))
- else None
+ else NONE
end)
- else None
- | _ => None))
+ else NONE
+ | _ => NONE))
val record_ex_sel_eq_simproc =
Simplifier.simproc (Theory.sign_of HOL.thy) "record_ex_sel_eq_simproc" ["Ex t"]
@@ -1092,14 +1092,14 @@
addsimprocs [record_split_simproc (K ~1)]) 1)));
fun mkeq (lr,Teq,(sel,Tsel),x) i =
- (case get_selectors sg sel of Some () =>
+ (case get_selectors sg sel of SOME () =>
let val x' = if not (loose_bvar1 (x,0))
then Free ("x" ^ string_of_int i, range_type Tsel)
else raise TERM ("",[x]);
val sel' = Const (sel,Tsel)$Bound 0;
val (l,r) = if lr then (sel',x') else (x',sel');
in Const ("op =",Teq)$l$r end
- | None => raise TERM ("",[Const (sel,Tsel)]));
+ | NONE => raise TERM ("",[Const (sel,Tsel)]));
fun dest_sel_eq (Const ("op =",Teq)$(Const (sel,Tsel)$Bound 0)$X) =
(true,Teq,(sel,Tsel),X)
@@ -1114,9 +1114,9 @@
val prop = list_all ([("r",T)],
Logic.mk_equals (Const ("Ex",Tex)$Abs(s,T,eq),
HOLogic.true_const));
- in Some (prove prop) end
- handle TERM _ => None
- | _ => None)
+ in SOME (prove prop) end
+ handle TERM _ => NONE
+ | _ => NONE)
end)
@@ -1161,16 +1161,16 @@
fun split_free_tac P i (free as Free (n,T)) =
(case rec_id (~1) T of
- "" => None
+ "" => NONE
| name => let val split = P free
in if split <> 0 then
(case get_splits sg (rec_id split T) of
- None => None
- | Some (_,_,_,induct_thm)
- => Some (mk_split_free_tac free induct_thm i))
- else None
+ NONE => NONE
+ | SOME (_,_,_,induct_thm)
+ => SOME (mk_split_free_tac free induct_thm i))
+ else NONE
end)
- | split_free_tac _ _ _ = None;
+ | split_free_tac _ _ _ = NONE;
val split_frees_tacs = mapfilter (split_free_tac P i) frees;
@@ -1217,14 +1217,14 @@
(*note: read_raw_typ avoids expanding type abbreviations*)
fun read_raw_parent sign s =
- (case Sign.read_raw_typ (sign, K None) s handle TYPE (msg, _, _) => error msg of
+ (case Sign.read_raw_typ (sign, K NONE) s handle TYPE (msg, _, _) => error msg of
Type (name, Ts) => (Ts, name)
| _ => error ("Bad parent record specification: " ^ quote s));
fun read_typ sign (env, s) =
let
fun def_sort (x, ~1) = assoc (env, x)
- | def_sort _ = None;
+ | def_sort _ = NONE;
val T = Type.no_tvars (Sign.read_typ (sign, def_sort) s) handle TYPE (msg, _, _) => error msg;
in (Term.add_typ_tfrees (T, env), T) end;
@@ -1264,8 +1264,8 @@
| [x] => (head_of x, false));
val rule'' = cterm_instantiate (map (pairself cert) (case (rev params) of
[] => (case assoc (map dest_Free (term_frees (prop_of st)), s) of
- None => sys_error "try_param_tac: no such variable"
- | Some T => [(P, if ca then concl else lambda (Free (s, T)) concl),
+ NONE => sys_error "try_param_tac: no such variable"
+ | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl),
(x, Free (s, T))])
| (_, T) :: _ => [(P, list_abs (params, if ca then concl
else incr_boundvars 1 (Abs (s, T, concl)))),
@@ -1295,11 +1295,11 @@
let
val UNIV = HOLogic.mk_UNIV repT;
- val (thy',{set_def=Some def, Abs_induct = abs_induct,
+ val (thy',{set_def=SOME def, Abs_induct = abs_induct,
Abs_inject=abs_inject, Abs_inverse = abs_inverse,...}) =
thy |> setmp TypedefPackage.quiet_mode true
- (TypedefPackage.add_typedef_i true None
- (suffix ext_typeN (Sign.base_name name), alphas, Syntax.NoSyn) UNIV None
+ (TypedefPackage.add_typedef_i true NONE
+ (suffix ext_typeN (Sign.base_name name), alphas, Syntax.NoSyn) UNIV NONE
(Tactic.rtac UNIV_witness 1))
val rewrite_rule = Tactic.rewrite_rule [def, rec_UNIV_I, rec_True_simp];
in (thy',map rewrite_rule [abs_inject, abs_inverse, abs_induct])
@@ -1557,9 +1557,9 @@
fun obj_to_meta_all thm =
let
- fun E thm = case (Some (spec OF [thm]) handle THM _ => None) of
- Some thm' => E thm'
- | None => thm;
+ fun E thm = case (SOME (spec OF [thm]) handle THM _ => NONE) of
+ SOME thm' => E thm'
+ | NONE => thm;
val th1 = E thm;
val th2 = Drule.forall_intr_vars th1;
in th2 end;
@@ -2047,8 +2047,8 @@
val init_env =
(case parent of
- None => []
- | Some (types, _) => foldr Term.add_typ_tfrees (types, []));
+ NONE => []
+ | SOME (types, _) => foldr Term.add_typ_tfrees (types, []));
(* fields *)
--- a/src/HOL/Tools/refute.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/refute.ML Sun Feb 13 17:15:14 2005 +0100
@@ -213,8 +213,8 @@
fun interpret thy model args t =
(case get_first (fn (_, f) => f thy model args t) (#interpreters (RefuteData.get thy)) of
- None => raise REFUTE ("interpret", "unable to interpret term " ^ quote (Sign.string_of_term (sign_of thy) t))
- | Some x => x);
+ NONE => raise REFUTE ("interpret", "unable to interpret term " ^ quote (Sign.string_of_term (sign_of thy) t))
+ | SOME x => x);
(* ------------------------------------------------------------------------- *)
(* print: tries to convert the constant denoted by the term 't' into a term *)
@@ -225,8 +225,8 @@
fun print thy model t intr assignment =
(case get_first (fn (_, f) => f thy model t intr assignment) (#printers (RefuteData.get thy)) of
- None => Const ("<<no printer available>>", fastype_of t)
- | Some x => x);
+ NONE => Const ("<<no printer available>>", fastype_of t)
+ | SOME x => x);
(* ------------------------------------------------------------------------- *)
(* print_model: turns the model into a string, using a fixed interpretation *)
@@ -256,9 +256,9 @@
space_implode "\n" (mapfilter (fn (t,intr) =>
(* print constants only if 'show_consts' is true *)
if (!show_consts) orelse not (is_Const t) then
- Some (Sign.string_of_term (sign_of thy) t ^ ": " ^ Sign.string_of_term (sign_of thy) (print thy model t intr assignment))
+ SOME (Sign.string_of_term (sign_of thy) t ^ ": " ^ Sign.string_of_term (sign_of thy) (print thy model t intr assignment))
else
- None) terms) ^ "\n"
+ NONE) terms) ^ "\n"
in
typs_msg ^ show_consts_msg ^ terms_msg
end;
@@ -275,8 +275,8 @@
val {interpreters, printers, parameters} = RefuteData.get thy
in
case assoc (interpreters, name) of
- None => RefuteData.put {interpreters = (name, f) :: interpreters, printers = printers, parameters = parameters} thy
- | Some _ => error ("Interpreter " ^ name ^ " already declared")
+ NONE => RefuteData.put {interpreters = (name, f) :: interpreters, printers = printers, parameters = parameters} thy
+ | SOME _ => error ("Interpreter " ^ name ^ " already declared")
end;
(* string -> (theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term option) -> theory -> theory *)
@@ -286,8 +286,8 @@
val {interpreters, printers, parameters} = RefuteData.get thy
in
case assoc (printers, name) of
- None => RefuteData.put {interpreters = interpreters, printers = (name, f) :: printers, parameters = parameters} thy
- | Some _ => error ("Printer " ^ name ^ " already declared")
+ NONE => RefuteData.put {interpreters = interpreters, printers = (name, f) :: printers, parameters = parameters} thy
+ | SOME _ => error ("Printer " ^ name ^ " already declared")
end;
(* ------------------------------------------------------------------------- *)
@@ -302,9 +302,9 @@
val {interpreters, printers, parameters} = RefuteData.get thy
in
case Symtab.lookup (parameters, name) of
- None => RefuteData.put
+ NONE => RefuteData.put
{interpreters = interpreters, printers = printers, parameters = Symtab.extend (parameters, [(name, value)])} thy
- | Some _ => RefuteData.put
+ | SOME _ => RefuteData.put
{interpreters = interpreters, printers = printers, parameters = Symtab.update ((name, value), parameters)} thy
end;
@@ -339,15 +339,15 @@
(* (string * string) list * string -> int *)
fun read_int (parms, name) =
case assoc_string (parms, name) of
- Some s => (case Int.fromString s of
+ SOME s => (case Int.fromString s of
SOME i => i
| NONE => error ("parameter " ^ quote name ^ " (value is " ^ quote s ^ ") must be an integer value"))
- | None => error ("parameter " ^ quote name ^ " must be assigned a value")
+ | NONE => error ("parameter " ^ quote name ^ " must be assigned a value")
(* (string * string) list * string -> string *)
fun read_string (parms, name) =
case assoc_string (parms, name) of
- Some s => s
- | None => error ("parameter " ^ quote name ^ " must be assigned a value")
+ SOME s => s
+ | NONE => error ("parameter " ^ quote name ^ " must be assigned a value")
(* (string * string) list *)
val allparams = override @ (get_default_params thy) (* 'override' first, defaults last *)
(* int *)
@@ -363,7 +363,7 @@
(* whose name is one of the other parameters (e.g. 'maxvars') *)
(* (string * int) list *)
val sizes = mapfilter
- (fn (name,value) => (case Int.fromString value of SOME i => Some (name, i) | NONE => None))
+ (fn (name,value) => (case Int.fromString value of SOME i => SOME (name, i) | NONE => NONE))
(filter (fn (name,_) => name<>"minsize" andalso name<>"maxsize" andalso name<>"maxvars" andalso name<>"maxtime" andalso name<>"satsolver")
allparams)
in
@@ -431,27 +431,27 @@
let
fun find_typeSubs (Const (s', T')) =
(if s=s' then
- Some (Type.typ_match (Sign.tsig_of (sign_of thy)) (Vartab.empty, (T', T)))
+ SOME (Type.typ_match (Sign.tsig_of (sign_of thy)) (Vartab.empty, (T', T)))
else
- None
- handle Type.TYPE_MATCH => None)
- | find_typeSubs (Free _) = None
- | find_typeSubs (Var _) = None
- | find_typeSubs (Bound _) = None
+ NONE
+ handle Type.TYPE_MATCH => NONE)
+ | find_typeSubs (Free _) = NONE
+ | find_typeSubs (Var _) = NONE
+ | find_typeSubs (Bound _) = NONE
| find_typeSubs (Abs (_, _, body)) = find_typeSubs body
- | find_typeSubs (t1 $ t2) = (case find_typeSubs t1 of Some x => Some x | None => find_typeSubs t2)
+ | find_typeSubs (t1 $ t2) = (case find_typeSubs t1 of SOME x => SOME x | NONE => find_typeSubs t2)
val typeSubs = (case find_typeSubs t of
- Some x => x
- | None => raise REFUTE ("collect_axioms", "no type instantiation found for " ^ quote s ^ " in " ^ Sign.string_of_term (sign_of thy) t))
+ SOME x => x
+ | NONE => raise REFUTE ("collect_axioms", "no type instantiation found for " ^ quote s ^ " in " ^ Sign.string_of_term (sign_of thy) t))
in
map_term_types
(map_type_tvar
(fn (v,_) =>
case Vartab.lookup (typeSubs, v) of
- None =>
+ NONE =>
(* schematic type variable not instantiated *)
raise REFUTE ("collect_axioms", "term " ^ Sign.string_of_term (sign_of thy) t ^ " still has a polymorphic type (after instantiating type of " ^ quote s ^ ")")
- | Some typ =>
+ | SOME typ =>
typ))
t
end
@@ -462,10 +462,10 @@
map_term_types (map_type_tvar
(fn (v,_) =>
case Vartab.lookup (typeSubs, v) of
- None =>
+ NONE =>
(* schematic type variable not instantiated *)
raise ERROR
- | Some typ =>
+ | SOME typ =>
typ)) t
(* Term.term list * Term.typ -> Term.term list *)
fun collect_type_axioms (axs, T) =
@@ -479,30 +479,30 @@
(* look up the definition of a type, as created by "typedef" *)
(* (string * Term.term) list -> (string * Term.term) option *)
fun get_typedefn [] =
- None
+ NONE
| get_typedefn ((axname,ax)::axms) =
(let
(* Term.term -> Term.typ option *)
fun type_of_type_definition (Const (s', T')) =
if s'="Typedef.type_definition" then
- Some T'
+ SOME T'
else
- None
- | type_of_type_definition (Free _) = None
- | type_of_type_definition (Var _) = None
- | type_of_type_definition (Bound _) = None
+ NONE
+ | type_of_type_definition (Free _) = NONE
+ | type_of_type_definition (Var _) = NONE
+ | type_of_type_definition (Bound _) = NONE
| type_of_type_definition (Abs (_, _, body)) = type_of_type_definition body
- | type_of_type_definition (t1 $ t2) = (case type_of_type_definition t1 of Some x => Some x | None => type_of_type_definition t2)
+ | type_of_type_definition (t1 $ t2) = (case type_of_type_definition t1 of SOME x => SOME x | NONE => type_of_type_definition t2)
in
case type_of_type_definition ax of
- Some T' =>
+ SOME T' =>
let
val T'' = (domain_type o domain_type) T'
val typeSubs = Type.typ_match (Sign.tsig_of (sign_of thy)) (Vartab.empty, (T'', T))
in
- Some (axname, monomorphic_term typeSubs ax)
+ SOME (axname, monomorphic_term typeSubs ax)
end
- | None =>
+ | NONE =>
get_typedefn axms
end
handle ERROR => get_typedefn axms
@@ -510,19 +510,19 @@
| Type.TYPE_MATCH => get_typedefn axms)
in
case DatatypePackage.datatype_info thy s of
- Some info => (* inductive datatype *)
+ SOME info => (* inductive datatype *)
(* only collect relevant type axioms for the argument types *)
foldl collect_type_axioms (axs, Ts)
- | None =>
+ | NONE =>
(case get_typedefn axioms of
- Some (axname, ax) =>
+ SOME (axname, ax) =>
if mem_term (ax, axs) then
(* collect relevant type axioms for the argument types *)
foldl collect_type_axioms (axs, Ts)
else
(immediate_output (" " ^ axname);
collect_term_axioms (ax :: axs, ax))
- | None =>
+ | NONE =>
(* at least collect relevant type axioms for the argument types *)
foldl collect_type_axioms (axs, Ts))
end
@@ -579,7 +579,7 @@
(* look up the definition of a constant, as created by "constdefs" *)
(* string -> Term.typ -> (string * Term.term) list -> (string * Term.term) option *)
fun get_defn [] =
- None
+ NONE
| get_defn ((axname,ax)::axms) =
(let
val (lhs, _) = Logic.dest_equals ax (* equations only *)
@@ -590,7 +590,7 @@
let
val typeSubs = Type.typ_match (Sign.tsig_of (sign_of thy)) (Vartab.empty, (T', T))
in
- Some (axname, monomorphic_term typeSubs ax)
+ SOME (axname, monomorphic_term typeSubs ax)
end
else
get_defn axms
@@ -603,7 +603,7 @@
(case body_type T of
Type (s', _) =>
(case DatatypePackage.constrs_of thy s' of
- Some constrs =>
+ SOME constrs =>
Library.exists (fn c =>
(case c of
Const (cname, ctype) =>
@@ -611,7 +611,7 @@
| _ =>
raise REFUTE ("collect_axioms", "IDT constructor is not a constant")))
constrs
- | None =>
+ | NONE =>
false)
| _ =>
false)
@@ -622,8 +622,8 @@
((case last_elem (binder_types T) of
Type (s', _) =>
(case DatatypePackage.datatype_info thy s' of
- Some info => s mem (#rec_names info)
- | None => false) (* not an inductive datatype *)
+ SOME info => s mem (#rec_names info)
+ | NONE => false) (* not an inductive datatype *)
| _ => (* a (free or schematic) type variable *)
false)
handle LIST "last_elem" => false) (* not even a function type *)
@@ -641,14 +641,14 @@
collect_type_axioms (axs, T)
) else
(case get_defn axioms of
- Some (axname, ax) =>
+ SOME (axname, ax) =>
if mem_term (ax, axs) then
(* collect relevant type axioms *)
collect_type_axioms (axs, T)
else
(immediate_output (" " ^ axname);
collect_term_axioms (ax :: axs, ax))
- | None =>
+ | NONE =>
(* collect relevant type axioms *)
collect_type_axioms (axs, T))
end
@@ -698,7 +698,7 @@
| Type ("set", [T1]) => collect_types (T1, acc)
| Type (s, Ts) =>
(case DatatypePackage.datatype_info thy s of
- Some info => (* inductive datatype *)
+ SOME info => (* inductive datatype *)
let
val index = #index info
val descr = #descr info
@@ -723,7 +723,7 @@
in
acc_constrs
end
- | None => (* not an inductive datatype, e.g. defined via "typedef" or "typedecl" *)
+ | NONE => (* not an inductive datatype, e.g. defined via "typedef" or "typedecl" *)
T ins (foldr collect_types (Ts, acc)))
| TFree _ => T ins acc
| TVar _ => T ins acc)
@@ -756,8 +756,8 @@
let
fun size_of_typ T =
case assoc (sizes, string_of_typ T) of
- Some n => n
- | None => minsize
+ SOME n => n
+ | NONE => minsize
in
map (fn T => (T, size_of_typ T)) xs
end;
@@ -775,15 +775,15 @@
let
(* int -> int list -> int list option *)
fun add1 _ [] =
- None (* overflow *)
+ NONE (* overflow *)
| add1 max (x::xs) =
if x<max orelse max<0 then
- Some ((x+1)::xs) (* add 1 to the head *)
+ SOME ((x+1)::xs) (* add 1 to the head *)
else
apsome (fn xs' => 0 :: xs') (add1 max xs) (* carry-over *)
(* int -> int list * int list -> int list option *)
fun shift _ (_, []) =
- None
+ NONE
| shift max (zeros, x::xs) =
if x=0 then
shift max (0::zeros, xs)
@@ -794,9 +794,9 @@
(* int -> int -> int -> int list option *)
fun make_first 0 sum _ =
if sum=0 then
- Some []
+ SOME []
else
- None
+ NONE
| make_first len sum max =
if sum<=max orelse max<0 then
apsome (fn xs' => sum :: xs') (make_first (len-1) 0 max)
@@ -807,29 +807,29 @@
(* int -> int list -> int list option *)
fun next max xs =
(case shift max ([], xs) of
- Some xs' =>
- Some xs'
- | None =>
+ SOME xs' =>
+ SOME xs'
+ | NONE =>
let
val (len, sum) = foldl (fn ((l, s), x) => (l+1, s+x)) ((0, 0), xs)
in
make_first len (sum+1) max (* increment 'sum' by 1 *)
end)
(* only consider those types for which the size is not fixed *)
- val mutables = filter (fn (T, _) => assoc (sizes, string_of_typ T) = None) xs
+ val mutables = filter (fn (T, _) => assoc (sizes, string_of_typ T) = NONE) xs
(* subtract 'minsize' from every size (will be added again at the end) *)
val diffs = map (fn (_, n) => n-minsize) mutables
in
case next (maxsize-minsize) diffs of
- Some diffs' =>
+ SOME diffs' =>
(* merge with those types for which the size is fixed *)
- Some (snd (foldl_map (fn (ds, (T, _)) =>
+ SOME (snd (foldl_map (fn (ds, (T, _)) =>
case assoc (sizes, string_of_typ T) of
- Some n => (ds, (T, n)) (* return the fixed size *)
- | None => (tl ds, (T, minsize + (hd ds)))) (* consume the head of 'ds', add 'minsize' *)
+ SOME n => (ds, (T, n)) (* return the fixed size *)
+ | NONE => (tl ds, (T, minsize + (hd ds)))) (* consume the head of 'ds', add 'minsize' *)
(diffs', xs)))
- | None =>
- None
+ | NONE =>
+ NONE
end;
(* ------------------------------------------------------------------------- *)
@@ -901,12 +901,12 @@
immediate_output " invoking SAT solver...";
(case SatSolver.invoke_solver satsolver fm of
SatSolver.SATISFIABLE assignment =>
- writeln ("\n*** Model found: ***\n" ^ print_model thy model (fn i => case assignment i of Some b => b | None => true))
+ writeln ("\n*** Model found: ***\n" ^ print_model thy model (fn i => case assignment i of SOME b => b | NONE => true))
| _ => (* SatSolver.UNSATISFIABLE, SatSolver.UNKNOWN *)
(immediate_output " no model found.\n";
case next_universe universe sizes minsize maxsize of
- Some universe' => find_model_loop universe'
- | None => writeln "Search terminated, no larger universe within the given limits."))
+ SOME universe' => find_model_loop universe'
+ | NONE => writeln "Search terminated, no larger universe within the given limits."))
handle SatSolver.NOT_CONFIGURED =>
error ("SAT solver " ^ quote satsolver ^ " is not configured.")
end handle MAXVARS_EXCEEDED =>
@@ -1207,7 +1207,7 @@
val wf = (if size=1 then True else if size=2 then True else exactly_one_true fms)
in
(* extend the model, increase 'next_idx', add well-formedness condition *)
- Some (intr, (typs, (t, intr)::terms), {maxvars = maxvars, next_idx = next, bounds = bounds, wellformed = SAnd (wellformed, wf)})
+ SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars, next_idx = next, bounds = bounds, wellformed = SAnd (wellformed, wf)})
end
in
case T of
@@ -1235,7 +1235,7 @@
val intr = Node copies
in
(* extend the model, increase 'next_idx', add well-formedness condition *)
- Some (intr, (typs, (t, intr)::terms), {maxvars = maxvars, next_idx = next, bounds = bounds, wellformed = SAnd (wellformed, wf)})
+ SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars, next_idx = next, bounds = bounds, wellformed = SAnd (wellformed, wf)})
end
| Type _ => interpret_groundtype ()
| TFree _ => interpret_groundtype ()
@@ -1243,10 +1243,10 @@
end
in
case assoc (terms, t) of
- Some intr =>
+ SOME intr =>
(* return an existing interpretation *)
- Some (intr, model, args)
- | None =>
+ SOME (intr, model, args)
+ | NONE =>
(case t of
Const (_, T) =>
interpret_groundterm T
@@ -1255,7 +1255,7 @@
| Var (_, T) =>
interpret_groundterm T
| Bound i =>
- Some (nth_elem (i, #bounds args), model, args)
+ SOME (nth_elem (i, #bounds args), model, args)
| Abs (x, T, body) =>
let
(* create all constants of type 'T' *)
@@ -1273,7 +1273,7 @@
end)
((model, args), constants)
in
- Some (Node bodies, model', args')
+ SOME (Node bodies, model', args')
end
| t1 $ t2 =>
let
@@ -1321,7 +1321,7 @@
val (intr1, model1, args1) = interpret thy model args t1
val (intr2, model2, args2) = interpret thy model1 args1 t2
in
- Some (interpretation_apply (intr1,intr2), model2, args2)
+ SOME (interpretation_apply (intr1,intr2), model2, args2)
end)
end;
@@ -1339,7 +1339,7 @@
val fmTrue = PropLogic.all (map toTrue xs)
val fmFalse = PropLogic.exists (map toFalse xs)
in
- Some (Leaf [fmTrue, fmFalse], m, a)
+ SOME (Leaf [fmTrue, fmFalse], m, a)
end
| _ =>
raise REFUTE ("Pure_interpreter", "\"all\" is not followed by a function")
@@ -1349,11 +1349,11 @@
val (i1, m1, a1) = interpret thy model args t1
val (i2, m2, a2) = interpret thy m1 a1 t2
in
- Some (make_equality (i1, i2), m2, a2)
+ SOME (make_equality (i1, i2), m2, a2)
end
| Const ("==>", _) => (* simpler than translating 'Const ("==>", _) $ t1 $ t2' *)
- Some (Node [Node [TT, FF], Node [TT, TT]], model, args)
- | _ => None;
+ SOME (Node [Node [TT, FF], Node [TT, TT]], model, args)
+ | _ => NONE;
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
@@ -1366,13 +1366,13 @@
(* ------------------------------------------------------------------------- *)
case t of
Const ("Trueprop", _) =>
- Some (Node [TT, FF], model, args)
+ SOME (Node [TT, FF], model, args)
| Const ("Not", _) =>
- Some (Node [FF, TT], model, args)
+ SOME (Node [FF, TT], model, args)
| Const ("True", _) => (* redundant, since 'True' is also an IDT constructor *)
- Some (TT, model, args)
+ SOME (TT, model, args)
| Const ("False", _) => (* redundant, since 'False' is also an IDT constructor *)
- Some (FF, model, args)
+ SOME (FF, model, args)
| Const ("All", _) $ t1 =>
(* if "All" occurs without an argument (i.e. as argument to a higher-order *)
(* function or predicate), it is handled by the 'stlc_interpreter' (i.e. *)
@@ -1386,7 +1386,7 @@
val fmTrue = PropLogic.all (map toTrue xs)
val fmFalse = PropLogic.exists (map toFalse xs)
in
- Some (Leaf [fmTrue, fmFalse], m, a)
+ SOME (Leaf [fmTrue, fmFalse], m, a)
end
| _ =>
raise REFUTE ("HOLogic_interpreter", "\"All\" is followed by a non-function")
@@ -1404,7 +1404,7 @@
val fmTrue = PropLogic.exists (map toTrue xs)
val fmFalse = PropLogic.all (map toFalse xs)
in
- Some (Leaf [fmTrue, fmFalse], m, a)
+ SOME (Leaf [fmTrue, fmFalse], m, a)
end
| _ =>
raise REFUTE ("HOLogic_interpreter", "\"Ex\" is followed by a non-function")
@@ -1414,19 +1414,19 @@
val (i1, m1, a1) = interpret thy model args t1
val (i2, m2, a2) = interpret thy m1 a1 t2
in
- Some (make_equality (i1, i2), m2, a2)
+ SOME (make_equality (i1, i2), m2, a2)
end
| Const ("op =", _) $ t1 =>
- Some (interpret thy model args (eta_expand t 1))
+ SOME (interpret thy model args (eta_expand t 1))
| Const ("op =", _) =>
- Some (interpret thy model args (eta_expand t 2))
+ SOME (interpret thy model args (eta_expand t 2))
| Const ("op &", _) =>
- Some (Node [Node [TT, FF], Node [FF, FF]], model, args)
+ SOME (Node [Node [TT, FF], Node [FF, FF]], model, args)
| Const ("op |", _) =>
- Some (Node [Node [TT, TT], Node [TT, FF]], model, args)
+ SOME (Node [Node [TT, TT], Node [TT, FF]], model, args)
| Const ("op -->", _) =>
- Some (Node [Node [TT, FF], Node [TT, TT]], model, args)
- | _ => None;
+ SOME (Node [Node [TT, FF], Node [TT, TT]], model, args)
+ | _ => NONE;
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
@@ -1436,42 +1436,42 @@
val (typs, terms) = model
in
case assoc (terms, t) of
- Some intr =>
+ SOME intr =>
(* return an existing interpretation *)
- Some (intr, model, args)
- | None =>
+ SOME (intr, model, args)
+ | NONE =>
(case t of
Free (x, Type ("set", [T])) =>
let
val (intr, _, args') = interpret thy (typs, []) args (Free (x, T --> HOLogic.boolT))
in
- Some (intr, (typs, (t, intr)::terms), args')
+ SOME (intr, (typs, (t, intr)::terms), args')
end
| Var ((x,i), Type ("set", [T])) =>
let
val (intr, _, args') = interpret thy (typs, []) args (Var ((x,i), T --> HOLogic.boolT))
in
- Some (intr, (typs, (t, intr)::terms), args')
+ SOME (intr, (typs, (t, intr)::terms), args')
end
| Const (s, Type ("set", [T])) =>
let
val (intr, _, args') = interpret thy (typs, []) args (Const (s, T --> HOLogic.boolT))
in
- Some (intr, (typs, (t, intr)::terms), args')
+ SOME (intr, (typs, (t, intr)::terms), args')
end
(* 'Collect' == identity *)
| Const ("Collect", _) $ t1 =>
- Some (interpret thy model args t1)
+ SOME (interpret thy model args t1)
| Const ("Collect", _) =>
- Some (interpret thy model args (eta_expand t 1))
+ SOME (interpret thy model args (eta_expand t 1))
(* 'op :' == application *)
| Const ("op :", _) $ t1 $ t2 =>
- Some (interpret thy model args (t2 $ t1))
+ SOME (interpret thy model args (t2 $ t1))
| Const ("op :", _) $ t1 =>
- Some (interpret thy model args (eta_expand t 1))
+ SOME (interpret thy model args (eta_expand t 1))
| Const ("op :", _) =>
- Some (interpret thy model args (eta_expand t 2))
- | _ => None)
+ SOME (interpret thy model args (eta_expand t 2))
+ | _ => NONE)
end;
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
@@ -1482,15 +1482,15 @@
(* Term.typ -> (interpretation * model * arguments) option *)
fun interpret_variable (Type (s, Ts)) =
(case DatatypePackage.datatype_info thy s of
- Some info => (* inductive datatype *)
+ SOME info => (* inductive datatype *)
let
val (typs, terms) = model
(* int option -- only recursive IDTs have an associated depth *)
val depth = assoc (typs, Type (s, Ts))
in
- if depth = (Some 0) then (* termination condition to avoid infinite recursion *)
+ if depth = (SOME 0) then (* termination condition to avoid infinite recursion *)
(* return a leaf of size 0 *)
- Some (Leaf [], model, args)
+ SOME (Leaf [], model, args)
else
let
val index = #index info
@@ -1505,7 +1505,7 @@
else
())
(* if the model specifies a depth for the current type, decrement it to avoid infinite recursion *)
- val typs' = (case depth of None => typs | Some n => overwrite (typs, (Type (s, Ts), n-1)))
+ val typs' = (case depth of NONE => typs | SOME n => overwrite (typs, (Type (s, Ts), n-1)))
(* recursively compute the size of the datatype *)
val size = size_of_dtyp thy typs' descr typ_assoc constrs
val next_idx = #next_idx args
@@ -1525,19 +1525,19 @@
val wf = (if size=1 then True else if size=2 then True else exactly_one_true fms)
in
(* extend the model, increase 'next_idx', add well-formedness condition *)
- Some (intr, (typs, (t, intr)::terms), {maxvars = #maxvars args, next_idx = next, bounds = #bounds args, wellformed = SAnd (#wellformed args, wf)})
+ SOME (intr, (typs, (t, intr)::terms), {maxvars = #maxvars args, next_idx = next, bounds = #bounds args, wellformed = SAnd (#wellformed args, wf)})
end
end
- | None => (* not an inductive datatype *)
- None)
+ | NONE => (* not an inductive datatype *)
+ NONE)
| interpret_variable _ = (* a (free or schematic) type variable *)
- None
+ NONE
in
case assoc (terms, t) of
- Some intr =>
+ SOME intr =>
(* return an existing interpretation *)
- Some (intr, model, args)
- | None =>
+ SOME (intr, model, args)
+ | NONE =>
(case t of
Free (_, T) => interpret_variable T
| Var (_, T) => interpret_variable T
@@ -1549,7 +1549,7 @@
(case body_type T of
Type (s', Ts') =>
(case DatatypePackage.datatype_info thy s' of
- Some info => (* body type is an inductive datatype *)
+ SOME info => (* body type is an inductive datatype *)
let
val index = #index info
val descr = #descr info
@@ -1570,12 +1570,12 @@
case constrs2 of
[] =>
(* 'Const (s, T)' is not a constructor of this datatype *)
- None
+ NONE
| c::cs =>
let
(* int option -- only recursive IDTs have an associated depth *)
val depth = assoc (typs, Type (s', Ts'))
- val typs' = (case depth of None => typs | Some n => overwrite (typs, (Type (s', Ts'), n-1)))
+ val typs' = (case depth of NONE => typs | SOME n => overwrite (typs, (Type (s', Ts'), n-1)))
(* constructors before 'Const (s, T)' generate elements of the datatype *)
val offset = size_of_dtyp thy typs' descr typ_assoc constrs1
(* 'Const (s, T)' and constructors after it generate elements of the datatype *)
@@ -1620,19 +1620,19 @@
(new_offset, Node (intrs @ (replicate (size-size') (make_partial ds))))
end
in
- Some ((snd o make_constr) (offset, snd c), model, args)
+ SOME ((snd o make_constr) (offset, snd c), model, args)
end
end
- | None => (* body type is not an inductive datatype *)
- None)
+ | NONE => (* body type is not an inductive datatype *)
+ NONE)
| _ => (* body type is a (free or schematic) type variable *)
- None)
+ NONE)
in
case interpret_constructor () of
- Some x => Some x
- | None => interpret_variable T
+ SOME x => SOME x
+ | NONE => interpret_variable T
end
- | _ => None)
+ | _ => NONE)
end;
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
@@ -1667,10 +1667,10 @@
Leaf (replicate size_nat False)
end
in
- Some (Node (map card constants), model, args)
+ SOME (Node (map card constants), model, args)
end
| _ =>
- None;
+ NONE;
(* ------------------------------------------------------------------------- *)
@@ -1682,10 +1682,10 @@
fun stlc_printer thy model t intr assignment =
let
(* Term.term -> Term.typ option *)
- fun typeof (Free (_, T)) = Some T
- | typeof (Var (_, T)) = Some T
- | typeof (Const (_, T)) = Some T
- | typeof _ = None
+ fun typeof (Free (_, T)) = SOME T
+ | typeof (Var (_, T)) = SOME T
+ | typeof (Const (_, T)) = SOME T
+ | typeof _ = NONE
(* string -> string *)
fun strip_leading_quote s =
(implode o (fn ss => case ss of [] => [] | x::xs => if x="'" then xs else ss) o explode) s
@@ -1707,7 +1707,7 @@
raise REFUTE ("stlc_printer", "interpretation for ground type is not a leaf")
in
case typeof t of
- Some T =>
+ SOME T =>
(case T of
Type ("fun", [T1, T2]) =>
let
@@ -1731,18 +1731,18 @@
val HOLogic_empty_set = Const ("{}", HOLogic_setT)
val HOLogic_insert = Const ("insert", HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
in
- Some (foldr (fn (pair, acc) => HOLogic_insert $ pair $ acc) (pairs, HOLogic_empty_set))
+ SOME (foldr (fn (pair, acc) => HOLogic_insert $ pair $ acc) (pairs, HOLogic_empty_set))
end
| Type ("prop", []) =>
(case index_from_interpretation intr of
- 0 => Some (HOLogic.mk_Trueprop HOLogic.true_const)
- | 1 => Some (HOLogic.mk_Trueprop HOLogic.false_const)
+ 0 => SOME (HOLogic.mk_Trueprop HOLogic.true_const)
+ | 1 => SOME (HOLogic.mk_Trueprop HOLogic.false_const)
| _ => raise REFUTE ("stlc_interpreter", "illegal interpretation for a propositional value"))
- | Type _ => Some (Const (string_of_typ T ^ string_of_int (index_from_interpretation intr), T))
- | TFree _ => Some (Const (string_of_typ T ^ string_of_int (index_from_interpretation intr), T))
- | TVar _ => Some (Const (string_of_typ T ^ string_of_int (index_from_interpretation intr), T)))
- | None =>
- None
+ | Type _ => SOME (Const (string_of_typ T ^ string_of_int (index_from_interpretation intr), T))
+ | TFree _ => SOME (Const (string_of_typ T ^ string_of_int (index_from_interpretation intr), T))
+ | TVar _ => SOME (Const (string_of_typ T ^ string_of_int (index_from_interpretation intr), T)))
+ | NONE =>
+ NONE
end;
(* theory -> model -> Term.term -> interpretation -> (int -> bool) -> string option *)
@@ -1750,13 +1750,13 @@
fun set_printer thy model t intr assignment =
let
(* Term.term -> Term.typ option *)
- fun typeof (Free (_, T)) = Some T
- | typeof (Var (_, T)) = Some T
- | typeof (Const (_, T)) = Some T
- | typeof _ = None
+ fun typeof (Free (_, T)) = SOME T
+ | typeof (Var (_, T)) = SOME T
+ | typeof (Const (_, T)) = SOME T
+ | typeof _ = NONE
in
case typeof t of
- Some (Type ("set", [T])) =>
+ SOME (Type ("set", [T])) =>
let
(* create all constants of type 'T' *)
val (i, _, _) = interpret thy model {maxvars=0, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
@@ -1770,9 +1770,9 @@
case result of
Leaf [fmTrue, fmFalse] =>
if PropLogic.eval assignment fmTrue then
- Some (print thy model (Free ("dummy", T)) arg assignment)
+ SOME (print thy model (Free ("dummy", T)) arg assignment)
else if PropLogic.eval assignment fmFalse then
- None
+ NONE
else
raise REFUTE ("set_printer", "illegal interpretation: no value assigned (SAT solver unsound?)")
| _ =>
@@ -1784,10 +1784,10 @@
val HOLogic_empty_set = Const ("{}", HOLogic_setT)
val HOLogic_insert = Const ("insert", T --> HOLogic_setT --> HOLogic_setT)
in
- Some (foldl (fn (acc, elem) => HOLogic_insert $ elem $ acc) (HOLogic_empty_set, elements))
+ SOME (foldl (fn (acc, elem) => HOLogic_insert $ elem $ acc) (HOLogic_empty_set, elements))
end
| _ =>
- None
+ NONE
end;
(* theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term option *)
@@ -1795,15 +1795,15 @@
fun IDT_printer thy model t intr assignment =
let
(* Term.term -> Term.typ option *)
- fun typeof (Free (_, T)) = Some T
- | typeof (Var (_, T)) = Some T
- | typeof (Const (_, T)) = Some T
- | typeof _ = None
+ fun typeof (Free (_, T)) = SOME T
+ | typeof (Var (_, T)) = SOME T
+ | typeof (Const (_, T)) = SOME T
+ | typeof _ = NONE
in
case typeof t of
- Some (Type (s, Ts)) =>
+ SOME (Type (s, Ts)) =>
(case DatatypePackage.datatype_info thy s of
- Some info => (* inductive datatype *)
+ SOME info => (* inductive datatype *)
let
val (typs, _) = model
val index = #index info
@@ -1824,7 +1824,7 @@
val _ = (if element<0 then raise REFUTE ("IDT_printer", "invalid interpretation (no value assigned)") else ())
(* int option -- only recursive IDTs have an associated depth *)
val depth = assoc (typs, Type (s, Ts))
- val typs' = (case depth of None => typs | Some n => overwrite (typs, (Type (s, Ts), n-1)))
+ val typs' = (case depth of NONE => typs | SOME n => overwrite (typs, (Type (s, Ts), n-1)))
(* int -> DatatypeAux.dtyp list -> Term.term list *)
fun make_args n [] =
if n<>0 then
@@ -1860,12 +1860,12 @@
make_term (n-c_size) cs
end
in
- Some (make_term element constrs)
+ SOME (make_term element constrs)
end
- | None => (* not an inductive datatype *)
- None)
+ | NONE => (* not an inductive datatype *)
+ NONE)
| _ => (* a (free or schematic) type variable *)
- None
+ NONE
end;
--- a/src/HOL/Tools/res_axioms.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/res_axioms.ML Sun Feb 13 17:15:14 2005 +0100
@@ -292,8 +292,8 @@
end;
-fun sk_lookup [] t = None
- | sk_lookup ((tm,sk_tm)::tms) t = if (t = tm) then Some (sk_tm) else (sk_lookup tms t);
+fun sk_lookup [] t = NONE
+ | sk_lookup ((tm,sk_tm)::tms) t = if (t = tm) then SOME (sk_tm) else (sk_lookup tms t);
@@ -301,8 +301,8 @@
fun get_skolem epss t =
let val sk_fun = sk_lookup epss t
in
- case sk_fun of None => get_new_skolem epss t
- | Some sk => (sk,epss)
+ case sk_fun of NONE => get_new_skolem epss t
+ | SOME sk => (sk,epss)
end;
--- a/src/HOL/Tools/res_clause.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/res_clause.ML Sun Feb 13 17:15:14 2005 +0100
@@ -19,7 +19,7 @@
val make_axiom_arity_clause :
string * (string * string list list) -> arityClause
val make_axiom_classrelClause :
- string * string Library.option -> classrelClause
+ string * string option -> classrelClause
val make_axiom_clause : Term.term -> string * int -> clause
val make_axiom_clause_thm : Thm.thm -> string * int -> clause
val make_conjecture_clause : Term.term -> clause
@@ -105,8 +105,8 @@
fun lookup_const c =
case Symtab.lookup (const_trans_table,c) of
- Some c' => c'
- | None => make_fixed_const c;
+ SOME c' => c'
+ | NONE => make_fixed_const c;
@@ -501,7 +501,7 @@
datatype classrelClause =
ClassrelClause of {clause_id: clause_id,
subclass: class,
- superclass: class Library.option};
+ superclass: class option};
fun make_classrelClause (clause_id,subclass,superclass) =
ClassrelClause {clause_id = clause_id,subclass = subclass, superclass = superclass};
@@ -510,8 +510,8 @@
fun make_axiom_classrelClause (subclass,superclass) =
let val cls_id = generate_id()
val sub = make_type_class subclass
- val sup = case superclass of None => None
- | Some s => Some (make_type_class s)
+ val sup = case superclass of NONE => NONE
+ | SOME s => SOME (make_type_class s)
in
make_classrelClause(cls_id,sub,sup)
end;
@@ -519,11 +519,11 @@
fun classrelClauses_of_aux (sub,[]) = []
- | classrelClauses_of_aux (sub,(sup::sups)) = make_axiom_classrelClause(sub,Some sup) :: (classrelClauses_of_aux (sub,sups));
+ | classrelClauses_of_aux (sub,(sup::sups)) = make_axiom_classrelClause(sub,SOME sup) :: (classrelClauses_of_aux (sub,sups));
fun classrelClauses_of (sub,sups) =
- case sups of [] => [make_axiom_classrelClause (sub,None)]
+ case sups of [] => [make_axiom_classrelClause (sub,NONE)]
| _ => classrelClauses_of_aux (sub, sups);
@@ -673,8 +673,8 @@
fun tptp_classrelLits sub sup =
let val tvar = "(T)"
in
- case sup of None => "[++" ^ sub ^ tvar ^ "]"
- | (Some supcls) => "[--" ^ sub ^ tvar ^ ",++" ^ supcls ^ tvar ^ "]"
+ case sup of NONE => "[++" ^ sub ^ tvar ^ "]"
+ | (SOME supcls) => "[--" ^ sub ^ tvar ^ ",++" ^ supcls ^ tvar ^ "]"
end;
--- a/src/HOL/Tools/rewrite_hol_proof.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/rewrite_hol_proof.ML Sun Feb 13 17:15:14 2005 +0100
@@ -286,10 +286,10 @@
(** Replace congruence rules by substitution rules **)
-fun strip_cong ps (PThm (("HOL.cong", _), _, _, _) % _ % _ % Some x % Some y %%
+fun strip_cong ps (PThm (("HOL.cong", _), _, _, _) % _ % _ % SOME x % SOME y %%
prf1 %% prf2) = strip_cong (((x, y), prf2) :: ps) prf1
- | strip_cong ps (PThm (("HOL.refl", _), _, _, _) % Some f) = Some (f, ps)
- | strip_cong _ _ = None;
+ | strip_cong ps (PThm (("HOL.refl", _), _, _, _) % SOME f) = SOME (f, ps)
+ | strip_cong _ _ = NONE;
val subst_prf = fst (strip_combt (#2 (#der (rep_thm subst))));
val sym_prf = fst (strip_combt (#2 (#der (rep_thm sym))));
@@ -298,7 +298,7 @@
| make_subst Ts prf xs (f, ((x, y), prf') :: ps) =
let val T = fastype_of1 (Ts, x)
in if x aconv y then make_subst Ts prf (xs @ [x]) (f, ps)
- else change_type (Some [T]) subst_prf %> x %> y %>
+ else change_type (SOME [T]) subst_prf %> x %> y %>
Abs ("z", T, list_comb (incr_boundvars 1 f,
map (incr_boundvars 1) xs @ Bound 0 ::
map (incr_boundvars 1 o snd o fst) ps)) %% prf' %%
@@ -306,7 +306,7 @@
end;
fun make_sym Ts ((x, y), prf) =
- ((y, x), change_type (Some [fastype_of1 (Ts, x)]) sym_prf %> x %> y %% prf);
+ ((y, x), change_type (SOME [fastype_of1 (Ts, x)]) sym_prf %> x %> y %% prf);
fun mk_AbsP P t = AbsP ("H", P, t);
@@ -321,6 +321,6 @@
| elim_cong Ts (PThm (("HOL.iffD2", _), _, _, _) % _ % P %% prf) =
apsome (mk_AbsP P o make_subst Ts (PBound 0) [] o
apsnd (map (make_sym Ts))) (strip_cong [] (incr_pboundvars 1 0 prf))
- | elim_cong _ _ = None;
+ | elim_cong _ _ = NONE;
end;
--- a/src/HOL/Tools/sat_solver.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/sat_solver.ML Sun Feb 13 17:15:14 2005 +0100
@@ -43,7 +43,7 @@
exception NOT_CONFIGURED;
(* ------------------------------------------------------------------------- *)
-(* type of partial (satisfying) assignments: 'a i = None' means that 'a' is *)
+(* type of partial (satisfying) assignments: 'a i = NONE' means that 'a' is *)
(* a satisfying assigment regardless of the value of variable 'i' *)
(* ------------------------------------------------------------------------- *)
@@ -186,20 +186,20 @@
fun parse_std_result_file path (satisfiable, assignment_prefix, unsatisfiable) =
let
- (* 'a option -> 'a Library.option *)
+ (* 'a option -> 'a option *)
fun option (SOME a) =
- Some a
+ SOME a
| option NONE =
- None
+ NONE
(* string -> int list *)
fun int_list_from_string s =
mapfilter (option o Int.fromString) (space_explode " " s)
(* int list -> assignment *)
fun assignment_from_list [] i =
- None (* the SAT solver didn't provide a value for this variable *)
+ NONE (* the SAT solver didn't provide a value for this variable *)
| assignment_from_list (x::xs) i =
- if x=i then (Some true)
- else if x=(~i) then (Some false)
+ if x=i then (SOME true)
+ else if x=(~i) then (SOME false)
else assignment_from_list xs i
(* int list -> string list -> assignment *)
fun parse_assignment xs [] =
@@ -365,27 +365,27 @@
(* int list -> int list -> int list option *)
(* binary increment: list 'xs' of current bits, list 'ys' of all bits (lower bits first) *)
fun next_list _ ([]:int list) =
- None
+ NONE
| next_list [] (y::ys) =
- Some [y]
+ SOME [y]
| next_list (x::xs) (y::ys) =
if x=y then
(* reset the bit, continue *)
next_list xs ys
else
(* set the lowest bit that wasn't set before, keep the higher bits *)
- Some (y::x::xs)
+ SOME (y::x::xs)
(* int list -> int -> bool *)
fun assignment_from_list xs i =
i mem xs
(* int list -> SatSolver.result *)
fun solver_loop xs =
if PropLogic.eval (assignment_from_list xs) fm then
- SatSolver.SATISFIABLE (Some o (assignment_from_list xs))
+ SatSolver.SATISFIABLE (SOME o (assignment_from_list xs))
else
(case next_list xs indices of
- Some xs' => solver_loop xs'
- | None => SatSolver.UNSATISFIABLE)
+ SOME xs' => solver_loop xs'
+ | NONE => SatSolver.UNSATISFIABLE)
in
(* start with the "first" assignment (all variables are interpreted as 'false') *)
solver_loop []
@@ -454,29 +454,29 @@
val (xs', fm') = eval_and_force xs fm
in
case fm' of
- True => Some xs'
- | False => None
+ True => SOME xs'
+ | False => NONE
| _ =>
let
val x = the (fresh_var xs') (* a fresh variable must exist since 'fm' did not evaluate to 'True'/'False' *)
in
case dpll (x::xs') fm' of (* passing fm' rather than fm should save some simplification work *)
- Some xs'' => Some xs''
- | None => dpll ((~x)::xs') fm' (* now try interpreting 'x' as 'False' *)
+ SOME xs'' => SOME xs''
+ | NONE => dpll ((~x)::xs') fm' (* now try interpreting 'x' as 'False' *)
end
end
(* int list -> assignment *)
fun assignment_from_list [] i =
- None (* the DPLL procedure didn't provide a value for this variable *)
+ NONE (* the DPLL procedure didn't provide a value for this variable *)
| assignment_from_list (x::xs) i =
- if x=i then (Some true)
- else if x=(~i) then (Some false)
+ if x=i then (SOME true)
+ else if x=(~i) then (SOME false)
else assignment_from_list xs i
in
(* initially, no variable is interpreted yet *)
case dpll [] fm' of
- Some assignment => SatSolver.SATISFIABLE (assignment_from_list assignment)
- | None => SatSolver.UNSATISFIABLE
+ SOME assignment => SatSolver.SATISFIABLE (assignment_from_list assignment)
+ | NONE => SatSolver.UNSATISFIABLE
end
end (* local *)
in
--- a/src/HOL/Tools/specification_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/specification_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -86,8 +86,8 @@
in
fun proc_exprop axiomatic cos arg =
case axiomatic of
- Some axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg)
- | None => mk_definitional cos arg
+ SOME axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg)
+ | NONE => mk_definitional cos arg
end
fun add_specification_i axiomatic cos arg =
@@ -109,8 +109,8 @@
val fmap' = map Library.swap fmap
fun unthaw (f as (a,S)) =
(case assoc (fmap',a) of
- None => TVar f
- | Some b => TFree (b,S))
+ NONE => TVar f
+ | SOME b => TFree (b,S))
in
map_term_types (map_type_tvar unthaw) t
end
@@ -190,7 +190,7 @@
let
val cv = cterm_of sg v
val cT = ctyp_of_term cv
- val spec' = instantiate' [Some cT] [None,Some cv] spec
+ val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec
in
thm RS spec'
end
@@ -257,7 +257,7 @@
OuterSyntax.command "specification" "define constants by specification" K.thy_goal
(specification_decl >> (fn (cos,alt_props) =>
Toplevel.print o (Toplevel.theory_to_proof
- (process_spec None cos alt_props))))
+ (process_spec NONE cos alt_props))))
val ax_specification_decl =
P.name --
@@ -268,7 +268,7 @@
OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal
(ax_specification_decl >> (fn (axname,(cos,alt_props)) =>
Toplevel.print o (Toplevel.theory_to_proof
- (process_spec (Some axname) cos alt_props))))
+ (process_spec (SOME axname) cos alt_props))))
val _ = OuterSyntax.add_parsers [specificationP,ax_specificationP]
--- a/src/HOL/Tools/split_rule.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/split_rule.ML Sun Feb 13 17:15:14 2005 +0100
@@ -2,7 +2,7 @@
ID: $Id$
Author: Stefan Berghofer, David von Oheimb, and Markus Wenzel, TU Muenchen
-Some tools for managing tupled arguments and abstractions in rules.
+SOME tools for managing tupled arguments and abstractions in rules.
*)
signature BASIC_SPLIT_RULE =
--- a/src/HOL/Tools/typedef_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Tools/typedef_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -101,7 +101,7 @@
fun prove_nonempty thy cset goal (witn1_tac, witn_names, witn_thms, witn2_tac) =
let
val is_def = Logic.is_equals o #prop o Thm.rep_thm;
- val thms = PureThy.get_thmss thy (map (rpair None) witn_names) @ witn_thms;
+ val thms = PureThy.get_thmss thy (map (rpair NONE) witn_names) @ witn_thms;
val tac =
witn1_tac THEN
TRY (rewrite_goals_tac (filter is_def thms)) THEN
@@ -116,7 +116,7 @@
(* prepare_typedef *)
fun read_term sg used s =
- #1 (Thm.read_def_cterm (sg, K None, K None) used true (s, HOLogic.typeT));
+ #1 (Thm.read_def_cterm (sg, K NONE, K NONE) used true (s, HOLogic.typeT));
fun cert_term sg _ t = Thm.cterm_of sg t handle TERM (msg, _) => error msg;
@@ -172,9 +172,9 @@
((if def then [(name, setT, NoSyn)] else []) @
[(Rep_name, newT --> oldT, NoSyn),
(Abs_name, oldT --> newT, NoSyn)])
- |> (if def then (apsnd (Some o hd) oo (PureThy.add_defs_i false o map Thm.no_attributes))
+ |> (if def then (apsnd (SOME o hd) oo (PureThy.add_defs_i false o map Thm.no_attributes))
[Logic.mk_defpair (setC, set)]
- else rpair None)
+ else rpair NONE)
|>>> PureThy.add_axioms_i [((typedef_name, typedef_prop),
[apsnd (fn cond_axm => Drule.standard (nonempty RS cond_axm))])]
|>> Theory.add_finals_i false [RepC, AbsC]
@@ -251,10 +251,10 @@
fun sane_typedef prep_term def opt_name typ set opt_morphs tac =
gen_typedef prep_term def
- (if_none opt_name (#1 typ)) typ set opt_morphs all_tac [] [] (Some tac);
+ (if_none opt_name (#1 typ)) typ set opt_morphs all_tac [] [] (SOME tac);
fun add_typedef_x name typ set names thms tac =
- #1 o gen_typedef read_term true name typ set None (Tactic.rtac exI 1) names thms tac;
+ #1 o gen_typedef read_term true name typ set NONE (Tactic.rtac exI 1) names thms tac;
val add_typedef = sane_typedef read_term;
val add_typedef_i = sane_typedef cert_term;
@@ -284,7 +284,7 @@
val (gr'', ps) =
foldl_map (Codegen.invoke_codegen thy dep true) (gr', ts);
val id = Codegen.mk_const_id (sign_of thy) s
- in Some (gr'', Codegen.mk_app brack (Pretty.str id) ps) end;
+ in SOME (gr'', Codegen.mk_app brack (Pretty.str id) ps) end;
fun get_name (Type (tname, _)) = tname
| get_name _ = "";
fun lookup f T = if_none (apsome f (Symtab.lookup
@@ -298,8 +298,8 @@
else if lookup #3 U = s andalso
is_none (Codegen.get_assoc_type thy (get_name U))
then mk_fun s U ts
- else None
- | _ => None)
+ else NONE
+ | _ => NONE)
end;
fun mk_tyexpr [] s = Pretty.str s
@@ -308,9 +308,9 @@
fun typedef_tycodegen thy gr dep brack (Type (s, Ts)) =
(case Symtab.lookup (TypedefData.get thy, s) of
- None => None
- | Some (newT as Type (tname, Us), oldT, Abs_name, Rep_name) =>
- if is_some (Codegen.get_assoc_type thy tname) then None else
+ NONE => NONE
+ | SOME (newT as Type (tname, Us), oldT, Abs_name, Rep_name) =>
+ if is_some (Codegen.get_assoc_type thy tname) then NONE else
let
val sg = sign_of thy;
val Abs_id = Codegen.mk_const_id sg Abs_name;
@@ -323,7 +323,7 @@
val (gr'', p :: ps) = foldl_map
(Codegen.invoke_tycodegen thy Abs_id false)
(Graph.add_edge (Abs_id, dep)
- (Graph.new_node (Abs_id, (None, "")) gr'), oldT :: Us);
+ (Graph.new_node (Abs_id, (NONE, "")) gr'), oldT :: Us);
val s =
Pretty.string_of (Pretty.block [Pretty.str "datatype ",
mk_tyexpr ps ty_id,
@@ -351,11 +351,11 @@
Codegen.mk_gen sg false [] "" oldT, Pretty.brk 1,
Pretty.str "i);"]]) ^ "\n\n"
else "")
- in Graph.map_node Abs_id (K (None, s)) gr'' end
+ in Graph.map_node Abs_id (K (NONE, s)) gr'' end
in
- Some (gr'', mk_tyexpr qs ty_id)
+ SOME (gr'', mk_tyexpr qs ty_id)
end)
- | typedef_tycodegen thy gr dep brack _ = None;
+ | typedef_tycodegen thy gr dep brack _ = NONE;
val setup =
[TypedefData.init,
@@ -376,8 +376,8 @@
val typedef_proof_decl =
Scan.optional (P.$$$ "(" |-- P.!!!
- (((P.$$$ "open" >> K false) -- Scan.option P.name || P.name >> (fn s => (true, Some s)))
- --| P.$$$ ")")) (true, None) --
+ (((P.$$$ "open" >> K false) -- Scan.option P.name || P.name >> (fn s => (true, SOME s)))
+ --| P.$$$ ")")) (true, NONE) --
(P.type_args -- P.name) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
Scan.option (P.$$$ "morphisms" |-- P.!!! (P.name -- P.name));
--- a/src/HOL/Transitive_Closure.thy Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/Transitive_Closure.thy Sun Feb 13 17:15:14 2005 +0100
@@ -470,8 +470,8 @@
| decr (Const ("Transitive_Closure.trancl", _ ) $ r) = (r,"r+")
| decr r = (r,"r");
val (rel,r) = decr rel;
- in Some (a,b,rel,r) end
- | dec _ = None
+ in SOME (a,b,rel,r) end
+ | dec _ = NONE
in dec t end;
end); (* struct *)
--- a/src/HOL/UNITY/Comp/Alloc.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/UNITY/Comp/Alloc.ML Sun Feb 13 17:15:14 2005 +0100
@@ -715,7 +715,7 @@
qed "System_correct";
-(** Some lemmas no longer used **)
+(** SOME lemmas no longer used **)
Goal "non_dummy = (% (g,a,r). (| giv = g, ask = a, rel = r |)) o \
\ (funPair giv (funPair ask rel))";
--- a/src/HOL/antisym_setup.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/antisym_setup.ML Sun Feb 13 17:15:14 2005 +0100
@@ -12,30 +12,30 @@
val less = Const("op <",T);
val t = HOLogic.mk_Trueprop(le $ s $ r);
in case Library.find_first (prp t) prems of
- None =>
+ NONE =>
let val t = HOLogic.mk_Trueprop(HOLogic.not_const $ (less $ r $ s))
in case Library.find_first (prp t) prems of
- None => None
- | Some thm => Some(mk_meta_eq(thm RS linorder_antisym_conv1))
+ NONE => NONE
+ | SOME thm => SOME(mk_meta_eq(thm RS linorder_antisym_conv1))
end
- | Some thm => Some(mk_meta_eq(thm RS order_antisym_conv))
+ | SOME thm => SOME(mk_meta_eq(thm RS order_antisym_conv))
end
- handle THM _ => None;
+ handle THM _ => NONE;
fun prove_antisym_less sg ss (NotC $ ((less as Const(_,T)) $ r $ s)) =
let val prems = prems_of_ss ss;
val le = Const("op <=",T);
val t = HOLogic.mk_Trueprop(le $ r $ s);
in case Library.find_first (prp t) prems of
- None =>
+ NONE =>
let val t = HOLogic.mk_Trueprop(NotC $ (less $ s $ r))
in case Library.find_first (prp t) prems of
- None => None
- | Some thm => Some(mk_meta_eq(thm RS linorder_antisym_conv3))
+ NONE => NONE
+ | SOME thm => SOME(mk_meta_eq(thm RS linorder_antisym_conv3))
end
- | Some thm => Some(mk_meta_eq(thm RS linorder_antisym_conv2))
+ | SOME thm => SOME(mk_meta_eq(thm RS linorder_antisym_conv2))
end
- handle THM _ => None;
+ handle THM _ => NONE;
val antisym_le = Simplifier.simproc (Theory.sign_of (the_context()))
"antisym le" ["(x::'a::order) <= y"] prove_antisym_le;
--- a/src/HOL/arith_data.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/arith_data.ML Sun Feb 13 17:15:14 2005 +0100
@@ -38,11 +38,11 @@
if HOLogic.is_zero tm then []
else
(case try HOLogic.dest_Suc tm of
- Some t => one :: dest_sum t
- | None =>
+ SOME t => one :: dest_sum t
+ | NONE =>
(case try dest_plus tm of
- Some (t, u) => dest_sum t @ dest_sum u
- | None => [tm]));
+ SOME (t, u) => dest_sum t @ dest_sum u
+ | NONE => [tm]));
(** generic proof tools **)
@@ -96,7 +96,7 @@
end;
fun gen_uncancel_tac rule ct =
- rtac (instantiate' [] [None, Some ct] (rule RS subst_equals)) 1;
+ rtac (instantiate' [] [NONE, SOME ct] (rule RS subst_equals)) 1;
(* nat eq *)
@@ -207,7 +207,7 @@
val name = "HOL/arith";
type T = {splits: thm list, inj_consts: (string * typ)list, discrete: string list, presburger: (int -> tactic) option};
- val empty = {splits = [], inj_consts = [], discrete = [], presburger = None};
+ val empty = {splits = [], inj_consts = [], discrete = [], presburger = NONE};
val copy = I;
val prep_ext = I;
fun merge ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1, presburger= presburger1},
@@ -215,7 +215,7 @@
{splits = Drule.merge_rules (splits1, splits2),
inj_consts = merge_lists inj_consts1 inj_consts2,
discrete = merge_lists discrete1 discrete2,
- presburger = (case presburger1 of None => presburger2 | p => p)};
+ presburger = (case presburger1 of NONE => presburger2 | p => p)};
fun print _ _ = ();
end;
@@ -239,8 +239,8 @@
fun nT (Type("fun",[N,_])) = N = HOLogic.natT
| nT _ = false;
-fun add_atom(t,m,(p,i)) = (case assoc(p,t) of None => ((t,m)::p,i)
- | Some n => (overwrite(p,(t,ratadd(n,m))), i));
+fun add_atom(t,m,(p,i)) = (case assoc(p,t) of NONE => ((t,m)::p,i)
+ | SOME n => (overwrite(p,(t,ratadd(n,m))), i));
exception Zero;
@@ -282,19 +282,19 @@
| demult(atom as Const("HOL.divide",_) $ t $ (Const("Numeral.number_of",_)$dent), m) =
(let val den = HOLogic.dest_binum dent
in if den = 0 then raise Zero else demult(t,ratmul(m, ratinv(rat_of_int den))) end
- handle TERM _ => (Some atom,m))
- | demult(Const("0",_),m) = (None, rat_of_int 0)
- | demult(Const("1",_),m) = (None, m)
+ handle TERM _ => (SOME atom,m))
+ | demult(Const("0",_),m) = (NONE, rat_of_int 0)
+ | demult(Const("1",_),m) = (NONE, m)
| demult(t as Const("Numeral.number_of",_)$n,m) =
- ((None,ratmul(m,rat_of_int(HOLogic.dest_binum n)))
- handle TERM _ => (Some t,m))
+ ((NONE,ratmul(m,rat_of_int(HOLogic.dest_binum n)))
+ handle TERM _ => (SOME t,m))
| demult(Const("uminus",_)$t, m) = demult(t,ratmul(m,rat_of_int(~1)))
| demult(t as Const f $ x, m) =
- (if f mem inj_consts then Some x else Some t,m)
- | demult(atom,m) = (Some atom,m)
+ (if f mem inj_consts then SOME x else SOME t,m)
+ | demult(atom,m) = (SOME atom,m)
-and atomult(mC,atom,t,m) = (case demult(t,m) of (None,m') => (Some atom,m')
- | (Some t',m') => (Some(mC $ atom $ t'),m'))
+and atomult(mC,atom,t,m) = (case demult(t,m) of (NONE,m') => (SOME atom,m')
+ | (SOME t',m') => (SOME(mC $ atom $ t'),m'))
in demult end;
fun decomp2 inj_consts (rel,lhs,rhs) =
@@ -310,12 +310,12 @@
| poly(Const("Suc",_)$t, m, (p,i)) = poly(t, m, (p,ratadd(i,m)))
| poly(t as Const("op *",_) $ _ $ _, m, pi as (p,i)) =
(case demult inj_consts (t,m) of
- (None,m') => (p,ratadd(i,m))
- | (Some u,m') => add_atom(u,m',pi))
+ (NONE,m') => (p,ratadd(i,m))
+ | (SOME u,m') => add_atom(u,m',pi))
| poly(t as Const("HOL.divide",_) $ _ $ _, m, pi as (p,i)) =
(case demult inj_consts (t,m) of
- (None,m') => (p,ratadd(i,m'))
- | (Some u,m') => add_atom(u,m',pi))
+ (NONE,m') => (p,ratadd(i,m'))
+ | (SOME u,m') => add_atom(u,m',pi))
| poly(all as (Const("Numeral.number_of",_)$t,m,(p,i))) =
((p,ratadd(i,ratmul(m,rat_of_int(HOLogic.dest_binum t))))
handle TERM _ => add_atom all)
@@ -327,14 +327,14 @@
and (q,j) = poly(rhs,rat_of_int 1,([],rat_of_int 0))
in case rel of
- "op <" => Some(p,i,"<",q,j)
- | "op <=" => Some(p,i,"<=",q,j)
- | "op =" => Some(p,i,"=",q,j)
- | _ => None
- end handle Zero => None;
+ "op <" => SOME(p,i,"<",q,j)
+ | "op <=" => SOME(p,i,"<=",q,j)
+ | "op =" => SOME(p,i,"=",q,j)
+ | _ => NONE
+ end handle Zero => NONE;
-fun negate(Some(x,i,rel,y,j,d)) = Some(x,i,"~"^rel,y,j,d)
- | negate None = None;
+fun negate(SOME(x,i,rel,y,j,d)) = SOME(x,i,"~"^rel,y,j,d)
+ | negate NONE = NONE;
fun of_lin_arith_sort sg U =
Type.of_sort (Sign.tsig_of sg) (U,["Ring_and_Field.ordered_idom"])
@@ -350,15 +350,15 @@
(case T of
Type("fun",[U,_]) =>
(case allows_lin_arith sg discrete U of
- (true,d) => (case decomp2 inj_consts xxx of None => None
- | Some(p,i,rel,q,j) => Some(p,i,rel,q,j,d))
- | (false,_) => None)
- | _ => None);
+ (true,d) => (case decomp2 inj_consts xxx of NONE => NONE
+ | SOME(p,i,rel,q,j) => SOME(p,i,rel,q,j,d))
+ | (false,_) => NONE)
+ | _ => NONE);
fun decomp2 data (_$(Const(rel,T)$lhs$rhs)) = decomp1 data (T,(rel,lhs,rhs))
| decomp2 data (_$(Const("Not",_)$(Const(rel,T)$lhs$rhs))) =
negate(decomp1 data (T,(rel,lhs,rhs)))
- | decomp2 data _ = None
+ | decomp2 data _ = NONE
fun decomp sg =
let val {discrete, inj_consts, ...} = ArithTheoryData.get_sg sg
@@ -463,7 +463,7 @@
fun presburger_tac i st =
(case ArithTheoryData.get_sg (sign_of_thm st) of
- {presburger = Some tac, ...} =>
+ {presburger = SOME tac, ...} =>
(tracing "Simple arithmetic decision procedure failed.\nNow trying full Presburger arithmetic..."; tac i st)
| _ => no_tac st);
@@ -502,26 +502,26 @@
let val r' = Tr $ (c $ t $ s)
in
case Library.find_first (prp r') prems of
- None =>
+ NONE =>
let val r' = Tr $ (HOLogic.not_const $ (Const("op <",T) $ s $ t))
in case Library.find_first (prp r') prems of
- None => []
- | Some thm' => [(thm' RS not_lessD) RS (thm RS antisym)]
+ NONE => []
+ | SOME thm' => [(thm' RS not_lessD) RS (thm RS antisym)]
end
- | Some thm' => [thm' RS (thm RS antisym)]
+ | SOME thm' => [thm' RS (thm RS antisym)]
end
| Tr $ (Const("Not",_) $ (Const("op <",T) $ s $ t)) =>
let val r' = Tr $ (Const("op <=",T) $ s $ t)
in
case Library.find_first (prp r') prems of
- None =>
+ NONE =>
let val r' = Tr $ (HOLogic.not_const $ (Const("op <",T) $ t $ s))
in case Library.find_first (prp r') prems of
- None => []
- | Some thm' =>
+ NONE => []
+ | SOME thm' =>
[(thm' RS not_lessD) RS ((thm RS not_lessD) RS antisym)]
end
- | Some thm' => [thm' RS ((thm RS not_lessD) RS antisym)]
+ | SOME thm' => [thm' RS ((thm RS not_lessD) RS antisym)]
end
| _ => []
end
--- a/src/HOL/eqrule_HOL_data.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/eqrule_HOL_data.ML Sun Feb 13 17:15:14 2005 +0100
@@ -25,9 +25,9 @@
val conjunct1 = thm "conjunct1";
fun mk_eq th = case concl_of th of
- Const("==",_)$_$_ => Some (th)
- | _$(Const("op =",_)$_$_) => Some (th RS eq_reflection)
- | _ => None;
+ Const("==",_)$_$_ => SOME (th)
+ | _$(Const("op =",_)$_$_) => SOME (th RS eq_reflection)
+ | _ => NONE;
val tranformation_pairs =
[("op -->", [mp]), ("op &", [conjunct1,conjunct2]),
@@ -45,8 +45,8 @@
(case Term.head_of p of
Const(a,_) =>
(case Library.assoc(pairs,a) of
- Some(rls) => flat (map atoms ([th] RL rls))
- | None => [th])
+ SOME(rls) => flat (map atoms ([th] RL rls))
+ | NONE => [th])
| _ => [th])
| _ => [th])
in atoms end;
--- a/src/HOL/ex/SVC_Oracle.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/ex/SVC_Oracle.ML Sun Feb 13 17:15:14 2005 +0100
@@ -39,8 +39,8 @@
Free _ => t (*but not existing Vars, lest the names clash*)
| Bound _ => t
| _ => (case gen_assoc Pattern.aeconv (!pairs, t) of
- Some v => v
- | None => insert t)
+ SOME v => v
+ | NONE => insert t)
(*abstraction of a numeric literal*)
fun lit (t as Const("0", _)) = t
| lit (t as Const("1", _)) = t
--- a/src/HOL/ex/mesontest2.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/ex/mesontest2.ML Sun Feb 13 17:15:14 2005 +0100
@@ -193,7 +193,7 @@
(*
* Original timings for John Harrison's MESON_TAC.
* Timings below on a 600MHz Pentium III (perch)
- * Some timings below refer to griffon, which is a dual 2.5GHz Power Mac G5.
+ * SOME timings below refer to griffon, which is a dual 2.5GHz Power Mac G5.
*
* A few variable names have been changed to avoid clashing with constants.
*
--- a/src/HOL/ex/svc_funcs.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/ex/svc_funcs.ML Sun Feb 13 17:15:14 2005 +0100
@@ -76,8 +76,8 @@
">/dev/null 2>&1"))
val svc_output =
(case Library.try File.read svc_output_file of
- Some out => out
- | None => error "SVC returned no output");
+ SOME out => out
+ | NONE => error "SVC returned no output");
in
if ! trace then tracing ("SVC Returns:\n" ^ svc_output)
else (File.rm svc_input_file; File.rm svc_output_file);
--- a/src/HOL/ex/svc_test.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/ex/svc_test.ML Sun Feb 13 17:15:14 2005 +0100
@@ -18,7 +18,7 @@
by (svc_tac 1);
qed "";
-(** Some big tautologies supplied by John Harrison **)
+(** SOME big tautologies supplied by John Harrison **)
(*Tautology name: puz013_1. Auto_tac manages; Blast_tac and Fast_tac take
a minute or more.*)
--- a/src/HOL/simpdata.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/simpdata.ML Sun Feb 13 17:15:14 2005 +0100
@@ -96,12 +96,12 @@
structure Quantifier1 = Quantifier1Fun
(struct
(*abstract syntax*)
- fun dest_eq((c as Const("op =",_)) $ s $ t) = Some(c,s,t)
- | dest_eq _ = None;
- fun dest_conj((c as Const("op &",_)) $ s $ t) = Some(c,s,t)
- | dest_conj _ = None;
- fun dest_imp((c as Const("op -->",_)) $ s $ t) = Some(c,s,t)
- | dest_imp _ = None;
+ fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t)
+ | dest_eq _ = NONE;
+ fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t)
+ | dest_conj _ = NONE;
+ fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t)
+ | dest_imp _ = NONE;
val conj = HOLogic.conj
val imp = HOLogic.imp
(*rules*)
@@ -150,9 +150,9 @@
Simplifier.simproc (Theory.sign_of (the_context ())) "let_simp" ["Let x f"]
(fn sg => fn ss => fn t =>
(case t of (Const ("Let",_)$x$f) => (* x and f are already in normal form *)
- if not (!use_let_simproc) then None
+ if not (!use_let_simproc) then NONE
else if is_Free x orelse is_Bound x orelse is_Const x
- then Some Let_def
+ then SOME Let_def
else
let
val n = case f of (Abs (x,_,_)) => x | _ => "x";
@@ -166,8 +166,8 @@
then
let
val rl = cterm_instantiate [(f_Let_unfold,cf)] Let_unfold;
- in Some (rl OF [fx_g]) end
- else if betapply (f,x) aconv g then None (* avoid identity conversion *)
+ in SOME (rl OF [fx_g]) end
+ else if betapply (f,x) aconv g then NONE (* avoid identity conversion *)
else let
val abs_g'= Abs (n,xT,g');
val g'x = abs_g'$x;
@@ -176,9 +176,9 @@
[(f_Let_folded,cterm_of sg f),
(g_Let_folded,cterm_of sg abs_g')]
Let_folded;
- in Some (rl OF [transitive fx_g g_g'x]) end)
+ in SOME (rl OF [transitive fx_g g_g'x]) end)
end
- | _ => None))
+ | _ => NONE))
end
(*** Case splitting ***)
@@ -196,7 +196,7 @@
(* Expects Trueprop(.) if not == *)
fun mk_eq_True r =
- Some (r RS meta_eq_to_obj_eq RS Eq_TrueI) handle Thm.THM _ => None;
+ SOME (r RS meta_eq_to_obj_eq RS Eq_TrueI) handle Thm.THM _ => NONE;
(*Congruence rules for = (instead of ==)*)
fun mk_meta_cong rl =
@@ -253,8 +253,8 @@
(case head_of p of
Const(a,_) =>
(case assoc(pairs,a) of
- Some(rls) => flat (map atoms ([th] RL rls))
- | None => [th])
+ SOME(rls) => flat (map atoms ([th] RL rls))
+ | NONE => [th])
| _ => [th])
| _ => [th])
in atoms end;
--- a/src/HOL/thy_syntax.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOL/thy_syntax.ML Sun Feb 13 17:15:14 2005 +0100
@@ -23,7 +23,7 @@
end;
val typedef_decl =
- optional ("(" $$-- name --$$ ")" >> Some) None --
+ optional ("(" $$-- name --$$ ")" >> SOME) NONE --
type_args -- name -- opt_infix --$$ "=" -- string -- opt_witness
>> mk_typedef_decl;
@@ -33,7 +33,7 @@
val record_decl =
(type_args -- name >> (mk_pair o apfst mk_list)) --$$ "="
- -- optional (typ --$$ "+" >> (parens o cat "Some")) "None"
+ -- optional (typ --$$ "+" >> (parens o cat "SOME")) "NONE"
-- repeat1 ((name --$$ "::" -- !! (typ -- opt_mixfix)) >> mk_triple2)
>> (fn ((x, y), zs) => cat_lines [x, y, mk_big_list zs]);
@@ -95,8 +95,8 @@
fun mk_bind_thms_string names =
(case find_first (not o ThmDatabase.is_ml_identifier) names of
- Some id => (warning (id ^ " is not a valid identifier"); "")
- | None =>
+ SOME id => (warning (id ^ " is not a valid identifier"); "")
+ | NONE =>
let
fun mk_dt_struct (s, (id, i)) =
s ^ "structure " ^ id ^ " =\n\
@@ -156,10 +156,10 @@
\ case_thms, split_thms, induction, size, simps}) =\n\
\ DatatypePackage.rep_datatype " ^
(case names of
- Some names => "(Some [" ^ Library.commas_quote names ^ "])\n " ^
+ SOME names => "(SOME [" ^ Library.commas_quote names ^ "])\n " ^
mk_thmss distinct ^ "\n " ^ mk_thmss inject ^
"\n " ^ mk_thm induct ^ " thy;\nin\n" ^ mk_bind_thms_string names
- | None => "None\n " ^ mk_thmss distinct ^ "\n " ^ mk_thmss inject ^
+ | NONE => "NONE\n " ^ mk_thmss distinct ^ "\n " ^ mk_thmss inject ^
"\n " ^ mk_thm induct ^ " thy;\nin\n") ^
"val thy = thy;\nend;\nval thy = thy\n";
@@ -180,7 +180,7 @@
simple_typ || ("(" $$-- complex_typ --$$ ")"));
val constructor = name -- opt_typs -- opt_mixfix >> (fn ((n, Ts), mx) =>
parens (n ^ ", " ^ brackets (Library.commas_quote Ts) ^ ", " ^ mx));
- val opt_name = optional ("(" $$-- name --$$ ")" >> Some) None
+ val opt_name = optional ("(" $$-- name --$$ ")" >> SOME) NONE
fun optlistlist s = optional (s $$-- enum "and" (list name)) [[]]
@@ -189,7 +189,7 @@
enum1 "and" (opt_name -- type_args -- name -- opt_infix --$$ "=" --
enum1 "|" constructor) >> mk_dt_string;
val rep_datatype_decl =
- ((optional ((repeat1 (name >> unenclose)) >> Some) None) --
+ ((optional ((repeat1 (name >> unenclose)) >> SOME) NONE) --
optlistlist "distinct" -- optlistlist "inject" --
("induct" $$-- name)) >> mk_rep_dt_string;
end;
--- a/src/HOLCF/IOA/meta_theory/CompoTraces.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOLCF/IOA/meta_theory/CompoTraces.ML Sun Feb 13 17:15:14 2005 +0100
@@ -5,7 +5,7 @@
Compositionality on Trace level.
*)
-simpset_ref () := simpset() setmksym (K None);
+simpset_ref () := simpset() setmksym (K NONE);
(* ---------------------------------------------------------------- *)
section "mksch rewrite rules";
@@ -1225,4 +1225,4 @@
qed"compositionality_tr_modules";
-simpset_ref () := simpset() setmksym (Some o symmetric_fun);
+simpset_ref () := simpset() setmksym (SOME o symmetric_fun);
--- a/src/HOLCF/IOA/meta_theory/ioa_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOLCF/IOA/meta_theory/ioa_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -98,7 +98,7 @@
fun unqualify s = implode(rev(afpl(rev(explode s))));
val q_atypstr = act_name thy atyp;
val uq_atypstr = unqualify q_atypstr;
-val prem = prems_of (get_thm thy (uq_atypstr ^ ".induct", None));
+val prem = prems_of (get_thm thy (uq_atypstr ^ ".induct", NONE));
in
extract_constrs thy prem
handle malformed =>
@@ -284,7 +284,7 @@
let
fun left_of (( _ $ left) $ _) = left |
left_of _ = raise malformed;
-val aut_def = concl_of (get_thm thy (aut_name ^ "_def", None));
+val aut_def = concl_of (get_thm thy (aut_name ^ "_def", NONE));
in
(#T(rep_cterm(cterm_of (sign_of thy) (left_of aut_def))))
handle malformed => error ("malformed_theorem : " ^ aut_name ^ "_def")
--- a/src/HOLCF/Up2.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOLCF/Up2.ML Sun Feb 13 17:15:14 2005 +0100
@@ -77,7 +77,7 @@
qed "monofun_Ifup2";
(* ------------------------------------------------------------------------ *)
-(* Some kind of surjectivity lemma *)
+(* SOME kind of surjectivity lemma *)
(* ------------------------------------------------------------------------ *)
Goal "z=Iup(x) ==> Iup(Ifup(LAM x. x)(z)) = z";
--- a/src/HOLCF/adm.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOLCF/adm.ML Sun Feb 13 17:15:14 2005 +0100
@@ -141,15 +141,15 @@
(*** the admissibility tactic ***)
fun try_dest_adm (Const _ $ (Const (name, _) $ Abs abs)) =
- if name = adm_name then Some abs else None
- | try_dest_adm _ = None;
+ if name = adm_name then SOME abs else NONE
+ | try_dest_adm _ = NONE;
fun adm_tac tac i state =
state |>
let val goali = nth_subgoal i state in
(case try_dest_adm (Logic.strip_assums_concl goali) of
- None => no_tac
- | Some (s, T, t) =>
+ NONE => no_tac
+ | SOME (s, T, t) =>
let
val sign = sign_of_thm state;
val prems = Logic.strip_assums_hyp goali;
--- a/src/HOLCF/domain/extender.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOLCF/domain/extender.ML Sun Feb 13 17:15:14 2005 +0100
@@ -13,7 +13,7 @@
Exception-
TERM
("typ_of_term: bad encoding of type",
- [Abs ("uu", "_", Const ("None", "_"))]) raised
+ [Abs ("uu", "_", Const ("NONE", "_"))]) raised
but this works fine:
domain Empty = silly Empty
@@ -67,15 +67,15 @@
| rm_sorts (TVar(s,_)) = TVar(s,[])
and remove_sorts l = map rm_sorts l;
fun analyse indirect (TFree(v,s)) = (case assoc_string(tvars,v) of
- None => error ("Free type variable " ^ quote v ^ " on rhs.")
- | Some sort => if eq_set_string (s,defaultS) orelse
+ NONE => error ("Free type variable " ^ quote v ^ " on rhs.")
+ | SOME sort => if eq_set_string (s,defaultS) orelse
eq_set_string (s,sort )
then TFree(distinct_name v,sort)
else error ("Inconsistent sort constraint" ^
" for type variable " ^ quote v))
| analyse indirect (t as Type(s,typl)) = (case assoc_string(dtnvs,s)of
- None => Type(s,map (analyse true) typl)
- | Some typevars => if indirect
+ NONE => Type(s,map (analyse true) typl)
+ | SOME typevars => if indirect
then error ("Indirect recursion of type " ^
quote (string_of_typ sg t))
else if dname <> s orelse (** BUG OR FEATURE?:
@@ -160,7 +160,7 @@
val domains_decl =
Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.and_list1 domain_decl
>> (fn (opt_name, doms) =>
- (case opt_name of None => space_implode "_" (map (#1 o #1) doms) | Some s => s, doms));
+ (case opt_name of NONE => space_implode "_" (map (#1 o #1) doms) | SOME s => s, doms));
val domainP =
OuterSyntax.command "domain" "define recursive domains (HOLCF)" K.thy_decl
--- a/src/HOLCF/domain/interface.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOLCF/domain/interface.ML Sun Feb 13 17:15:14 2005 +0100
@@ -12,10 +12,10 @@
(* --- generation of bindings for axioms and theorems in .thy.ML - *)
- fun get dname name = "get_thm thy (" ^ Library.quote (dname ^ "." ^ name) ^ ", None)";
+ fun get dname name = "get_thm thy (" ^ Library.quote (dname ^ "." ^ name) ^ ", NONE)";
fun gen_vals dname name = "val "^ name ^ " = get_thm" ^
(if hd (rev (Symbol.explode name)) = "s" then "s" else "")^
- " thy (" ^ Library.quote (dname ^ "." ^ name) ^ ", None);";
+ " thy (" ^ Library.quote (dname ^ "." ^ name) ^ ", NONE);";
fun gen_vall name l = "val "^name^" = "^ mk_list l^";";
val rews = "iso_rews @ when_rews @\n\
\ con_rews @ sel_rews @ dis_rews @ dist_les @ dist_eqs @\n\
@@ -49,9 +49,9 @@
val comp_dname = implode (separate "_" dnames);
val conss' = ListPair.map
(fn (dnam,cons'') => let fun sel n m = upd_second
- (fn None => dnam^"_sel_"^(string_of_int n)^
+ (fn NONE => dnam^"_sel_"^(string_of_int n)^
"_"^(string_of_int m)
- | Some s => s)
+ | SOME s => s)
fun fill_sels n con = upd_third (mapn (sel n) 1) con;
in mapn fill_sels 1 cons'' end)
(dnames, map snd eqs'');
@@ -119,10 +119,10 @@
(* here ident may be used instead of name'
to avoid ambiguity with standard mixfix option *)
val domain_arg = "(" $$-- (optional ($$ "lazy" >> K true) false)
- -- (optional ((name' >> Some) --$$ "::") None)
+ -- (optional ((name' >> SOME) --$$ "::") NONE)
-- typ --$$ ")"
>> (fn ((lazy,sel),tp) => (lazy,sel,tp))
- || typ >> (fn x => (false,None,x))
+ || typ >> (fn x => (false,NONE,x))
val domain_cons = name' -- !! (repeat domain_arg)
-- ThyParse.opt_mixfix
>> (fn ((con,args),syn) => (con,syn,args));
--- a/src/HOLCF/domain/library.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOLCF/domain/library.ML Sun Feb 13 17:15:14 2005 +0100
@@ -64,10 +64,10 @@
fun nonreserved s = if s mem ["n","x","f","P"] then s^"'" else s;
fun index_vnames(vn::vns,occupied) =
(case assoc(occupied,vn) of
- None => if vn mem vns
+ NONE => if vn mem vns
then (vn^"1") :: index_vnames(vns,(vn,1) ::occupied)
else vn :: index_vnames(vns, occupied)
- | Some(i) => (vn^(string_of_int (i+1)))
+ | SOME(i) => (vn^(string_of_int (i+1)))
:: index_vnames(vns,(vn,i+1)::occupied))
| index_vnames([],occupied) = [];
in index_vnames(map nonreserved ids, [("O",0),("o",0)]) end;
--- a/src/HOLCF/domain/theorems.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/HOLCF/domain/theorems.ML Sun Feb 13 17:15:14 2005 +0100
@@ -18,7 +18,7 @@
(* ----- general proof facilities ------------------------------------------- *)
fun inferT sg pre_tm =
- #1 (Sign.infer_types (Sign.pp sg) sg (K None) (K None) [] true ([pre_tm],propT));
+ #1 (Sign.infer_types (Sign.pp sg) sg (K NONE) (K NONE) [] true ([pre_tm],propT));
fun pg'' thy defs t = let val sg = sign_of thy;
val ct = Thm.cterm_of sg (inferT sg t);
@@ -65,7 +65,7 @@
(* ----- getting the axioms and definitions --------------------------------- *)
-local fun ga s dn = get_thm thy (dn ^ "." ^ s, None) in
+local fun ga s dn = get_thm thy (dn ^ "." ^ s, NONE) in
val ax_abs_iso = ga "abs_iso" dname;
val ax_rep_iso = ga "rep_iso" dname;
val ax_when_def = ga "when_def" dname;
@@ -341,7 +341,7 @@
(* ----- getting the composite axiom and definitions ------------------------ *)
-local fun ga s dn = get_thm thy (dn ^ "." ^ s, None) in
+local fun ga s dn = get_thm thy (dn ^ "." ^ s, NONE) in
val axs_reach = map (ga "reach" ) dnames;
val axs_take_def = map (ga "take_def" ) dnames;
val axs_finite_def = map (ga "finite_def") dnames;
@@ -349,8 +349,8 @@
val ax_bisim_def = ga "bisim_def" comp_dnam;
end; (* local *)
-local fun gt s dn = get_thm thy (dn ^ "." ^ s, None);
- fun gts s dn = get_thms thy (dn ^ "." ^ s, None) in
+local fun gt s dn = get_thm thy (dn ^ "." ^ s, NONE);
+ fun gts s dn = get_thms thy (dn ^ "." ^ s, NONE) in
val cases = map (gt "casedist" ) dnames;
val con_rews = flat (map (gts "con_rews" ) dnames);
val copy_rews = flat (map (gts "copy_rews") dnames);
--- a/src/LCF/ex/ROOT.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/LCF/ex/ROOT.ML Sun Feb 13 17:15:14 2005 +0100
@@ -3,7 +3,7 @@
Author: Tobias Nipkow
Copyright 1991 University of Cambridge
-Some examples from Lawrence Paulson's book Logic and Computation.
+SOME examples from Lawrence Paulson's book Logic and Computation.
*)
time_use_thy "Ex1";
--- a/src/Provers/Arith/abel_cancel.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/abel_cancel.ML Sun Feb 13 17:15:14 2005 +0100
@@ -119,8 +119,8 @@
simp_tac prepare_ss 1 THEN
IF_UNSOLVED (simp_tac cancel_ss 1) THEN
IF_UNSOLVED (simp_tac inverse_ss 1))
- in Some thm end
- handle Cancel => None;
+ in SOME thm end
+ handle Cancel => NONE;
val sum_conv =
@@ -178,8 +178,8 @@
simp_tac prepare_ss 1 THEN
simp_tac sum_cancel_ss 1 THEN
IF_UNSOLVED (simp_tac add_ac_ss 1))
- in Some thm end
- handle Cancel => None;
+ in SOME thm end
+ handle Cancel => NONE;
val rel_conv =
Simplifier.simproc_i (Sign.deref sg_ref) "cancel_relations"
--- a/src/Provers/Arith/abstract_numerals.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/abstract_numerals.ML Sun Feb 13 17:15:14 2005 +0100
@@ -61,8 +61,8 @@
[Data.norm_tac [Data.numeral_0_eq_0, Data.numeral_1_eq_1]] sg
(t, t'))
end
- handle TERM _ => None
- | TYPE _ => None; (*Typically (if thy doesn't include Numeral)
+ handle TERM _ => NONE
+ | TYPE _ => NONE; (*Typically (if thy doesn't include Numeral)
Undeclared type constructor "Numeral.bin"*)
end;
--- a/src/Provers/Arith/assoc_fold.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/assoc_fold.ML Sun Feb 13 17:15:14 2005 +0100
@@ -57,8 +57,8 @@
val th = Tactic.prove sg [] [] (Logic.mk_equals (lhs, rhs))
(fn _ => rtac Data.eq_reflection 1 THEN
simp_tac assoc_ss 1)
- in Some th end
- handle Assoc_fail => None;
+ in SOME th end
+ handle Assoc_fail => NONE;
end;
--- a/src/Provers/Arith/cancel_div_mod.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/cancel_div_mod.ML Sun Feb 13 17:15:14 2005 +0100
@@ -73,10 +73,10 @@
fun proc sg _ t =
let val (divs,mods) = coll_div_mod t ([],[])
- in if null divs orelse null mods then None
+ in if null divs orelse null mods then NONE
else case divs inter mods of
- pq :: _ => Some(cancel sg t pq)
- | [] => None
+ pq :: _ => SOME(cancel sg t pq)
+ | [] => NONE
end
end
\ No newline at end of file
--- a/src/Provers/Arith/cancel_factor.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/cancel_factor.ML Sun Feb 13 17:15:14 2005 +0100
@@ -56,18 +56,18 @@
fun proc sg _ t =
(case try Data.dest_bal t of
- None => None
- | Some bal =>
+ NONE => NONE
+ | SOME bal =>
(case pairself Data.dest_sum bal of
- ([_], _) => None
- | (_, [_]) => None
+ ([_], _) => NONE
+ | (_, [_]) => NONE
| ts_us =>
let
val (tks, uks) = pairself count_terms ts_us;
val d = gcds (gcds (0, map snd tks), map snd uks);
in
- if d = 0 orelse d = 1 then None
- else Some
+ if d = 0 orelse d = 1 then NONE
+ else SOME
(Data.prove_conv (Data.multiply_tac d) Data.norm_tac sg
(t, Data.mk_bal (div_sum d tks, div_sum d uks)))
end));
--- a/src/Provers/Arith/cancel_numeral_factor.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/cancel_numeral_factor.ML Sun Feb 13 17:15:14 2005 +0100
@@ -80,8 +80,8 @@
[Data.trans_tac reshape, rtac Data.cancel 1,
Data.numeral_simp_tac] sg hyps xs (t', rhs))
end
- handle TERM _ => None
- | TYPE _ => None; (*Typically (if thy doesn't include Numeral)
+ handle TERM _ => NONE
+ | TYPE _ => NONE; (*Typically (if thy doesn't include Numeral)
Undeclared type constructor "Numeral.bin"*)
end;
--- a/src/Provers/Arith/cancel_numerals.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/cancel_numerals.ML Sun Feb 13 17:15:14 2005 +0100
@@ -105,8 +105,8 @@
(t', Data.mk_bal (Data.mk_sum T terms1',
newshape(n2-n1,terms2'))))
end
- handle TERM _ => None
- | TYPE _ => None; (*Typically (if thy doesn't include Numeral)
+ handle TERM _ => NONE
+ | TYPE _ => NONE; (*Typically (if thy doesn't include Numeral)
Undeclared type constructor "Numeral.bin"*)
end;
--- a/src/Provers/Arith/cancel_sums.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/cancel_sums.ML Sun Feb 13 17:15:14 2005 +0100
@@ -60,14 +60,14 @@
fun proc sg _ t =
(case try Data.dest_bal t of
- None => None
- | Some bal =>
+ NONE => NONE
+ | SOME bal =>
let
val (ts, us) = pairself (sort Term.term_ord o Data.dest_sum) bal;
val (ts', us', vs) = cancel ts us [];
in
- if null vs then None
- else Some
+ if null vs then NONE
+ else SOME
(Data.prove_conv (uncancel_sums_tac sg vs) Data.norm_tac sg
(t, Data.mk_bal (Data.mk_sum ts', Data.mk_sum us')))
end);
--- a/src/Provers/Arith/combine_numerals.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/combine_numerals.ML Sun Feb 13 17:15:14 2005 +0100
@@ -47,21 +47,21 @@
raise TERM("combine_numerals: remove", [])
| remove (m, u, t::terms) =
case try Data.dest_coeff t of
- Some(n,v) => if m=n andalso u aconv v then terms
+ SOME(n,v) => if m=n andalso u aconv v then terms
else t :: remove (m, u, terms)
- | None => t :: remove (m, u, terms);
+ | NONE => t :: remove (m, u, terms);
(*a left-to-right scan of terms, seeking another term of the form #n*u, where
#m*u is already in terms for some m*)
fun find_repeated (tab, _, []) = raise TERM("find_repeated", [])
| find_repeated (tab, past, t::terms) =
case try Data.dest_coeff t of
- Some(n,u) =>
+ SOME(n,u) =>
(case Termtab.lookup (tab, u) of
- Some m => (u, m, n, rev (remove (m,u,past)) @ terms)
- | None => find_repeated (Termtab.update ((u,n), tab),
+ SOME m => (u, m, n, rev (remove (m,u,past)) @ terms)
+ | NONE => find_repeated (Termtab.update ((u,n), tab),
t::past, terms))
- | None => find_repeated (tab, t::past, terms);
+ | NONE => find_repeated (tab, t::past, terms);
(*the simplification procedure*)
fun proc sg _ t =
@@ -84,8 +84,8 @@
Data.numeral_simp_tac] sg xs
(t', Data.mk_sum T (Data.mk_coeff(Data.add(m,n), u) :: terms)))
end
- handle TERM _ => None
- | TYPE _ => None; (*Typically (if thy doesn't include Numeral)
+ handle TERM _ => NONE
+ | TYPE _ => NONE; (*Typically (if thy doesn't include Numeral)
Undeclared type constructor "Numeral.bin"*)
end;
--- a/src/Provers/Arith/extract_common_term.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/extract_common_term.ML Sun Feb 13 17:15:14 2005 +0100
@@ -70,8 +70,8 @@
in
apsome Data.simplify_meta_eq reshape
end
- handle TERM _ => None
- | TYPE _ => None; (*Typically (if thy doesn't include Numeral)
+ handle TERM _ => NONE
+ | TYPE _ => NONE; (*Typically (if thy doesn't include Numeral)
Undeclared type constructor "Numeral.bin"*)
end;
--- a/src/Provers/Arith/fast_lin_arith.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/Arith/fast_lin_arith.ML Sun Feb 13 17:15:14 2005 +0100
@@ -340,9 +340,9 @@
flat(map (fn x => map (fn y => f x y) ys) xs);
fun extract_first p =
- let fun extract xs (y::ys) = if p y then (Some y,xs@ys)
+ let fun extract xs (y::ys) = if p y then (SOME y,xs@ys)
else extract (y::xs) ys
- | extract xs [] = (None,xs)
+ | extract xs [] = (NONE,xs)
in extract [] end;
fun print_ineqs ineqs =
@@ -361,8 +361,8 @@
val (triv,nontriv) = partition is_trivial ineqs in
if not(null triv)
then case Library.find_first is_answer triv of
- None => elim(nontriv,hist)
- | Some(Lineq(_,_,_,j)) => Success j
+ NONE => elim(nontriv,hist)
+ | SOME(Lineq(_,_,_,j)) => Success j
else
if null nontriv then Failure(hist)
else
@@ -372,7 +372,7 @@
val sclist = sort (fn (x,y) => int_ord(abs(x),abs(y)))
(filter (fn i => i<>0) clist)
val c = hd sclist
- val (Some(eq as Lineq(_,_,ceq,_)),othereqs) =
+ val (SOME(eq as Lineq(_,_,ceq,_)),othereqs) =
extract_first (fn Lineq(_,_,l,_) => c mem l) eqs
val v = find_index_eq c ceq
val (ioth,roth) = partition (fn (Lineq(_,_,l,_)) => el v l = 0)
@@ -424,41 +424,41 @@
map fst lhs union (map fst rhs union ats))
([], mapfilter (fn thm => if Thm.no_prems thm
then LA_Data.decomp sg (concl_of thm)
- else None) asms)
+ else NONE) asms)
fun add2 thm1 thm2 =
let val conj = thm1 RS (thm2 RS LA_Logic.conjI)
- in get_first (fn th => Some(conj RS th) handle THM _ => None) add_mono_thms
+ in get_first (fn th => SOME(conj RS th) handle THM _ => NONE) add_mono_thms
end;
- fun try_add [] _ = None
+ fun try_add [] _ = NONE
| try_add (thm1::thm1s) thm2 = case add2 thm1 thm2 of
- None => try_add thm1s thm2 | some => some;
+ NONE => try_add thm1s thm2 | some => some;
fun addthms thm1 thm2 =
case add2 thm1 thm2 of
- None => (case try_add ([thm1] RL inj_thms) thm2 of
- None => ( the(try_add ([thm2] RL inj_thms) thm1)
+ NONE => (case try_add ([thm1] RL inj_thms) thm2 of
+ NONE => ( the(try_add ([thm2] RL inj_thms) thm1)
handle OPTION =>
(trace_thm "" thm1; trace_thm "" thm2;
sys_error "Lin.arith. failed to add thms")
)
- | Some thm => thm)
- | Some thm => thm;
+ | SOME thm => thm)
+ | SOME thm => thm;
fun multn(n,thm) =
let fun mul(i,th) = if i=1 then th else mul(i-1, addthms thm th)
in if n < 0 then mul(~n,thm) RS LA_Logic.sym else mul(n,thm) end;
(*
fun multn2(n,thm) =
- let val Some(mth,cv) =
- get_first (fn (th,cv) => Some(thm RS th,cv) handle THM _ => None) mult_mono_thms
+ let val SOME(mth,cv) =
+ get_first (fn (th,cv) => SOME(thm RS th,cv) handle THM _ => NONE) mult_mono_thms
val ct = cterm_of sg (LA_Data.number_of(n,#T(rep_cterm cv)))
in instantiate ([],[(cv,ct)]) mth end
*)
fun multn2(n,thm) =
- let val Some(mth) =
- get_first (fn th => Some(thm RS th) handle THM _ => None) mult_mono_thms
+ let val SOME(mth) =
+ get_first (fn th => SOME(thm RS th) handle THM _ => NONE) mult_mono_thms
fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
val cv = cvar(mth, hd(prems_of mth));
val ct = cterm_of sg (LA_Data.number_of(n,#T(rep_cterm cv)))
@@ -494,7 +494,7 @@
end
end;
-fun coeff poly atom = case assoc(poly,atom) of None => 0 | Some i => i;
+fun coeff poly atom = case assoc(poly,atom) of NONE => 0 | SOME i => i;
fun lcms is = foldl lcm (1,is);
@@ -563,12 +563,12 @@
fun mknat pTs ixs (atom,i) =
if LA_Logic.is_nat(pTs,atom)
then let val l = map (fn j => if j=i then 1 else 0) ixs
- in Some(Lineq(0,Le,l,Nat(i))) end
- else None
+ in SOME(Lineq(0,Le,l,Nat(i))) end
+ else NONE
(* This code is tricky. It takes a list of premises in the order they occur
-in the subgoal. Numerical premises are coded as Some(tuple), non-numerical
-ones as None. Going through the premises, each numeric one is converted into
+in the subgoal. Numerical premises are coded as SOME(tuple), non-numerical
+ones as NONE. Going through the premises, each numeric one is converted into
a Lineq. The tricky bit is to convert ~= which is split into two cases < and
>. Thus split_items returns a list of equation systems. This may blow up if
there are many ~=, but in practice it does not seem to happen. The really
@@ -580,10 +580,10 @@
*)
fun split_items(items) =
let fun elim_neq front _ [] = [rev front]
- | elim_neq front n (None::ineqs) = elim_neq front (n+1) ineqs
- | elim_neq front n (Some(ineq as (l,i,rel,r,j,d))::ineqs) =
- if rel = "~=" then elim_neq front n (ineqs @ [Some(l,i,"<",r,j,d)]) @
- elim_neq front n (ineqs @ [Some(r,j,"<",l,i,d)])
+ | elim_neq front n (NONE::ineqs) = elim_neq front (n+1) ineqs
+ | elim_neq front n (SOME(ineq as (l,i,rel,r,j,d))::ineqs) =
+ if rel = "~=" then elim_neq front n (ineqs @ [SOME(l,i,"<",r,j,d)]) @
+ elim_neq front n (ineqs @ [SOME(r,j,"<",l,i,d)])
else elim_neq ((ineq,n) :: front) (n+1) ineqs
in elim_neq [] 0 items end;
@@ -611,9 +611,9 @@
| Failure(hist) =>
(if not ex then ()
else trace_ex(sg,params,atoms,discr initems,n,hist);
- None)
+ NONE)
end
- | refute [] js = Some js
+ | refute [] js = SOME js
in refute end;
fun refute sg ps ex items = refutes sg ps ex (split_items items) [];
@@ -637,15 +637,15 @@
fun prove sg ps ex Hs concl =
let val Hitems = map (LA_Data.decomp sg) Hs
-in if count (fn None => false | Some(_,_,r,_,_,_) => r = "~=") Hitems
- > !fast_arith_neq_limit then None
+in if count (fn NONE => false | SOME(_,_,r,_,_,_) => r = "~=") Hitems
+ > !fast_arith_neq_limit then NONE
else
case LA_Data.decomp sg concl of
- None => refute sg ps ex (Hitems@[None])
- | Some(citem as (r,i,rel,l,j,d)) =>
+ NONE => refute sg ps ex (Hitems@[NONE])
+ | SOME(citem as (r,i,rel,l,j,d)) =>
let val neg::rel0 = explode rel
val nrel = if neg = "~" then implode rel0 else "~"^rel
- in refute sg ps ex (Hitems @ [Some(r,i,nrel,l,j,d)]) end
+ in refute sg ps ex (Hitems @ [SOME(r,i,nrel,l,j,d)]) end
end;
(*
@@ -659,8 +659,8 @@
val concl = Logic.strip_assums_concl A
in trace_thm ("Trying to refute subgoal " ^ string_of_int i) st;
case prove (Thm.sign_of_thm st) (pTs,params) ex Hs concl of
- None => (trace_msg "Refutation failed."; no_tac)
- | Some js => (trace_msg "Refutation succeeded."; refute_tac(i,js))
+ NONE => (trace_msg "Refutation failed."; no_tac)
+ | SOME js => (trace_msg "Refutation succeeded."; refute_tac(i,js))
end) i st;
fun cut_lin_arith_tac thms i = cut_facts_tac thms i THEN lin_arith_tac false i;
@@ -702,7 +702,7 @@
val thm1' = implies_intr ct1 thm1
val thm2' = implies_intr ct2 thm2
in (thm2' COMP (thm1' COMP thm), js2) end;
-(* needs handle THM _ => None ? *)
+(* needs handle THM _ => NONE ? *)
fun prover sg thms Tconcl js pos =
let val nTconcl = LA_Logic.neg_prop Tconcl
@@ -711,9 +711,9 @@
val tree = splitasms (thms @ [nTconclthm])
val (thm,_) = fwdproof sg tree js
val contr = if pos then LA_Logic.ccontr else LA_Logic.notI
-in Some(LA_Logic.mk_Eq((implies_intr cnTconcl thm) COMP contr)) end
+in SOME(LA_Logic.mk_Eq((implies_intr cnTconcl thm) COMP contr)) end
(* in case concl contains ?-var, which makes assume fail: *)
-handle THM _ => None;
+handle THM _ => NONE;
(* PRE: concl is not negated!
This assumption is OK because
@@ -727,11 +727,11 @@
val Hs = map (#prop o rep_thm) thms
val Tconcl = LA_Logic.mk_Trueprop concl
in case prove sg ([],[]) false Hs Tconcl of (* concl provable? *)
- Some js => prover sg thms Tconcl js true
- | None => let val nTconcl = LA_Logic.neg_prop Tconcl
+ SOME js => prover sg thms Tconcl js true
+ | NONE => let val nTconcl = LA_Logic.neg_prop Tconcl
in case prove sg ([],[]) false Hs nTconcl of (* ~concl provable? *)
- Some js => prover sg thms nTconcl js false
- | None => None
+ SOME js => prover sg thms nTconcl js false
+ | NONE => NONE
end
end;
--- a/src/Provers/blast.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/blast.ML Sun Feb 13 17:15:14 2005 +0100
@@ -24,7 +24,7 @@
"Recursive" chains of rules can sometimes exclude other unsafe formulae
from expansion. This happens because newly-created formulae always
have priority over existing ones. But obviously recursive rules
- such as transitivity are treated specially to prevent this. Sometimes
+ such as transitivity are treated specially to prevent this. SOMEtimes
the formulae get into the wrong order (see WRONG below).
With substition for equalities (hyp_subst_tac):
@@ -174,10 +174,10 @@
| fromType alist (Term.TFree(a,_)) = Free a
| fromType alist (Term.TVar (ixn,_)) =
(case (assoc_string_int(!alist,ixn)) of
- None => let val t' = Var(ref None)
+ NONE => let val t' = Var(ref NONE)
in alist := (ixn, t') :: !alist; t'
end
- | Some v => v)
+ | SOME v => v)
(*Monomorphic constants used in blast, plus binary propositional connectives.*)
val standard_consts =
@@ -192,8 +192,8 @@
known to be monomorphic, which promotes efficiency. *)
fun fromConst alist (a,T) =
case Symtab.lookup(standard_const_table, a) of
- None => TConst(a, fromType alist T)
- | Some() => Const(a)
+ NONE => TConst(a, fromType alist T)
+ | SOME() => Const(a)
(*Tests whether 2 terms are alpha-convertible; chases instantiations*)
@@ -201,8 +201,8 @@
| (TConst (a,ta)) aconv (TConst (b,tb)) = a=b andalso ta aconv tb
| (Skolem (a,_)) aconv (Skolem (b,_)) = a=b (*arglists must then be equal*)
| (Free a) aconv (Free b) = a=b
- | (Var(ref(Some t))) aconv u = t aconv u
- | t aconv (Var(ref(Some u))) = t aconv u
+ | (Var(ref(SOME t))) aconv u = t aconv u
+ | t aconv (Var(ref(SOME u))) = t aconv u
| (Var v) aconv (Var w) = v=w (*both Vars are un-assigned*)
| (Bound i) aconv (Bound j) = i=j
| (Abs(_,t)) aconv (Abs(_,u)) = t aconv u
@@ -225,8 +225,8 @@
(*Accumulates the Vars in the term, suppressing duplicates*)
fun add_term_vars (Skolem(a,args), vars) = add_vars_vars(args,vars)
- | add_term_vars (Var (v as ref None), vars) = ins_var (v, vars)
- | add_term_vars (Var (ref (Some u)), vars) = add_term_vars(u,vars)
+ | add_term_vars (Var (v as ref NONE), vars) = ins_var (v, vars)
+ | add_term_vars (Var (ref (SOME u)), vars) = add_term_vars(u,vars)
| add_term_vars (TConst (_,t), vars) = add_term_vars(t,vars)
| add_term_vars (Abs (_,body), vars) = add_term_vars(body,vars)
| add_term_vars (f$t, vars) = add_term_vars (f, add_term_vars(t, vars))
@@ -236,9 +236,9 @@
| add_terms_vars (t::ts, vars) = add_terms_vars (ts, add_term_vars(t,vars))
(*Var list version.*)
and add_vars_vars ([], vars) = vars
- | add_vars_vars (ref (Some u) :: vs, vars) =
+ | add_vars_vars (ref (SOME u) :: vs, vars) =
add_vars_vars (vs, add_term_vars(u,vars))
- | add_vars_vars (v::vs, vars) = (*v must be a ref None*)
+ | add_vars_vars (v::vs, vars) = (*v must be a ref NONE*)
add_vars_vars (vs, ins_var (v, vars));
@@ -285,8 +285,8 @@
fun norm t = case t of
Skolem (a,args) => Skolem(a, vars_in_vars args)
| TConst(a,aT) => TConst(a, norm aT)
- | (Var (ref None)) => t
- | (Var (ref (Some u))) => norm u
+ | (Var (ref NONE)) => t
+ | (Var (ref (SOME u))) => norm u
| (f $ u) => (case norm f of
Abs(_,body) => norm (subst_bound (u, body))
| nf => nf $ norm u)
@@ -296,8 +296,8 @@
(*Weak (one-level) normalize for use in unification*)
fun wkNormAux t = case t of
(Var v) => (case !v of
- Some u => wkNorm u
- | None => t)
+ SOME u => wkNorm u
+ | NONE => t)
| (f $ u) => (case wkNormAux f of
Abs(_,body) => wkNorm (subst_bound (u, body))
| nf => nf $ u)
@@ -324,8 +324,8 @@
| occL lev (u::us) = occ lev u orelse occL lev us
and occ lev (Var w) =
v=w orelse
- (case !w of None => false
- | Some u => occ lev u)
+ (case !w of NONE => false
+ | SOME u => occ lev u)
| occ lev (Skolem(_,args)) = occL lev (map Var args)
(*ignore TConst, since term variables can't occur in types (?) *)
| occ lev (Bound i) = lev <= i
@@ -343,7 +343,7 @@
(*Restore the trail to some previous state: for backtracking*)
fun clearTo n =
while !ntrail<>n do
- (hd(!trail) := None;
+ (hd(!trail) := NONE;
trail := tl (!trail);
ntrail := !ntrail - 1);
@@ -357,11 +357,11 @@
fun update (t as Var v, u) =
if t aconv u then ()
else if varOccur v u then raise UNIFY
- else if mem_var(v, vars) then v := Some u
+ else if mem_var(v, vars) then v := SOME u
else (*avoid updating Vars in the branch if possible!*)
if is_Var u andalso mem_var(dest_Var u, vars)
- then dest_Var u := Some t
- else (v := Some u;
+ then dest_Var u := SOME t
+ else (v := SOME u;
trail := v :: !trail; ntrail := !ntrail + 1)
fun unifyAux (t,u) =
case (wkNorm t, wkNorm u) of
@@ -387,10 +387,10 @@
| from (Term.Bound i) = Bound i
| from (Term.Var (ixn,T)) =
(case (assoc_string_int(!alistVar,ixn)) of
- None => let val t' = Var(ref None)
+ NONE => let val t' = Var(ref NONE)
in alistVar := (ixn, t') :: !alistVar; t'
end
- | Some v => v)
+ | SOME v => v)
| from (Term.Abs (a,_,u)) =
(case from u of
u' as (f $ Bound 0) =>
@@ -406,13 +406,13 @@
let val name = ref "a"
val updated = ref []
fun inst (TConst(a,t)) = inst t
- | inst (Var(v as ref None)) = (updated := v :: (!updated);
- v := Some (Free ("?" ^ !name));
+ | inst (Var(v as ref NONE)) = (updated := v :: (!updated);
+ v := SOME (Free ("?" ^ !name));
name := Symbol.bump_string (!name))
| inst (Abs(a,t)) = inst t
| inst (f $ u) = (inst f; inst u)
| inst _ = ()
- fun revert() = seq (fn v => v:=None) (!updated)
+ fun revert() = seq (fn v => v:=NONE) (!updated)
in inst t; revert end;
@@ -436,9 +436,9 @@
If we don't find it then the premise is ill-formed and could cause
PROOF FAILED*)
fun delete_concl [] = raise ElimBadPrem
- | delete_concl (Const "*Goal*" $ (Var (ref (Some (Const"*False*")))) :: Ps) =
+ | delete_concl (Const "*Goal*" $ (Var (ref (SOME (Const"*False*")))) :: Ps) =
Ps
- | delete_concl (Const "Not" $ (Var (ref (Some (Const"*False*")))) :: Ps) =
+ | delete_concl (Const "Not" $ (Var (ref (SOME (Const"*False*")))) :: Ps) =
Ps
| delete_concl (P::Ps) = P :: delete_concl Ps;
@@ -453,7 +453,7 @@
fun convertRule vars t =
let val (P::Ps) = strip_imp_prems t
val Var v = strip_imp_concl t
- in v := Some (Const"*False*");
+ in v := SOME (Const"*False*");
(P, map (convertPrem o skoPrem vars) Ps)
end
handle Bind => raise ElimBadConcl;
@@ -595,12 +595,12 @@
| _ => dummyT);
(*Display top-level overloading if any*)
-fun topType (TConst(a,t)) = Some (showType t)
+fun topType (TConst(a,t)) = SOME (showType t)
| topType (Abs(a,t)) = topType t
| topType (f $ u) = (case topType f of
- None => topType u
+ NONE => topType u
| some => some)
- | topType _ = None;
+ | topType _ = NONE;
(*Convert from prototerms to ordinary terms with dummy types for tracing*)
@@ -609,8 +609,8 @@
| showTerm d (Skolem(a,_)) = Term.Const (a,dummyT)
| showTerm d (Free a) = Term.Free (a,dummyT)
| showTerm d (Bound i) = Term.Bound i
- | showTerm d (Var(ref(Some u))) = showTerm d u
- | showTerm d (Var(ref None)) = dummyVar2
+ | showTerm d (Var(ref(SOME u))) = showTerm d u
+ | showTerm d (Var(ref NONE)) = dummyVar2
| showTerm d (Abs(a,t)) = if d=0 then dummyVar
else Term.Abs(a, dummyT, showTerm (d-1) t)
| showTerm d (f $ u) = if d=0 then dummyVar
@@ -623,8 +623,8 @@
val stm = string_of sign 8 t'
in
case topType t' of
- None => stm (*no type to attach*)
- | Some T => stm ^ "\t:: " ^ Sign.string_of_typ sign T
+ NONE => stm (*no type to attach*)
+ | SOME T => stm ^ "\t:: " ^ Sign.string_of_typ sign T
end;
@@ -672,7 +672,7 @@
(*Replace the ATOMIC term "old" by "new" in t*)
fun subst_atomic (old,new) t =
- let fun subst (Var(ref(Some u))) = subst u
+ let fun subst (Var(ref(SOME u))) = subst u
| subst (Abs(a,body)) = Abs(a, subst body)
| subst (f$t) = subst f $ subst t
| subst t = if t aconv old then new else t
@@ -708,7 +708,7 @@
Skolem(_,vars) => vars
| _ => []
fun occEq u = (t aconv u) orelse occ u
- and occ (Var(ref(Some u))) = occEq u
+ and occ (Var(ref(SOME u))) = occEq u
| occ (Var v) = not (mem_var (v, vars))
| occ (Abs(_,u)) = occEq u
| occ (f$u) = occEq u orelse occEq f
@@ -785,14 +785,14 @@
(*Try to unify complementary literals and return the corresponding tactic. *)
fun tryClose (Const"*Goal*" $ G, L) =
- if unify([],G,L) then Some eAssume_tac else None
+ if unify([],G,L) then SOME eAssume_tac else NONE
| tryClose (G, Const"*Goal*" $ L) =
- if unify([],G,L) then Some eAssume_tac else None
+ if unify([],G,L) then SOME eAssume_tac else NONE
| tryClose (Const"Not" $ G, L) =
- if unify([],G,L) then Some eContr_tac else None
+ if unify([],G,L) then SOME eContr_tac else NONE
| tryClose (G, Const"Not" $ L) =
- if unify([],G,L) then Some eContr_tac else None
- | tryClose _ = None;
+ if unify([],G,L) then SOME eContr_tac else NONE
+ | tryClose _ = NONE;
(*Were there Skolem terms in the premise? Must NOT chase Vars*)
@@ -1006,8 +1006,8 @@
fun closeF [] = raise CLOSEF
| closeF (L::Ls) =
case tryClose(G,L) of
- None => closeF Ls
- | Some tac =>
+ NONE => closeF Ls
+ | SOME tac =>
let val choices' =
(if !trace then (immediate_output"branch closed";
traceVars sign ntrl)
@@ -1222,10 +1222,10 @@
| Term.Bound i => apply (Bound i)
| Term.Var (ix,_) =>
(case (assoc_string_int(!alistVar,ix)) of
- None => (alistVar := (ix, (ref None, bounds ts))
+ NONE => (alistVar := (ix, (ref NONE, bounds ts))
:: !alistVar;
Var (hdvar(!alistVar)))
- | Some(v,is) => if is=bounds ts then Var v
+ | SOME(v,is) => if is=bounds ts then Var v
else raise TRANS
("Discrepancy among occurrences of "
^ Syntax.string_of_vname ix))
@@ -1267,7 +1267,7 @@
let val start = startTiming()
in
case Seq.pull(EVERY' (rev tacs) i st) of
- None => (writeln ("PROOF FAILED for depth " ^
+ NONE => (writeln ("PROOF FAILED for depth " ^
Int.toString lim);
if !trace then writeln "************************\n"
else ();
@@ -1333,8 +1333,8 @@
Method.bang_sectioned_args'
Data.cla_modifiers (Scan.lift (Scan.option Args.nat)) m;
-fun blast_meth None = Data.cla_meth' blast_tac
- | blast_meth (Some lim) = Data.cla_meth' (fn cs => depth_tac cs lim);
+fun blast_meth NONE = Data.cla_meth' blast_tac
+ | blast_meth (SOME lim) = Data.cla_meth' (fn cs => depth_tac cs lim);
val setup = [Method.add_methods [("blast", blast_args blast_meth,
"tableau prover")]];
--- a/src/Provers/clasimp.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/clasimp.ML Sun Feb 13 17:15:14 2005 +0100
@@ -331,8 +331,8 @@
fun auto_args m =
Method.bang_sectioned_args' clasimp_modifiers (Scan.lift (Scan.option (Args.nat -- Args.nat))) m;
-fun auto_meth None = clasimp_meth (CHANGED_PROP o auto_tac)
- | auto_meth (Some (m, n)) = clasimp_meth (CHANGED_PROP o (fn css => mk_auto_tac css m n));
+fun auto_meth NONE = clasimp_meth (CHANGED_PROP o auto_tac)
+ | auto_meth (SOME (m, n)) = clasimp_meth (CHANGED_PROP o (fn css => mk_auto_tac css m n));
(* theory setup *)
--- a/src/Provers/classical.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/classical.ML Sun Feb 13 17:15:14 2005 +0100
@@ -224,7 +224,7 @@
(case try (fn () =>
rule_by_tactic (TRYALL (etac revcut_rl))
(th RSN (2, revcut_rl) |> assumption 2 |> Seq.hd)) () of
- Some th' => th'
+ SOME th' => th'
| _ => error ("Bad format for elimination rule\n" ^ string_of_thm th));
--- a/src/Provers/hypsubst.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/hypsubst.ML Sun Feb 13 17:15:14 2005 +0100
@@ -204,8 +204,8 @@
(r+1, imp_intr_tac i)
in
case Seq.pull(tac st) of
- None => Seq.single(st)
- | Some(st',_) => imptac (r',hyps) st'
+ NONE => Seq.single(st)
+ | SOME(st',_) => imptac (r',hyps) st'
end handle _ => error "?? in blast_hyp_subst_tac"
in imptac (0, rev hyps) end;
--- a/src/Provers/induct_method.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/induct_method.ML Sun Feb 13 17:15:14 2005 +0100
@@ -43,19 +43,19 @@
fun prep_inst align cert tune (tm, ts) =
let
- fun prep_var (x, Some t) =
+ fun prep_var (x, SOME t) =
let
val cx = cert x;
val {T = xT, sign, ...} = Thm.rep_cterm cx;
val ct = cert (tune t);
in
- if Sign.typ_instance sign (#T (Thm.rep_cterm ct), xT) then Some (cx, ct)
+ if Sign.typ_instance sign (#T (Thm.rep_cterm ct), xT) then SOME (cx, ct)
else raise ERROR_MESSAGE (Pretty.string_of (Pretty.block
[Pretty.str "Ill-typed instantiation:", Pretty.fbrk,
Display.pretty_cterm ct, Pretty.str " ::", Pretty.brk 1,
Display.pretty_ctyp (#T (Thm.crep_cterm ct))]))
end
- | prep_var (_, None) = None;
+ | prep_var (_, NONE) = NONE;
val xs = InductAttrib.vars_of tm;
in
align "Rule has fewer variables than instantiations given" xs ts
@@ -82,7 +82,7 @@
|> Seq.map (rpair (make (Thm.sign_of_thm rule, Thm.prop_of rule) cases)))
|> Seq.flat;
-fun find_casesT ctxt ((Some t :: _) :: _) = InductAttrib.find_casesT ctxt (fastype_of t)
+fun find_casesT ctxt ((SOME t :: _) :: _) = InductAttrib.find_casesT ctxt (fastype_of t)
| find_casesT _ _ = [];
fun find_casesS ctxt (fact :: _) = InductAttrib.find_casesS ctxt fact
@@ -101,16 +101,16 @@
val ruleq =
(case opt_rule of
- None =>
+ NONE =>
let val rules = find_casesS ctxt facts @ find_casesT ctxt insts @ [Data.cases_default] in
Method.trace ctxt rules;
Seq.flat (Seq.map (Seq.try inst_rule) (Seq.of_list rules))
end
- | Some r => Seq.single (inst_rule r));
+ | SOME r => Seq.single (inst_rule r));
fun prep_rule (th, (cases, n)) = Seq.map (apsnd (rpair (drop (n, facts))) o rpair cases)
(Method.multi_resolves (take (n, facts)) [th]);
- in resolveq_cases_tac (RuleCases.make is_open None) (Seq.flat (Seq.map prep_rule ruleq)) end;
+ in resolveq_cases_tac (RuleCases.make is_open NONE) (Seq.flat (Seq.map prep_rule ruleq)) end;
in
@@ -229,7 +229,7 @@
|> (Method.insert_tac more_facts THEN' atomize_tac) i
|> Seq.map (fn st' => divinate_inst (internalize k rule) i st' |> Seq.map (fn rule' =>
st' |> Tactic.rtac rule' i
- |> Seq.map (rpair (make is_open (Some (Thm.prop_of rule')) (rulified_term rule') cases)))
+ |> Seq.map (rpair (make is_open (SOME (Thm.prop_of rule')) (rulified_term rule') cases)))
|> Seq.flat)
|> Seq.flat)
|> Seq.flat;
@@ -245,7 +245,7 @@
(* rename all outermost !!-bound vars of type T in all premises of thm to x,
possibly indexed to avoid clashes *)
-fun rename [[Some(Free(x,Type(T,_)))]] thm =
+fun rename [[SOME(Free(x,Type(T,_)))]] thm =
let
fun index i [] = []
| index i (y::ys) = if x=y then x^string_of_int i :: index (i+1) ys
@@ -268,7 +268,7 @@
| rename _ thm = thm;
fun find_inductT ctxt insts =
- foldr multiply (insts |> mapfilter (fn [] => None | ts => last_elem ts)
+ foldr multiply (insts |> mapfilter (fn [] => NONE | ts => last_elem ts)
|> map (InductAttrib.find_inductT ctxt o fastype_of), [[]])
|> map join_rules |> flat |> map (rename insts);
@@ -284,7 +284,7 @@
val cert = Thm.cterm_of sg;
fun rule_versions r = Seq.cons (r, Seq.filter (not o curry Thm.eq_thm r)
- (Seq.make (fn () => Some (localize r, Seq.empty))))
+ (Seq.make (fn () => SOME (localize r, Seq.empty))))
|> Seq.map (rpair (RuleCases.get r));
val inst_rule = apfst (fn r =>
@@ -296,13 +296,13 @@
val ruleq =
(case opt_rule of
- None =>
+ NONE =>
let val rules = find_inductS ctxt facts @ find_inductT ctxt insts in
conditional (null rules) (fn () => error "Unable to figure out induct rule");
Method.trace ctxt rules;
rules |> Seq.THEN (Seq.of_list, Seq.THEN (rule_versions, Seq.try inst_rule))
end
- | Some r => r |> Seq.THEN (rule_versions, Seq.single o inst_rule));
+ | SOME r => r |> Seq.THEN (rule_versions, Seq.single o inst_rule));
fun prep_rule (th, (cases, n)) =
Seq.map (rpair (cases, n - length facts, drop (n, facts)))
@@ -327,8 +327,8 @@
local
fun check k get name =
- (case get name of Some x => x
- | None => error ("No rule for " ^ k ^ " " ^ quote name));
+ (case get name of SOME x => x
+ | NONE => error ("No rule for " ^ k ^ " " ^ quote name));
fun spec k = (Args.$$$ k -- Args.colon) |-- Args.!!! Args.name;
@@ -350,7 +350,7 @@
-- Args.colon;
val term = Scan.unless (Scan.lift kind_inst) Args.local_term;
val term_dummy = Scan.unless (Scan.lift kind_inst)
- (Scan.lift (Args.$$$ "_") >> K None || Args.local_term >> Some);
+ (Scan.lift (Args.$$$ "_") >> K NONE || Args.local_term >> SOME);
val instss = Args.and_list (Scan.repeat term_dummy);
--- a/src/Provers/order.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/order.ML Sun Feb 13 17:15:14 2005 +0100
@@ -126,7 +126,7 @@
fun mkasm_partial sign (t, n) =
case Less.decomp_part sign t of
- Some (x, rel, y) => (case rel of
+ SOME (x, rel, y) => (case rel of
"<" => if (x aconv y) then raise Contr (Thm ([Asm n], Less.less_reflE))
else [Less (x, y, Asm n)]
| "~<" => []
@@ -140,7 +140,7 @@
NotEq (y, x,Thm ( [Asm n], thm "not_sym"))]
| _ => error ("partial_tac: unknown relation symbol ``" ^ rel ^
"''returned by decomp_part."))
- | None => [];
+ | NONE => [];
(* ************************************************************************ *)
(* *)
@@ -154,7 +154,7 @@
fun mkasm_linear sign (t, n) =
case Less.decomp_lin sign t of
- Some (x, rel, y) => (case rel of
+ SOME (x, rel, y) => (case rel of
"<" => if (x aconv y) then raise Contr (Thm ([Asm n], Less.less_reflE))
else [Less (x, y, Asm n)]
| "~<" => [Le (y, x, Thm ([Asm n], Less.not_lessD))]
@@ -170,7 +170,7 @@
NotEq (y, x,Thm ( [Asm n], thm "not_sym"))]
| _ => error ("linear_tac: unknown relation symbol ``" ^ rel ^
"''returned by decomp_lin."))
- | None => [];
+ | NONE => [];
(* ************************************************************************ *)
(* *)
@@ -183,14 +183,14 @@
fun mkconcl_partial sign t =
case Less.decomp_part sign t of
- Some (x, rel, y) => (case rel of
+ SOME (x, rel, y) => (case rel of
"<" => ([Less (x, y, Asm ~1)], Asm 0)
| "<=" => ([Le (x, y, Asm ~1)], Asm 0)
| "=" => ([Le (x, y, Asm ~1), Le (y, x, Asm ~1)],
Thm ([Asm 0, Asm 1], Less.eqI))
| "~=" => ([NotEq (x,y, Asm ~1)], Asm 0)
| _ => raise Cannot)
- | None => raise Cannot;
+ | NONE => raise Cannot;
(* ************************************************************************ *)
(* *)
@@ -203,7 +203,7 @@
fun mkconcl_linear sign t =
case Less.decomp_lin sign t of
- Some (x, rel, y) => (case rel of
+ SOME (x, rel, y) => (case rel of
"<" => ([Less (x, y, Asm ~1)], Asm 0)
| "~<" => ([Le (y, x, Asm ~1)], Thm ([Asm 0], Less.not_lessI))
| "<=" => ([Le (x, y, Asm ~1)], Asm 0)
@@ -212,7 +212,7 @@
Thm ([Asm 0, Asm 1], Less.eqI))
| "~=" => ([NotEq (x,y, Asm ~1)], Asm 0)
| _ => raise Cannot)
- | None => raise Cannot;
+ | NONE => raise Cannot;
(* ******************************************************************* *)
(* *)
@@ -312,16 +312,16 @@
(* ******************************************************************* *)
(* *)
-(* triv_solv less1 : less -> proof Library.option *)
+(* triv_solv less1 : less -> proof option *)
(* *)
(* Solves trivial goal x <= x. *)
(* *)
(* ******************************************************************* *)
fun triv_solv (Le (x, x', _)) =
- if x aconv x' then Some (Thm ([], Less.le_refl))
- else None
-| triv_solv _ = None;
+ if x aconv x' then SOME (Thm ([], Less.le_refl))
+ else NONE
+| triv_solv _ = NONE;
(* ********************************************************************* *)
(* Graph functions *)
@@ -692,12 +692,12 @@
val xypath = (completeComponentPath x u xupred)@[e]@(completeComponentPath v y vypred)
val xyLesss = transPath (tl xypath, hd xypath)
in
- if xyLesss subsumes subgoal then Some (getprf xyLesss)
- else None
+ if xyLesss subsumes subgoal then SOME (getprf xyLesss)
+ else NONE
end)
- else None
+ else NONE
end)
- else None
+ else NONE
end
| _ =>
let val (uvfound, uvpred) = dfs (op =) sccGraph (getIndex u ntc) (getIndex v ntc)
@@ -717,16 +717,16 @@
val xypath = (completeComponentPath x u xupred)@[uvLesss]@(completeComponentPath v y vypred)
val xyLesss = transPath (tl xypath, hd xypath)
in
- if xyLesss subsumes subgoal then Some (getprf xyLesss)
- else None
+ if xyLesss subsumes subgoal then SOME (getprf xyLesss)
+ else NONE
end )
- else None
+ else NONE
end)
- else None
+ else NONE
end )
- else None
+ else NONE
end )
- ) else None
+ ) else NONE
end;
@@ -769,7 +769,7 @@
(* for all edges u ~= v or u < v check if they are part of path x < y *)
fun processNeqEdges [] = raise Cannot
| processNeqEdges (e::es) =
- case (findLess e x y xi yi xreachable yreachable) of (Some prf) => prf
+ case (findLess e x y xi yi xreachable yreachable) of (SOME prf) => prf
| _ => processNeqEdges es
in
@@ -783,10 +783,10 @@
(* if there aren't any edges that are candidate for <= then just search a edge in neqE that implies the subgoal *)
else if sccSubgraphs = [] then (
(case (Library.find_first (fn fact => fact subsumes subgoal) neqE, subgoal) of
- ( Some (NotEq (x, y, p)), NotEq (x', y', _)) =>
+ ( SOME (NotEq (x, y, p)), NotEq (x', y', _)) =>
if (x aconv x' andalso y aconv y') then p
else Thm ([p], thm "not_sym")
- | ( Some (Less (x, y, p)), NotEq (x', y', _)) =>
+ | ( SOME (Less (x, y, p)), NotEq (x', y', _)) =>
if x aconv x' andalso y aconv y' then (Thm ([p], Less.less_imp_neq))
else (Thm ([(Thm ([p], Less.less_imp_neq))], thm "not_sym"))
| _ => raise Cannot)
@@ -864,14 +864,14 @@
(* there exists a path x < y or y < x such that
x ~= y may be concluded *)
case (findLess e x y xi yi xreachable yreachable) of
- (Some prf) => (Thm ([prf], Less.less_imp_neq))
- | None => (
+ (SOME prf) => (Thm ([prf], Less.less_imp_neq))
+ | NONE => (
let
val yr = dfs_int_reachable sccGraph yi
val xr = dfs_int_reachable sccGraph_transpose xi
in
case (findLess e y x yi xi yr xr) of
- (Some prf) => (Thm ([(Thm ([prf], Less.less_imp_neq))], thm "not_sym"))
+ (SOME prf) => (Thm ([(Thm ([prf], Less.less_imp_neq))], thm "not_sym"))
| _ => processNeqEdges es
end)
) end)
@@ -949,12 +949,12 @@
val neqEdges = map snd (flat (map (adjacent (op aconv) neqG) comp));
(* find an edge leading to a contradiction *)
- fun findContr [] = None
+ fun findContr [] = NONE
| findContr (e::es) =
let val ui = (getIndex (lower e) ntc)
val vi = (getIndex (upper e) ntc)
in
- if ui = vi then Some e
+ if ui = vi then SOME e
else findContr es
end;
@@ -997,7 +997,7 @@
val subGraph = sccSubGraph intern [];
in
- case findContr neqEdges of Some e => handleContr e subGraph
+ case findContr neqEdges of SOME e => handleContr e subGraph
| _ => ((k, (reOrganizeEdges (extern) [])), subGraph)
end;
@@ -1025,8 +1025,8 @@
let
val (subgoals, prf) = mkconcl_partial sign concl
fun solve facts less =
- (case triv_solv less of None => findProof (sccGraph, neqE, ntc, sccSubgraphs ) less
- | Some prf => prf )
+ (case triv_solv less of NONE => findProof (sccGraph, neqE, ntc, sccSubgraphs ) less
+ | SOME prf => prf )
in
map (solve asms) subgoals
end
@@ -1051,8 +1051,8 @@
let
val (subgoals, prf) = mkconcl_linear sign concl
fun solve facts less =
- (case triv_solv less of None => findProof (sccGraph, neqE, ntc, sccSubgraphs) less
- | Some prf => prf )
+ (case triv_solv less of NONE => findProof (sccGraph, neqE, ntc, sccSubgraphs) less
+ | SOME prf => prf )
in
map (solve asms) subgoals
end
--- a/src/Provers/quantifier1.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/quantifier1.ML Sun Feb 13 17:15:14 2005 +0100
@@ -72,34 +72,34 @@
fun def xs eq =
let val n = length xs
in case dest_eq eq of
- Some(c,s,t) =>
+ SOME(c,s,t) =>
s = Bound n andalso not(loose_bvar1(t,n)) orelse
t = Bound n andalso not(loose_bvar1(s,n))
- | None => false
+ | NONE => false
end;
-fun extract_conj xs t = case dest_conj t of None => None
- | Some(conj,P,Q) =>
- (if def xs P then Some(xs,P,Q) else
- if def xs Q then Some(xs,Q,P) else
+fun extract_conj xs t = case dest_conj t of NONE => NONE
+ | SOME(conj,P,Q) =>
+ (if def xs P then SOME(xs,P,Q) else
+ if def xs Q then SOME(xs,Q,P) else
(case extract_conj xs P of
- Some(xs,eq,P') => Some(xs,eq, conj $ P' $ Q)
- | None => (case extract_conj xs Q of
- Some(xs,eq,Q') => Some(xs,eq,conj $ P $ Q')
- | None => None)));
+ SOME(xs,eq,P') => SOME(xs,eq, conj $ P' $ Q)
+ | NONE => (case extract_conj xs Q of
+ SOME(xs,eq,Q') => SOME(xs,eq,conj $ P $ Q')
+ | NONE => NONE)));
-fun extract_imp xs t = case dest_imp t of None => None
- | Some(imp,P,Q) => if def xs P then Some(xs,P,Q)
+fun extract_imp xs t = case dest_imp t of NONE => NONE
+ | SOME(imp,P,Q) => if def xs P then SOME(xs,P,Q)
else (case extract_conj xs P of
- Some(xs,eq,P') => Some(xs, eq, imp $ P' $ Q)
- | None => (case extract_imp xs Q of
- None => None
- | Some(xs,eq,Q') =>
- Some(xs,eq,imp$P$Q')));
+ SOME(xs,eq,P') => SOME(xs, eq, imp $ P' $ Q)
+ | NONE => (case extract_imp xs Q of
+ NONE => NONE
+ | SOME(xs,eq,Q') =>
+ SOME(xs,eq,imp$P$Q')));
fun extract_quant extract q =
let fun exqu xs ((qC as Const(qa,_)) $ Abs(x,T,Q)) =
- if qa = q then exqu ((qC,x,T)::xs) Q else None
+ if qa = q then exqu ((qC,x,T)::xs) Q else NONE
| exqu xs P = extract xs P
in exqu end;
@@ -147,33 +147,33 @@
fun rearrange_all sg _ (F as (all as Const(q,_)) $ Abs(x,T, P)) =
(case extract_quant extract_imp q [] P of
- None => None
- | Some(xs,eq,Q) =>
+ NONE => NONE
+ | SOME(xs,eq,Q) =>
let val R = quantify all x T xs (imp $ eq $ Q)
- in Some(prove_conv prove_one_point_all_tac sg (F,R)) end)
- | rearrange_all _ _ _ = None;
+ in SOME(prove_conv prove_one_point_all_tac sg (F,R)) end)
+ | rearrange_all _ _ _ = NONE;
fun rearrange_ball tac sg _ (F as Ball $ A $ Abs(x,T,P)) =
(case extract_imp [] P of
- None => None
- | Some(xs,eq,Q) => if not(null xs) then None else
+ NONE => NONE
+ | SOME(xs,eq,Q) => if not(null xs) then NONE else
let val R = imp $ eq $ Q
- in Some(prove_conv tac sg (F,Ball $ A $ Abs(x,T,R))) end)
- | rearrange_ball _ _ _ _ = None;
+ in SOME(prove_conv tac sg (F,Ball $ A $ Abs(x,T,R))) end)
+ | rearrange_ball _ _ _ _ = NONE;
fun rearrange_ex sg _ (F as (ex as Const(q,_)) $ Abs(x,T,P)) =
(case extract_quant extract_conj q [] P of
- None => None
- | Some(xs,eq,Q) =>
+ NONE => NONE
+ | SOME(xs,eq,Q) =>
let val R = quantify ex x T xs (conj $ eq $ Q)
- in Some(prove_conv prove_one_point_ex_tac sg (F,R)) end)
- | rearrange_ex _ _ _ = None;
+ in SOME(prove_conv prove_one_point_ex_tac sg (F,R)) end)
+ | rearrange_ex _ _ _ = NONE;
fun rearrange_bex tac sg _ (F as Bex $ A $ Abs(x,T,P)) =
(case extract_conj [] P of
- None => None
- | Some(xs,eq,Q) => if not(null xs) then None else
- Some(prove_conv tac sg (F,Bex $ A $ Abs(x,T,conj$eq$Q))))
- | rearrange_bex _ _ _ _ = None;
+ NONE => NONE
+ | SOME(xs,eq,Q) => if not(null xs) then NONE else
+ SOME(prove_conv tac sg (F,Bex $ A $ Abs(x,T,conj$eq$Q))))
+ | rearrange_bex _ _ _ _ = NONE;
end;
--- a/src/Provers/quasi.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/quasi.ML Sun Feb 13 17:15:14 2005 +0100
@@ -57,7 +57,7 @@
val less_imp_neq : thm (* x < y ==> x ~= y *)
(* Analysis of premises and conclusion *)
- (* decomp_x (`x Rel y') should yield Some (x, Rel, y)
+ (* decomp_x (`x Rel y') should yield SOME (x, Rel, y)
where Rel is one of "<", "<=", "=" and "~=",
other relation symbols cause an error message *)
(* decomp_trans is used by trans_tac, it may only return Rel = "<=" *)
@@ -127,12 +127,12 @@
fun mkasm_trans sign (t, n) =
case Less.decomp_trans sign t of
- Some (x, rel, y) =>
+ SOME (x, rel, y) =>
(case rel of
"<=" => [Le (x, y, Asm n)]
| _ => error ("trans_tac: unknown relation symbol ``" ^ rel ^
"''returned by decomp_trans."))
- | None => [];
+ | NONE => [];
(* ************************************************************************ *)
(* *)
@@ -146,7 +146,7 @@
fun mkasm_quasi sign (t, n) =
case Less.decomp_quasi sign t of
- Some (x, rel, y) => (case rel of
+ SOME (x, rel, y) => (case rel of
"<" => if (x aconv y) then raise Contr (Thm ([Asm n], Less.less_reflE))
else [Less (x, y, Asm n)]
| "<=" => [Le (x, y, Asm n)]
@@ -158,7 +158,7 @@
NotEq (y, x,Thm ( [Asm n], thm "not_sym"))]
| _ => error ("quasi_tac: unknown relation symbol ``" ^ rel ^
"''returned by decomp_quasi."))
- | None => [];
+ | NONE => [];
(* ************************************************************************ *)
@@ -173,10 +173,10 @@
fun mkconcl_trans sign t =
case Less.decomp_trans sign t of
- Some (x, rel, y) => (case rel of
+ SOME (x, rel, y) => (case rel of
"<=" => (Le (x, y, Asm ~1), Asm 0)
| _ => raise Cannot)
- | None => raise Cannot;
+ | NONE => raise Cannot;
(* ************************************************************************ *)
@@ -190,12 +190,12 @@
fun mkconcl_quasi sign t =
case Less.decomp_quasi sign t of
- Some (x, rel, y) => (case rel of
+ SOME (x, rel, y) => (case rel of
"<" => ([Less (x, y, Asm ~1)], Asm 0)
| "<=" => ([Le (x, y, Asm ~1)], Asm 0)
| "~=" => ([NotEq (x,y, Asm ~1)], Asm 0)
| _ => raise Cannot)
-| None => raise Cannot;
+| NONE => raise Cannot;
(* ******************************************************************* *)
@@ -266,16 +266,16 @@
(* ******************************************************************* *)
(* *)
-(* triv_solv less1 : less -> proof Library.option *)
+(* triv_solv less1 : less -> proof option *)
(* *)
(* Solves trivial goal x <= x. *)
(* *)
(* ******************************************************************* *)
fun triv_solv (Le (x, x', _)) =
- if x aconv x' then Some (Thm ([], Less.le_refl))
- else None
-| triv_solv _ = None;
+ if x aconv x' then SOME (Thm ([], Less.le_refl))
+ else NONE
+| triv_solv _ = NONE;
(* ********************************************************************* *)
(* Graph functions *)
@@ -463,15 +463,15 @@
end )
| (Less (x,y,_)) => (
let
- fun findParallelNeq [] = None
+ fun findParallelNeq [] = NONE
| findParallelNeq (e::es) =
- if (x aconv (lower e) andalso y aconv (upper e)) then Some e
- else if (y aconv (lower e) andalso x aconv (upper e)) then Some (NotEq (x,y, (Thm ([getprf e], thm "not_sym"))))
+ if (x aconv (lower e) andalso y aconv (upper e)) then SOME e
+ else if (y aconv (lower e) andalso x aconv (upper e)) then SOME (NotEq (x,y, (Thm ([getprf e], thm "not_sym"))))
else findParallelNeq es ;
in
(* test if there is a edge x ~= y respectivly y ~= x and *)
(* if it possible to find a path x <= y in leG, thus we can conclude x < y *)
- (case findParallelNeq neqE of (Some e) =>
+ (case findParallelNeq neqE of (SOME e) =>
let
val (xyLeFound,xyLePath) = findPath x y leG
in
@@ -487,7 +487,7 @@
| (NotEq (x,y,_)) =>
(* First check if a single premiss is sufficient *)
(case (Library.find_first (fn fact => fact subsumes subgoal) neqE, subgoal) of
- (Some (NotEq (x, y, p)), NotEq (x', y', _)) =>
+ (SOME (NotEq (x, y, p)), NotEq (x', y', _)) =>
if (x aconv x' andalso y aconv y') then p
else Thm ([p], thm "not_sym")
| _ => raise Cannot
@@ -540,8 +540,8 @@
let
val (subgoals, prf) = mkconcl_quasi sign concl
fun solve facts less =
- (case triv_solv less of None => findQuasiProof (leG, neqE) less
- | Some prf => prf )
+ (case triv_solv less of NONE => findQuasiProof (leG, neqE) less
+ | SOME prf => prf )
in map (solve asms) subgoals end
end;
--- a/src/Provers/simp.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/simp.ML Sun Feb 13 17:15:14 2005 +0100
@@ -130,8 +130,8 @@
(*Applies tactic and returns the first resulting state, FAILS if none!*)
fun one_result(tac,thm) = case Seq.pull(tac thm) of
- Some(thm',_) => thm'
- | None => raise THM("Simplifier: could not continue", 0, [thm]);
+ SOME(thm',_) => thm'
+ | NONE => raise THM("Simplifier: could not continue", 0, [thm]);
fun res1(thm,thms,i) = one_result(resolve_tac thms i,thm);
@@ -174,8 +174,8 @@
Abs (_,_,body) => add_vars(body,vars)
| r$s => (case head_of tm of
Const(c,T) => (case assoc(new_asms,c) of
- None => add_vars(r,add_vars(s,vars))
- | Some(al) => add_list(tm,al,vars))
+ NONE => add_vars(r,add_vars(s,vars))
+ | SOME(al) => add_list(tm,al,vars))
| _ => add_vars(r,add_vars(s,vars)))
| _ => vars
in add_vars end;
@@ -205,7 +205,7 @@
| Free _ => resolve_tac congs 1 ORELSE refl1_tac
| _ => refl1_tac))
val add_norm_tac = DEPTH_FIRST (has_fewer_prems nops) norm_step_tac
- val Some(thm'',_) = Seq.pull(add_norm_tac thm')
+ val SOME(thm'',_) = Seq.pull(add_norm_tac thm')
in thm'' end;
fun add_norm_tags congs =
@@ -444,14 +444,14 @@
if var_lhs(thm,i) then (ss,thm,anet,ats,cs) else
if lhs_is_NORM(thm,i) then (ss, res1(thm,trans_norms,i), anet,ats,cs)
else case Seq.pull(cong_tac i thm) of
- Some(thm',_) =>
+ SOME(thm',_) =>
let val ps = prems_of thm and ps' = prems_of thm';
val n = length(ps')-length(ps);
val a = length(strip_assums_hyp(nth_elem(i-1,ps)))
val l = map (fn p => length(strip_assums_hyp(p)))
(take(n,drop(i-1,ps')));
in (simp_refl(rev(l),a,REW::ss),thm',anet,ats,cs) end
- | None => (REW::ss,thm,anet,ats,cs);
+ | NONE => (REW::ss,thm,anet,ats,cs);
(*NB: the "Adding rewrites:" trace will look strange because assumptions
are represented by rules, generalized over their parameters*)
@@ -468,22 +468,22 @@
end;
fun rew(seq,thm,ss,anet,ats,cs, more) = case Seq.pull seq of
- Some(thm',seq') =>
+ SOME(thm',seq') =>
let val n = (nprems_of thm') - (nprems_of thm)
in pr_rew(i,n,thm,thm',more);
if n=0 then (SIMP_LHS::ss, thm', anet, ats, cs)
else ((replicate n PROVE) @ (POP_CS::SIMP_LHS::ss),
thm', anet, ats, (ss,thm,anet,ats,seq',more)::cs)
end
- | None => if more
+ | NONE => if more
then rew(tapply(lhs_net_tac anet i THEN assume_tac i,thm),
thm,ss,anet,ats,cs,false)
else (ss,thm,anet,ats,cs);
fun try_true(thm,ss,anet,ats,cs) =
case Seq.pull(auto_tac i thm) of
- Some(thm',_) => (ss,thm',anet,ats,cs)
- | None => let val (ss0,thm0,anet0,ats0,seq,more)::cs0 = cs
+ SOME(thm',_) => (ss,thm',anet,ats,cs)
+ | NONE => let val (ss0,thm0,anet0,ats0,seq,more)::cs0 = cs
in if !tracing
then (writeln"*** Failed to prove precondition. Normal form:";
pr_goal_concl i thm; writeln"")
@@ -494,8 +494,8 @@
fun split(thm,ss,anet,ats,cs) =
case Seq.pull(tapply(split_tac
(cong_tac i,splits,split_consts) i,thm)) of
- Some(thm',_) => (SIMP_LHS::SPLIT::ss,thm',anet,ats,cs)
- | None => (ss,thm,anet,ats,cs);
+ SOME(thm',_) => (SIMP_LHS::SPLIT::ss,thm',anet,ats,cs)
+ | NONE => (ss,thm,anet,ats,cs);
fun step(s::ss, thm, anet, ats, cs) = case s of
MK_EQ => (ss, res1(thm,[red2],i), anet, ats, cs)
@@ -575,10 +575,10 @@
let val (Const(_,cT), va, vb) = dest_red(hd(prems_of thm));
val [P] = term_vars(concl_of thm) \\ [va,vb]
val eqT::_ = binder_types cT
- in if Type.typ_instance tsig (T,eqT) then Some(thm,va,vb,P)
+ in if Type.typ_instance tsig (T,eqT) then SOME(thm,va,vb,P)
else find thms
end
- | find [] = None
+ | find [] = NONE
in find subst_thms end;
fun mk_cong sg (f,aTs,rT) (refl,eq) =
@@ -595,8 +595,8 @@
fun mk(c,T::Ts,i,yik) =
let val si = radixstring(26,"a",i)
in case find_subst tsig T of
- None => mk(c,Ts,i-1,eta_Var(("X"^si,0),T)::yik)
- | Some s => let val c' = c RSN (2,ri(s,i,si,T,yik))
+ NONE => mk(c,Ts,i-1,eta_Var(("X"^si,0),T)::yik)
+ | SOME s => let val c' = c RSN (2,ri(s,i,si,T,yik))
in mk(c',Ts,i-1,eta_Var(("Y"^si,0),T)::yik) end
end
| mk(c,[],_,_) = c;
@@ -608,17 +608,17 @@
fun find_refl(r::rs) =
let val (Const(eq,eqT),_,_) = dest_red(concl_of r)
in if Type.typ_instance tsig (rT, hd(binder_types eqT))
- then Some(r,(eq,body_type eqT)) else find_refl rs
+ then SOME(r,(eq,body_type eqT)) else find_refl rs
end
- | find_refl([]) = None;
+ | find_refl([]) = NONE;
in case find_refl refl_thms of
- None => [] | Some(refl) => [mk_cong sg (f,aTs,rT) refl]
+ NONE => [] | SOME(refl) => [mk_cong sg (f,aTs,rT) refl]
end;
fun mk_cong_thy thy f =
let val sg = sign_of thy;
val T = case Sign.const_type sg f of
- None => error(f^" not declared") | Some(T) => T;
+ NONE => error(f^" not declared") | SOME(T) => T;
val T' = incr_tvar 9 T;
in mk_cong_type sg (Const(f,T'),T') end;
@@ -628,9 +628,9 @@
let val sg = sign_of thy;
val S0 = Sign.defaultS sg;
fun readfT(f,s) =
- let val T = incr_tvar 9 (Sign.read_typ(sg,K(Some(S0))) s);
+ let val T = incr_tvar 9 (Sign.read_typ(sg,K(SOME(S0))) s);
val t = case Sign.const_type sg f of
- Some(_) => Const(f,T) | None => Free(f,T)
+ SOME(_) => Const(f,T) | NONE => Free(f,T)
in (t,T) end
in flat o map (mk_cong_type sg o readfT) end;
--- a/src/Provers/simplifier.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/simplifier.ML Sun Feb 13 17:15:14 2005 +0100
@@ -158,7 +158,7 @@
fun add_solver add (ContextSolver ((name, f), _)) simpset =
add (simpset, mk_solver name (f ctxt));
in
- ((case subgoal_tac of None => ss | Some tac => ss setsubgoaler tac ctxt)
+ ((case subgoal_tac of NONE => ss | SOME tac => ss setsubgoaler tac ctxt)
addsimprocs map prep_simproc simprocs)
|> fold_rev add_loop loop_tacs
|> fold_rev (add_solver (op addSolver)) unsafe_solvers
@@ -169,7 +169,7 @@
ContextSS {simprocs = simprocs, subgoal_tac = subgoal_tac, loop_tacs = loop_tacs,
unsafe_solvers = unsafe_solvers, solvers = solvers};
-val empty_context_ss = make_context_ss ([], None, [], [], []);
+val empty_context_ss = make_context_ss ([], NONE, [], [], []);
fun merge_context_ss (ctxt_ss1, ctxt_ss2) =
let
@@ -179,7 +179,7 @@
unsafe_solvers = unsafe_solvers2, solvers = solvers2} = ctxt_ss2;
val simprocs' = merge_context_simprocs simprocs1 simprocs2;
- val subgoal_tac' = (case subgoal_tac1 of None => subgoal_tac2 | some => some);
+ val subgoal_tac' = (case subgoal_tac1 of NONE => subgoal_tac2 | some => some);
val loop_tacs' = merge_alists loop_tacs1 loop_tacs2;
val unsafe_solvers' = merge_context_solvers unsafe_solvers1 unsafe_solvers2;
val solvers' = merge_context_solvers solvers1 solvers2;
@@ -272,11 +272,11 @@
fun set_context_subgoaler tac =
map_context_ss (fn (simprocs, _, loop_tacs, unsafe_solvers, solvers) =>
- (simprocs, Some tac, loop_tacs, unsafe_solvers, solvers));
+ (simprocs, SOME tac, loop_tacs, unsafe_solvers, solvers));
val reset_context_subgoaler =
map_context_ss (fn (simprocs, _, loop_tacs, unsafe_solvers, solvers) =>
- (simprocs, None, loop_tacs, unsafe_solvers, solvers));
+ (simprocs, NONE, loop_tacs, unsafe_solvers, solvers));
fun add_context_looper (name, tac) =
map_context_ss (fn (simprocs, subgoal_tac, loop_tacs, unsafe_solvers, solvers) =>
--- a/src/Provers/splitter.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/splitter.ML Sun Feb 13 17:15:14 2005 +0100
@@ -332,8 +332,8 @@
(case strip_comb lhs of (Const(a,aT),args) =>
let val info = (aT,lhs,thm,fastype_of t,length args)
in case assoc(cmap,a) of
- Some infos => overwrite(cmap,(a,info::infos))
- | None => (a,[info])::cmap
+ SOME infos => overwrite(cmap,(a,info::infos))
+ | NONE => (a,[info])::cmap
end
| _ => split_format_err())
| _ => split_format_err())
--- a/src/Provers/trancl.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Provers/trancl.ML Sun Feb 13 17:15:14 2005 +0100
@@ -57,8 +57,8 @@
Returns one of the following:
- None if not an instance of a relation,
- Some (x, y, r, s) if instance of a relation, where
+ NONE if not an instance of a relation,
+ SOME (x, y, r, s) if instance of a relation, where
x: left hand side argument, y: right hand side argument,
r: the relation,
s: the kind of closure, one of
@@ -121,7 +121,7 @@
fun mkasm_trancl Rel (t, n) =
case Cls.decomp t of
- Some (x, y, rel,r) => if rel aconv Rel then
+ SOME (x, y, rel,r) => if rel aconv Rel then
(case r of
"r" => [Trans (x,y, Thm([Asm n], Cls.r_into_trancl))]
@@ -129,7 +129,7 @@
| "r*" => []
| _ => error ("trancl_tac: unknown relation symbol"))
else []
- | None => [];
+ | NONE => [];
(* ************************************************************************ *)
(* *)
@@ -144,14 +144,14 @@
fun mkasm_rtrancl Rel (t, n) =
case Cls.decomp t of
- Some (x, y, rel, r) => if rel aconv Rel then
+ SOME (x, y, rel, r) => if rel aconv Rel then
(case r of
"r" => [ Trans (x,y, Thm([Asm n], Cls.r_into_trancl))]
| "r+" => [ Trans (x,y, Asm n)]
| "r*" => [ RTrans(x,y, Asm n)]
| _ => error ("rtrancl_tac: unknown relation symbol" ))
else []
- | None => [];
+ | NONE => [];
(* ************************************************************************ *)
(* *)
@@ -168,18 +168,18 @@
fun mkconcl_trancl t =
case Cls.decomp t of
- Some (x, y, rel, r) => (case r of
+ SOME (x, y, rel, r) => (case r of
"r+" => (rel, Trans (x,y, Asm ~1), Asm 0)
| _ => raise Cannot)
- | None => raise Cannot;
+ | NONE => raise Cannot;
fun mkconcl_rtrancl t =
case Cls.decomp t of
- Some (x, y, rel,r ) => (case r of
+ SOME (x, y, rel,r ) => (case r of
"r+" => (rel, Trans (x,y, Asm ~1), Asm 0)
| "r*" => (rel, RTrans (x,y, Asm ~1), Asm 0)
| _ => raise Cannot)
- | None => raise Cannot;
+ | NONE => raise Cannot;
(* ************************************************************************ *)
(* *)
--- a/src/Pure/General/file.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/file.ML Sun Feb 13 17:15:14 2005 +0100
@@ -92,8 +92,8 @@
fun info path =
let val name = sysify_path path in
(case file_info name of
- "" => None
- | s => Some (Info s))
+ "" => NONE
+ | s => SOME (Info s))
end;
val exists = is_some o info;
--- a/src/Pure/General/graph.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/graph.ML Sun Feb 13 17:15:14 2005 +0100
@@ -92,8 +92,8 @@
fun get_entry (Graph tab) x =
(case Table.lookup (tab, x) of
- Some entry => entry
- | None => raise UNDEF x);
+ SOME entry => entry
+ | NONE => raise UNDEF x);
fun map_entry x f (G as Graph tab) = Graph (Table.update ((x, f (get_entry G x)), tab));
@@ -154,8 +154,8 @@
fun del_nodes xs (Graph tab) =
let
fun del (x, (i, (preds, succs))) =
- if x mem_key xs then None
- else Some (x, (i, (foldl op del_key (preds, xs), foldl op del_key (succs, xs))));
+ if x mem_key xs then NONE
+ else SOME (x, (i, (foldl op del_key (preds, xs), foldl op del_key (succs, xs))));
in Graph (Table.make (mapfilter del (Table.dest tab))) end;
@@ -177,7 +177,7 @@
fun diff_edges G1 G2 =
flat (dest G1 |> map (fn (x, ys) => ys |> mapfilter (fn y =>
- if is_edge G2 (x, y) then None else Some (x, y))));
+ if is_edge G2 (x, y) then NONE else SOME (x, y))));
fun edges G = diff_edges G empty;
--- a/src/Pure/General/history.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/history.ML Sun Feb 13 17:15:14 2005 +0100
@@ -45,9 +45,9 @@
| redo (History (x, (lim, len, undo_list, r :: redo_list))) =
History (r, (lim, len + 1, x :: undo_list, redo_list));
-fun push None _ x xs = x :: xs
- | push (Some 0) _ _ _ = []
- | push (Some n) len x xs = if len < n then x :: xs else take (n, x :: xs);
+fun push NONE _ x xs = x :: xs
+ | push (SOME 0) _ _ _ = []
+ | push (SOME n) len x xs = if len < n then x :: xs else take (n, x :: xs);
fun apply_copy cp f (History (x, (lim, len, undo_list, _))) =
let val x' = cp x
--- a/src/Pure/General/lazy_scan.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/lazy_scan.ML Sun Feb 13 17:15:14 2005 +0100
@@ -97,8 +97,8 @@
fun one p toks =
case LazySeq.getItem toks of
- None => raise SyntaxError
- | Some(t,toks) => if p t
+ NONE => raise SyntaxError
+ | SOME(t,toks) => if p t
then (t,toks)
else raise SyntaxError
@@ -106,8 +106,8 @@
fun any p toks =
case LazySeq.getItem toks of
- None => ([],toks)
- | Some(x,toks2) => if p x
+ NONE => ([],toks)
+ | SOME(x,toks2) => if p x
then
let
val (xs,toks3) = any p toks2
@@ -125,7 +125,7 @@
end
fun optional sc def = sc || succeed def
-fun option sc = (sc >> Some) || succeed None
+fun option sc = (sc >> SOME) || succeed NONE
(*
fun repeat sc =
@@ -152,9 +152,9 @@
fun repeat sc =
let
fun R xs toks =
- case Some (sc toks) handle SyntaxError => None of
- Some (x,toks2) => R (x::xs) toks2
- | None => (List.rev xs,toks)
+ case SOME (sc toks) handle SyntaxError => NONE of
+ SOME (x,toks2) => R (x::xs) toks2
+ | NONE => (List.rev xs,toks)
in
R []
end
@@ -187,12 +187,12 @@
let
fun F toks =
if LazySeq.null toks
- then None
+ then NONE
else
let
val (x,toks') = sc toks
in
- Some(x,LazySeq.make (fn () => F toks'))
+ SOME(x,LazySeq.make (fn () => F toks'))
end
in
LazySeq.make (fn () => F toks)
--- a/src/Pure/General/lazy_seq.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/lazy_seq.ML Sun Feb 13 17:15:14 2005 +0100
@@ -86,42 +86,42 @@
fun null s = is_some (getItem s)
local
- fun F n None = n
- | F n (Some(_,s)) = F (n+1) (getItem s)
+ fun F n NONE = n
+ | F n (SOME(_,s)) = F (n+1) (getItem s)
in
fun length s = F 0 (getItem s)
end
fun s1 @ s2 =
let
- fun F None = getItem s2
- | F (Some(x,s1')) = Some(x,F' s1')
+ fun F NONE = getItem s2
+ | F (SOME(x,s1')) = SOME(x,F' s1')
and F' s = make (fn () => F (getItem s))
in
F' s1
end
local
- fun F None = raise Empty
- | F (Some arg) = arg
+ fun F NONE = raise Empty
+ | F (SOME arg) = arg
in
fun hd s = #1 (F (getItem s))
fun tl s = #2 (F (getItem s))
end
local
- fun F x None = x
- | F _ (Some(x,s)) = F x (getItem s)
- fun G None = raise Empty
- | G (Some(x,s)) = F x (getItem s)
+ fun F x NONE = x
+ | F _ (SOME(x,s)) = F x (getItem s)
+ fun G NONE = raise Empty
+ | G (SOME(x,s)) = F x (getItem s)
in
fun last s = G (getItem s)
end
local
- fun F None _ = raise Subscript
- | F (Some(x,_)) 0 = x
- | F (Some(_,s)) n = F (getItem s) (n-1)
+ fun F NONE _ = raise Subscript
+ | F (SOME(x,_)) 0 = x
+ | F (SOME(_,s)) n = F (getItem s) (n-1)
in
fun nth(s,n) =
if n < 0
@@ -130,9 +130,9 @@
end
local
- fun F None _ = raise Subscript
- | F (Some(x,s)) n = Some(x,F' s (n-1))
- and F' s 0 = Seq (value None)
+ fun F NONE _ = raise Subscript
+ | F (SOME(x,s)) n = SOME(x,F' s (n-1))
+ and F' s 0 = Seq (value NONE)
| F' s n = make (fn () => F (getItem s) n)
in
fun take (s,n) =
@@ -143,8 +143,8 @@
local
fun F s 0 = s
- | F None _ = raise Subscript
- | F (Some(_,s)) n = F (getItem s) (n-1)
+ | F NONE _ = raise Subscript
+ | F (SOME(_,s)) n = F (getItem s) (n-1)
in
fun drop (s,0) = s
| drop (s,n) =
@@ -154,25 +154,25 @@
end
local
- fun F s None = s
- | F s (Some(x,s')) = F (Some(x, Seq (value s))) (getItem s')
+ fun F s NONE = s
+ | F s (SOME(x,s')) = F (SOME(x, Seq (value s))) (getItem s')
in
-fun rev s = make (fn () => F None (getItem s))
+fun rev s = make (fn () => F NONE (getItem s))
end
local
- fun F s None = getItem s
- | F s (Some(x,s')) = F (Seq (value (Some(x,s)))) (getItem s')
+ fun F s NONE = getItem s
+ | F s (SOME(x,s')) = F (Seq (value (SOME(x,s)))) (getItem s')
in
fun revAppend (s1,s2) = make (fn () => F s2 (getItem s1))
end
local
- fun F None = None
- | F (Some(s1,s2)) =
+ fun F NONE = NONE
+ | F (SOME(s1,s2)) =
let
- fun G None = F (getItem s2)
- | G (Some(x,s)) = Some(x,make (fn () => G (getItem s)))
+ fun G NONE = F (getItem s2)
+ | G (SOME(x,s)) = SOME(x,make (fn () => G (getItem s)))
in
G (getItem s1)
end
@@ -182,8 +182,8 @@
fun app f =
let
- fun F None = ()
- | F (Some(x,s)) =
+ fun F NONE = ()
+ | F (SOME(x,s)) =
(f x;
F (getItem s))
in
@@ -192,8 +192,8 @@
fun map f =
let
- fun F None = None
- | F (Some(x,s)) = Some(f x,F' s)
+ fun F NONE = NONE
+ | F (SOME(x,s)) = SOME(f x,F' s)
and F' s = make (fn() => F (getItem s))
in
F'
@@ -201,11 +201,11 @@
fun mapPartial f =
let
- fun F None = None
- | F (Some(x,s)) =
+ fun F NONE = NONE
+ | F (SOME(x,s)) =
(case f x of
- Some y => Some(y,F' s)
- | None => F (getItem s))
+ SOME y => SOME(y,F' s)
+ | NONE => F (getItem s))
and F' s = make (fn()=> F (getItem s))
in
F'
@@ -213,23 +213,23 @@
fun find P =
let
- fun F None = None
- | F (Some(x,s)) =
+ fun F NONE = NONE
+ | F (SOME(x,s)) =
if P x
- then Some x
+ then SOME x
else F (getItem s)
in
F o getItem
end
-(*fun filter p = mapPartial (fn x => if p x then Some x else None)*)
+(*fun filter p = mapPartial (fn x => if p x then SOME x else NONE)*)
fun filter P =
let
- fun F None = None
- | F (Some(x,s)) =
+ fun F NONE = NONE
+ | F (SOME(x,s)) =
if P x
- then Some(x,F' s)
+ then SOME(x,F' s)
else F (getItem s)
and F' s = make (fn () => F (getItem s))
in
@@ -240,22 +240,22 @@
let
val s' = map (fn x => (x,f x)) s
in
- (mapPartial (fn (x,true) => Some x | _ => None) s',
- mapPartial (fn (x,false) => Some x | _ => None) s')
+ (mapPartial (fn (x,true) => SOME x | _ => NONE) s',
+ mapPartial (fn (x,false) => SOME x | _ => NONE) s')
end
fun exists P =
let
- fun F None = false
- | F (Some(x,s)) = P x orelse F (getItem s)
+ fun F NONE = false
+ | F (SOME(x,s)) = P x orelse F (getItem s)
in
F o getItem
end
fun all P =
let
- fun F None = true
- | F (Some(x,s)) = P x andalso F (getItem s)
+ fun F NONE = true
+ | F (SOME(x,s)) = P x andalso F (getItem s)
in
F o getItem
end
@@ -264,16 +264,16 @@
fun tabulate (n,f) =
let
- fun F n = make (fn () => Some(f n,F (n+1)))
+ fun F n = make (fn () => SOME(f n,F (n+1)))
in
F n
end
fun collate c (s1,s2) =
let
- fun F None _ = LESS
- | F _ None = GREATER
- | F (Some(x,s1)) (Some(y,s2)) =
+ fun F NONE _ = LESS
+ | F _ NONE = GREATER
+ | F (SOME(x,s1)) (SOME(y,s2)) =
(case c (x,y) of
LESS => LESS
| GREATER => GREATER
@@ -283,13 +283,13 @@
F' s1 s2
end
-fun empty _ = Seq (value None)
-fun single x = Seq(value (Some(x,Seq (value None))))
-fun cons a = Seq(value (Some a))
+fun empty _ = Seq (value NONE)
+fun single x = Seq(value (SOME(x,Seq (value NONE))))
+fun cons a = Seq(value (SOME a))
fun cycle seqfn =
let
- val knot = ref (Seq (value None))
+ val knot = ref (Seq (value NONE))
in
knot := seqfn (fn () => !knot);
!knot
@@ -297,15 +297,15 @@
fun iterates f =
let
- fun F x = make (fn() => Some(x,F (f x)))
+ fun F x = make (fn() => SOME(x,F (f x)))
in
F
end
fun interleave (s1,s2) =
let
- fun F None = getItem s2
- | F (Some(x,s1')) = Some(x,interleave(s2,s1'))
+ fun F NONE = getItem s2
+ | F (SOME(x,s1')) = SOME(x,interleave(s2,s1'))
in
make (fn()=> F (getItem s1))
end
@@ -313,8 +313,8 @@
(* val force_all = app ignore *)
local
- fun F None = ()
- | F (Some(x,s)) = F (getItem s)
+ fun F NONE = ()
+ | F (SOME(x,s)) = F (getItem s)
in
fun force_all s = F (getItem s)
end
@@ -322,15 +322,15 @@
fun of_function f =
let
fun F () = case f () of
- Some x => Some(x,make F)
- | None => None
+ SOME x => SOME(x,make F)
+ | NONE => NONE
in
make F
end
local
- fun F [] = None
- | F (x::xs) = Some(x,F' xs)
+ fun F [] = NONE
+ | F (x::xs) = SOME(x,F' xs)
and F' xs = make (fn () => F xs)
in
fun of_list xs = F' xs
@@ -344,18 +344,18 @@
fun get_input () =
case !buffer of
(c::cs) => (buffer:=cs;
- Some c)
+ SOME c)
| [] => (case String.explode (TextIO.input is) of
- [] => None
+ [] => NONE
| (c::cs) => (buffer := cs;
- Some c))
+ SOME c))
in
of_function get_input
end
local
- fun F k None = k []
- | F k (Some(x,s)) = F (fn xs => k (x::xs)) (getItem s)
+ fun F k NONE = k []
+ | F k (SOME(x,s)) = F (fn xs => k (x::xs)) (getItem s)
in
fun list_of s = F (fn x => x) (getItem s)
end
@@ -363,17 +363,17 @@
(* Adapted from seq.ML *)
val succeed = single
-fun fail _ = Seq (value None)
+fun fail _ = Seq (value NONE)
(* fun op THEN (f, g) x = flat (map g (f x)) *)
fun op THEN (f, g) =
let
- fun F None = None
- | F (Some(x,xs)) =
+ fun F NONE = NONE
+ | F (SOME(x,xs)) =
let
- fun G None = F (getItem xs)
- | G (Some(y,ys)) = Some(y,make (fn () => G (getItem ys)))
+ fun G NONE = F (getItem xs)
+ | G (SOME(y,ys)) = SOME(y,make (fn () => G (getItem ys)))
in
G (getItem (g x))
end
@@ -384,15 +384,15 @@
fun op ORELSE (f, g) x =
make (fn () =>
case getItem (f x) of
- None => getItem (g x)
+ NONE => getItem (g x)
| some => some)
fun op APPEND (f, g) x =
let
fun copy s =
case getItem s of
- None => getItem (g x)
- | Some(x,xs) => Some(x,make (fn () => copy xs))
+ NONE => getItem (g x)
+ | SOME(x,xs) => SOME(x,make (fn () => copy xs))
in
make (fn () => copy (f x))
end
@@ -403,20 +403,20 @@
fun TRY f x =
make (fn () =>
case getItem (f x) of
- None => Some(x,Seq (value None))
+ NONE => SOME(x,Seq (value NONE))
| some => some)
fun REPEAT f =
let
fun rep qs x =
case getItem (f x) of
- None => Some(x, make (fn () => repq qs))
- | Some (x', q) => rep (q :: qs) x'
- and repq [] = None
+ NONE => SOME(x, make (fn () => repq qs))
+ | SOME (x', q) => rep (q :: qs) x'
+ and repq [] = NONE
| repq (q :: qs) =
case getItem q of
- None => repq qs
- | Some (x, q) => rep (q :: qs) x
+ NONE => repq qs
+ | SOME (x, q) => rep (q :: qs) x
in
fn x => make (fn () => rep [] x)
end
@@ -434,15 +434,15 @@
fun DETERM f x =
make (fn () =>
case getItem (f x) of
- None => None
- | Some (x', _) => Some(x',Seq (value None)))
+ NONE => NONE
+ | SOME (x', _) => SOME(x',Seq (value NONE)))
(*partial function as procedure*)
fun try f x =
make (fn () =>
case Library.try f x of
- Some y => Some(y,Seq (value None))
- | None => None)
+ SOME y => SOME(y,Seq (value NONE))
+ | NONE => NONE)
(*functional to print a sequence, up to "count" elements;
the function prelem should print the element number and also the element*)
@@ -453,8 +453,8 @@
then ()
else
case getItem xq of
- None => ()
- | Some (x, xq') =>
+ NONE => ()
+ | SOME (x, xq') =>
(prelem k x;
writeln "";
pr (k + 1) xq')
@@ -468,8 +468,8 @@
fun its s =
make (fn () =>
case getItem s of
- None => getItem yq
- | Some (a, s') => getItem(f (a, its s')))
+ NONE => getItem yq
+ | SOME (a, s') => getItem(f (a, its s')))
in
its xq
end
@@ -477,8 +477,8 @@
(*map over a sequence s1, append the sequence s2*)
fun mapp f s1 s2 =
let
- fun F None = getItem s2
- | F (Some(x,s1')) = Some(f x,F' s1')
+ fun F NONE = getItem s2
+ | F (SOME(x,s1')) = SOME(f x,F' s1')
and F' s = make (fn () => F (getItem s))
in
F' s1
@@ -486,22 +486,22 @@
(*turn a list of sequences into a sequence of lists*)
local
- fun F [] = Some([],Seq (value None))
+ fun F [] = SOME([],Seq (value NONE))
| F (xq :: xqs) =
case getItem xq of
- None => None
- | Some (x, xq') =>
+ NONE => NONE
+ | SOME (x, xq') =>
(case F xqs of
- None => None
- | Some (xs, xsq) =>
+ NONE => NONE
+ | SOME (xs, xsq) =>
let
fun G s =
make (fn () =>
case getItem s of
- None => F (xq' :: xqs)
- | Some(ys,ysq) => Some(x::ys,G ysq))
+ NONE => F (xq' :: xqs)
+ | SOME(ys,ysq) => SOME(x::ys,G ysq))
in
- Some (x :: xs, G xsq)
+ SOME (x :: xs, G xsq)
end)
in
fun commute xqs = make (fn () => F xqs)
@@ -514,21 +514,21 @@
then ([], xq)
else
case getItem xq of
- None => ([], xq)
- | Some (x, xq') => apfst (Library.cons x) (chop (n - 1, xq'))
+ NONE => ([], xq)
+ | SOME (x, xq') => apfst (Library.cons x) (chop (n - 1, xq'))
fun foldl f e s =
let
- fun F k None = k e
- | F k (Some(x,s)) = F (fn y => k (f(x,y))) (getItem s)
+ fun F k NONE = k e
+ | F k (SOME(x,s)) = F (fn y => k (f(x,y))) (getItem s)
in
F (fn x => x) (getItem s)
end
fun foldr f e s =
let
- fun F e None = e
- | F e (Some(x,s)) = F (f(x,e)) (getItem s)
+ fun F e NONE = e
+ | F e (SOME(x,s)) = F (f(x,e)) (getItem s)
in
F e (getItem s)
end
--- a/src/Pure/General/name_space.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/name_space.ML Sun Feb 13 17:15:14 2005 +0100
@@ -133,11 +133,11 @@
fun lookup (NameSpace tab) xname =
(case Symtab.lookup (tab, xname) of
- None => (xname, true)
- | Some ([name], _) => (name, true)
- | Some (name :: _, _) => (name, false)
- | Some ([], []) => (xname, true)
- | Some ([], name' :: _) => (hidden name', true));
+ NONE => (xname, true)
+ | SOME ([name], _) => (name, true)
+ | SOME (name :: _, _) => (name, false)
+ | SOME ([], []) => (xname, true)
+ | SOME ([], name' :: _) => (hidden name', true));
fun intern spc xname = #1 (lookup spc xname);
@@ -170,8 +170,8 @@
(* dest *)
-fun dest_entry (xname, ([], _)) = None
- | dest_entry (xname, (name :: _, _)) = Some (name, xname);
+fun dest_entry (xname, ([], _)) = NONE
+ | dest_entry (xname, (name :: _, _)) = SOME (name, xname);
fun dest (NameSpace tab) =
map (apsnd (sort (int_ord o pairself size)))
--- a/src/Pure/General/output.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/output.ML Sun Feb 13 17:15:14 2005 +0100
@@ -90,10 +90,10 @@
fun lookup_mode name = Symtab.lookup (! modes, name);
fun get_mode () =
- (case Library.get_first lookup_mode (! print_mode) of Some p => p
- | None =>
- (case lookup_mode "" of Some p => p
- | None => raise MISSING_DEFAULT_OUTPUT)); (*sys_error would require output again!*)
+ (case Library.get_first lookup_mode (! print_mode) of SOME p => p
+ | NONE =>
+ (case lookup_mode "" of SOME p => p
+ | NONE => raise MISSING_DEFAULT_OUTPUT)); (*sys_error would require output again!*)
fun output_width x = #output_width (get_mode ()) x;
val output = #1 o output_width;
@@ -171,11 +171,11 @@
Error of string |
OK of 'a;
-fun get_error (Error msg) = Some msg
- | get_error _ = None;
+fun get_error (Error msg) = SOME msg
+ | get_error _ = NONE;
-fun get_ok (OK x) = Some x
- | get_ok _ = None;
+fun get_ok (OK x) = SOME x
+ | get_ok _ = NONE;
fun handle_error f x =
let
--- a/src/Pure/General/path.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/path.ML Sun Feb 13 17:15:14 2005 +0100
@@ -130,7 +130,7 @@
fun split_path f (path as Path xs) =
(case try split_last xs of
- Some (prfx, Basic s) => f (prfx, s)
+ SOME (prfx, Basic s) => f (prfx, s)
| _ => error ("Cannot split path into dir/base: " ^ quote (pack path)));
val dir = split_path (fn (prfx, _) => Path prfx);
--- a/src/Pure/General/position.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/position.ML Sun Feb 13 17:15:14 2005 +0100
@@ -25,24 +25,24 @@
datatype T =
Pos of int option * string option;
-val none = Pos (None, None);
-fun line n = Pos (Some n, None);
-fun name s = Pos (None, Some s);
-fun line_name n s = Pos (Some n, Some s);
+val none = Pos (NONE, NONE);
+fun line n = Pos (SOME n, NONE);
+fun name s = Pos (NONE, SOME s);
+fun line_name n s = Pos (SOME n, SOME s);
(* increment *)
-fun inc (pos as Pos (None, _)) = pos
- | inc (Pos (Some n, opt_s)) = Pos (Some (n + 1), opt_s);
+fun inc (pos as Pos (NONE, _)) = pos
+ | inc (Pos (SOME n, opt_s)) = Pos (SOME (n + 1), opt_s);
(* str_of *)
-fun str_of (Pos (None, None)) = ""
- | str_of (Pos (Some n, None)) = " (line " ^ string_of_int n ^ ")"
- | str_of (Pos (None, Some s)) = " (" ^ s ^ ")"
- | str_of (Pos (Some n, Some s)) = " (line " ^ string_of_int n ^ " of " ^ s ^ ")";
+fun str_of (Pos (NONE, NONE)) = ""
+ | str_of (Pos (SOME n, NONE)) = " (line " ^ string_of_int n ^ ")"
+ | str_of (Pos (NONE, SOME s)) = " (" ^ s ^ ")"
+ | str_of (Pos (SOME n, SOME s)) = " (line " ^ string_of_int n ^ " of " ^ s ^ ")";
end;
--- a/src/Pure/General/scan.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/scan.ML Sun Feb 13 17:15:14 2005 +0100
@@ -139,29 +139,29 @@
(* generic scanners *)
-fun fail _ = raise FAIL None;
-fun fail_with msg_of xs = raise FAIL (Some (msg_of xs));
+fun fail _ = raise FAIL NONE;
+fun fail_with msg_of xs = raise FAIL (SOME (msg_of xs));
fun succeed y xs = (y, xs);
-fun one _ [] = raise MORE None
+fun one _ [] = raise MORE NONE
| one pred (x :: xs) =
- if pred x then (x, xs) else raise FAIL None;
+ if pred x then (x, xs) else raise FAIL NONE;
-fun $$ _ [] = raise MORE None
+fun $$ _ [] = raise MORE NONE
| $$ a (x :: xs) =
- if a = x then (x, xs) else raise FAIL None;
+ if a = x then (x, xs) else raise FAIL NONE;
fun this ys xs =
let
fun drop_prefix [] xs = xs
- | drop_prefix (_ :: _) [] = raise MORE None
+ | drop_prefix (_ :: _) [] = raise MORE NONE
| drop_prefix (y :: ys) (x :: xs) =
- if y = x then drop_prefix ys xs else raise FAIL None;
+ if y = x then drop_prefix ys xs else raise FAIL NONE;
in (ys, drop_prefix ys xs) end;
fun this_string s = this (explode s) >> K s; (*primitive string -- no symbols yet!*)
-fun any _ [] = raise MORE None
+fun any _ [] = raise MORE NONE
| any pred (lst as x :: xs) =
if pred x then apfst (cons x) (any pred xs)
else ([], lst);
@@ -173,11 +173,11 @@
in (x :: xs, toks3) end;
fun optional scan def = scan || succeed def
-fun option scan = (scan >> Some) || succeed None
+fun option scan = (scan >> SOME) || succeed NONE
fun repeat scan =
- let fun rep ys xs = (case (Some (scan xs) handle FAIL _ => None) of
- None => (rev ys, xs) | Some (y, xs') => rep (y :: ys) xs')
+ let fun rep ys xs = (case (SOME (scan xs) handle FAIL _ => NONE) of
+ NONE => (rev ys, xs) | SOME (y, xs') => rep (y :: ys) xs')
in rep [] end;
fun repeat1 scan toks =
@@ -188,16 +188,16 @@
fun max leq scan1 scan2 xs =
(case (option scan1 xs, option scan2 xs) of
- ((None, _), (None, _)) => raise FAIL None (*looses FAIL msg!*)
- | ((Some tok1, xs'), (None, _)) => (tok1, xs')
- | ((None, _), (Some tok2, xs')) => (tok2, xs')
- | ((Some tok1, xs1'), (Some tok2, xs2')) =>
+ ((NONE, _), (NONE, _)) => raise FAIL NONE (*looses FAIL msg!*)
+ | ((SOME tok1, xs'), (NONE, _)) => (tok1, xs')
+ | ((NONE, _), (SOME tok2, xs')) => (tok2, xs')
+ | ((SOME tok1, xs1'), (SOME tok2, xs2')) =>
if leq (tok2, tok1) then (tok1, xs1') else (tok2, xs2'));
fun ahead scan xs = (fst (scan xs), xs);
fun unless test scan =
- ahead (option test) :-- (fn None => scan | _ => fail) >> #2;
+ ahead (option test) :-- (fn NONE => scan | _ => fail) >> #2;
fun first [] = fail
| first (scan :: scans) = scan || first scans;
@@ -227,10 +227,10 @@
(* exception handling *)
fun !! err scan xs = scan xs handle FAIL msg => raise ABORT (err (xs, msg));
-fun try scan xs = scan xs handle MORE _ => raise FAIL None | ABORT _ => raise FAIL None;
-fun force scan xs = scan xs handle MORE _ => raise FAIL None;
-fun prompt str scan xs = scan xs handle MORE None => raise MORE (Some str);
-fun catch scan xs = scan xs handle ABORT msg => raise FAIL (Some msg);
+fun try scan xs = scan xs handle MORE _ => raise FAIL NONE | ABORT _ => raise FAIL NONE;
+fun force scan xs = scan xs handle MORE _ => raise FAIL NONE;
+fun prompt str scan xs = scan xs handle MORE NONE => raise MORE (SOME str);
+fun catch scan xs = scan xs handle ABORT msg => raise FAIL (SOME msg);
fun error scan xs = scan xs handle ABORT msg => Output.error msg;
@@ -256,8 +256,8 @@
fun read stopper scan xs =
(case error (finite stopper (option scan)) xs of
- (y as Some _, []) => y
- | _ => None);
+ (y as SOME _, []) => y
+ | _ => NONE);
(* infinite scans -- draining state-based source *)
@@ -279,8 +279,8 @@
val ((ys, (state', xs')), src') =
(case (get def_prmpt src, opt_recover) of
(([], s), _) => (([], (state, [])), s)
- | ((xs, s), None) => drain_with (error scanner) ((state, xs), s)
- | ((xs, s), Some r) => drain_loop (unless (lift (one (#2 stopper))) r) ((state, xs), s));
+ | ((xs, s), NONE) => drain_with (error scanner) ((state, xs), s)
+ | ((xs, s), SOME r) => drain_loop (unless (lift (one (#2 stopper))) r) ((state, xs), s));
in (ys, (state', unget (xs', src'))) end;
fun source def_prmpt get unget stopper scan opt_recover src =
@@ -364,16 +364,16 @@
fun literal lex chrs =
let
fun lit Empty res _ = res
- | lit (Branch _) _ [] = raise MORE None
+ | lit (Branch _) _ [] = raise MORE NONE
| lit (Branch (d, a, lt, eq, gt)) res (chs as c :: cs) =
(case String.compare (c, d) of
LESS => lit lt res chs
- | EQUAL => lit eq (if a = no_literal then res else Some (a, cs)) cs
+ | EQUAL => lit eq (if a = no_literal then res else SOME (a, cs)) cs
| GREATER => lit gt res chs);
in
- (case lit lex None chrs of
- None => raise FAIL None
- | Some res => res)
+ (case lit lex NONE chrs of
+ NONE => raise FAIL NONE
+ | SOME res => res)
end;
end;
--- a/src/Pure/General/seq.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/seq.ML Sun Feb 13 17:15:14 2005 +0100
@@ -56,17 +56,17 @@
(*the abstraction for making a sequence*)
val make = Seq;
-(*return next sequence element as None or Some (x, xq)*)
+(*return next sequence element as NONE or SOME (x, xq)*)
fun pull (Seq f) = f ();
(*the empty sequence*)
-val empty = Seq (fn () => None);
+val empty = Seq (fn () => NONE);
(*prefix an element to the sequence -- use cons (x, xq) only if
evaluation of xq need not be delayed, otherwise use
- make (fn () => Some (x, xq))*)
-fun cons x_xq = make (fn () => Some x_xq);
+ make (fn () => SOME (x, xq))*)
+fun cons x_xq = make (fn () => SOME x_xq);
fun single x = cons (x, empty);
@@ -77,8 +77,8 @@
(*partial function as procedure*)
fun try f x =
(case Library.try f x of
- Some y => single y
- | None => empty);
+ SOME y => single y
+ | NONE => empty);
(*the list of the first n elements, paired with rest of sequence;
@@ -87,14 +87,14 @@
if n <= 0 then ([], xq)
else
(case pull xq of
- None => ([], xq)
- | Some (x, xq') => apfst (Library.cons x) (chop (n - 1, xq')));
+ NONE => ([], xq)
+ | SOME (x, xq') => apfst (Library.cons x) (chop (n - 1, xq')));
(*conversion from sequence to list*)
fun list_of xq =
(case pull xq of
- None => []
- | Some (x, xq') => x :: list_of xq');
+ NONE => []
+ | SOME (x, xq') => x :: list_of xq');
(*conversion from list to sequence*)
fun of_list xs = foldr cons (xs, empty);
@@ -104,8 +104,8 @@
fun map f xq =
make (fn () =>
(case pull xq of
- None => None
- | Some (x, xq') => Some (f x, map f xq')));
+ NONE => NONE
+ | SOME (x, xq') => SOME (f x, map f xq')));
(*map over a sequence xq, append the sequence yq*)
fun mapp f xq yq =
@@ -113,8 +113,8 @@
fun copy s =
make (fn () =>
(case pull s of
- None => pull yq
- | Some (x, s') => Some (f x, copy s')))
+ NONE => pull yq
+ | SOME (x, s') => SOME (f x, copy s')))
in copy xq end;
(*sequence append: put the elements of xq in front of those of yq*)
@@ -123,8 +123,8 @@
fun copy s =
make (fn () =>
(case pull s of
- None => pull yq
- | Some (x, s') => Some (x, copy s')))
+ NONE => pull yq
+ | SOME (x, s') => SOME (x, copy s')))
in copy xq end;
(*filter sequence by predicate*)
@@ -133,23 +133,23 @@
fun copy s =
make (fn () =>
(case pull s of
- None => None
- | Some (x, s') => if pred x then Some (x, copy s') else pull (copy s')));
+ NONE => NONE
+ | SOME (x, s') => if pred x then SOME (x, copy s') else pull (copy s')));
in copy xq end;
(*flatten a sequence of sequences to a single sequence*)
fun flat xqq =
make (fn () =>
(case pull xqq of
- None => None
- | Some (xq, xqq') => pull (append (xq, flat xqq'))));
+ NONE => NONE
+ | SOME (xq, xqq') => pull (append (xq, flat xqq'))));
(*interleave elements of xq with those of yq -- fairer than append*)
fun interleave (xq, yq) =
make (fn () =>
(case pull xq of
- None => pull yq
- | Some (x, xq') => Some (x, interleave (yq, xq'))));
+ NONE => pull yq
+ | SOME (x, xq') => SOME (x, interleave (yq, xq'))));
(*functional to print a sequence, up to "count" elements;
@@ -160,8 +160,8 @@
if k > count then ()
else
(case pull xq of
- None => ()
- | Some (x, xq') => (prelem k x; writeln ""; pr (k + 1, xq')))
+ NONE => ()
+ | SOME (x, xq') => (prelem k x; writeln ""; pr (k + 1, xq')))
in pr (1, seq) end;
(*accumulating a function over a sequence; this is lazy*)
@@ -170,8 +170,8 @@
fun its s =
make (fn () =>
(case pull s of
- None => pull yq
- | Some (a, s') => pull (f (a, its s'))))
+ NONE => pull yq
+ | SOME (a, s') => pull (f (a, its s'))))
in its xq end;
(*turn a list of sequences into a sequence of lists*)
@@ -179,12 +179,12 @@
| commute (xq :: xqs) =
make (fn () =>
(case pull xq of
- None => None
- | Some (x, xq') =>
+ NONE => NONE
+ | SOME (x, xq') =>
(case pull (commute xqs) of
- None => None
- | Some (xs, xsq) =>
- Some (x :: xs, append (map (Library.cons x) xsq, commute (xq' :: xqs))))));
+ NONE => NONE
+ | SOME (xs, xsq) =>
+ SOME (x :: xs, append (map (Library.cons x) xsq, commute (xq' :: xqs))))));
@@ -197,7 +197,7 @@
fun op ORELSE (f, g) x =
(case pull (f x) of
- None => g x
+ NONE => g x
| some => make (fn () => some));
fun op APPEND (f, g) x =
@@ -212,13 +212,13 @@
let
fun rep qs x =
(case pull (f x) of
- None => Some (x, make (fn () => repq qs))
- | Some (x', q) => rep (q :: qs) x')
- and repq [] = None
+ NONE => SOME (x, make (fn () => repq qs))
+ | SOME (x', q) => rep (q :: qs) x')
+ and repq [] = NONE
| repq (q :: qs) =
(case pull q of
- None => repq qs
- | Some (x, q) => rep (q :: qs) x);
+ NONE => repq qs
+ | SOME (x, q) => rep (q :: qs) x);
in fn x => make (fn () => rep [] x) end;
fun REPEAT1 f = THEN (f, REPEAT f);
@@ -229,7 +229,7 @@
fun DETERM f x =
(case pull (f x) of
- None => empty
- | Some (x', _) => cons (x', empty));
+ NONE => empty
+ | SOME (x', _) => cons (x', empty));
end;
--- a/src/Pure/General/source.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/source.ML Sun Feb 13 17:15:14 2005 +0100
@@ -72,8 +72,8 @@
fun get_single src =
(case get src of
- ([], _) => None
- | (x :: xs, src') => Some (x, unget (xs, src')));
+ ([], _) => NONE
+ | (x :: xs, src') => SOME (x, unget (xs, src')));
fun exhaust src =
(case get src of
@@ -93,7 +93,7 @@
end;
fun mapfilter f src = make_source [] src default_prompt (drain_mapfilter f);
-fun filter pred = mapfilter (fn x => if pred x then Some x else None);
+fun filter pred = mapfilter (fn x => if pred x then SOME x else NONE);
--- a/src/Pure/General/symbol.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/symbol.ML Sun Feb 13 17:15:14 2005 +0100
@@ -361,13 +361,13 @@
fun scanner msg scan chs =
let
- fun message (cs, None) = msg ^ ": " ^ quote (beginning 10 cs)
- | message (cs, Some msg') = msg ^ ", " ^ msg' ^ ": " ^ quote (beginning 10 cs);
+ fun message (cs, NONE) = msg ^ ": " ^ quote (beginning 10 cs)
+ | message (cs, SOME msg') = msg ^ ", " ^ msg' ^ ": " ^ quote (beginning 10 cs);
val fin_scan = Scan.error (Scan.finite stopper (!! message scan));
in
(case fin_scan chs of
(result, []) => result
- | (_, rest) => error (message (rest, None)))
+ | (_, rest) => error (message (rest, NONE)))
end;
@@ -404,7 +404,7 @@
Scan.any ((not o is_blank) andf not_equal "\"" andf not_eof) >> K [malformed];
fun source do_recover src =
- Source.source stopper (Scan.bulk scan) (if do_recover then Some recover else None) src;
+ Source.source stopper (Scan.bulk scan) (if do_recover then SOME recover else NONE) src;
(* explode *)
--- a/src/Pure/General/table.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/General/table.ML Sun Feb 13 17:15:14 2005 +0100
@@ -89,30 +89,30 @@
fun keys tab = rev (foldl_table (fn (rev_ks, (k, _)) => k :: rev_ks) ([], tab));
fun exists P tab = foldl_table (fn (false, e) => P e | (b, _) => b) (false, tab);
-fun min_key Empty = None
- | min_key (Branch2 (left, (k, _), _)) = Some (if_none (min_key left) k)
- | min_key (Branch3 (left, (k, _), _, _, _)) = Some (if_none (min_key left) k);
+fun min_key Empty = NONE
+ | min_key (Branch2 (left, (k, _), _)) = SOME (if_none (min_key left) k)
+ | min_key (Branch3 (left, (k, _), _, _, _)) = SOME (if_none (min_key left) k);
-fun max_key Empty = None
- | max_key (Branch2 (_, (k, _), right)) = Some (if_none (max_key right) k)
- | max_key (Branch3 (_, _, _, (k,_), right)) = Some (if_none (max_key right) k);
+fun max_key Empty = NONE
+ | max_key (Branch2 (_, (k, _), right)) = SOME (if_none (max_key right) k)
+ | max_key (Branch3 (_, _, _, (k,_), right)) = SOME (if_none (max_key right) k);
(* lookup *)
-fun lookup (Empty, _) = None
+fun lookup (Empty, _) = NONE
| lookup (Branch2 (left, (k, x), right), key) =
(case Key.ord (key, k) of
LESS => lookup (left, key)
- | EQUAL => Some x
+ | EQUAL => SOME x
| GREATER => lookup (right, key))
| lookup (Branch3 (left, (k1, x1), mid, (k2, x2), right), key) =
(case Key.ord (key, k1) of
LESS => lookup (left, key)
- | EQUAL => Some x1
+ | EQUAL => SOME x1
| GREATER =>
(case Key.ord (key, k2) of
LESS => lookup (mid, key)
- | EQUAL => Some x2
+ | EQUAL => SOME x2
| GREATER => lookup (right, key)));
@@ -175,7 +175,7 @@
let
fun add ((tab, dups), (key, x)) =
(case lookup (tab, key) of
- None => (update ((key, x), tab), dups)
+ NONE => (update ((key, x), tab), dups)
| _ => (tab, key :: dups));
in
(case foldl add ((table, []), pairs) of
@@ -188,17 +188,17 @@
(* delete *)
-fun compare' None (k2, _) = LESS
- | compare' (Some k1) (k2, _) = Key.ord (k1, k2);
+fun compare' NONE (k2, _) = LESS
+ | compare' (SOME k1) (k2, _) = Key.ord (k1, k2);
fun if_eq EQUAL x y = x
| if_eq _ x y = y;
exception UNDEF of key;
-fun del (Some k) Empty = raise UNDEF k
- | del None (Branch2 (Empty, p, Empty)) = (p, (true, Empty))
- | del None (Branch3 (Empty, p, Empty, q, Empty)) =
+fun del (SOME k) Empty = raise UNDEF k
+ | del NONE (Branch2 (Empty, p, Empty)) = (p, (true, Empty))
+ | del NONE (Branch3 (Empty, p, Empty, q, Empty)) =
(p, (false, Branch2 (Empty, q, Empty)))
| del k (Branch2 (Empty, p, Empty)) = (case compare' k p of
EQUAL => (p, (true, Empty)) | _ => raise UNDEF (the k))
@@ -215,7 +215,7 @@
(true, Branch3 (l', p, rl, rp, rr))
| Branch3 (rl, rp, rm, rq, rr) => (false, Branch2
(Branch2 (l', p, rl), rp, Branch2 (rm, rq, rr)))))
- | ord => (case del (if_eq ord None k) r of
+ | ord => (case del (if_eq ord NONE k) r of
(p', (false, r')) => (p', (false, Branch2 (l, if_eq ord p' p, r')))
| (p', (true, r')) => (p', case l of
Branch2 (ll, lp, lr) =>
@@ -234,7 +234,7 @@
| (Branch2 (ml, mp, mr), Branch3 (rl, rp, rm, rq, rr)) =>
Branch3 (Branch2 (l', p, ml), mp, Branch2 (mr, q, rl), rp,
Branch2 (rm, rq, rr)))))
- | ord => (case del (if_eq ord None k) m of
+ | ord => (case del (if_eq ord NONE k) m of
(p', (false, m')) =>
(p', (false, Branch3 (l, if_eq ord p' p, m', q, r)))
| (p', (true, m')) => (p', (false, case (l, r) of
@@ -246,7 +246,7 @@
| (_, Branch3 (rl, rp, rm, rq, rr)) =>
Branch3 (l, if_eq ord p' p, Branch2 (m', q, rl), rp,
Branch2 (rm, rq, rr))))))
- | ord => (case del (if_eq ord None k) r of
+ | ord => (case del (if_eq ord NONE k) r of
(q', (false, r')) =>
(q', (false, Branch3 (l, p, m, if_eq ord q' q, r')))
| (q', (true, r')) => (q', (false, case (l, m) of
@@ -259,7 +259,7 @@
Branch3 (Branch2 (ll, lp, lm), lq, Branch2 (lr, p, ml), mp,
Branch2 (mr, if_eq ord q' q, r'))))));
-fun delete k t = snd (snd (del (Some k) t));
+fun delete k t = snd (snd (del (SOME k) t));
(* join and merge *)
@@ -268,18 +268,18 @@
let
fun add ((tab, dups), (key, x)) =
(case lookup (tab, key) of
- None => (update ((key, x), tab), dups)
- | Some y =>
+ NONE => (update ((key, x), tab), dups)
+ | SOME y =>
(case f (y, x) of
- Some z => (update ((key, z), tab), dups)
- | None => (tab, key :: dups)));
+ SOME z => (update ((key, z), tab), dups)
+ | NONE => (tab, key :: dups)));
in
(case foldl_table add ((table1, []), table2) of
(table', []) => table'
| (_, dups) => raise DUPS (rev dups))
end;
-fun merge eq tabs = join (fn (y, x) => if eq (y, x) then Some y else None) tabs;
+fun merge eq tabs = join (fn (y, x) => if eq (y, x) then SOME y else NONE) tabs;
(* tables with multiple entries per key (preserving order) *)
@@ -289,8 +289,8 @@
fun make_multi pairs = foldr update_multi (pairs, empty);
fun dest_multi tab = flat (map (fn (key, xs) => map (Library.pair key) xs) (dest tab));
-fun merge_multi eq tabs = join (fn (xs, xs') => Some (gen_merge_lists eq xs xs')) tabs;
-fun merge_multi' eq tabs = join (fn (xs, xs') => Some (gen_merge_lists' eq xs xs')) tabs;
+fun merge_multi eq tabs = join (fn (xs, xs') => SOME (gen_merge_lists eq xs xs')) tabs;
+fun merge_multi' eq tabs = join (fn (xs, xs') => SOME (gen_merge_lists' eq xs xs')) tabs;
(*final declarations of this structure!*)
--- a/src/Pure/IsaPlanner/focus_term_lib.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/IsaPlanner/focus_term_lib.ML Sun Feb 13 17:15:14 2005 +0100
@@ -263,27 +263,27 @@
handle out_of_term_exception _=> ft;
(* focus tools for working directly with the focus representation *)
- fun focus_up_appr (focus(t, appl(l,m))) = None
- | focus_up_appr (focus(t, appr(r,m))) = Some (focus(t$r, m))
- | focus_up_appr (focus(t, abs(s,ty,m))) = None
- | focus_up_appr (focus(t, root)) = None;
+ fun focus_up_appr (focus(t, appl(l,m))) = NONE
+ | focus_up_appr (focus(t, appr(r,m))) = SOME (focus(t$r, m))
+ | focus_up_appr (focus(t, abs(s,ty,m))) = NONE
+ | focus_up_appr (focus(t, root)) = NONE;
- fun focus_up_appl (focus(t, appl(l,m))) = Some (focus(l$t, m))
- | focus_up_appl (focus(t, appr(r,m))) = None
- | focus_up_appl (focus(t, abs(s,ty,m))) = None
- | focus_up_appl (focus(t, root)) = None;
+ fun focus_up_appl (focus(t, appl(l,m))) = SOME (focus(l$t, m))
+ | focus_up_appl (focus(t, appr(r,m))) = NONE
+ | focus_up_appl (focus(t, abs(s,ty,m))) = NONE
+ | focus_up_appl (focus(t, root)) = NONE;
- fun focus_up_abs (focus(t, appl(l,m))) = None
- | focus_up_abs (focus(t, appr(r,m))) = None
+ fun focus_up_abs (focus(t, appl(l,m))) = NONE
+ | focus_up_abs (focus(t, appr(r,m))) = NONE
| focus_up_abs (focus(t, abs(s,ty,m))) =
- Some (focus_up (focus(Abs(s,ty,t), m)))
- | focus_up_abs (focus(t, root)) = None;
+ SOME (focus_up (focus(Abs(s,ty,t), m)))
+ | focus_up_abs (focus(t, root)) = NONE;
- fun focus_up_abs_or_appr (focus(t, appl(l,m))) = None
- | focus_up_abs_or_appr (focus(t, appr(r,m))) = Some (focus(t$r, m))
+ fun focus_up_abs_or_appr (focus(t, appl(l,m))) = NONE
+ | focus_up_abs_or_appr (focus(t, appr(r,m))) = SOME (focus(t$r, m))
| focus_up_abs_or_appr (focus(t, abs(s,ty,m))) =
- Some (focus_up (focus(Abs(s,ty,t), m)))
- | focus_up_abs_or_appr (focus(t, root)) = None;
+ SOME (focus_up (focus(Abs(s,ty,t), m)))
+ | focus_up_abs_or_appr (focus(t, root)) = NONE;
fun tyenv_of_focus (focus(t, u)) = tyenv_of_upterm u;
@@ -310,8 +310,8 @@
fun next_leaf_of_fcterm_seq_aux t () =
let val nextt = next_leaf_fcterm t
in
- Some (nextt, Seq.make (next_leaf_of_fcterm_seq_aux nextt))
- end handle out_of_term_exception _ => None;
+ SOME (nextt, Seq.make (next_leaf_of_fcterm_seq_aux nextt))
+ end handle out_of_term_exception _ => NONE;
(* sequence of next upterms from start upterm...
ie a sequence of all leafs after this one*)
--- a/src/Pure/IsaPlanner/isa_fterm.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/IsaPlanner/isa_fterm.ML Sun Feb 13 17:15:14 2005 +0100
@@ -31,7 +31,7 @@
Term.term ->
(
((Term.indexname * Term.typ) list * (Term.indexname * Term.term) list)
- * (string * Type) list * Term.term) Library.option
+ * (string * Type) list * Term.term) option
val clean_unify_ft :
Sign.sg ->
int ->
@@ -97,7 +97,7 @@
open BasicIsaFTerm;
-(* Some general search within a focus term... *)
+(* SOME general search within a focus term... *)
(* Note: only upterms with a free or constant are going to yeald a
match, thus if we get anything else (bound or var) skip it! This is
@@ -122,8 +122,8 @@
fun mk_match_up_seq ft =
case focus_up_abs_or_appr ft of
- Some ft' => Seq.append(f ft, mk_match_up_seq ft')
- | None => f ft
+ SOME ft' => Seq.append(f ft, mk_match_up_seq ft')
+ | NONE => f ft
in
Seq.flat (Seq.map mk_match_up_seq fts)
end;
@@ -134,8 +134,8 @@
fun mk_match_up_seq ft =
case focus_up_abs_or_appr ft of
- Some ft' => Seq.append(f ft, mk_match_up_seq ft')
- | None => f ft
+ SOME ft' => Seq.append(f ft, mk_match_up_seq ft')
+ | NONE => f ft
in
Seq.flat (Seq.map mk_match_up_seq fts)
end;
@@ -294,8 +294,8 @@
fun clean_match_ft tsig pat ft =
let val (t, (Ts,absterm)) = prepmatch ft in
case TermLib.clean_match tsig (pat, t) of
- None => None
- | Some insts => Some (insts, Ts, absterm) end;
+ NONE => NONE
+ | SOME insts => SOME (insts, Ts, absterm) end;
(* ix = max var index *)
fun clean_unify_ft sgn ix pat ft =
let val (t, (Ts,absterm)) = prepmatch ft in
--- a/src/Pure/IsaPlanner/isand.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/IsaPlanner/isand.ML Sun Feb 13 17:15:14 2005 +0100
@@ -175,9 +175,9 @@
(* lookup type of a free var name from a list *)
fun lookupfree vs vn =
case Library.find_first (fn (n,ty) => n = vn) vs of
- None => raise ERROR_MESSAGE ("prepare_goal_export:lookupfree: "
+ NONE => raise ERROR_MESSAGE ("prepare_goal_export:lookupfree: "
^ vn ^ " does not occur in the term")
- | Some x => x;
+ | SOME x => x;
in
fun export_back (export {fixes = vs, assumes = hprems,
sgid = i, gth = gth}) newth =
--- a/src/Pure/IsaPlanner/isaplib.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/IsaPlanner/isaplib.ML Sun Feb 13 17:15:14 2005 +0100
@@ -20,7 +20,7 @@
val NTH : int -> ('a -> 'b Seq.seq) -> 'a -> 'b Seq.seq
val all_but_last_of_seq : 'a Seq.seq -> 'a Seq.seq
val nat_seq : int Seq.seq
- val nth_of_seq : int -> 'a Seq.seq -> 'a Library.option
+ val nth_of_seq : int -> 'a Seq.seq -> 'a option
val pair_seq : 'a Seq.seq -> 'b Seq.seq -> ('a * 'b) Seq.seq
val seq_is_empty : 'a Seq.seq -> bool
val number_seq : 'a Seq.seq -> (int * 'a) Seq.seq
@@ -51,9 +51,9 @@
val string_of_list : ('a -> string) -> 'a list -> string
(* options *)
- val aptosome : 'a Library.option -> ('a -> 'b) -> 'b Library.option
- val seq_mapfilter : ('a -> 'b Library.option) -> 'a Seq.seq -> 'b Seq.seq
- val seq_map_to_some_filter : ('a -> 'b) -> 'a Library.option Seq.seq
+ val aptosome : 'a option -> ('a -> 'b) -> 'b option
+ val seq_mapfilter : ('a -> 'b option) -> 'a Seq.seq -> 'b Seq.seq
+ val seq_map_to_some_filter : ('a -> 'b) -> 'a option Seq.seq
-> 'b Seq.seq
end;
@@ -70,20 +70,20 @@
let
fun recf s () =
case (Seq.pull s) of
- None => None
- | Some (None,s') => recf s' ()
- | Some (Some d, s') =>
- Some (f d, Seq.make (recf s'))
+ NONE => NONE
+ | SOME (NONE,s') => recf s' ()
+ | SOME (SOME d, s') =>
+ SOME (f d, Seq.make (recf s'))
in Seq.make (recf s0) end;
fun seq_mapfilter f s0 =
let
fun recf s () =
case (Seq.pull s) of
- None => None
- | Some (a,s') =>
- (case f a of None => recf s' ()
- | Some b => Some (b, Seq.make (recf s')))
+ NONE => NONE
+ | SOME (a,s') =>
+ (case f a of NONE => recf s' ()
+ | SOME b => SOME (b, Seq.make (recf s')))
in Seq.make (recf s0) end;
@@ -98,7 +98,7 @@
(* the sequence of natural numbers *)
val nat_seq =
- let fun nseq i () = Some (i, Seq.make (nseq (i+1)))
+ let fun nseq i () = SOME (i, Seq.make (nseq (i+1)))
in Seq.make (nseq 1)
end;
@@ -119,12 +119,12 @@
let
fun pseq s1 s2 () =
case Seq.pull s1 of
- None => None
- | Some (s1h, s1t) =>
+ NONE => NONE
+ | SOME (s1h, s1t) =>
case Seq.pull s2 of
- None => None
- | Some (s2h, s2t) =>
- Some ((s1h, s2h), Seq.make (pseq s1t s2t))
+ NONE => NONE
+ | SOME (s2h, s2t) =>
+ SOME ((s1h, s2h), Seq.make (pseq s1t s2t))
in
Seq.make (pseq seq1 seq2)
end;
@@ -137,12 +137,12 @@
let
fun f () =
case Seq.pull s of
- None => None
- | Some (a, s2) =>
+ NONE => NONE
+ | SOME (a, s2) =>
(case Seq.pull s2 of
- None => None
- | Some (a2,s3) =>
- Some (a, all_but_last_of_seq (Seq.cons (a2,s3))))
+ NONE => NONE
+ | SOME (a2,s3) =>
+ SOME (a, all_but_last_of_seq (Seq.cons (a2,s3))))
in
Seq.make f
end;
@@ -152,8 +152,8 @@
(* nth elem for sequenes, return none if out of bounds *)
fun nth_of_seq i l =
- if (seq_is_empty l) then None
- else if i <= 1 then Some (Seq.hd l)
+ if (seq_is_empty l) then NONE
+ else if i <= 1 then SOME (Seq.hd l)
else nth_of_seq (i-1) (Seq.tl l);
(* for use with tactics... gives no_tac if element isn't there *)
@@ -290,7 +290,7 @@
(* options *)
- fun aptosome None f = None
- | aptosome (Some x) f = Some (f x);
+ fun aptosome NONE f = NONE
+ | aptosome (SOME x) f = SOME (f x);
end;
--- a/src/Pure/IsaPlanner/rw_tools.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/IsaPlanner/rw_tools.ML Sun Feb 13 17:15:14 2005 +0100
@@ -166,8 +166,8 @@
fun rename_term_bvars ns (Abs(s,ty,t)) =
let val s2opt = Library.find_first (fn (x,y) => s = x) ns
in case s2opt of
- None => (Abs(s,ty,rename_term_bvars ns t))
- | Some (_,s2) => Abs(s2,ty, rename_term_bvars ns t) end
+ NONE => (Abs(s,ty,rename_term_bvars ns t))
+ | SOME (_,s2) => Abs(s2,ty, rename_term_bvars ns t) end
| rename_term_bvars ns (a$b) =
(rename_term_bvars ns a) $ (rename_term_bvars ns b)
| rename_term_bvars _ x = x;
--- a/src/Pure/IsaPlanner/term_lib.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/IsaPlanner/term_lib.ML Sun Feb 13 17:15:14 2005 +0100
@@ -23,7 +23,7 @@
(* Matching/unification with exceptions handled *)
val clean_match : Type.tsig -> Term.term * Term.term
-> ((Term.indexname * Term.typ) list
- * (Term.indexname * Term.term) list) Library.option
+ * (Term.indexname * Term.term) list) option
val clean_unify : Sign.sg -> int -> Term.term * Term.term
-> ((Term.indexname * Term.typ) list
* (Term.indexname * Term.term) list) Seq.seq
@@ -75,7 +75,7 @@
Term.term ->
Term.term ->
(Term.term * Term.term) list ->
- (Term.term * Term.term) list Library.option
+ (Term.term * Term.term) list option
(* is the typ a function or is it atomic *)
val is_fun_typ : Term.typ -> bool
@@ -84,7 +84,7 @@
(* variable analysis *)
val is_some_kind_of_var : Term.term -> bool
val same_var_check :
- ''a -> ''b -> (''a * ''b) list -> (''a * ''b) list Library.option
+ ''a -> ''b -> (''a * ''b) list -> (''a * ''b) list option
val has_new_vars : Term.term * Term.term -> bool
val has_new_typ_vars : Term.term * Term.term -> bool
(* checks to see if the lhs -> rhs is a invalid rewrite rule *)
@@ -143,7 +143,7 @@
(* Higher order matching with exception handled *)
(* returns optional instantiation *)
fun clean_match tsig (a as (pat, t)) =
- Some (Pattern.match tsig a) handle Pattern.MATCH => None;
+ SOME (Pattern.match tsig a) handle Pattern.MATCH => NONE;
(* Higher order unification with exception handled, return the instantiations *)
(* given Signature, max var index, pat, tgt *)
(* returns Seq of instantiations *)
@@ -156,12 +156,12 @@
(* is it OK to ignore the type instantiation info?
or should I be using it? *)
val typs_unify =
- (Some (Type.unify (Sign.tsig_of sgn) (Term.Vartab.empty, ix)
+ (SOME (Type.unify (Sign.tsig_of sgn) (Term.Vartab.empty, ix)
(pat_ty, tgt_ty)))
- handle Type.TUNIFY => None;
+ handle Type.TUNIFY => NONE;
in
case typs_unify of
- Some (typinsttab, ix2) =>
+ SOME (typinsttab, ix2) =>
let
(* is it right to throw away teh flexes?
or should I be using them somehow? *)
@@ -174,14 +174,14 @@
| Term.TERM _ => Seq.empty;
fun clean_unify' useq () =
(case (Seq.pull useq) of
- None => None
- | Some (h,t) => Some (mk_insts h, Seq.make (clean_unify' t)))
- handle Library.LIST _ => None
- | Term.TERM _ => None;
+ NONE => NONE
+ | SOME (h,t) => SOME (mk_insts h, Seq.make (clean_unify' t)))
+ handle Library.LIST _ => NONE
+ | Term.TERM _ => NONE;
in
(Seq.make (clean_unify' useq))
end
- | None => Seq.empty
+ | NONE => Seq.empty
end;
fun asm_mk t = (assume (cterm_of (Theory.sign_of (the_context())) t));
@@ -295,15 +295,15 @@
Note frees and vars must all have unique names. *)
fun same_var_check a b L =
let
- fun bterm_from t [] = None
+ fun bterm_from t [] = NONE
| bterm_from t ((a,b)::m) =
- if t = a then Some b else bterm_from t m
+ if t = a then SOME b else bterm_from t m
val bl_opt = bterm_from a L
in
case bterm_from a L of
- None => Some ((a,b)::L)
- | Some b2 => if b2 = b then Some L else None
+ NONE => SOME ((a,b)::L)
+ | SOME b2 => if b2 = b then SOME L else NONE
end;
(* FIXME: make more efficient, only require a single new var! *)
@@ -334,11 +334,11 @@
fun myadd (n,ty) (L as (renames, names)) =
let val n' = Syntax.dest_skolem n in
case (Library.find_first (fn (_,n2) => (n2 = n)) renames) of
- None =>
+ NONE =>
let val renamedn = Term.variant names n' in
(renamedn, (((renamedn, ty), n) :: renames,
renamedn :: names)) end
- | (Some ((renamedn, _), _)) => (renamedn, L)
+ | (SOME ((renamedn, _), _)) => (renamedn, L)
end
handle LIST _ => (n, L);
@@ -363,7 +363,7 @@
let
fun fmap fv =
let (n,ty) = (Term.dest_Free fv) in
- Some (fv, Free (Syntax.dest_skolem n, ty)) handle LIST _ => None
+ SOME (fv, Free (Syntax.dest_skolem n, ty)) handle LIST _ => NONE
end
val substfrees = mapfilter fmap (Term.term_frees t)
in
@@ -389,7 +389,7 @@
let
val sgn = Theory.sign_of (the_context())
in
- Sign.simple_read_term sgn (Sign.read_typ (sgn, K None) tystr) tstr
+ Sign.simple_read_term sgn (Sign.read_typ (sgn, K NONE) tystr) tstr
end;
@@ -402,7 +402,7 @@
fun term_name_eq (Abs(_,ty1,t1)) (Abs(_,ty2,t2)) l =
if ty1 = ty2 then term_name_eq t1 t2 l
- else None
+ else NONE
| term_name_eq (ah $ at) (bh $ bt) l =
let
@@ -411,12 +411,12 @@
if is_some lopt then
term_name_eq at bt (the lopt)
else
- None
+ NONE
end
| term_name_eq (Const(a,T)) (Const(b,U)) l =
- if a=b andalso T=U then Some l
- else None
+ if a=b andalso T=U then SOME l
+ else NONE
| term_name_eq (a as Free(s1,ty1)) (b as Free(s2,ty2)) l =
same_var_check a b l
@@ -431,14 +431,14 @@
same_var_check a b l
| term_name_eq (Bound i) (Bound j) l =
- if i = j then Some l else None
+ if i = j then SOME l else NONE
- | term_name_eq a b l = ((*writeln ("unchecked case:\n" ^ "a:\n" ^ (pretty_print_term a) ^ "\nb:\n" ^ (pretty_print_term b));*) None);
+ | term_name_eq a b l = ((*writeln ("unchecked case:\n" ^ "a:\n" ^ (pretty_print_term a) ^ "\nb:\n" ^ (pretty_print_term b));*) NONE);
(* checks to see if the term in a name-equivalent member of the list of terms. *)
- fun get_name_eq_member a [] = None
+ fun get_name_eq_member a [] = NONE
| get_name_eq_member a (h :: t) =
- if is_some (term_name_eq a h []) then Some h
+ if is_some (term_name_eq a h []) then SOME h
else get_name_eq_member a t;
fun name_eq_member a [] = false
@@ -474,13 +474,13 @@
pattern in the variables. *)
fun term_contains1 (Bs, FVs) (Abs(s,ty,t)) (Abs(s2,ty2,t2)) =
if ty = ty2 then
- term_contains1 ((Some(s,s2,ty)::Bs), FVs) t t2
+ term_contains1 ((SOME(s,s2,ty)::Bs), FVs) t t2
else []
| term_contains1 T t1 (Abs(s2,ty2,t2)) = []
| term_contains1 (Bs, FVs) (Abs(s,ty,t)) t2 =
- term_contains1 (None::Bs, FVs) t t2
+ term_contains1 (NONE::Bs, FVs) t t2
| term_contains1 T (ah $ at) (bh $ bt) =
(term_contains1 T ah (bh $ bt)) @
@@ -499,9 +499,9 @@
list mapping, if it is then it checks that the pair are mapped to
each other, if so returns the current mapping list, else none. *)
let
- fun bterm_from t [] = None
+ fun bterm_from t [] = NONE
| bterm_from t ((a,b)::m) =
- if t = a then Some b else bterm_from t m
+ if t = a then SOME b else bterm_from t m
(* check to see if, w.r.t. the variable mapping, two terms are leaf
terms and are mapped to each other. Note constants are only mapped
@@ -509,11 +509,11 @@
fun same_leaf_check (T as (Bs,FVs)) (Bound i) (Bound j) =
let
fun aux_chk (i1,i2) [] = false
- | aux_chk (0,0) ((Some _) :: bnds) = true
- | aux_chk (i1,0) (None :: bnds) = false
- | aux_chk (i1,i2) ((Some _) :: bnds) =
+ | aux_chk (0,0) ((SOME _) :: bnds) = true
+ | aux_chk (i1,0) (NONE :: bnds) = false
+ | aux_chk (i1,i2) ((SOME _) :: bnds) =
aux_chk (i1 - 1,i2 - 1) bnds
- | aux_chk (i1,i2) (None :: bnds) =
+ | aux_chk (i1,i2) (NONE :: bnds) =
aux_chk (i1,i2 - 1) bnds
in
if (aux_chk (i,j) Bs) then [T]
@@ -523,16 +523,16 @@
(a as (Free (an,aty)))
(b as (Free (bn,bty))) =
(case bterm_from an Fs of
- Some b2n => if bn = b2n then [T]
+ SOME b2n => if bn = b2n then [T]
else [] (* conflict of var name *)
- | None => [(Bs,((an,bn)::Fs,Vs))])
+ | NONE => [(Bs,((an,bn)::Fs,Vs))])
| same_leaf_check (T as (Bs,(Fs,Vs)))
(a as (Var (an,aty)))
(b as (Var (bn,bty))) =
(case bterm_from an Vs of
- Some b2n => if bn = b2n then [T]
+ SOME b2n => if bn = b2n then [T]
else [] (* conflict of var name *)
- | None => [(Bs,(Fs,(an,bn)::Vs))])
+ | NONE => [(Bs,(Fs,(an,bn)::Vs))])
| same_leaf_check T (a as (Const _)) (b as (Const _)) =
if a = b then [T] else []
| same_leaf_check T _ _ = []
@@ -585,13 +585,13 @@
unique name *)
fun unique_namify_aux (ntab,(Abs(s,ty,t))) =
(case (fvartabS.lookup (ntab,s)) of
- None =>
+ NONE =>
let
val (ntab2,t2) = unique_namify_aux ((fvartabS.update ((s,s),ntab)), t)
in
(ntab2,Abs(s,ty,t2))
end
- | Some s2 =>
+ | SOME s2 =>
let
val s_new = (Term.variant (fvartabS.keys ntab) s)
val (ntab2,t2) =
@@ -609,12 +609,12 @@
| unique_namify_aux (nt as (ntab,Const x)) = nt
| unique_namify_aux (nt as (ntab,f as Free (s,ty))) =
(case (fvartabS.lookup (ntab,s)) of
- None => ((fvartabS.update ((s,s),ntab)), f)
- | Some _ => nt)
+ NONE => ((fvartabS.update ((s,s),ntab)), f)
+ | SOME _ => nt)
| unique_namify_aux (nt as (ntab,v as Var ((s,i),ty))) =
(case (fvartabS.lookup (ntab,s)) of
- None => ((fvartabS.update ((s,s),ntab)), v)
- | Some _ => nt)
+ NONE => ((fvartabS.update ((s,s),ntab)), v)
+ | SOME _ => nt)
| unique_namify_aux (nt as (ntab, Bound i)) = nt;
fun unique_namify t =
@@ -628,14 +628,14 @@
fun bounds_to_frees_aux T (ntab,(Abs(s,ty,t))) =
(case (fvartabS.lookup (ntab,s)) of
- None =>
+ NONE =>
let
val (ntab2,t2) = bounds_to_frees_aux ((s,ty)::T)
((fvartabS.update ((s,s),ntab)), t)
in
(ntab2,Abs(s,ty,t2))
end
- | Some s2 =>
+ | SOME s2 =>
let
val s_new = (Term.variant (fvartabS.keys ntab) s)
val (ntab2,t2) =
@@ -654,12 +654,12 @@
| bounds_to_frees_aux T (nt as (ntab,Const x)) = nt
| bounds_to_frees_aux T (nt as (ntab,f as Free (s,ty))) =
(case (fvartabS.lookup (ntab,s)) of
- None => ((fvartabS.update ((s,s),ntab)), f)
- | Some _ => nt)
+ NONE => ((fvartabS.update ((s,s),ntab)), f)
+ | SOME _ => nt)
| bounds_to_frees_aux T (nt as (ntab,v as Var ((s,i),ty))) =
(case (fvartabS.lookup (ntab,s)) of
- None => ((fvartabS.update ((s,s),ntab)), v)
- | Some _ => nt)
+ NONE => ((fvartabS.update ((s,s),ntab)), v)
+ | SOME _ => nt)
| bounds_to_frees_aux T (nt as (ntab, Bound i)) =
let
val (s,ty) = Library.nth_elem (i,T)
--- a/src/Pure/Isar/args.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/args.ML Sun Feb 13 17:15:14 2005 +0100
@@ -112,8 +112,8 @@
fun get_pos [] = " (past end-of-text!)"
| get_pos (Arg (_, (_, pos)) :: _) = Position.str_of pos;
- fun err (args, None) = "Argument syntax error" ^ get_pos args
- | err (args, Some msg) = "Argument syntax error" ^ get_pos args ^ ": " ^ msg;
+ fun err (args, NONE) = "Argument syntax error" ^ get_pos args
+ | err (args, SOME msg) = "Argument syntax error" ^ get_pos args ^ ": " ^ msg;
in Scan.!! err scan end;
@@ -135,14 +135,14 @@
fun mode s = Scan.lift (Scan.optional (parens ($$$$ s) >> K true) false);
val name = Scan.one (fn Arg (k, (x, _)) => k = Ident orelse k = String) >> val_of;
-val name_dummy = $$$ "_" >> K None || name >> Some;
+val name_dummy = $$$ "_" >> K NONE || name >> SOME;
val keyword_symid =
Scan.one (fn Arg (k, (x, _)) => k = Keyword andalso OuterLex.is_sid x) >> val_of;
fun kind f = Scan.one (K true) :--
(fn Arg (Ident, (x, _)) =>
- (case f x of Some y => Scan.succeed y | _ => Scan.fail)
+ (case f x of SOME y => Scan.succeed y | _ => Scan.fail)
| _ => Scan.fail) >> #2;
val nat = kind Syntax.read_nat;
@@ -150,9 +150,9 @@
(*read variable name; leading '?' may be omitted if name contains no dot*)
val var = kind (apsome #1 o (fn s =>
- Some (Term.dest_Var (Syntax.read_var s))
- handle _ => Some (Term.dest_Var (Syntax.read_var ("?" ^ s)))
- handle _ => None));
+ SOME (Term.dest_Var (Syntax.read_var s))
+ handle _ => SOME (Term.dest_Var (Syntax.read_var ("?" ^ s)))
+ handle _ => NONE));
(* enumerations *)
@@ -234,7 +234,7 @@
fun syntax kind scan (src as Src ((s, args), pos)) st =
(case Scan.error (Scan.finite' stopper (Scan.option scan)) (st, args) of
- (Some x, (st', [])) => (st', x)
+ (SOME x, (st', [])) => (st', x)
| (_, (_, args')) => err_in_src kind "bad arguments" (Src ((s, args'), pos)));
--- a/src/Pure/Isar/attrib.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/attrib.ML Sun Feb 13 17:15:14 2005 +0100
@@ -89,8 +89,8 @@
val name = NameSpace.intern space raw_name;
in
(case Symtab.lookup (attrs, name) of
- None => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
- | Some ((p, _), _) => transform_failure (curry ATTRIB_FAIL (name, pos)) (which p src))
+ NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
+ | SOME ((p, _), _) => transform_failure (curry ATTRIB_FAIL (name, pos)) (which p src))
end;
in attr end;
@@ -228,7 +228,7 @@
Syntax.string_of_vname xi);
val (rtypes, rsorts) = types_sorts thm;
fun readT (xi, s) =
- let val S = case rsorts xi of Some S => S | None => absent xi;
+ let val S = case rsorts xi of SOME S => S | NONE => absent xi;
val T = ProofContext.read_typ ctxt s;
in if Sign.typ_instance sign (T, TVar (xi, S)) then (xi, T)
else error
@@ -244,8 +244,8 @@
val vars = Drule.vars_of thm';
fun get_typ xi =
(case assoc (vars, xi) of
- Some T => T
- | None => error ("No such variable in theorem: " ^ Syntax.string_of_vname xi));
+ SOME T => T
+ | NONE => error ("No such variable in theorem: " ^ Syntax.string_of_vname xi));
val (xs, ss) = Library.split_list tinsts;
val Ts = map get_typ xs;
@@ -256,7 +256,7 @@
are distinct from those occuring in the theorem. *)
val (ts, envT) =
- ProofContext.read_termTs ctxt (K false) (K None) (K None) used (ss ~~ Ts);
+ ProofContext.read_termTs ctxt (K false) (K NONE) (K NONE) used (ss ~~ Ts);
val cenvT = map (apsnd (Thm.ctyp_of sign)) envT;
val cenv =
@@ -283,8 +283,8 @@
| read_instantiate' context_of (args, concl_args) x thm =
let
fun zip_vars _ [] = []
- | zip_vars (_ :: xs) (None :: opt_ts) = zip_vars xs opt_ts
- | zip_vars ((x, _) :: xs) (Some t :: opt_ts) = (x, t) :: zip_vars xs opt_ts
+ | zip_vars (_ :: xs) (NONE :: opt_ts) = zip_vars xs opt_ts
+ | zip_vars ((x, _) :: xs) (SOME t :: opt_ts) = (x, t) :: zip_vars xs opt_ts
| zip_vars [] _ = error "More instantiations than variables in theorem";
val insts =
zip_vars (Drule.vars_of_terms [Thm.prop_of thm]) args @
--- a/src/Pure/Isar/auto_bind.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/auto_bind.ML Sun Feb 13 17:15:14 2005 +0100
@@ -23,21 +23,21 @@
fun strip_judgment sg = ObjectLogic.drop_judgment sg o Logic.strip_assums_concl;
fun statement_binds sg name prop =
- [((name, 0), Some (Term.list_abs (Logic.strip_params prop, strip_judgment sg prop)))];
+ [((name, 0), SOME (Term.list_abs (Logic.strip_params prop, strip_judgment sg prop)))];
(* goal *)
fun goal sg [prop] = statement_binds sg thesisN prop
- | goal _ _ = [((thesisN, 0), None)];
+ | goal _ _ = [((thesisN, 0), NONE)];
(* facts *)
fun get_arg sg prop =
(case strip_judgment sg prop of
- _ $ t => Some (Term.list_abs (Logic.strip_params prop, t))
- | _ => None);
+ _ $ t => SOME (Term.list_abs (Logic.strip_params prop, t))
+ | _ => NONE);
fun facts _ [] = []
| facts sg props =
--- a/src/Pure/Isar/calculation.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/calculation.ML Sun Feb 13 17:15:14 2005 +0100
@@ -64,7 +64,7 @@
val name = "Isar/calculation";
type T = (thm NetRules.T * thm list) * (thm list * int) option;
- fun init thy = (GlobalCalculation.get thy, None);
+ fun init thy = (GlobalCalculation.get thy, NONE);
fun print ctxt (rs, _) = print_rules ProofContext.pretty_thm ctxt rs;
end;
@@ -77,15 +77,15 @@
fun get_calculation state =
(case #2 (LocalCalculation.get (Proof.context_of state)) of
- None => None
- | Some (thms, lev) => if lev = Proof.level state then Some thms else None);
+ NONE => NONE
+ | SOME (thms, lev) => if lev = Proof.level state then SOME thms else NONE);
fun put_calculation thms state =
Proof.map_context
- (LocalCalculation.put (get_local_rules state, Some (thms, Proof.level state))) state;
+ (LocalCalculation.put (get_local_rules state, SOME (thms, Proof.level state))) state;
fun reset_calculation state =
- Proof.map_context (LocalCalculation.put (get_local_rules state, None)) state;
+ Proof.map_context (LocalCalculation.put (get_local_rules state, NONE)) state;
@@ -101,9 +101,9 @@
val trans_add_local = local_att (apfst o NetRules.insert);
val trans_del_local = local_att (apfst o NetRules.delete);
-val sym_add_global = global_att (apsnd o Drule.add_rule) o ContextRules.elim_query_global None;
+val sym_add_global = global_att (apsnd o Drule.add_rule) o ContextRules.elim_query_global NONE;
val sym_del_global = global_att (apsnd o Drule.del_rule) o ContextRules.rule_del_global;
-val sym_add_local = local_att (apsnd o Drule.add_rule) o ContextRules.elim_query_local None;
+val sym_add_local = local_att (apsnd o Drule.add_rule) o ContextRules.elim_query_local NONE;
val sym_del_local = local_att (apsnd o Drule.del_rule) o ContextRules.rule_del_local;
@@ -164,8 +164,8 @@
fun projection ths th = Library.exists (Library.curry eq_prop th) ths;
fun combine ths =
- (case opt_rules of Some rules => rules
- | None =>
+ (case opt_rules of SOME rules => rules
+ | NONE =>
(case ths of [] => NetRules.rules (#1 (get_local_rules state))
| th :: _ => NetRules.retrieve (#1 (get_local_rules state)) (strip_assums_concl th)))
|> Seq.of_list |> Seq.map (Method.multi_resolve ths) |> Seq.flat
@@ -174,8 +174,8 @@
val facts = Proof.the_facts (assert_sane final state);
val (initial, calculations) =
(case get_calculation state of
- None => (true, Seq.single facts)
- | Some calc => (false, Seq.map single (combine (calc @ facts))));
+ NONE => (true, Seq.single facts)
+ | SOME calc => (false, Seq.map single (combine (calc @ facts))));
in
err_if state (initial andalso final) "No calculation yet";
err_if state (initial andalso is_some opt_rules) "Initial calculation -- no rules to be given";
@@ -194,8 +194,8 @@
val facts = Proof.the_facts (assert_sane final state);
val (initial, thms) =
(case get_calculation state of
- None => (true, [])
- | Some thms => (false, thms));
+ NONE => (true, [])
+ | SOME thms => (false, thms));
val calc = thms @ facts;
in
err_if state (initial andalso final) "No calculation yet";
--- a/src/Pure/Isar/constdefs.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/constdefs.ML Sun Feb 13 17:15:14 2005 +0100
@@ -40,13 +40,13 @@
val ctxt =
ProofContext.init thy |> ProofContext.add_fixes
- (flat (map (fn (As, T) => map (fn A => (A, T, None)) As) structs));
+ (flat (map (fn (As, T) => map (fn A => (A, T, NONE)) As) structs));
val (ctxt', d, mx) =
- (case decl of None => (ctxt, None, Syntax.NoSyn) | Some (x, raw_T, mx) =>
+ (case decl of NONE => (ctxt, NONE, Syntax.NoSyn) | SOME (x, raw_T, mx) =>
let
val x' = Syntax.const_name x mx and mx' = Syntax.fix_mixfix x mx;
val T = apsome (prep_typ ctxt) raw_T;
- in (ProofContext.add_fixes_liberal [(x', T, Some mx')] ctxt, Some x', mx') end);
+ in (ProofContext.add_fixes_liberal [(x', T, SOME mx')] ctxt, SOME x', mx') end);
val prop = prep_term ctxt' raw_prop;
val concl = Logic.strip_imp_concl prop;
@@ -55,7 +55,7 @@
val head = Term.head_of lhs;
val (c, T) = Term.dest_Free head handle TERM _ =>
err "Locally fixed variable required as head of definition:" [head];
- val _ = (case d of None => () | Some c' =>
+ val _ = (case d of NONE => () | SOME c' =>
if c <> c' then
err ("Head of definition " ^ quote c ^ " differs from declaration " ^ quote c') []
else ());
--- a/src/Pure/Isar/context_rules.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/context_rules.ML Sun Feb 13 17:15:14 2005 +0100
@@ -95,7 +95,7 @@
Rules {next = next, rules = rules, netpairs = netpairs, wrappers = wrappers};
fun add_rule (i, b) opt_w th (Rules {next, rules, netpairs, wrappers}) =
- let val w = (case opt_w of Some w => w | None => Tactic.subgoals_of_brl (b, th)) in
+ let val w = (case opt_w of SOME w => w | NONE => Tactic.subgoals_of_brl (b, th)) in
make_rules (next - 1) ((w, ((i, b), th)) :: rules)
(map_nth_elem i (curry insert_tagged_brl ((w, next), (b, th))) netpairs) wrappers
end;
@@ -113,7 +113,7 @@
let
fun prt_kind (i, b) =
Pretty.big_list (the (assoc (kind_names, (i, b))) ^ ":")
- (mapfilter (fn (_, (k, th)) => if k = (i, b) then Some (prt x th) else None)
+ (mapfilter (fn (_, (k, th)) => if k = (i, b) then SOME (prt x th) else NONE)
(sort (Int.compare o pairself fst) rules));
in Pretty.writeln (Pretty.chunks (map prt_kind rule_kinds)) end;
@@ -254,14 +254,14 @@
fun modifier att (x, ths) =
#1 (Thm.applys_attributes ((x, rev ths), [att]));
-val addXIs_global = modifier (intro_query_global None);
-val addXEs_global = modifier (elim_query_global None);
-val addXDs_global = modifier (dest_query_global None);
+val addXIs_global = modifier (intro_query_global NONE);
+val addXEs_global = modifier (elim_query_global NONE);
+val addXDs_global = modifier (dest_query_global NONE);
val delrules_global = modifier rule_del_global;
-val addXIs_local = modifier (intro_query_local None);
-val addXEs_local = modifier (elim_query_local None);
-val addXDs_local = modifier (dest_query_local None);
+val addXIs_local = modifier (intro_query_local NONE);
+val addXEs_local = modifier (elim_query_local NONE);
+val addXDs_local = modifier (dest_query_local NONE);
val delrules_local = modifier rule_del_local;
--- a/src/Pure/Isar/delta_data.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/delta_data.ML Sun Feb 13 17:15:14 2005 +0100
@@ -31,8 +31,8 @@
let val delta_tab = ProofContext.get_delta ctxt
val delta_data = Symtab.lookup(delta_tab,key)
in
- case delta_data of (Some (Data d)) => d
- | None => (Args.empty)
+ case delta_data of (SOME (Data d)) => d
+ | NONE => (Args.empty)
end;
fun put delta_data ctxt =
--- a/src/Pure/Isar/isar_cmd.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/isar_cmd.ML Sun Feb 13 17:15:14 2005 +0100
@@ -114,8 +114,8 @@
fun with_modes modes e =
Library.setmp print_mode (modes @ ! print_mode) e ();
-fun set_limit _ None = ()
- | set_limit r (Some n) = r := n;
+fun set_limit _ NONE = ()
+ | set_limit r (SOME n) = r := n;
fun pr (ms, (lim1, lim2)) = Toplevel.keep (fn state =>
(set_limit goals_limit lim1; set_limit ProofContext.prems_limit lim2; Toplevel.quiet := false;
@@ -280,7 +280,7 @@
fun string_of_prfs full state ms arg = with_modes ms (fn () =>
Pretty.string_of (case arg of (* FIXME context syntax!? *)
- None =>
+ NONE =>
let
val (_, (_, thm)) = Proof.get_goal state;
val {sign, prop, der = (_, prf), ...} = rep_thm thm;
@@ -289,7 +289,7 @@
ProofSyntax.pretty_proof sign
(if full then Reconstruct.reconstruct_proof sign prop prf' else prf')
end
- | Some args => Pretty.chunks
+ | SOME args => Pretty.chunks
(map (ProofSyntax.pretty_proof_of full) (IsarThy.get_thmss args state))));
fun string_of_prop state ms s =
@@ -337,7 +337,7 @@
fun present_text assert_proof present (s, pos) = Toplevel.keep' (fn int => fn state =>
(if can Toplevel.proof_of state = assert_proof then () else raise Toplevel.UNDEF;
- Context.setmp (Some (Toplevel.theory_of state)) present s;
+ Context.setmp (SOME (Toplevel.theory_of state)) present s;
OuterSyntax.check_text (s, pos) int state;
raise Toplevel.UNDEF));
--- a/src/Pure/Isar/isar_output.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/isar_output.ML Sun Feb 13 17:15:14 2005 +0100
@@ -76,14 +76,14 @@
fun command src =
let val ((name, _), pos) = Args.dest_src src in
(case Symtab.lookup (! global_commands, name) of
- None => error ("Unknown antiquotation command: " ^ quote name ^ Position.str_of pos)
- | Some f => transform_failure (curry Antiquote.ANTIQUOTE_FAIL (name, pos)) (f src))
+ NONE => error ("Unknown antiquotation command: " ^ quote name ^ Position.str_of pos)
+ | SOME f => transform_failure (curry Antiquote.ANTIQUOTE_FAIL (name, pos)) (f src))
end;
fun option (name, s) f () =
(case Symtab.lookup (! global_options, name) of
- None => error ("Unknown antiquotation option: " ^ quote name)
- | Some opt => opt s f ());
+ NONE => error ("Unknown antiquotation option: " ^ quote name)
+ | SOME opt => opt s f ());
fun options [] f = f
| options (opt :: opts) f = option opt (options opts f);
@@ -124,7 +124,7 @@
let
val ctxt0 =
if ! locale = "" then Toplevel.context_of state
- else #3 (Locale.read_context_statement (Some (! locale)) [] []
+ else #3 (Locale.read_context_statement (SOME (! locale)) [] []
(ProofContext.init (Toplevel.theory_of state)));
val (ctxt, x) = syntax scan src ctxt0;
in f src ctxt x end;
@@ -145,7 +145,7 @@
|> Symbol.source false
|> T.source false (K (keyword_lexicon, Scan.empty_lexicon)) pos
|> T.source_proper
- |> Source.source T.stopper (Scan.error (Scan.bulk (P.!!! antiq))) None
+ |> Source.source T.stopper (Scan.error (Scan.bulk (P.!!! antiq))) NONE
|> Source.exhaust;
in
@@ -201,16 +201,16 @@
| is_hidden _ = false;
fun filter_newlines toks = (case mapfilter
- (fn (tk, i) => if is_newline tk then Some tk else None) toks of
+ (fn (tk, i) => if is_newline tk then SOME tk else NONE) toks of
[] => [] | [tk] => [tk] | _ :: toks' => toks');
fun present_tokens lex (flag, toks) state state' =
- Buffer.add (case flag of None => "" | Some b => Latex.flag_markup b) o
+ Buffer.add (case flag of NONE => "" | SOME b => Latex.flag_markup b) o
((if !hide_commands andalso exists (is_hidden o fst) toks then []
else if !hide_commands andalso is_proof state then
if is_proof state' then [] else filter_newlines toks
else mapfilter (fn (tk, i) =>
- if i > ! interest_level then None else Some tk) toks)
+ if i > ! interest_level then NONE else SOME tk) toks)
|> map (apsnd (eval_antiquote lex state'))
|> Latex.output_tokens
|> Buffer.add);
@@ -234,7 +234,7 @@
(opt_newline -- check_level i) >> pair (i - 1));
val ignore_cmd = Scan.state -- ignore
- >> (fn (i, (x, pos)) => (false, (None, ((Latex.Basic x, ("", pos)), i))));
+ >> (fn (i, (x, pos)) => (false, (NONE, ((Latex.Basic x, ("", pos)), i))));
val is_improper = not o (T.is_proper orf T.is_begin_ignore orf T.is_end_ignore);
@@ -249,7 +249,7 @@
val stopper =
- ((false, (None, ((Latex.Basic (#1 T.stopper), ("", Position.none)), 0))),
+ ((false, (NONE, ((Latex.Basic (#1 T.stopper), ("", Position.none)), 0))),
fn (_, (_, ((Latex.Basic x, _), _))) => (#2 T.stopper x) | _ => false);
in
@@ -259,15 +259,15 @@
val text = P.position P.text;
val token = Scan.depend (fn i =>
(improper |-- markup Markup -- P.!!!! (improper |-- text --| improper_end)
- >> (fn (x, y) => (true, (Some true, ((Latex.Markup (T.val_of x), y), i)))) ||
+ >> (fn (x, y) => (true, (SOME true, ((Latex.Markup (T.val_of x), y), i)))) ||
improper |-- markup MarkupEnv -- P.!!!! (improper |-- text --| improper_end)
- >> (fn (x, y) => (true, (Some true, ((Latex.MarkupEnv (T.val_of x), y), i)))) ||
+ >> (fn (x, y) => (true, (SOME true, ((Latex.MarkupEnv (T.val_of x), y), i)))) ||
(P.$$$ "--" |-- P.!!!! (improper |-- text))
- >> (fn y => (false, (None, ((Latex.Markup "cmt", y), i)))) ||
+ >> (fn y => (false, (NONE, ((Latex.Markup "cmt", y), i)))) ||
(improper -- markup Verbatim) |-- P.!!!! (improper |-- text --| improper_end)
- >> (fn y => (true, (None, ((Latex.Verbatim, y), i)))) ||
+ >> (fn y => (true, (NONE, ((Latex.Verbatim, y), i)))) ||
P.position (Scan.one T.not_eof)
- >> (fn (x, pos) => (T.is_kind T.Command x, (Some false, ((Latex.Basic x, ("", pos)), i)))))
+ >> (fn (x, pos) => (T.is_kind T.Command x, (SOME false, ((Latex.Basic x, ("", pos)), i)))))
>> pair i);
val body = Scan.any (not o fst andf not o #2 stopper);
@@ -277,8 +277,8 @@
val cmds =
src
|> Source.filter (not o T.is_semicolon)
- |> Source.source' 0 T.stopper (Scan.error (Scan.bulk (ignore_cmd || token))) None
- |> Source.source stopper (Scan.error (Scan.bulk section)) None
+ |> Source.source' 0 T.stopper (Scan.error (Scan.bulk (ignore_cmd || token))) NONE
+ |> Source.source stopper (Scan.error (Scan.bulk section)) NONE
|> Source.exhaust;
in
if length cmds = length trs then
--- a/src/Pure/Isar/isar_syn.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/isar_syn.ML Sun Feb 13 17:15:14 2005 +0100
@@ -186,9 +186,9 @@
Scan.optional ((P.$$$ "(" -- P.$$$ "structure") |-- P.!!! (vars --| P.$$$ ")")) [];
val constdecl =
- (P.name --| P.$$$ "where") >> (fn x => (x, None, Syntax.NoSyn)) ||
- P.name -- (P.$$$ "::" |-- P.!!! P.typ >> Some) -- P.opt_mixfix' >> P.triple1 ||
- P.name -- (P.mixfix' >> pair None) >> P.triple2;
+ (P.name --| P.$$$ "where") >> (fn x => (x, NONE, Syntax.NoSyn)) ||
+ P.name -- (P.$$$ "::" |-- P.!!! P.typ >> SOME) -- P.opt_mixfix' >> P.triple1 ||
+ P.name -- (P.mixfix' >> pair NONE) >> P.triple2;
val constdef = Scan.option constdecl -- (P.opt_thm_name ":" -- P.prop);
--- a/src/Pure/Isar/isar_thy.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/isar_thy.ML Sun Feb 13 17:15:14 2005 +0100
@@ -74,7 +74,7 @@
-> Proof.context attribute Locale.element_i Locale.elem_expr list
-> ((bstring * Proof.context attribute list) * (term * (term list * term list)) list) list
-> bool -> theory -> ProofHistory.T
- val smart_multi_theorem: string -> xstring Library.option
+ val smart_multi_theorem: string -> xstring option
-> bstring * Args.src list -> Args.src Locale.element Locale.elem_expr list
-> ((bstring * Args.src list) * (string * (string list * string list)) list) list
-> bool -> theory -> ProofHistory.T
@@ -172,7 +172,7 @@
fun gen_hide intern (kind, xnames) thy =
(case assoc (kinds, kind) of
- Some check =>
+ SOME check =>
let
val sg = Theory.sign_of thy;
val names = map (intern sg kind) xnames;
@@ -181,7 +181,7 @@
if null bads then Theory.hide_space_i true (kind, names) thy
else error ("Attempt to hide undeclared item(s): " ^ commas_quote bads)
end
- | None => error ("Bad name space specification: " ^ quote kind));
+ | NONE => error ("Bad name space specification: " ^ quote kind));
in
@@ -223,7 +223,7 @@
fun present_thmss kind (thy, named_thmss) =
(conditional (kind <> "" andalso kind <> Drule.internalK) (fn () =>
- Context.setmp (Some thy) (Present.results (kind ^ "s")) named_thmss);
+ Context.setmp (SOME thy) (Present.results (kind ^ "s")) named_thmss);
(thy, named_thmss));
in
@@ -241,8 +241,8 @@
fun smart_theorems k opt_loc args thy = thy
|> (case opt_loc of
- None => PureThy.note_thmss (Drule.kind k) (global_attribs thy args)
- | Some loc => Locale.note_thmss k loc (local_attribs thy args))
+ NONE => PureThy.note_thmss (Drule.kind k) (global_attribs thy args)
+ | SOME loc => Locale.note_thmss k loc (local_attribs thy args))
|> present_thmss k;
fun declare_theorems opt_loc args = #1 o smart_theorems "" opt_loc [(("", []), args)];
@@ -332,20 +332,20 @@
fun check_goal false state = state
| check_goal true state =
let
- val rule = ref (None: thm option);
+ val rule = ref (NONE: thm option);
fun fail exn =
(["Problem! Local statement will fail to solve any pending goal"] @
- (case exn of None => [] | Some e => [Toplevel.exn_message e]) @
- (case ! rule of None => [] | Some thm =>
+ (case exn of NONE => [] | SOME e => [Toplevel.exn_message e]) @
+ (case ! rule of NONE => [] | SOME thm =>
[Pretty.string_of (pretty_rule "Failed attempt" (Proof.context_of state) thm)]))
|> cat_lines |> error;
val check =
(fn () => Seq.pull (SkipProof.local_skip_proof (K (K ()),
- fn _ => fn thm => rule := Some thm) true state))
+ fn _ => fn thm => rule := SOME thm) true state))
|> setmp proofs 0
|> transform_error;
- val result = (check (), None) handle Interrupt => raise Interrupt | e => (None, Some e);
- in (case result of (Some _, None) => () | (_, exn) => fail exn); state end;
+ val result = (check (), NONE) handle Interrupt => raise Interrupt | e => (NONE, SOME e);
+ in (case result of (SOME _, NONE) => () | (_, exn) => fail exn); state end;
end;
@@ -371,27 +371,27 @@
in
fun multi_theorem k a elems concl int thy =
- global_statement (Proof.multi_theorem k None (apsnd (map (Attrib.global_attribute thy)) a)
+ global_statement (Proof.multi_theorem k NONE (apsnd (map (Attrib.global_attribute thy)) a)
(map (Locale.map_attrib_elem_or_expr (Attrib.local_attribute thy)) elems)) concl int thy;
-fun multi_theorem_i k a = global_statement_i o Proof.multi_theorem_i k None a;
+fun multi_theorem_i k a = global_statement_i o Proof.multi_theorem_i k NONE a;
fun locale_multi_theorem k locale (name, atts) elems concl int thy =
global_statement (Proof.multi_theorem k
- (Some (locale, (map (Attrib.local_attribute thy) atts,
+ (SOME (locale, (map (Attrib.local_attribute thy) atts,
map (map (Attrib.local_attribute thy) o snd o fst) concl)))
(name, []) (map (Locale.map_attrib_elem_or_expr (Attrib.local_attribute thy)) elems))
(map (apfst (apsnd (K []))) concl) int thy;
fun locale_multi_theorem_i k locale (name, atts) elems concl =
- global_statement_i (Proof.multi_theorem_i k (Some (locale, (atts, map (snd o fst) concl)))
+ global_statement_i (Proof.multi_theorem_i k (SOME (locale, (atts, map (snd o fst) concl)))
(name, []) elems) (map (apfst (apsnd (K []))) concl);
fun theorem k (a, t) = multi_theorem k a [] [(("", []), [t])];
fun theorem_i k (a, t) = multi_theorem_i k a [] [(("", []), [t])];
-fun smart_multi_theorem k None a elems = multi_theorem k a elems
- | smart_multi_theorem k (Some locale) a elems = locale_multi_theorem k locale a elems;
+fun smart_multi_theorem k NONE a elems = multi_theorem k a elems
+ | smart_multi_theorem k (SOME locale) a elems = locale_multi_theorem k locale a elems;
val assume = local_statement Proof.assume I;
val assume_i = local_statement_i Proof.assume_i I;
@@ -466,7 +466,7 @@
in
if kind = "" orelse kind = Drule.internalK then ()
else (print_result (Proof.context_of state) ((kind, name), res);
- Context.setmp (Some thy) (Present.results kind) res);
+ Context.setmp (SOME thy) (Present.results kind) res);
thy
end);
@@ -612,9 +612,9 @@
(* context switch *)
fun fetch_context f x =
- (case Context.pass None f x of
- ((), None) => raise Toplevel.UNDEF
- | ((), Some thy) => thy);
+ (case Context.pass NONE f x of
+ ((), NONE) => raise Toplevel.UNDEF
+ | ((), SOME thy) => thy);
fun init_context f x = Toplevel.init_theory (fn _ => fetch_context f x) (K ()) (K ());
--- a/src/Pure/Isar/locale.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/locale.ML Sun Feb 13 17:15:14 2005 +0100
@@ -67,7 +67,7 @@
val add_locale: bool -> bstring -> expr -> context attribute element list -> theory -> theory
val add_locale_i: bool -> bstring -> expr -> context attribute element_i list
-> theory -> theory
- val smart_note_thmss: string -> (string * 'a) Library.option ->
+ val smart_note_thmss: string -> (string * 'a) option ->
((bstring * theory attribute list) * (thm list * theory attribute list) list) list ->
theory -> theory * (bstring * thm list) list
val note_thmss: string -> xstring ->
@@ -138,7 +138,7 @@
(*joining of locale elements: only facts may be added later!*)
fun join ({predicate, import, elems, params}: locale, {elems = elems', ...}: locale) =
- Some {predicate = predicate, import = import, elems = gen_merge_lists eq_snd elems elems',
+ SOME {predicate = predicate, import = import, elems = gen_merge_lists eq_snd elems elems',
params = params};
fun merge ((space1, locs1), (space2, locs2)) =
(NameSpace.merge (space1, space2), Symtab.join join (locs1, locs2));
@@ -165,8 +165,8 @@
fun the_locale thy name =
(case get_locale thy name of
- Some loc => loc
- | None => error ("Unknown locale " ^ quote name));
+ SOME loc => loc
+ | NONE => error ("Unknown locale " ^ quote name));
(* import hierarchy
@@ -178,8 +178,8 @@
val sign = sign_of thy;
fun imps (Locale name) low = (name = low) orelse
(case get_locale thy name of
- None => false
- | Some {import, ...} => imps import low)
+ NONE => false
+ | SOME {import, ...} => imps import low)
| imps (Rename (expr, _)) low = imps expr low
| imps (Merge es) low = exists (fn e => imps e low) es;
in
@@ -239,7 +239,7 @@
fun rename_elem ren (Fixes fixes) = Fixes (fixes |> map (fn (x, T, mx) =>
let val x' = rename ren x in
if x = x' then (x, T, mx)
- else (x', T, if mx = None then mx else Some Syntax.NoSyn) (*drop syntax*)
+ else (x', T, if mx = NONE then mx else SOME Syntax.NoSyn) (*drop syntax*)
end))
| rename_elem ren (Assumes asms) = Assumes (map (apsnd (map (fn (t, (ps, qs)) =>
(rename_term ren t, (map (rename_term ren) ps, map (rename_term ren) qs))))) asms)
@@ -317,14 +317,14 @@
fun unify_frozen ctxt maxidx Ts Us =
let
- fun paramify (i, None) = (i, None)
- | paramify (i, Some T) = apsnd Some (TypeInfer.paramify_dummies (i, T));
+ fun paramify (i, NONE) = (i, NONE)
+ | paramify (i, SOME T) = apsnd SOME (TypeInfer.paramify_dummies (i, T));
val (maxidx', Ts') = foldl_map paramify (maxidx, Ts);
val (maxidx'', Us') = foldl_map paramify (maxidx', Us);
val tsig = Sign.tsig_of (ProofContext.sign_of ctxt);
- fun unify (env, (Some T, Some U)) = (Type.unify tsig env (U, T)
+ fun unify (env, (SOME T, SOME U)) = (Type.unify tsig env (U, T)
handle Type.TUNIFY =>
raise TYPE ("unify_frozen: failed to unify types", [U, T], []))
| unify (env, _) = env;
@@ -337,8 +337,8 @@
fun params_of' elemss = gen_distinct eq_fst (flat (map (snd o fst o fst) elemss));
(* CB: param_types has the following type:
- ('a * 'b Library.option) list -> ('a * 'b) list *)
-fun param_types ps = mapfilter (fn (_, None) => None | (x, Some T) => Some (x, T)) ps;
+ ('a * 'b option) list -> ('a * 'b) list *)
+fun param_types ps = mapfilter (fn (_, NONE) => NONE | (x, SOME T) => SOME (x, T)) ps;
(* flatten expressions *)
@@ -357,15 +357,15 @@
|> Symtab.make_multi |> Symtab.dest;
in
(case find_first (fn (_, ids) => length ids > 1) param_decls of
- Some (q, ids) => err_in_locale ctxt ("Multiple declaration of parameter " ^ quote q)
+ SOME (q, ids) => err_in_locale ctxt ("Multiple declaration of parameter " ^ quote q)
(map (apsnd (map fst)) ids)
- | None => map (apfst (apfst (apsnd #1))) elemss)
+ | NONE => map (apfst (apfst (apsnd #1))) elemss)
end;
(* CB: unify_parms has the following type:
ProofContext.context ->
(string * Term.typ) list ->
- (string * Term.typ Library.option) list list ->
+ (string * Term.typ option) list list ->
((string * Term.sort) * Term.typ) list list *)
fun unify_parms ctxt fixed_parms raw_parmss =
@@ -398,7 +398,7 @@
foldr Term.add_typ_tfrees (mapfilter snd ps, [])
|> mapfilter (fn (a, S) =>
let val T = Envir.norm_type unifier' (TVar ((a, i), S))
- in if T = TFree (a, S) then None else Some ((a, S), T) end)
+ in if T = TFree (a, S) then NONE else SOME ((a, S), T) end)
in map inst_parms idx_parmss end;
in
@@ -429,11 +429,11 @@
let
val thy = ProofContext.theory_of ctxt;
- fun renaming (Some x :: xs) (y :: ys) = (y, x) :: renaming xs ys
- | renaming (None :: xs) (y :: ys) = renaming xs ys
+ fun renaming (SOME x :: xs) (y :: ys) = (y, x) :: renaming xs ys
+ | renaming (NONE :: xs) (y :: ys) = renaming xs ys
| renaming [] _ = []
| renaming xs [] = raise ERROR_MESSAGE ("Too many arguments in renaming: " ^
- commas (map (fn None => "_" | Some x => quote x) xs));
+ commas (map (fn NONE => "_" | SOME x => quote x) xs));
fun rename_parms top ren ((name, ps), (parms, axs)) =
let val ps' = map (rename ren) ps in
@@ -460,7 +460,7 @@
(foldl_map (fn (axs, ((name, parms), _)) => let
val {elems, ...} = the_locale thy name;
val ts = flat (mapfilter (fn (Assumes asms, _) =>
- Some (flat (map (map #1 o #2) asms)) | _ => None) elems);
+ SOME (flat (map (map #1 o #2) asms)) | _ => NONE) elems);
val (axs1, axs2) = splitAt (length ts, axs);
in (axs2, ((name, parms), (ps, axs1))) end) (axioms, ids''))
else ids'';
@@ -508,7 +508,7 @@
elemss;
fun inst_ax th = let
val {hyps, prop, ...} = Thm.rep_thm th;
- val ps = map (apsnd Some) (foldl Term.add_frees ([], prop :: hyps));
+ val ps = map (apsnd SOME) (foldl Term.add_frees ([], prop :: hyps));
val [env] = unify_parms ctxt all_params [ps];
val th' = inst_thm ctxt env th;
in th' end;
@@ -653,13 +653,13 @@
information (for expr it is the identifier, where parameters additionially
contain type information (extracted from the locale record), for a Fixes
element, it is an identifier with name = "" and parameters with type
- information None, for other elements it is simply ("", [])).
+ information NONE, for other elements it is simply ("", [])).
The implementation of activate_facts relies on identifier names being
empty strings for external elements.
TODO: correct this comment wrt axioms. *)
fun flatten _ (ids, Elem (Fixes fixes)) =
- (ids, [((("", map (rpair None o #1) fixes), []), Ext (Fixes fixes))])
+ (ids, [((("", map (rpair NONE o #1) fixes), []), Ext (Fixes fixes))])
| flatten _ (ids, Elem elem) = (ids, [((("", []), []), Ext elem)])
| flatten (ctxt, prep_expr) (ids, Expr expr) =
apsnd (map (apsnd Int)) (flatten_expr ctxt (ids, prep_expr expr));
@@ -700,7 +700,7 @@
ProofContext.declare_terms (map Free fixed_params) ctxt;
val int_elemss =
raw_elemss
- |> mapfilter (fn (id, Int es) => Some (id, es) | _ => None)
+ |> mapfilter (fn (id, Int es) => SOME (id, es) | _ => NONE)
|> unify_elemss ctxt_with_fixed fixed_params;
val (_, raw_elemss') =
foldl_map (fn ((_, es) :: elemss, (id, Int _)) => (elemss, (id, Int es)) | x => x)
@@ -789,7 +789,7 @@
| finish_ext_elem _ _ (Notes facts, _) = Notes facts;
(* CB: finish_parms introduces type info from parms to identifiers *)
-(* CB: only needed for types that have been None so far???
+(* CB: only needed for types that have been NONE so far???
If so, which are these??? *)
fun finish_parms parms (((name, ps), axs), elems) =
@@ -958,8 +958,8 @@
val thy = ProofContext.theory_of ctxt;
val locale = apsome (prep_locale (Theory.sign_of thy)) raw_locale;
val (target_stmt, fixed_params, import) =
- (case locale of None => ([], [], empty)
- | Some name =>
+ (case locale of NONE => ([], [], empty)
+ | SOME name =>
let val {predicate = (stmt, _), params = (ps, _), ...} =
the_locale thy name
in (stmt, param_types ps, Locale name) end);
@@ -1008,20 +1008,20 @@
(** analyse the instantiation theorem inst **)
val inst = case raw_inst of
- None => if null ints
- then None
+ NONE => if null ints
+ then NONE
else error "Locale has assumptions but no chained fact was found"
- | Some [] => if null ints
- then None
+ | SOME [] => if null ints
+ then NONE
else error "Locale has assumptions but no chained fact was found"
- | Some [thm] => if null ints
- then (warning "Locale has no assumptions: fact ignored"; None)
- else Some thm
- | Some _ => error "Multiple facts are not allowed";
+ | SOME [thm] => if null ints
+ then (warning "Locale has no assumptions: fact ignored"; NONE)
+ else SOME thm
+ | SOME _ => error "Multiple facts are not allowed";
val args = case inst of
- None => []
- | Some thm => thm |> prop_of |> ObjectLogic.drop_judgment sign
+ NONE => []
+ | SOME thm => thm |> prop_of |> ObjectLogic.drop_judgment sign
|> Term.strip_comb
|>> (fn t as (Const (s, _)) => if (intern sign loc_name = s)
then t
@@ -1097,15 +1097,15 @@
val hyps = #hyps (rep_thm thm);
val ass = map (fn ax => (prop_of ax, ax)) axioms;
val axs' = foldl (fn (axs, hyp) =>
- (case gen_assoc (op aconv) (ass, hyp) of None => axs
- | Some ax => axs @ [ax])) ([], hyps);
+ (case gen_assoc (op aconv) (ass, hyp) of NONE => axs
+ | SOME ax => axs @ [ax])) ([], hyps);
val thm' = Drule.satisfy_hyps axs' thm;
(* instantiate types *)
val thm'' = inst_thm ctxt Tenv' thm';
(* substitute arguments and discharge hypotheses *)
val thm''' = case inst of
- None => thm''
- | Some inst_thm => let
+ NONE => thm''
+ | SOME inst_thm => let
val hyps = #hyps (rep_thm thm'');
val th = thm'' |> implies_intr_hyps
|> forall_intr_list (cparams' @ cdefined')
@@ -1168,11 +1168,11 @@
val prt_thm = Pretty.quote o ProofContext.pretty_thm ctxt;
fun prt_syn syn =
- let val s = (case syn of None => "(structure)" | Some mx => Syntax.string_of_mixfix mx)
+ let val s = (case syn of NONE => "(structure)" | SOME mx => Syntax.string_of_mixfix mx)
in if s = "" then [] else [Pretty.brk 2, Pretty.str s] end;
- fun prt_fix (x, Some T, syn) = Pretty.block (Pretty.str (x ^ " ::") :: Pretty.brk 1 ::
+ fun prt_fix (x, SOME T, syn) = Pretty.block (Pretty.str (x ^ " ::") :: Pretty.brk 1 ::
prt_typ T :: Pretty.brk 1 :: prt_syn syn)
- | prt_fix (x, None, syn) = Pretty.block (Pretty.str x :: Pretty.brk 1 :: prt_syn syn);
+ | prt_fix (x, NONE, syn) = Pretty.block (Pretty.str x :: Pretty.brk 1 :: prt_syn syn);
fun prt_name "" = [Pretty.brk 1]
| prt_name name = [Pretty.str (ProofContext.cond_extern ctxt name ^ ":"), Pretty.brk 1];
@@ -1210,8 +1210,8 @@
|>> hide_bound_names (map (#1 o #1) args)
|>> Theory.parent_path;
-fun smart_note_thmss kind None = PureThy.note_thmss_i (Drule.kind kind)
- | smart_note_thmss kind (Some (loc, _)) = note_thmss_qualified kind loc;
+fun smart_note_thmss kind NONE = PureThy.note_thmss_i (Drule.kind kind)
+ | smart_note_thmss kind (SOME (loc, _)) = note_thmss_qualified kind loc;
(* CB: only used in Proof.finish_global. *)
end;
@@ -1231,7 +1231,7 @@
val thy_ctxt = ProofContext.init thy;
val loc = prep_locale (Theory.sign_of thy) raw_loc;
val (_, (stmt, _), loc_ctxt, _, _) =
- cert_context_statement (Some loc) [] [] thy_ctxt;
+ cert_context_statement (SOME loc) [] [] thy_ctxt;
val args = map (apsnd (map (apfst (get_thms loc_ctxt)))) raw_args;
val export = ProofContext.export_standard stmt loc_ctxt thy_ctxt;
val results = map (map export o #2) (#2 (ProofContext.note_thmss_i args loc_ctxt));
@@ -1367,10 +1367,10 @@
fun define_preds bname (parms, ((exts, exts'), (ints, ints')), defs) elemss thy =
let
val (thy', (elemss', more_ts)) =
- if Library.null exts then (thy, (elemss, []))
+ if null exts then (thy, (elemss, []))
else
let
- val aname = if Library.null ints then bname else bname ^ "_" ^ axiomsN;
+ val aname = if null ints then bname else bname ^ "_" ^ axiomsN;
val (def_thy, (statement, intro, axioms)) =
thy |> def_pred aname parms defs exts exts';
val elemss' = change_elemss axioms elemss @
@@ -1381,7 +1381,7 @@
|> #1 |> rpair (elemss', [statement])
end;
val (thy'', predicate) =
- if Library.null ints then (thy', ([], []))
+ if null ints then (thy', ([], []))
else
let
val (def_thy, (statement, intro, axioms)) =
@@ -1424,7 +1424,7 @@
fun axiomify axioms elemss =
(axioms, elemss) |> foldl_map (fn (axs, (id, elems)) => let
val ts = flat (mapfilter (fn (Assumes asms) =>
- Some (flat (map (map #1 o #2) asms)) | _ => None) elems);
+ SOME (flat (map (map #1 o #2) asms)) | _ => NONE) elems);
val (axs1, axs2) = splitAt (length ts, axs);
in (axs2, ((id, axs1), elems)) end)
|> snd;
@@ -1455,7 +1455,7 @@
val setup =
[LocalesData.init,
- add_locale_i true "var" empty [Fixes [(Syntax.internal "x", None, Some Syntax.NoSyn)]],
- add_locale_i true "struct" empty [Fixes [(Syntax.internal "S", None, None)]]];
+ add_locale_i true "var" empty [Fixes [(Syntax.internal "x", NONE, SOME Syntax.NoSyn)]],
+ add_locale_i true "struct" empty [Fixes [(Syntax.internal "S", NONE, NONE)]]];
end;
--- a/src/Pure/Isar/method.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/method.ML Sun Feb 13 17:15:14 2005 +0100
@@ -349,7 +349,7 @@
("No such variable in theorem: " ^ Syntax.string_of_vname xi);
val (rtypes, rsorts) = types_sorts thm;
fun readT (xi, s) =
- let val S = case rsorts xi of Some S => S | None => absent xi;
+ let val S = case rsorts xi of SOME S => S | NONE => absent xi;
val T = Sign.read_typ (sign, sorts) s;
in if Sign.typ_instance sign (T, TVar (xi, S)) then (xi, T)
else error
@@ -359,8 +359,8 @@
(* Preprocess rule: extract vars and their types, apply Tinsts *)
fun get_typ xi =
(case rtypes xi of
- Some T => typ_subst_TVars Tinsts_env T
- | None => absent xi);
+ SOME T => typ_subst_TVars Tinsts_env T
+ | NONE => absent xi);
val (xis, ss) = Library.split_list tinsts;
val Ts = map get_typ xis;
val (_, _, Bi, _) = dest_state(st,i)
@@ -370,7 +370,7 @@
(* as they are printed: bound variables with *)
(* the same name are renamed during printing *)
fun types' (a, ~1) = (case assoc (params, a) of
- None => types (a, ~1)
+ NONE => types (a, ~1)
| some => some)
| types' xi = types xi;
fun internal x = is_some (types' (x, ~1));
@@ -487,12 +487,12 @@
fun tactic txt ctxt = METHOD (fn facts =>
(Context.use_mltext
("let fun tactic (ctxt: PureIsar.Proof.context) (facts: thm list) : tactic = \
- \let val thm = PureIsar.ProofContext.get_thm_closure ctxt o rpair None\n\
- \ and thms = PureIsar.ProofContext.get_thms_closure ctxt o rpair None in\n"
+ \let val thm = PureIsar.ProofContext.get_thm_closure ctxt o rpair NONE\n\
+ \ and thms = PureIsar.ProofContext.get_thms_closure ctxt o rpair NONE in\n"
^ txt ^
"\nend in PureIsar.Method.set_tactic tactic end")
- false None;
- Context.setmp (Some (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts));
+ false NONE;
+ Context.setmp (SOME (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts));
@@ -543,8 +543,8 @@
val name = NameSpace.intern space raw_name;
in
(case Symtab.lookup (meths, name) of
- None => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
- | Some ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src))
+ NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
+ | SOME ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src))
end;
in meth end;
@@ -633,7 +633,7 @@
val ruleN = "rule";
fun modifier name kind kind' att =
- Args.$$$ name |-- (kind >> K None || kind' |-- Args.nat --| Args.colon >> Some)
+ Args.$$$ name |-- (kind >> K NONE || kind' |-- Args.nat --| Args.colon >> SOME)
>> (pair (I: Proof.context -> Proof.context) o att);
val rules_modifiers =
@@ -726,8 +726,8 @@
fun close_text asm = Basic (fn ctxt => METHOD (K
(FILTER Thm.no_prems ((if asm then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac))));
-fun finish_text asm None = close_text asm
- | finish_text asm (Some txt) = Then [txt, close_text asm];
+fun finish_text asm NONE = close_text asm
+ | finish_text asm (SOME txt) = Then [txt, close_text asm];
fun proof opt_text state =
state
@@ -738,10 +738,10 @@
fun local_qed asm opt_text = Proof.local_qed (refine (finish_text asm opt_text));
fun local_terminal_proof (text, opt_text) pr =
- Seq.THEN (proof (Some text), local_qed true opt_text pr);
-val local_default_proof = local_terminal_proof (default_text, None);
-val local_immediate_proof = local_terminal_proof (this_text, None);
-fun local_done_proof pr = Seq.THEN (proof (Some done_text), local_qed false None pr);
+ Seq.THEN (proof (SOME text), local_qed true opt_text pr);
+val local_default_proof = local_terminal_proof (default_text, NONE);
+val local_immediate_proof = local_terminal_proof (this_text, NONE);
+fun local_done_proof pr = Seq.THEN (proof (SOME done_text), local_qed false NONE pr);
fun global_qeds asm opt_text = Proof.global_qed (refine (finish_text asm opt_text));
@@ -754,16 +754,16 @@
fun global_term_proof asm (text, opt_text) state =
state
- |> proof (Some text)
+ |> proof (SOME text)
|> Proof.check_result "Terminal proof method failed" state
|> (Seq.flat o Seq.map (global_qeds asm opt_text))
|> Proof.check_result "Failed to finish proof (after successful terminal method)" state
|> Seq.hd;
val global_terminal_proof = global_term_proof true;
-val global_default_proof = global_terminal_proof (default_text, None);
-val global_immediate_proof = global_terminal_proof (this_text, None);
-val global_done_proof = global_term_proof false (done_text, None);
+val global_default_proof = global_terminal_proof (default_text, NONE);
+val global_immediate_proof = global_terminal_proof (this_text, NONE);
+val global_done_proof = global_term_proof false (done_text, NONE);
@@ -814,7 +814,7 @@
val setup =
[MethodsData.init, add_methods pure_methods,
- (#1 o PureThy.add_thms [(("", Drule.equal_intr_rule), [ContextRules.intro_query_global None])])];
+ (#1 o PureThy.add_thms [(("", Drule.equal_intr_rule), [ContextRules.intro_query_global NONE])])];
end;
--- a/src/Pure/Isar/object_logic.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/object_logic.ML Sun Feb 13 17:15:14 2005 +0100
@@ -42,12 +42,12 @@
val name = "Pure/object-logic";
type T = string option * (thm list * thm list);
- val empty = (None, ([], [Drule.norm_hhf_eq]));
+ val empty = (NONE, ([], [Drule.norm_hhf_eq]));
val copy = I;
val prep_ext = I;
- fun merge_judgment (Some x, Some y) =
- if x = y then Some x else error "Attempt to merge different object-logics"
+ fun merge_judgment (SOME x, SOME y) =
+ if x = y then SOME x else error "Attempt to merge different object-logics"
| merge_judgment (j1, j2) = if is_some j1 then j1 else j2;
fun merge ((judgment1, (atomize1, rulify1)), (judgment2, (atomize2, rulify2))) =
@@ -66,14 +66,14 @@
local
-fun new_judgment name (None, rules) = (Some name, rules)
- | new_judgment _ (Some _, _) = error "Attempt to redeclare object-logic judgment";
+fun new_judgment name (NONE, rules) = (SOME name, rules)
+ | new_judgment _ (SOME _, _) = error "Attempt to redeclare object-logic judgment";
fun add_final name thy =
let
val typ = case Sign.const_type (sign_of thy) name of
- Some T => T
- | None => error "Internal error in ObjectLogic.gen_add_judgment";
+ SOME T => T
+ | NONE => error "Internal error in ObjectLogic.gen_add_judgment";
in
Theory.add_finals_i false [Const(name,typ)] thy
end;
@@ -100,7 +100,7 @@
fun judgment_name sg =
(case ObjectLogicData.get_sg sg of
- (Some name, _) => name
+ (SOME name, _) => name
| _ => raise TERM ("Unknown object-logic judgment", []));
fun is_judgment sg (Const (c, _) $ _) = c = judgment_name sg
@@ -166,7 +166,7 @@
rewrite_goal_tac (get_atomize (Thm.sign_of_thm st)) i st;
fun atomize_goal i st =
- (case Seq.pull (atomize_tac i st) of None => st | Some (st', _) => st');
+ (case Seq.pull (atomize_tac i st) of NONE => st | SOME (st', _) => st');
(* rulify *)
--- a/src/Pure/Isar/obtain.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/obtain.ML Sun Feb 13 17:15:14 2005 +0100
@@ -97,7 +97,7 @@
val raw_parms = map occs_var xs;
val parms = mapfilter I raw_parms;
val parm_names =
- mapfilter (fn (Some (Free a), x) => Some (a, x) | _ => None) (raw_parms ~~ xs);
+ mapfilter (fn (SOME (Free a), x) => SOME (a, x) | _ => NONE) (raw_parms ~~ xs);
val that_prop =
Term.list_all_free (map #1 parm_names, Logic.list_implies (asm_props, bound_thesis))
@@ -107,7 +107,7 @@
Term.list_all_free ([bound_thesis_var], Logic.mk_implies (that_prop, bound_thesis)));
fun after_qed st = st
- |> Method.local_qed false None print
+ |> Method.local_qed false NONE print
|> Seq.map (fn st' => st'
|> Proof.fix_i vars
|> Proof.assm_i (export_obtain state parms (Proof.the_fact st')) asms);
@@ -115,9 +115,9 @@
state
|> Proof.enter_forward
|> Proof.have_i Seq.single true [(("", []), [(obtain_prop, ([], []))])]
- |> Method.proof (Some (Method.Basic (K Method.succeed))) |> Seq.hd
- |> Proof.fix_i [([thesisN], None)]
- |> Proof.assume_i [((thatN, [ContextRules.intro_query_local None]), [(that_prop, ([], []))])]
+ |> Method.proof (SOME (Method.Basic (K Method.succeed))) |> Seq.hd
+ |> Proof.fix_i [([thesisN], NONE)]
+ |> Proof.assume_i [((thatN, [ContextRules.intro_query_local NONE]), [(that_prop, ([], []))])]
|> (fn state' =>
state'
|> Proof.from_facts chain_facts
--- a/src/Pure/Isar/outer_lex.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/outer_lex.ML Sun Feb 13 17:15:14 2005 +0100
@@ -186,8 +186,8 @@
fun is_symid str =
(case try Symbol.explode str of
- Some [s] => Symbol.is_symbolic s orelse is_sym_char s
- | Some ss => forall is_sym_char ss
+ SOME [s] => Symbol.is_symbolic s orelse is_sym_char s
+ | SOME ss => forall is_sym_char ss
| _ => false);
val is_sid = is_symid orf Syntax.is_identifier;
@@ -282,7 +282,7 @@
fun source do_recover get_lex pos src =
Source.source' pos Symbol.stopper (Scan.bulk (fn xs => scan (get_lex ()) xs))
- (if do_recover then Some recover else None) src;
+ (if do_recover then SOME recover else NONE) src;
fun source_proper src = src |> Source.filter is_proper;
--- a/src/Pure/Isar/outer_parse.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/outer_parse.ML Sun Feb 13 17:15:14 2005 +0100
@@ -102,8 +102,8 @@
fun get_pos [] = " (past end-of-file!)"
| get_pos (tok :: _) = T.pos_of tok;
- fun err (toks, None) = kind ^ get_pos toks
- | err (toks, Some msg) = kind ^ get_pos toks ^ ": " ^ msg;
+ fun err (toks, NONE) = kind ^ get_pos toks
+ | err (toks, SOME msg) = kind ^ get_pos toks ^ ": " ^ msg;
in Scan.!! err scan end;
fun !!! scan = cut "Outer syntax error" scan;
@@ -179,7 +179,7 @@
val text = group "text" (short_ident || long_ident || sym_ident || string || number || verbatim);
val path = group "file name/path specification" name >> Path.unpack;
-val uname = underscore >> K None || name >> Some;
+val uname = underscore >> K NONE || name >> SOME;
(* sorts and arities *)
@@ -308,7 +308,7 @@
local
-val loc_mixfix = $$$ "(" -- $$$ "structure" -- !!! ($$$ ")") >> K None || opt_mixfix >> Some;
+val loc_mixfix = $$$ "(" -- $$$ "structure" -- !!! ($$$ ")") >> K NONE || opt_mixfix >> SOME;
val loc_keyword = $$$ "fixes" || $$$ "assumes" || $$$ "defines" || $$$ "notes" || $$$ "includes";
val loc_element =
--- a/src/Pure/Isar/outer_syntax.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/outer_syntax.ML Sun Feb 13 17:15:14 2005 +0100
@@ -132,18 +132,18 @@
fun body cmd trc (name, _) =
(case cmd name of
- Some (int_only, parse) =>
+ SOME (int_only, parse) =>
P.!!! (Scan.prompt (name ^ "# ") (trace trc parse >> pair int_only))
- | None => sys_error ("no parser for outer syntax command " ^ quote name));
+ | NONE => sys_error ("no parser for outer syntax command " ^ quote name));
in
fun command do_terminate do_trace cmd =
- P.semicolon >> K None ||
- P.sync >> K None ||
+ P.semicolon >> K NONE ||
+ P.sync >> K NONE ||
(P.position P.command :-- body cmd do_trace) --| terminate do_terminate
>> (fn ((name, pos), (int_only, f)) =>
- Some (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
+ SOME (Toplevel.empty |> Toplevel.name name |> Toplevel.position pos |>
Toplevel.interactive int_only |> f));
end;
@@ -167,7 +167,7 @@
| bads => error ("Clash of outer syntax commands and keywords: " ^ commas_quote bads))
end;
-fun get_markup (ms, (name, (_, Some m))) = (name, m) :: ms
+fun get_markup (ms, (name, (_, SOME m))) = (name, m) :: ms
| get_markup (ms, _) = ms;
fun make_markups () = global_markups := Symtab.foldl get_markup ([], ! global_parsers);
@@ -186,7 +186,7 @@
fun get_parser () = apsome (#2 o #1) o curry Symtab.lookup (! global_parsers);
fun is_markup kind name =
- (case assoc (! global_markups, name) of Some k => k = kind | None => false);
+ (case assoc (! global_markups, name) of SOME k => k = kind | NONE => false);
fun markup kind = Scan.one (T.is_kind T.Command andf is_markup kind o T.val_of);
@@ -241,16 +241,16 @@
let
val no_terminator =
Scan.unless P.semicolon (Scan.one (T.not_sync andf T.not_eof));
- fun recover x = (Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [None]) x;
+ fun recover x = (Scan.prompt "recover# " (Scan.repeat no_terminator) >> K [NONE]) x;
in
src
|> T.source_proper
|> Source.source T.stopper
- (Scan.bulk (P.$$$ "--" -- P.!!! P.text >> K None || P.not_eof >> Some))
- (if do_recover then Some recover else None)
+ (Scan.bulk (P.$$$ "--" -- P.!!! P.text >> K NONE || P.not_eof >> SOME))
+ (if do_recover then SOME recover else NONE)
|> Source.mapfilter I
|> Source.source T.stopper (Scan.bulk (fn xs => P.!!! (command term trc (cmd ())) xs))
- (if do_recover then Some recover else None)
+ (if do_recover then SOME recover else NONE)
|> Source.mapfilter I
end;
@@ -296,7 +296,7 @@
val pos = Path.position path;
val (name', parents, files) = ThyHeader.scan (src, pos);
val ml_path = ThyLoad.ml_path name;
- val ml_file = if ml andalso is_some (ThyLoad.check_file None ml_path) then [ml_path] else [];
+ val ml_file = if ml andalso is_some (ThyLoad.check_file NONE ml_path) then [ml_path] else [];
in
if name <> name' then
error ("Filename " ^ quote (Path.pack path) ^
@@ -315,7 +315,7 @@
val tr = Toplevel.imperative (fn () => ThyInfo.load_file time path);
val tr_name = if time then "time_use" else "use";
in
- if is_none (ThyLoad.check_file None path) then ()
+ if is_none (ThyLoad.check_file NONE path) then ()
else Toplevel.excursion [Toplevel.empty |> Toplevel.name tr_name |> tr]
end;
@@ -383,9 +383,9 @@
(*final declarations of this structure!*)
-val command = parser false None;
-val markup_command = parser false o Some;
-val improper_command = parser true None;
+val command = parser false NONE;
+val markup_command = parser false o SOME;
+val improper_command = parser true NONE;
end;
--- a/src/Pure/Isar/proof.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/proof.ML Sun Feb 13 17:15:14 2005 +0100
@@ -163,9 +163,9 @@
the proof state *)
fun kind_name _ (Theorem {kind = s, theory_spec = ((a, _), _),
- locale_spec = None, view = _}) = s ^ (if a = "" then "" else " " ^ a)
+ locale_spec = NONE, view = _}) = s ^ (if a = "" then "" else " " ^ a)
| kind_name sg (Theorem {kind = s, theory_spec = ((a, _), _),
- locale_spec = Some (name, _), view = _}) =
+ locale_spec = SOME (name, _), view = _}) =
s ^ " (in " ^ Locale.cond_extern sg name ^ ")" ^ (if a = "" then "" else " " ^ a)
| kind_name _ (Show _) = "show"
| kind_name _ (Have _) = "have";
@@ -208,15 +208,15 @@
fun check_result msg state sq =
(case Seq.pull sq of
- None => raise STATE (msg, state)
- | Some s_sq => Seq.cons s_sq);
+ NONE => raise STATE (msg, state)
+ | SOME s_sq => Seq.cons s_sq);
fun map_current f (State (Node {context, facts, mode, goal}, nodes)) =
State (make_node (f (context, facts, mode, goal)), nodes);
fun init_state thy =
- State (make_node (ProofContext.init thy, None, Forward, None), []);
+ State (make_node (ProofContext.init thy, NONE, Forward, NONE), []);
@@ -247,7 +247,7 @@
(* facts *)
-fun the_facts (State (Node {facts = Some facts, ...}, _)) = facts
+fun the_facts (State (Node {facts = SOME facts, ...}, _)) = facts
| the_facts state = raise STATE ("No current facts available", state);
fun the_fact state =
@@ -264,19 +264,19 @@
fun put_facts facts state =
state
|> map_current (fn (ctxt, _, mode, goal) => (ctxt, facts, mode, goal))
- |> (case facts of None => reset_thms thisN | Some ths => put_thms (thisN, ths));
+ |> (case facts of NONE => reset_thms thisN | SOME ths => put_thms (thisN, ths));
-val reset_facts = put_facts None;
+val reset_facts = put_facts NONE;
fun these_factss (state, named_factss) =
- state |> put_facts (Some (flat (map snd named_factss)));
+ state |> put_facts (SOME (flat (map snd named_factss)));
(* goal *)
local
- fun find i (state as State (Node {goal = Some goal, ...}, _)) = (context_of state, (i, goal))
- | find i (State (Node {goal = None, ...}, node :: nodes)) = find (i + 1) (State (node, nodes))
+ fun find i (state as State (Node {goal = SOME goal, ...}, _)) = (context_of state, (i, goal))
+ | find i (State (Node {goal = NONE, ...}, node :: nodes)) = find (i + 1) (State (node, nodes))
| find _ (state as State (_, [])) = err_malformed "find_goal" state;
in val find_goal = find 0 end;
@@ -287,12 +287,12 @@
(*
(* Jia: added here: get all goals from the state 10/01/05 *)
-fun find_all i (state as State (Node {goal = Some goal, ...}, node::nodes)) =
+fun find_all i (state as State (Node {goal = SOME goal, ...}, node::nodes)) =
let val others = find_all (i+1) (State (node, nodes))
in
(context_of state, (i, goal)) :: others
end
- | find_all i (State (Node {goal = None, ...}, node :: nodes)) = find_all (i + 1) (State (node, nodes))
+ | find_all i (State (Node {goal = NONE, ...}, node :: nodes)) = find_all (i + 1) (State (node, nodes))
| find_all _ (state as State (_, [])) = [];
val find_all_goals = find_all 0;
@@ -308,8 +308,8 @@
fun put_goal goal = map_current (fn (ctxt, facts, mode, _) => (ctxt, facts, mode, goal));
-fun map_goal f g (State (Node {context, facts, mode, goal = Some goal}, nodes)) =
- State (make_node (f context, facts, mode, Some (g goal)), nodes)
+fun map_goal f g (State (Node {context, facts, mode, goal = SOME goal}, nodes)) =
+ State (make_node (f context, facts, mode, SOME (g goal)), nodes)
| map_goal f g (State (nd, node :: nodes)) =
let val State (node', nodes') = map_goal f g (State (node, nodes))
in map_context f (State (nd, node' :: nodes')) end
@@ -381,7 +381,7 @@
fun new_block state =
state
|> open_block
- |> put_goal None;
+ |> put_goal NONE;
fun close_block (state as State (_, node :: nodes)) = State (node, nodes)
| close_block state = raise STATE ("Unbalanced block parentheses", state);
@@ -396,8 +396,8 @@
val pretty_goals_local = Display.pretty_goals_aux o ProofContext.pp;
-fun pretty_facts _ _ None = []
- | pretty_facts s ctxt (Some ths) =
+fun pretty_facts _ _ NONE = []
+ | pretty_facts s ctxt (SOME ths) =
[Pretty.big_list (s ^ "this:") (map (ProofContext.pretty_thm ctxt) ths), Pretty.str ""];
fun pretty_state nr (state as State (Node {context = ctxt, facts, mode, goal = _}, nodes)) =
@@ -419,7 +419,7 @@
fun prt_goal (_, (i, (((kind, names, _), (goal_facts, thm)), _))) =
pretty_facts "using " ctxt
- (if mode <> Backward orelse null goal_facts then None else Some goal_facts) @
+ (if mode <> Backward orelse null goal_facts then NONE else SOME goal_facts) @
[Pretty.str ("goal (" ^ kind_name (sign_of state) kind ^ prt_names names ^
levels_up (i div 2) ^ subgoals (Thm.nprems_of thm) ^ "):")] @
pretty_goals_local ctxt begin_goal (true, !show_main_goal) (! Display.goals_limit) thm;
@@ -517,12 +517,12 @@
fun transfer_facts inner_state state =
(case get_facts inner_state of
- None => Seq.single (reset_facts state)
- | Some thms =>
+ NONE => Seq.single (reset_facts state)
+ | SOME thms =>
let
val exp_tac = ProofContext.export false (context_of inner_state) (context_of state);
val thmqs = map exp_tac thms;
- in Seq.map (fn ths => put_facts (Some ths) state) (Seq.commute thmqs) end);
+ in Seq.map (fn ths => put_facts (SOME ths) state) (Seq.commute thmqs) end);
(* prepare result *)
@@ -634,7 +634,7 @@
(* locale instantiation *)
fun instantiate_locale (loc_name, prfx_attribs) state = let
- val facts = if is_chain state then get_facts state else None;
+ val facts = if is_chain state then get_facts state else NONE;
in
state
|> assert_forward_or_chain
@@ -689,7 +689,7 @@
val lhs = ProofContext.bind_skolem ctxt [x] (Free (x, T));
in
state
- |> fix [([x], None)]
+ |> fix [([x], NONE)]
|> match_bind_i [(pats, rhs)]
|> assm_i ProofContext.export_def [((name, atts), [(Logic.mk_equals (lhs, rhs), ([], []))])]
end;
@@ -731,7 +731,7 @@
fun from_facts facts state =
state
- |> put_facts (Some facts)
+ |> put_facts (SOME facts)
|> chain;
@@ -764,11 +764,11 @@
commas (map (ProofContext.string_of_term (context_of state')) vars));
state'
|> map_context (autofix props)
- |> put_goal (Some (((kind, names, propss), ([], goal)),
+ |> put_goal (SOME (((kind, names, propss), ([], goal)),
(after_qed o map_context gen_binds, pr)))
|> map_context (ProofContext.add_cases
(if length props = 1 then
- RuleCases.make true None (Thm.sign_of_thm goal, Thm.prop_of goal) [rule_contextN]
+ RuleCases.make true NONE (Thm.sign_of_thm goal, Thm.prop_of goal) [rule_contextN]
else [(rule_contextN, RuleCases.empty)]))
|> auto_bind_goal props
|> (if is_chain state then use_facts else reset_facts)
@@ -793,7 +793,7 @@
|> setup_goal I ProofContext.bind_propp_schematic_i ProofContext.fix_frees
(Theorem {kind = kind,
theory_spec = (a, map (snd o fst) concl),
- locale_spec = (case raw_locale of None => None | Some (_, x) => Some (the opt_name, x)),
+ locale_spec = (case raw_locale of NONE => NONE | SOME (_, x) => SOME (the opt_name, x)),
view = view})
Seq.single true (map (fst o fst) concl) propp
end;
@@ -823,12 +823,12 @@
(* current goal *)
-fun current_goal (State (Node {context, goal = Some goal, ...}, _)) = (context, goal)
+fun current_goal (State (Node {context, goal = SOME goal, ...}, _)) = (context, goal)
| current_goal state = raise STATE ("No current goal!", state);
-fun assert_current_goal true (state as State (Node {goal = None, ...}, _)) =
+fun assert_current_goal true (state as State (Node {goal = NONE, ...}, _)) =
raise STATE ("No goal in this block!", state)
- | assert_current_goal false (state as State (Node {goal = Some _, ...}, _)) =
+ | assert_current_goal false (state as State (Node {goal = SOME _, ...}, _)) =
raise STATE ("Goal present in this block!", state)
| assert_current_goal _ state = state;
@@ -907,7 +907,7 @@
val (theory', results') =
theory_of state
- |> (case locale_spec of None => I | Some (loc, (loc_atts, loc_attss)) => fn thy =>
+ |> (case locale_spec of NONE => I | SOME (loc, (loc_atts, loc_attss)) => fn thy =>
if length names <> length loc_attss then
raise THEORY ("Bad number of locale attributes", [thy])
else (thy, locale_ctxt)
--- a/src/Pure/Isar/proof_context.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/proof_context.ML Sun Feb 13 17:15:14 2005 +0100
@@ -290,13 +290,13 @@
val name = Object.name_of_kind kind;
in
(case Symtab.lookup (ProofDataData.get thy, name) of
- Some (k, meths) =>
+ SOME (k, meths) =>
if Object.eq_kind (kind, k) then
(case Symtab.lookup (data, name) of
- Some x => (x, meths)
- | None => err_undef ctxt name)
+ SOME x => (x, meths)
+ | NONE => err_undef ctxt name)
else err_access ctxt name
- | None => err_uninit ctxt name)
+ | NONE => err_uninit ctxt name)
end;
fun get_data kind f ctxt =
@@ -379,18 +379,18 @@
local
-fun mixfix x None mx = (fixedN ^ x, mixfix_type mx, Syntax.fix_mixfix x mx)
- | mixfix x (Some T) mx = (fixedN ^ x, T, Syntax.fix_mixfix x mx);
+fun mixfix x NONE mx = (fixedN ^ x, mixfix_type mx, Syntax.fix_mixfix x mx)
+ | mixfix x (SOME T) mx = (fixedN ^ x, T, Syntax.fix_mixfix x mx);
-fun prep_mixfix (_, _, None) = None
- | prep_mixfix (x, opt_T, Some mx) = Some (mixfix x opt_T mx);
+fun prep_mixfix (_, _, NONE) = NONE
+ | prep_mixfix (x, opt_T, SOME mx) = SOME (mixfix x opt_T mx);
-fun prep_mixfix' (_, _, None) = None
- | prep_mixfix' (x, _, Some Syntax.NoSyn) = None
- | prep_mixfix' (x, opt_T, _) = Some (x, mixfix x opt_T (Syntax.literal x));
+fun prep_mixfix' (_, _, NONE) = NONE
+ | prep_mixfix' (x, _, SOME Syntax.NoSyn) = NONE
+ | prep_mixfix' (x, opt_T, _) = SOME (x, mixfix x opt_T (Syntax.literal x));
-fun prep_struct (x, _, None) = Some x
- | prep_struct _ = None;
+fun prep_struct (x, _, NONE) = SOME x
+ | prep_struct _ = NONE;
in
@@ -450,11 +450,11 @@
fun def_type (Context {binds, defs = (types, _, _, _), ...}) pattern xi =
(case Vartab.lookup (types, xi) of
- None =>
- if pattern then None else
+ NONE =>
+ if pattern then NONE else
(case Vartab.lookup (binds, xi) of
- Some (Some (_, T)) => Some (TypeInfer.polymorphicT T)
- | _ => None)
+ SOME (SOME (_, T)) => SOME (TypeInfer.polymorphicT T)
+ | _ => NONE)
| some => some);
fun default_type (Context {defs = (types, _, _, _), ...}) x = Vartab.lookup (types, (x, ~1));
@@ -504,8 +504,8 @@
if internal x then t
else
(case lookup_skolem ctxt (no_skolem false ctxt x) of
- Some x' => Free (x', T)
- | None => t)
+ SOME x' => Free (x', T)
+ | NONE => t)
| intern (t $ u) = intern t $ intern u
| intern (Abs (x, T, t)) = Abs (x, T, intern t)
| intern a = a;
@@ -520,8 +520,8 @@
fun extern (t as Free (x, T)) =
(case assoc (rev_fixes, x) of
- Some x' => Free (if lookup_skolem ctxt x' = Some x then x' else NameSpace.hidden x', T)
- | None => t)
+ SOME x' => Free (if lookup_skolem ctxt x' = SOME x then x' else NameSpace.hidden x', T)
+ | NONE => t)
| extern (t $ u) = extern t $ extern u
| extern (Abs (x, T, t)) = Abs (x, T, extern t)
| extern a = a;
@@ -581,7 +581,7 @@
fun norm (t as Var (xi, T)) =
(case Vartab.lookup (binds, xi) of
- Some (Some (u, U)) =>
+ SOME (SOME (u, U)) =>
let
val env = unifyT ctxt (T, U) handle Type.TUNIFY =>
raise TYPE ("norm_term: ill-typed variable assignment", [T, U], [t, u]);
@@ -613,7 +613,7 @@
local
-fun append_env e1 e2 x = (case e2 x of None => e1 x | some => some);
+fun append_env e1 e2 x = (case e2 x of NONE => e1 x | some => some);
fun gen_read' read app pattern dummies schematic
ctxt internal more_types more_sorts more_used s =
@@ -631,7 +631,7 @@
end;
fun gen_read read app pattern dummies schematic ctxt =
- gen_read' read app pattern dummies schematic ctxt (K false) (K None) (K None) [];
+ gen_read' read app pattern dummies schematic ctxt (K false) (K NONE) (K NONE) [];
in
@@ -643,7 +643,7 @@
val read_prop_pats = read_term_pats propT;
fun read_term_liberal ctxt =
- gen_read' (read_term_sg true) I false false false ctxt (K true) (K None) (K None) [];
+ gen_read' (read_term_sg true) I false false false ctxt (K true) (K NONE) (K NONE) [];
val read_term = gen_read (read_term_sg true) I false false false;
val read_term_dummies = gen_read (read_term_sg true) I false true false;
@@ -700,8 +700,8 @@
fun ins_skolem def_ty = foldr
(fn ((x, x'), types) =>
(case def_ty x' of
- Some T => Vartab.update (((x, ~1), T), types)
- | None => types));
+ SOME T => Vartab.update (((x, ~1), T), types)
+ | NONE => types));
fun map_defaults f = map_context (fn (thy, syntax, data, asms, binds, thms, cases, defs, delta, delta_count) =>
(thy, syntax, data, asms, binds, thms, cases, f defs, delta, delta_count));
@@ -744,8 +744,8 @@
Library.gen_rems Library.eq_fst (Symtab.dest occ2, Symtab.dest occ1)
|> map (fn (a, ts) => map (pair a) (mapfilter (try (#1 o Term.dest_Free)) ts)) |> flat
|> mapfilter (fn (a, x) =>
- (case def_type ctxt1 false (x, ~1) of None => Some (a, x)
- | Some T => if known_tfree a T then None else Some (a, x)));
+ (case def_type ctxt1 false (x, ~1) of NONE => SOME (a, x)
+ | SOME T => if known_tfree a T then NONE else SOME (a, x)));
val tfrees = map #1 extras |> Library.sort_strings |> Library.unique_strings;
val frees = map #2 extras |> Library.sort_strings |> Library.unique_strings;
in
@@ -781,10 +781,10 @@
(** export theorems **)
-fun get_free x (None, t as Free (y, _)) = if x = y then Some t else None
+fun get_free x (NONE, t as Free (y, _)) = if x = y then SOME t else NONE
| get_free _ (opt, _) = opt;
-fun find_free t x = foldl_aterms (get_free x) (None, t);
+fun find_free t x = foldl_aterms (get_free x) (NONE, t);
fun export_view view is_goal inner outer =
let
@@ -815,8 +815,8 @@
let val exp = export_view view false inner outer in
fn th =>
(case Seq.pull (exp th) of
- Some (th', _) => th' |> Drule.local_standard
- | None => raise CONTEXT ("Internal failure while exporting theorem", inner))
+ SOME (th', _) => th' |> Drule.local_standard
+ | NONE => raise CONTEXT ("Internal failure while exporting theorem", inner))
end;
@@ -828,7 +828,7 @@
fun del_bind (ctxt, xi) =
ctxt
|> map_context (fn (thy, syntax, data, asms, binds, thms, cases, defs, delta, delta_count) =>
- (thy, syntax, data, asms, Vartab.update ((xi, None), binds), thms, cases, defs, delta, delta_count));
+ (thy, syntax, data, asms, Vartab.update ((xi, NONE), binds), thms, cases, defs, delta, delta_count));
fun upd_bind (ctxt, ((x, i), t)) =
let
@@ -840,11 +840,11 @@
ctxt
|> declare_term t'
|> map_context (fn (thy, syntax, data, asms, binds, thms, cases, defs, delta, delta_count) =>
- (thy, syntax, data, asms, Vartab.update (((x, i), Some (t', T)), binds), thms, cases, defs, delta, delta_count))
+ (thy, syntax, data, asms, Vartab.update (((x, i), SOME (t', T)), binds), thms, cases, defs, delta, delta_count))
end;
-fun del_upd_bind (ctxt, (xi, None)) = del_bind (ctxt, xi)
- | del_upd_bind (ctxt, (xi, Some t)) = upd_bind (ctxt, (xi, t));
+fun del_upd_bind (ctxt, (xi, NONE)) = del_bind (ctxt, xi)
+ | del_upd_bind (ctxt, (xi, SOME t)) = upd_bind (ctxt, (xi, t));
fun update_binds bs ctxt = foldl upd_bind (ctxt, bs);
fun delete_update_binds bs ctxt = foldl del_upd_bind (ctxt, bs);
@@ -863,14 +863,14 @@
map swap pairs); (*prefer assignment of variables from patterns*)
val env =
(case Seq.pull envs of
- None => fail ()
- | Some (env, _) => env); (*ignore further results*)
+ NONE => fail ()
+ | SOME (env, _) => env); (*ignore further results*)
val domain =
filter_out Term.is_replaced_dummy_pattern (map #1 (Drule.vars_of_terms (map #1 pairs)));
val _ = (*may not assign variables from text*)
if null (map #1 (Envir.alist_of env) inter (map #1 (Drule.vars_of_terms (map #2 pairs))))
then () else fail ();
- fun norm_bind (xi, t) = if xi mem domain then Some (xi, Envir.norm_term env t) else None;
+ fun norm_bind (xi, t) = if xi mem domain then SOME (xi, Envir.norm_term env t) else NONE;
in mapfilter norm_bind (Envir.alist_of env) end;
@@ -885,7 +885,7 @@
in
-fun drop_schematic (b as (xi, Some t)) = if null (Term.term_vars t) then b else (xi, None)
+fun drop_schematic (b as (xi, SOME t)) = if null (Term.term_vars t) then b else (xi, NONE)
| drop_schematic b = b;
val add_binds = gen_binds read_term;
@@ -917,7 +917,7 @@
val binds' =
if gen then map #1 binds ~~ generalize ctxt' ctxt (map #2 binds)
else binds;
- val binds'' = map (apsnd Some) binds';
+ val binds'' = map (apsnd SOME) binds';
in
warn_extra_tfrees ctxt
(if gen then ctxt (*sic!*) |> declare_terms (map #2 binds') |> add_binds_i binds''
@@ -959,8 +959,8 @@
(*generalize result: context evaluated now, binds added later*)
val gen = generalize ctxt' ctxt;
- fun gen_binds c = c |> add_binds_i (map #1 binds ~~ map Some (gen (map #2 binds)));
- in (ctxt' |> add_binds_i (map (apsnd Some) binds), (propss, gen_binds)) end;
+ fun gen_binds c = c |> add_binds_i (map #1 binds ~~ map SOME (gen (map #2 binds)));
+ in (ctxt' |> add_binds_i (map (apsnd SOME) binds), (propss, gen_binds)) end;
in
@@ -990,7 +990,7 @@
in
fn xnamei as (xname, _) =>
(case Symtab.lookup (tab, NameSpace.intern space xname) of
- Some (Some ths) => map (Thm.transfer_sg (Sign.deref sg_ref))
+ SOME (SOME ths) => map (Thm.transfer_sg (Sign.deref sg_ref))
(PureThy.select_thm xnamei ths)
| _ => get_from_thy xnamei) |> g xname
end;
@@ -1025,7 +1025,7 @@
if not q andalso NameSpace.is_qualified name then
raise CONTEXT ("Attempt to declare qualified name " ^ quote name, ctxt)
else (thy, syntax, data, asms, binds, (q, NameSpace.extend (space, [name]),
- Symtab.update ((name, Some ths), tab),
+ Symtab.update ((name, SOME ths), tab),
FactIndex.add (is_known ctxt) (index, (name, ths))), cases, defs, delta, delta_count));
fun put_thm (name, th) = put_thms (name, [th]);
@@ -1038,7 +1038,7 @@
fun reset_thms name =
map_context (fn (thy, syntax, data, asms, binds, (q, space, tab, index), cases, defs, delta, delta_count) =>
- (thy, syntax, data, asms, binds, (q, space, Symtab.update ((name, None), tab), index),
+ (thy, syntax, data, asms, binds, (q, space, Symtab.update ((name, NONE), tab), index),
cases, defs, delta,delta_count));
@@ -1192,7 +1192,7 @@
fun err_dups ctxt xs = raise CONTEXT ("Duplicate variable(s): " ^ commas_quote xs, ctxt);
val declare =
- declare_terms_syntax o mapfilter (fn (_, None) => None | (x, Some T) => Some (Free (x, T)));
+ declare_terms_syntax o mapfilter (fn (_, NONE) => NONE | (x, SOME T) => SOME (Free (x, T)));
fun add_vars xs Ts ctxt =
let val xs' = Term.variantlist (map Syntax.skolem xs, map #2 (fixes_of ctxt)) in
@@ -1221,7 +1221,7 @@
ctxt' |> add xs Ts
end;
-fun prep_type (x, None, Some mx) = ([x], Some (mixfix_type mx))
+fun prep_type (x, NONE, SOME mx) = ([x], SOME (mixfix_type mx))
| prep_type (x, opt_T, _) = ([x], opt_T);
in
@@ -1240,17 +1240,17 @@
fun fix_frees ts ctxt =
let
val frees = foldl Term.add_frees ([], ts);
- fun new (x, T) = if is_fixed ctxt x then None else Some ([x], Some T);
+ fun new (x, T) = if is_fixed ctxt x then NONE else SOME ([x], SOME T);
in fix_direct false (rev (mapfilter new frees)) ctxt end;
(*Note: improper use may result in variable capture / dynamic scoping!*)
fun bind_skolem ctxt xs =
let
- val ctxt' = ctxt |> fix_i [(xs, None)];
+ val ctxt' = ctxt |> fix_i [(xs, NONE)];
fun bind (t as Free (x, T)) =
if x mem_string xs then
- (case lookup_skolem ctxt' x of Some x' => Free (x', T) | None => t)
+ (case lookup_skolem ctxt' x of SOME x' => Free (x', T) | NONE => t)
else t
| bind (t $ u) = bind t $ bind u
| bind (Abs (x, T, t)) = Abs (x, T, bind t)
@@ -1275,8 +1275,8 @@
fun get_case (ctxt as Context {cases, ...}) name xs =
(case assoc (cases, name) of
- None => raise CONTEXT ("Unknown case: " ^ quote name, ctxt)
- | Some c => prep_case ctxt name xs c);
+ NONE => raise CONTEXT ("Unknown case: " ^ quote name, ctxt)
+ | SOME c => prep_case ctxt name xs c);
fun add_cases xs = map_context (fn (thy, syntax, data, asms, binds, thms, cases, defs, delta, delta_count) =>
@@ -1309,7 +1309,7 @@
(* term bindings *)
-val smash_option = fn (_, None) => None | (xi, Some b) => Some (xi, b);
+val smash_option = fn (_, NONE) => NONE | (xi, SOME b) => SOME (xi, b);
fun pretty_binds (ctxt as Context {binds, ...}) =
let
@@ -1336,7 +1336,7 @@
fun apply_case ({fixes, assumes, binds}: RuleCases.T) ctxt =
let
- fun bind (c, (x, T)) = (c |> fix_i [([x], Some T)], bind_skolem c [x] (Free (x, T)));
+ fun bind (c, (x, T)) = (c |> fix_i [([x], SOME T)], bind_skolem c [x] (Free (x, T)));
val (ctxt', xs) = foldl_map bind (ctxt, fixes);
fun app t = foldl Term.betapply (t, xs);
in (ctxt', (map (apsnd (apsome app)) binds, map (apsnd (map app)) assumes)) end;
@@ -1360,7 +1360,7 @@
(Pretty.str (name ^ ":") ::
prt_sect "fix" [] (Pretty.str o fst) fixes @
prt_sect "let" [Pretty.str "and"] prt_let
- (mapfilter (fn (xi, Some t) => Some (xi, t) | _ => None) lets) @
+ (mapfilter (fn (xi, SOME t) => SOME (xi, t) | _ => NONE) lets) @
(if forall (null o #2) asms then []
else prt_sect "assume" [Pretty.str "and"] prt_asm asms)));
@@ -1450,9 +1450,9 @@
fun lthms_containing (ctxt as Context {thms = (_, _, _, index), ...}) idx =
let
fun valid (name, ths) =
- (case try (transform_error (fn () => get_thms ctxt (name, None))) () of
- None => false
- | Some ths' => Library.equal_lists Thm.eq_thm (ths, ths'));
+ (case try (transform_error (fn () => get_thms ctxt (name, NONE))) () of
+ NONE => false
+ | SOME ths' => Library.equal_lists Thm.eq_thm (ths, ths'));
in gen_distinct eq_fst (filter valid (FactIndex.find idx index)) end;
@@ -1471,7 +1471,7 @@
val (cs, xs) = foldl (FactIndex.index_term (is_known ctxt))
(([], []), map (read_term_dummies ctxt) ss);
- val empty_idx = Library.null cs andalso Library.null xs;
+ val empty_idx = null cs andalso null xs;
val header =
if empty_idx then [Pretty.block [Pretty.str "Known facts:", Pretty.fbrk]]
@@ -1484,7 +1484,7 @@
val len = length facts;
val limit = if_none opt_limit (! thms_containing_limit);
in
- if empty_idx andalso not (Library.null ss) then
+ if empty_idx andalso not (null ss) then
warning "thms_containing: no consts/frees in specification"
else (header @ (if len <= limit then [] else [Pretty.str "..."]) @
map prt_fact (Library.drop (len - limit, facts))) |> Pretty.chunks |> Pretty.writeln
--- a/src/Pure/Isar/proof_history.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/proof_history.ML Sun Feb 13 17:15:14 2005 +0100
@@ -57,11 +57,11 @@
fun back_fun recur ((_, stq), nodes) =
(case Seq.pull stq of
- None =>
+ NONE =>
if recur andalso not (null nodes) then
(writeln "back: trying previous proof state ..."; back_fun recur (hd nodes, tl nodes))
else raise FAIL "back: no alternatives"
- | Some result => (result, nodes));
+ | SOME result => (result, nodes));
val back = hist_app o back_fun;
@@ -70,8 +70,8 @@
fun applys f = hist_app (fn (node as (st, _), nodes) =>
(case Seq.pull (f st) of
- None => raise FAIL "empty result sequence -- proof command failed"
- | Some results => (results, node :: nodes)));
+ NONE => raise FAIL "empty result sequence -- proof command failed"
+ | SOME results => (results, node :: nodes)));
fun apply f = applys (Seq.single o f);
--- a/src/Pure/Isar/rule_cases.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/rule_cases.ML Sun Feb 13 17:15:14 2005 +0100
@@ -36,26 +36,26 @@
fun lookup_consumes thm =
let fun err () = raise THM ("Malformed 'consumes' tag of theorem", 0, [thm]) in
(case Library.assoc (Thm.tags_of_thm thm, consumes_tagN) of
- None => None
- | Some [s] => (case Syntax.read_nat s of Some n => Some n | _ => err ())
+ NONE => NONE
+ | SOME [s] => (case Syntax.read_nat s of SOME n => SOME n | _ => err ())
| _ => err ())
end;
-fun put_consumes None th = th
- | put_consumes (Some n) th = th
+fun put_consumes NONE th = th
+ | put_consumes (SOME n) th = th
|> Drule.untag_rule consumes_tagN
|> Drule.tag_rule (consumes_tagN, [Library.string_of_int n]);
val save_consumes = put_consumes o lookup_consumes;
-fun consumes n x = Drule.rule_attribute (K (put_consumes (Some n))) x;
+fun consumes n x = Drule.rule_attribute (K (put_consumes (SOME n))) x;
fun consumes_default n x = if Library.is_some (lookup_consumes (#2 x)) then x else consumes n x;
(* case names *)
-fun put_case_names None thm = thm
- | put_case_names (Some names) thm =
+fun put_case_names NONE thm = thm
+ | put_case_names (SOME names) thm =
thm
|> Drule.untag_rule cases_tagN
|> Drule.tag_rule (cases_tagN, names);
@@ -63,7 +63,7 @@
fun lookup_case_names thm = Library.assoc (Thm.tags_of_thm thm, cases_tagN);
val save_case_names = put_case_names o lookup_case_names;
-val name = put_case_names o Some;
+val name = put_case_names o SOME;
fun case_names ss = Drule.rule_attribute (K (name ss));
@@ -76,8 +76,8 @@
val n = if_none (lookup_consumes thm) 0;
val ss =
(case lookup_case_names thm of
- None => map Library.string_of_int (1 upto (Thm.nprems_of thm - n))
- | Some ss => ss);
+ NONE => map Library.string_of_int (1 upto (Thm.nprems_of thm - n))
+ | SOME ss => ss);
in (ss, n) end;
fun add thm = (thm, get thm);
@@ -102,20 +102,20 @@
let
val Bi = Drule.norm_hhf sg (nth_subgoal(i,prop));
val all_xs = Logic.strip_params Bi
- val n = (case split of None => length all_xs
- | Some t => length (Logic.strip_params(nth_subgoal(i,t))))
+ val n = (case split of NONE => length all_xs
+ | SOME t => length (Logic.strip_params(nth_subgoal(i,t))))
val (ind_xs, goal_xs) = splitAt(n,all_xs)
val rename = if is_open then I else apfst Syntax.internal
val xs = map rename ind_xs @ goal_xs
val asms = map (curry Term.list_abs xs) (Logic.strip_assums_hyp Bi);
val named_asms =
- (case split of None => [("", asms)]
- | Some t =>
+ (case split of NONE => [("", asms)]
+ | SOME t =>
let val h = length (Logic.strip_assums_hyp(nth_subgoal(i,t)))
val (ps,qs) = splitAt(h, asms)
in [(hypsN, ps), (premsN, qs)] end);
val concl = Term.list_abs (xs, Logic.strip_assums_concl Bi);
- val bind = ((case_conclN, 0), Some (ObjectLogic.drop_judgment sg concl));
+ val bind = ((case_conclN, 0), SOME (ObjectLogic.drop_judgment sg concl));
in (name, {fixes = xs, assumes = named_asms, binds = [bind]}) end;
fun make is_open split (sg, prop) names =
--- a/src/Pure/Isar/session.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/session.ML Sun Feb 13 17:15:14 2005 +0100
@@ -25,7 +25,7 @@
val session = ref ([pure]: string list);
val session_path = ref ([]: string list);
val session_finished = ref false;
-val rpath = ref (None: Url.T option);
+val rpath = ref (NONE: Url.T option);
(* access path *)
@@ -76,7 +76,7 @@
if is_some (!rpath) then
error "Path for remote theory browsing information may only be set once"
else
- rpath := Some (Url.unpack rpath_str);
+ rpath := SOME (Url.unpack rpath_str);
(!rpath, rpath_str <> ""));
fun use_dir root_file build modes reset info doc doc_graph parent name dump rpath_str level verbose hide =
@@ -84,7 +84,7 @@
(Library.setmp Proofterm.proofs level (Library.setmp IsarOutput.hide_commands hide (fn () =>
(init reset parent name;
Present.init build info doc doc_graph (path ()) name
- (if dump = "" then None else Some (Path.unpack dump)) (get_rpath rpath_str) verbose;
+ (if dump = "" then NONE else SOME (Path.unpack dump)) (get_rpath rpath_str) verbose;
File.use (Path.basic root_file);
finish ())))) ()
handle exn => (writeln (Toplevel.exn_message exn); exit 1);
--- a/src/Pure/Isar/skip_proof.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/skip_proof.ML Sun Feb 13 17:15:14 2005 +0100
@@ -60,8 +60,8 @@
fun cheating int ctxt = Method.METHOD (K (setmp quick_and_dirty (int orelse ! quick_and_dirty)
(cheat_tac (ProofContext.theory_of ctxt))));
-fun local_skip_proof x int = Method.local_terminal_proof (Method.Basic (cheating int), None) x;
-fun global_skip_proof int = Method.global_terminal_proof (Method.Basic (cheating int), None);
+fun local_skip_proof x int = Method.local_terminal_proof (Method.Basic (cheating int), NONE) x;
+fun global_skip_proof int = Method.global_terminal_proof (Method.Basic (cheating int), NONE);
end;
--- a/src/Pure/Isar/thy_header.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/thy_header.ML Sun Feb 13 17:15:14 2005 +0100
@@ -28,11 +28,11 @@
|> Symbol.source false
|> T.source false (fn () => (get_lex (), Scan.empty_lexicon)) pos
|> T.source_proper
- |> Source.source T.stopper (Scan.single scan) None
+ |> Source.source T.stopper (Scan.single scan) NONE
|> Source.get_single;
in
- (case res of Some (x, _) => x
- | None => error ("Unexpected end of input" ^ Position.str_of pos))
+ (case res of SOME (x, _) => x
+ | NONE => error ("Unexpected end of input" ^ Position.str_of pos))
end;
--- a/src/Pure/Isar/toplevel.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Isar/toplevel.ML Sun Feb 13 17:15:14 2005 +0100
@@ -115,13 +115,13 @@
datatype state = State of (node History.T * ((node -> unit) * (node -> unit))) option;
-val toplevel = State None;
+val toplevel = State NONE;
-fun is_toplevel (State None) = true
+fun is_toplevel (State NONE) = true
| is_toplevel _ = false;
-fun str_of_state (State None) = "at top level"
- | str_of_state (State (Some (node, _))) = str_of_node (History.current node);
+fun str_of_state (State NONE) = "at top level"
+ | str_of_state (State (SOME (node, _))) = str_of_node (History.current node);
(* prompt_state hook *)
@@ -134,8 +134,8 @@
(* print state *)
-fun pretty_current_node _ (State None) = []
- | pretty_current_node prt (State (Some (node, _))) = prt (History.current node);
+fun pretty_current_node _ (State NONE) = []
+ | pretty_current_node prt (State (SOME (node, _))) = prt (History.current node);
val print_state_context =
Pretty.writeln o Pretty.chunks o pretty_current_node pretty_node_context;
@@ -157,8 +157,8 @@
exception UNDEF;
-fun node_history_of (State None) = raise UNDEF
- | node_history_of (State (Some (node, _))) = node;
+fun node_history_of (State NONE) = raise UNDEF
+ | node_history_of (State (SOME (node, _))) = node;
val node_of = History.current o node_history_of;
@@ -199,30 +199,30 @@
| copy_node _ node = node;
fun interruptible f x =
- let val y = ref (None: node History.T option);
- in raise_interrupt (fn () => y := Some (f x)) (); the (! y) end;
+ let val y = ref (NONE: node History.T option);
+ in raise_interrupt (fn () => y := SOME (f x)) (); the (! y) end;
val rollback = ERROR_MESSAGE "Stale signature encountered after succesful execution!";
-fun return (result, None) = result
- | return (result, Some exn) = raise FAILURE (result, exn);
+fun return (result, NONE) = result
+ | return (result, SOME exn) = raise FAILURE (result, exn);
in
-fun node_trans _ _ _ (State None) = raise UNDEF
- | node_trans int hist f (State (Some (node, term))) =
+fun node_trans _ _ _ (State NONE) = raise UNDEF
+ | node_trans int hist f (State (SOME (node, term))) =
let
- fun mk_state nd = State (Some (nd, term));
+ fun mk_state nd = State (SOME (nd, term));
val cont_node = History.map (checkpoint_node int) node;
val back_node = History.map (copy_node int) cont_node;
val trans = if hist then History.apply_copy (copy_node int) else History.map;
val (result, opt_exn) =
- (mk_state (transform_error (interruptible (trans (f int))) cont_node), None)
- handle exn => (mk_state cont_node, Some exn);
+ (mk_state (transform_error (interruptible (trans (f int))) cont_node), NONE)
+ handle exn => (mk_state cont_node, SOME exn);
in
- if is_stale result then return (mk_state back_node, Some (if_none opt_exn rollback))
+ if is_stale result then return (mk_state back_node, SOME (if_none opt_exn rollback))
else return (result, opt_exn)
end;
@@ -253,23 +253,23 @@
MapCurrent of bool -> node -> node | (*change node, bypassing history*)
AppCurrent of bool -> node -> node; (*change node, recording history*)
-fun undo_limit int = if int then None else Some 0;
+fun undo_limit int = if int then NONE else SOME 0;
local
fun apply_tr _ Reset _ = toplevel
- | apply_tr int (Init (f, term)) (State None) =
- State (Some (History.init (undo_limit int) (f int), term))
- | apply_tr _ (Init _ ) (State (Some _)) = raise UNDEF
- | apply_tr _ Exit (State None) = raise UNDEF
- | apply_tr _ Exit (State (Some (node, (exit, _)))) =
- (exit (History.current node); State None)
- | apply_tr _ Kill (State None) = raise UNDEF
- | apply_tr _ Kill (State (Some (node, (_, kill)))) =
- (kill (History.current node); State None)
+ | apply_tr int (Init (f, term)) (State NONE) =
+ State (SOME (History.init (undo_limit int) (f int), term))
+ | apply_tr _ (Init _ ) (State (SOME _)) = raise UNDEF
+ | apply_tr _ Exit (State NONE) = raise UNDEF
+ | apply_tr _ Exit (State (SOME (node, (exit, _)))) =
+ (exit (History.current node); State NONE)
+ | apply_tr _ Kill (State NONE) = raise UNDEF
+ | apply_tr _ Kill (State (SOME (node, (_, kill)))) =
+ (kill (History.current node); State NONE)
| apply_tr int (Keep f) state = (raise_interrupt (f int) state; state)
- | apply_tr _ (History _) (State None) = raise UNDEF
- | apply_tr _ (History f) (State (Some (node, term))) = State (Some (f node, term))
+ | apply_tr _ (History _) (State NONE) = raise UNDEF
+ | apply_tr _ (History f) (State (SOME (node, term))) = State (SOME (f node, term))
| apply_tr int (MapCurrent f) state = node_trans int false f state
| apply_tr int (AppCurrent f) state = node_trans int true f state;
@@ -283,8 +283,8 @@
in
-fun apply_trans int trs state = (apply_union int trs state, None)
- handle FAILURE (alt_state, exn) => (alt_state, Some exn) | exn => (state, Some exn);
+fun apply_trans int trs state = (apply_union int trs state, NONE)
+ handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn);
end;
@@ -307,7 +307,7 @@
fun map_transition f (Transition {name, pos, source, int_only, print, no_timing, trans}) =
make_transition (f (name, pos, source, int_only, print, no_timing, trans));
-val empty = make_transition ("<unknown>", Position.none, None, false, false, false, []);
+val empty = make_transition ("<unknown>", Position.none, NONE, false, false, false, []);
fun name_of (Transition {name, ...}) = name;
fun source_of (Transition {source, ...}) = source;
@@ -333,7 +333,7 @@
(name, pos, source, int_only, print, no_timing, trans));
fun source src = map_transition (fn (name, pos, _, int_only, print, no_timing, trans) =>
- (name, pos, Some src, int_only, print, no_timing, trans));
+ (name, pos, SOME src, int_only, print, no_timing, trans));
fun interactive int_only = map_transition (fn (name, pos, source, _, print, no_timing, trans) =>
(name, pos, source, int_only, print, no_timing, trans));
@@ -428,7 +428,7 @@
else msg
| msg_on_debug (TERM (msg, ts)) =
(case (!debug, Context.get_context ()) of
- (true, Some thy) =>
+ (true, SOME thy) =>
let val sg = Theory.sign_of thy in
raised_msg "TERM"
(cat_lines
@@ -437,7 +437,7 @@
| _ => raised_msg "TERM" msg)
| msg_on_debug (TYPE (msg, Ts, ts)) =
(case (!debug, Context.get_context ()) of
- (true, Some thy) =>
+ (true, SOME thy) =>
let val sg = Theory.sign_of thy in
raised_msg "TYPE"
(cat_lines
@@ -480,14 +480,14 @@
| exn_message (TYPE (msg, Ts, ts)) = msg_on_debug (TYPE (msg, Ts, ts))
| exn_message (TERM (msg, ts)) = msg_on_debug (TERM (msg, ts))
| exn_message (THM (msg, i, thms)) = msg_on_debug (THM (msg, i, thms))
- | exn_message Library.OPTION = raised "Library.OPTION"
+ | exn_message Option = raised "Option"
| exn_message (Library.LIST msg) = raised_msg "Library.LIST" msg
| exn_message exn = General.exnMessage exn
and fail_message kind ((name, pos), exn) =
"Error in " ^ kind ^ " " ^ quote name ^ Position.str_of pos ^ ":\n" ^ exn_message exn;
-fun print_exn None = ()
- | print_exn (Some (exn, s)) = error_msg (cat_lines [exn_message exn, s]);
+fun print_exn NONE = ()
+ | print_exn (SOME (exn, s)) = error_msg (cat_lines [exn_message exn, s]);
(* apply transitions *)
@@ -509,11 +509,11 @@
fun apply int tr st =
(case app int tr st of
- (_, Some TERMINATE) => None
- | (_, Some RESTART) => Some (toplevel, None)
- | (state', Some (EXCURSION_FAIL exn_info)) => Some (state', Some exn_info)
- | (state', Some exn) => Some (state', Some (exn, at_command tr))
- | (state', None) => Some (state', None));
+ (_, SOME TERMINATE) => NONE
+ | (_, SOME RESTART) => SOME (toplevel, NONE)
+ | (state', SOME (EXCURSION_FAIL exn_info)) => SOME (state', SOME exn_info)
+ | (state', SOME exn) => SOME (state', SOME (exn, at_command tr))
+ | (state', NONE) => SOME (state', NONE));
end;
@@ -525,17 +525,17 @@
fun excur [] x = x
| excur ((tr, f) :: trs) (st, res) =
(case apply false tr st of
- Some (st', None) =>
+ SOME (st', NONE) =>
excur trs (st', transform_error (fn () => f st st' res) () handle exn =>
raise EXCURSION_FAIL (exn, "Presentation failed\n" ^ at_command tr))
- | Some (st', Some exn_info) => raise EXCURSION_FAIL exn_info
- | None => raise EXCURSION_FAIL (TERMINATE, at_command tr));
+ | SOME (st', SOME exn_info) => raise EXCURSION_FAIL exn_info
+ | NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr));
in
fun excursion_result (trs, res) =
- (case excur trs (State None, res) of
- (State None, res') => res'
+ (case excur trs (State NONE, res) of
+ (State NONE, res') => res'
| _ => raise ERROR_MESSAGE "Unfinished development at end of input")
handle exn => error (exn_message exn);
@@ -550,9 +550,9 @@
(* the global state reference *)
-val global_state = ref (toplevel, None: (exn * string) option);
+val global_state = ref (toplevel, NONE: (exn * string) option);
-fun set_state state = global_state := (state, None);
+fun set_state state = global_state := (state, NONE);
fun get_state () = fst (! global_state);
fun exn () = snd (! global_state);
@@ -572,8 +572,8 @@
fun >> tr =
(case apply true tr (get_state ()) of
- None => false
- | Some (state', exn_info) =>
+ NONE => false
+ | SOME (state', exn_info) =>
(global_state := (state', exn_info);
check_stale state'; print_exn exn_info;
true));
@@ -583,13 +583,13 @@
(*Note: this is for Poly/ML only, we really do not intend to exhibit
interrupts here*)
-fun get_interrupt src = Some (Source.get_single src) handle Interrupt => None;
+fun get_interrupt src = SOME (Source.get_single src) handle Interrupt => NONE;
fun raw_loop src =
(case get_interrupt (Source.set_prompt (prompt_state (get_state ())) src) of
- None => (writeln "\nInterrupt."; raw_loop src)
- | Some None => ()
- | Some (Some (tr, src')) => if >> tr then raw_loop src' else ());
+ NONE => (writeln "\nInterrupt."; raw_loop src)
+ | SOME NONE => ()
+ | SOME (SOME (tr, src')) => if >> tr then raw_loop src' else ());
fun loop src = ignore_interrupt raw_loop src;
--- a/src/Pure/Proof/extraction.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Proof/extraction.ML Sun Feb 13 17:15:14 2005 +0100
@@ -54,22 +54,22 @@
Const ("Type", itselfT T --> Type ("Type", [])) $ Logic.mk_type T;
fun typeof_proc defaultS vs (Const ("typeof", _) $ u) =
- Some (mk_typ (case strip_comb u of
+ SOME (mk_typ (case strip_comb u of
(Var ((a, i), _), _) =>
if a mem vs then TFree ("'" ^ a ^ ":" ^ string_of_int i, defaultS)
else nullT
| (Free (a, _), _) =>
if a mem vs then TFree ("'" ^ a, defaultS) else nullT
| _ => nullT))
- | typeof_proc _ _ _ = None;
+ | typeof_proc _ _ _ = NONE;
-fun rlz_proc (Const ("realizes", Type (_, [Type ("Null", []), _])) $ r $ t) = Some t
+fun rlz_proc (Const ("realizes", Type (_, [Type ("Null", []), _])) $ r $ t) = SOME t
| rlz_proc (Const ("realizes", Type (_, [T, _])) $ r $ t) =
(case strip_comb t of
- (Var (ixn, U), ts) => Some (list_comb (Var (ixn, T --> U), r :: ts))
- | (Free (s, U), ts) => Some (list_comb (Free (s, T --> U), r :: ts))
- | _ => None)
- | rlz_proc _ = None;
+ (Var (ixn, U), ts) => SOME (list_comb (Var (ixn, T --> U), r :: ts))
+ | (Free (s, U), ts) => SOME (list_comb (Free (s, T --> U), r :: ts))
+ | _ => NONE)
+ | rlz_proc _ = NONE;
val unpack_ixn = apfst implode o apsnd (fst o read_int o tl) o
take_prefix (not o equal ":") o explode;
@@ -98,10 +98,10 @@
let
val cache = ref ([] : (term * term) list);
fun lookup f x = (case assoc (!cache, x) of
- None =>
+ NONE =>
let val y = f x
in (cache := (x, y) :: !cache; y) end
- | Some y => y);
+ | SOME y => y);
in
get_first (fn (_, (prems, (tm1, tm2))) =>
let
@@ -115,15 +115,15 @@
iTs = Vartab.make Tenv, asol = Vartab.make tenv};
val env'' = foldl (fn (env, p) =>
Pattern.unify (sign, env, [pairself (lookup rew) p])) (env', prems')
- in Some (Envir.norm_term env'' (inc (ren tm2)))
- end handle Pattern.MATCH => None | Pattern.Unif => None)
+ in SOME (Envir.norm_term env'' (inc (ren tm2)))
+ end handle Pattern.MATCH => NONE | Pattern.Unif => NONE)
(sort (Int.compare o pairself fst)
(Net.match_term rules (Pattern.eta_contract tm)))
end;
in rew end;
-val chtype = change_type o Some;
+val chtype = change_type o SOME;
fun add_prefix a b = NameSpace.pack (a :: NameSpace.unpack b);
@@ -145,7 +145,7 @@
fun forall_intr_prf (t, prf) =
let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
- in Abst (a, Some T, prf_abstract_over t prf) end;
+ in Abst (a, SOME T, prf_abstract_over t prf) end;
val mkabs = foldr (fn (v, t) => Abs ("x", fastype_of v, abstract_over (v, t)));
@@ -171,11 +171,11 @@
val fs = Term.add_frees ([], t)
in fn
Var (ixn, _) => (case assoc (Term.add_vars ([], t), ixn) of
- None => error "get_var_type: no such variable in term"
- | Some T => Var (ixn, T))
+ NONE => error "get_var_type: no such variable in term"
+ | SOME T => Var (ixn, T))
| Free (s, _) => (case assoc (Term.add_frees ([], t), s) of
- None => error "get_var_type: no such variable in term"
- | Some T => Free (s, T))
+ NONE => error "get_var_type: no such variable in term"
+ | SOME T => Free (s, T))
| _ => error "get_var_type: not a variable"
end;
@@ -204,7 +204,7 @@
realizers = Symtab.empty,
defs = [],
expand = [],
- prep = None};
+ prep = NONE};
val copy = I;
val prep_ext = I;
@@ -220,7 +220,7 @@
(realizers1, realizers2),
defs = gen_merge_lists eq_thm defs1 defs2,
expand = merge_lists expand1 expand2,
- prep = (case prep1 of None => prep2 | _ => prep1)};
+ prep = (case prep1 of NONE => prep2 | _ => prep1)};
fun print sg (x : T) = ();
end;
@@ -244,7 +244,7 @@
in
ExtractionData.put
{realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types,
- realizers = realizers, defs = defs, expand = expand, prep = Some prep} thy
+ realizers = realizers, defs = defs, expand = expand, prep = SOME prep} thy
end;
(** equations characterizing realizability **)
@@ -431,7 +431,7 @@
val rtypes = map fst types;
val typroc = typeof_proc (Sign.defaultS sg);
val prep = if_none prep (K I) sg o ProofRewriteRules.elim_defs sg false defs o
- Reconstruct.expand_proof sg (("", None) :: map (apsnd Some) expand);
+ Reconstruct.expand_proof sg (("", NONE) :: map (apsnd SOME) expand);
val rrews = Net.merge (#net realizes_eqns, #net typeof_eqns, K false);
fun find_inst prop Ts ts vs =
@@ -462,36 +462,36 @@
fun corr d defs vs ts Ts hs (PBound i) _ _ = (defs, PBound i)
- | corr d defs vs ts Ts hs (Abst (s, Some T, prf)) (Abst (_, _, prf')) t =
+ | corr d defs vs ts Ts hs (Abst (s, SOME T, prf)) (Abst (_, _, prf')) t =
let val (defs', corr_prf) = corr d defs vs [] (T :: Ts)
(dummyt :: hs) prf (incr_pboundvars 1 0 prf')
- (case t of Some (Abs (_, _, u)) => Some u | _ => None)
- in (defs', Abst (s, Some T, corr_prf)) end
+ (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE)
+ in (defs', Abst (s, SOME T, corr_prf)) end
- | corr d defs vs ts Ts hs (AbsP (s, Some prop, prf)) (AbsP (_, _, prf')) t =
+ | corr d defs vs ts Ts hs (AbsP (s, SOME prop, prf)) (AbsP (_, _, prf')) t =
let
val T = etype_of sg vs Ts prop;
val u = if T = nullT then
- (case t of Some u => Some (incr_boundvars 1 u) | None => None)
- else (case t of Some (Abs (_, _, u)) => Some u | _ => None);
+ (case t of SOME u => SOME (incr_boundvars 1 u) | NONE => NONE)
+ else (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE);
val (defs', corr_prf) = corr d defs vs [] (T :: Ts) (prop :: hs)
(incr_pboundvars 0 1 prf) (incr_pboundvars 0 1 prf') u;
val rlz = Const ("realizes", T --> propT --> propT)
in (defs',
if T = nullT then AbsP ("R",
- Some (app_rlz_rews Ts vs (rlz $ nullt $ prop)),
+ SOME (app_rlz_rews Ts vs (rlz $ nullt $ prop)),
prf_subst_bounds [nullt] corr_prf)
- else Abst (s, Some T, AbsP ("R",
- Some (app_rlz_rews (T :: Ts) vs
+ else Abst (s, SOME T, AbsP ("R",
+ SOME (app_rlz_rews (T :: Ts) vs
(rlz $ Bound 0 $ incr_boundvars 1 prop)), corr_prf)))
end
- | corr d defs vs ts Ts hs (prf % Some t) (prf' % _) t' =
+ | corr d defs vs ts Ts hs (prf % SOME t) (prf' % _) t' =
let
val (Us, T) = strip_type (fastype_of1 (Ts, t));
val (defs', corr_prf) = corr d defs vs (t :: ts) Ts hs prf prf'
(if tname_of T mem rtypes then t'
- else (case t' of Some (u $ _) => Some u | _ => None));
+ else (case t' of SOME (u $ _) => SOME u | _ => NONE));
val u = if not (tname_of T mem rtypes) then t else
let
val eT = etype_of sg vs Ts t;
@@ -500,21 +500,21 @@
val u = list_comb (incr_boundvars (length Us') t,
map Bound (length Us - 1 downto 0));
val u' = (case assoc (types, tname_of T) of
- Some ((_, Some f)) => f r eT u T
+ SOME ((_, SOME f)) => f r eT u T
| _ => Const ("realizes", eT --> T --> T) $ r $ u)
in app_rlz_rews Ts vs (list_abs (map (pair "x") Us', u')) end
- in (defs', corr_prf % Some u) end
+ in (defs', corr_prf % SOME u) end
| corr d defs vs ts Ts hs (prf1 %% prf2) (prf1' %% prf2') t =
let
val prop = Reconstruct.prop_of' hs prf2';
val T = etype_of sg vs Ts prop;
- val (defs1, f, u) = if T = nullT then (defs, t, None) else
+ val (defs1, f, u) = if T = nullT then (defs, t, NONE) else
(case t of
- Some (f $ u) => (defs, Some f, Some u)
+ SOME (f $ u) => (defs, SOME f, SOME u)
| _ =>
let val (defs1, u) = extr d defs vs [] Ts hs prf2'
- in (defs1, None, Some u) end)
+ in (defs1, NONE, SOME u) end)
val (defs2, corr_prf1) = corr d defs1 vs [] Ts hs prf1 prf1' f;
val (defs3, corr_prf2) = corr d defs2 vs [] Ts hs prf2 prf2' u;
in
@@ -522,7 +522,7 @@
(defs3, corr_prf1 % u %% corr_prf2)
end
- | corr d defs vs ts Ts hs (prf0 as PThm ((name, _), prf, prop, Some Ts')) _ _ =
+ | corr d defs vs ts Ts hs (prf0 as PThm ((name, _), prf, prop, SOME Ts')) _ _ =
let
val (vs', tye) = find_inst prop Ts ts vs;
val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye;
@@ -532,8 +532,8 @@
in
if T = nullT andalso realizes_null vs' prop aconv prop then (defs, prf0)
else case Symtab.lookup (realizers, name) of
- None => (case find vs' (find' name defs') of
- None =>
+ NONE => (case find vs' (find' name defs') of
+ NONE =>
let
val _ = assert (T = nullT) "corr: internal error";
val _ = msg d ("Building correctness proof for " ^ quote name ^
@@ -541,25 +541,25 @@
else " (relevant variables: " ^ commas_quote vs' ^ ")"));
val prf' = prep (Reconstruct.reconstruct_proof sg prop prf);
val (defs'', corr_prf) =
- corr (d + 1) defs' vs' [] [] [] prf' prf' None;
+ corr (d + 1) defs' vs' [] [] [] prf' prf' NONE;
val corr_prop = Reconstruct.prop_of corr_prf;
val corr_prf' = foldr forall_intr_prf
(map (get_var_type corr_prop) (vfs_of prop), proof_combt
(PThm ((corr_name name vs', []), corr_prf, corr_prop,
- Some (map TVar (term_tvars corr_prop))), vfs_of corr_prop))
+ SOME (map TVar (term_tvars corr_prop))), vfs_of corr_prop))
in
((name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'',
prf_subst_TVars tye' corr_prf')
end
- | Some (_, (_, prf')) => (defs', prf_subst_TVars tye' prf'))
- | Some rs => (case find vs' rs of
- Some (_, prf') => (defs', prf_subst_TVars tye' prf')
- | None => error ("corr: no realizer for instance of theorem " ^
+ | SOME (_, (_, prf')) => (defs', prf_subst_TVars tye' prf'))
+ | SOME rs => (case find vs' rs of
+ SOME (_, prf') => (defs', prf_subst_TVars tye' prf')
+ | NONE => error ("corr: no realizer for instance of theorem " ^
quote name ^ ":\n" ^ Sign.string_of_term sg (Envir.beta_norm
(Reconstruct.prop_of (proof_combt (prf0, ts))))))
end
- | corr d defs vs ts Ts hs (prf0 as PAxm (s, prop, Some Ts')) _ _ =
+ | corr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) _ _ =
let
val (vs', tye) = find_inst prop Ts ts vs;
val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
@@ -567,8 +567,8 @@
if etype_of sg vs' [] prop = nullT andalso
realizes_null vs' prop aconv prop then (defs, prf0)
else case find vs' (Symtab.lookup_multi (realizers, s)) of
- Some (_, prf) => (defs, prf_subst_TVars tye' prf)
- | None => error ("corr: no realizer for instance of axiom " ^
+ SOME (_, prf) => (defs, prf_subst_TVars tye' prf)
+ | NONE => error ("corr: no realizer for instance of axiom " ^
quote s ^ ":\n" ^ Sign.string_of_term sg (Envir.beta_norm
(Reconstruct.prop_of (proof_combt (prf0, ts)))))
end
@@ -577,12 +577,12 @@
and extr d defs vs ts Ts hs (PBound i) = (defs, Bound i)
- | extr d defs vs ts Ts hs (Abst (s, Some T, prf)) =
+ | extr d defs vs ts Ts hs (Abst (s, SOME T, prf)) =
let val (defs', t) = extr d defs vs []
(T :: Ts) (dummyt :: hs) (incr_pboundvars 1 0 prf)
in (defs', Abs (s, T, t)) end
- | extr d defs vs ts Ts hs (AbsP (s, Some t, prf)) =
+ | extr d defs vs ts Ts hs (AbsP (s, SOME t, prf)) =
let
val T = etype_of sg vs Ts t;
val (defs', t) = extr d defs vs [] (T :: Ts) (t :: hs)
@@ -591,7 +591,7 @@
if T = nullT then subst_bound (nullt, t) else Abs (s, T, t))
end
- | extr d defs vs ts Ts hs (prf % Some t) =
+ | extr d defs vs ts Ts hs (prf % SOME t) =
let val (defs', u) = extr d defs vs (t :: ts) Ts hs prf
in (defs',
if tname_of (body_type (fastype_of1 (Ts, t))) mem rtypes then u
@@ -609,14 +609,14 @@
in (defs'', f $ t) end
end
- | extr d defs vs ts Ts hs (prf0 as PThm ((s, _), prf, prop, Some Ts')) =
+ | extr d defs vs ts Ts hs (prf0 as PThm ((s, _), prf, prop, SOME Ts')) =
let
val (vs', tye) = find_inst prop Ts ts vs;
val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
in
case Symtab.lookup (realizers, s) of
- None => (case find vs' (find' s defs) of
- None =>
+ NONE => (case find vs' (find' s defs) of
+ NONE =>
let
val _ = msg d ("Extracting " ^ quote s ^
(if null vs' then ""
@@ -624,7 +624,7 @@
val prf' = prep (Reconstruct.reconstruct_proof sg prop prf);
val (defs', t) = extr (d + 1) defs vs' [] [] [] prf';
val (defs'', corr_prf) =
- corr (d + 1) defs' vs' [] [] [] prf' prf' (Some t);
+ corr (d + 1) defs' vs' [] [] [] prf' prf' (SOME t);
val nt = Envir.beta_norm t;
val args = filter_out (fn v => tname_of (body_type
@@ -649,32 +649,32 @@
(chtype [propT, T] combination_axm %> f %> f %> c %> t' %%
(chtype [T --> propT] reflexive_axm %> f) %%
PAxm (cname ^ "_def", eqn,
- Some (map TVar (term_tvars eqn))))) %% corr_prf;
+ SOME (map TVar (term_tvars eqn))))) %% corr_prf;
val corr_prop = Reconstruct.prop_of corr_prf';
val corr_prf'' = foldr forall_intr_prf
(map (get_var_type corr_prop) (vfs_of prop), proof_combt
(PThm ((corr_name s vs', []), corr_prf', corr_prop,
- Some (map TVar (term_tvars corr_prop))), vfs_of corr_prop));
+ SOME (map TVar (term_tvars corr_prop))), vfs_of corr_prop));
in
((s, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'',
subst_TVars tye' u)
end
- | Some ((_, u), _) => (defs, subst_TVars tye' u))
- | Some rs => (case find vs' rs of
- Some (t, _) => (defs, subst_TVars tye' t)
- | None => error ("extr: no realizer for instance of theorem " ^
+ | SOME ((_, u), _) => (defs, subst_TVars tye' u))
+ | SOME rs => (case find vs' rs of
+ SOME (t, _) => (defs, subst_TVars tye' t)
+ | NONE => error ("extr: no realizer for instance of theorem " ^
quote s ^ ":\n" ^ Sign.string_of_term sg (Envir.beta_norm
(Reconstruct.prop_of (proof_combt (prf0, ts))))))
end
- | extr d defs vs ts Ts hs (prf0 as PAxm (s, prop, Some Ts')) =
+ | extr d defs vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) =
let
val (vs', tye) = find_inst prop Ts ts vs;
val tye' = (map fst (term_tvars prop) ~~ Ts') @ tye
in
case find vs' (Symtab.lookup_multi (realizers, s)) of
- Some (t, _) => (defs, subst_TVars tye' t)
- | None => error ("extr: no realizer for instance of axiom " ^
+ SOME (t, _) => (defs, subst_TVars tye' t)
+ | NONE => error ("extr: no realizer for instance of axiom " ^
quote s ^ ":\n" ^ Sign.string_of_term sg (Envir.beta_norm
(Reconstruct.prop_of (proof_combt (prf0, ts)))))
end
@@ -696,7 +696,7 @@
fun add_def ((s, (vs, ((t, u), (prf, _)))), thy) =
(case Sign.const_type (sign_of thy) (extr_name s vs) of
- None =>
+ NONE =>
let
val corr_prop = Reconstruct.prop_of prf;
val ft = fst (Type.freeze_thaw t);
@@ -712,7 +712,7 @@
(ProofChecker.thm_of_proof thy'
(fst (Proofterm.freeze_thaw_prf prf)))))), []) thy')
end
- | Some _ => thy);
+ | SOME _ => thy);
in thy |>
Theory.absolute_path |>
@@ -734,7 +734,7 @@
(Scan.repeat1 (P.xname -- parse_vars --| P.$$$ ":" -- P.string -- P.string) >>
(fn xs => Toplevel.theory (fn thy => add_realizers
(map (fn (((a, vs), s1), s2) =>
- (PureThy.get_thm thy (a, None), (vs, s1, s2))) xs) thy)));
+ (PureThy.get_thm thy (a, NONE), (vs, s1, s2))) xs) thy)));
val realizabilityP =
OuterSyntax.command "realizability"
@@ -749,14 +749,14 @@
val extractP =
OuterSyntax.command "extract" "extract terms from proofs" K.thy_decl
(Scan.repeat1 (P.xname -- parse_vars) >> (fn xs => Toplevel.theory
- (fn thy => extract (map (apfst (PureThy.get_thm thy o rpair None)) xs) thy)));
+ (fn thy => extract (map (apfst (PureThy.get_thm thy o rpair NONE)) xs) thy)));
val parsers = [realizersP, realizabilityP, typeofP, extractP];
val setup =
[ExtractionData.init,
- add_types [("prop", ([], None))],
+ add_types [("prop", ([], NONE))],
add_typeof_eqns
["(typeof (PROP P)) == (Type (TYPE(Null))) ==> \
--- a/src/Pure/Proof/proof_rewrite_rules.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Proof/proof_rewrite_rules.ML Sun Feb 13 17:15:14 2005 +0100
@@ -22,81 +22,81 @@
fun rew b =
let
- fun ? x = if b then Some x else None;
+ fun ? x = if b then SOME x else NONE;
fun ax (prf as PAxm (s, prop, _)) Ts =
- if b then PAxm (s, prop, Some Ts) else prf;
+ if b then PAxm (s, prop, SOME Ts) else prf;
fun ty T = if b then
let val Type (_, [Type (_, [U, _]), _]) = T
- in Some U end
- else None;
+ in SOME U end
+ else NONE;
val equal_intr_axm = ax equal_intr_axm [];
val equal_elim_axm = ax equal_elim_axm [];
val symmetric_axm = ax symmetric_axm [propT];
fun rew' _ (PThm (("ProtoPure.rev_triv_goal", _), _, _, _) % _ %%
- (PThm (("ProtoPure.triv_goal", _), _, _, _) % _ %% prf)) = Some prf
+ (PThm (("ProtoPure.triv_goal", _), _, _, _) % _ %% prf)) = SOME prf
| rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
- (PAxm ("ProtoPure.equal_intr", _, _) % _ % _ %% prf %% _)) = Some prf
+ (PAxm ("ProtoPure.equal_intr", _, _) % _ % _ %% prf %% _)) = SOME prf
| rew' _ (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
(PAxm ("ProtoPure.equal_intr", _, _) % A % B %% prf1 %% prf2)) =
- Some (equal_intr_axm % B % A %% prf2 %% prf1)
+ SOME (equal_intr_axm % B % A %% prf2 %% prf1)
- | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A) % Some (_ $ B) %%
- (PAxm ("ProtoPure.combination", _, _) % Some (Const ("Goal", _)) %
+ | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %%
+ (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("Goal", _)) %
_ % _ % _ %% (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %%
((tg as PThm (("ProtoPure.triv_goal", _), _, _, _)) % _ %% prf2)) =
- Some (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2))
+ SOME (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2))
- | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A) % Some (_ $ B) %%
+ | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %%
(PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
- (PAxm ("ProtoPure.combination", _, _) % Some (Const ("Goal", _)) %
+ (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("Goal", _)) %
_ % _ % _ %% (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1)) %%
((tg as PThm (("ProtoPure.triv_goal", _), _, _, _)) % _ %% prf2)) =
- Some (tg %> B %% (equal_elim_axm %> A %> B %%
+ SOME (tg %> B %% (equal_elim_axm %> A %> B %%
(symmetric_axm % ? B % ? A %% prf1) %% prf2))
- | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
+ | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
(PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
- (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==>", _)) % _ % _ % _ %%
+ (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==>", _)) % _ % _ % _ %%
(PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) =
let
val _ $ A $ C = Envir.beta_norm X;
val _ $ B $ D = Envir.beta_norm Y
- in Some (AbsP ("H1", ? X, AbsP ("H2", ? B,
+ in SOME (AbsP ("H1", ? X, AbsP ("H2", ? B,
equal_elim_axm %> C %> D %% incr_pboundvars 2 0 prf2 %%
(PBound 1 %% (equal_elim_axm %> B %> A %%
(symmetric_axm % ? A % ? B %% incr_pboundvars 2 0 prf1) %% PBound 0)))))
end
- | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
+ | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
(PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
(PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
- (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==>", _)) % _ % _ % _ %%
+ (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==>", _)) % _ % _ % _ %%
(PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2))) =
let
val _ $ A $ C = Envir.beta_norm Y;
val _ $ B $ D = Envir.beta_norm X
- in Some (AbsP ("H1", ? X, AbsP ("H2", ? A,
+ in SOME (AbsP ("H1", ? X, AbsP ("H2", ? A,
equal_elim_axm %> D %> C %%
(symmetric_axm % ? C % ? D %% incr_pboundvars 2 0 prf2)
%% (PBound 1 %% (equal_elim_axm %> A %> B %% incr_pboundvars 2 0 prf1 %% PBound 0)))))
end
- | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
- (PAxm ("ProtoPure.combination", _, _) % Some (Const ("all", _)) % _ % _ % _ %%
+ | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
+ (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("all", _)) % _ % _ % _ %%
(PAxm ("ProtoPure.reflexive", _, _) % _) %%
(PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf))) =
let
val Const (_, T) $ P = Envir.beta_norm X;
val _ $ Q = Envir.beta_norm Y;
- in Some (AbsP ("H", ? X, Abst ("x", ty T,
+ in SOME (AbsP ("H", ? X, Abst ("x", ty T,
equal_elim_axm %> incr_boundvars 1 P $ Bound 0 %> incr_boundvars 1 Q $ Bound 0 %%
(incr_pboundvars 1 1 prf %> Bound 0) %% (PBound 0 %> Bound 0))))
end
- | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some X % Some Y %%
+ | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % SOME X % SOME Y %%
(PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
- (PAxm ("ProtoPure.combination", _, _) % Some (Const ("all", _)) % _ % _ % _ %%
+ (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("all", _)) % _ % _ % _ %%
(PAxm ("ProtoPure.reflexive", _, _) % _) %%
(PAxm ("ProtoPure.abstract_rule", _, _) % _ % _ %% prf)))) =
let
@@ -104,81 +104,81 @@
val _ $ Q = Envir.beta_norm Y;
val t = incr_boundvars 1 P $ Bound 0;
val u = incr_boundvars 1 Q $ Bound 0
- in Some (AbsP ("H", ? X, Abst ("x", ty T,
+ in SOME (AbsP ("H", ? X, Abst ("x", ty T,
equal_elim_axm %> t %> u %%
(symmetric_axm % ? u % ? t %% (incr_pboundvars 1 1 prf %> Bound 0))
%% (PBound 0 %> Bound 0))))
end
- | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some A % Some C %%
- (PAxm ("ProtoPure.transitive", _, _) % _ % Some B % _ %% prf1 %% prf2) %% prf3) =
- Some (equal_elim_axm %> B %> C %% prf2 %%
+ | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % SOME A % SOME C %%
+ (PAxm ("ProtoPure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2) %% prf3) =
+ SOME (equal_elim_axm %> B %> C %% prf2 %%
(equal_elim_axm %> A %> B %% prf1 %% prf3))
- | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % Some A % Some C %%
+ | rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % SOME A % SOME C %%
(PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
- (PAxm ("ProtoPure.transitive", _, _) % _ % Some B % _ %% prf1 %% prf2)) %% prf3) =
- Some (equal_elim_axm %> B %> C %% (symmetric_axm % ? C % ? B %% prf1) %%
+ (PAxm ("ProtoPure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2)) %% prf3) =
+ SOME (equal_elim_axm %> B %> C %% (symmetric_axm % ? C % ? B %% prf1) %%
(equal_elim_axm %> A %> B %% (symmetric_axm % ? B % ? A %% prf2) %% prf3))
| rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
- (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf) = Some prf
+ (PAxm ("ProtoPure.reflexive", _, _) % _) %% prf) = SOME prf
| rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
(PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
- (PAxm ("ProtoPure.reflexive", _, _) % _)) %% prf) = Some prf
+ (PAxm ("ProtoPure.reflexive", _, _) % _)) %% prf) = SOME prf
| rew' _ (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
- (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %% prf)) = Some prf
+ (PAxm ("ProtoPure.symmetric", _, _) % _ % _ %% prf)) = SOME prf
| rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
- (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A $ C) % Some (_ $ B $ D) %%
+ (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %%
(PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
- (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
+ (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
(PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3) %% prf4) =
- Some (equal_elim_axm %> C %> D %% prf2 %%
+ SOME (equal_elim_axm %> C %> D %% prf2 %%
(equal_elim_axm %> A %> C %% prf3 %%
(equal_elim_axm %> B %> A %% (symmetric_axm % ? A % ? B %% prf1) %% prf4)))
| rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
(PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
- (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ A $ C) % Some (_ $ B $ D) %%
+ (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %%
(PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
- (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
+ (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
(PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3)) %% prf4) =
- Some (equal_elim_axm %> A %> B %% prf1 %%
+ SOME (equal_elim_axm %> A %> B %% prf1 %%
(equal_elim_axm %> C %> A %% (symmetric_axm % ? A % ? C %% prf3) %%
(equal_elim_axm %> D %> C %% (symmetric_axm % ? C % ? D %% prf2) %% prf4)))
| rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
- (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ B $ D) % Some (_ $ A $ C) %%
+ (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %%
(PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
(PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
- (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
+ (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
(PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3) %% prf4) =
- Some (equal_elim_axm %> D %> C %% (symmetric_axm % ? C % ? D %% prf2) %%
+ SOME (equal_elim_axm %> D %> C %% (symmetric_axm % ? C % ? D %% prf2) %%
(equal_elim_axm %> B %> D %% prf3 %%
(equal_elim_axm %> A %> B %% prf1 %% prf4)))
| rew' _ (PAxm ("ProtoPure.equal_elim", _, _) % _ % _ %%
(PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
- (PAxm ("ProtoPure.equal_elim", _, _) % Some (_ $ B $ D) % Some (_ $ A $ C) %%
+ (PAxm ("ProtoPure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %%
(PAxm ("ProtoPure.symmetric", _, _) % _ % _ %%
(PAxm ("ProtoPure.combination", _, _) % _ % _ % _ % _ %%
- (PAxm ("ProtoPure.combination", _, _) % Some (Const ("==", _)) % _ % _ % _ %%
+ (PAxm ("ProtoPure.combination", _, _) % SOME (Const ("==", _)) % _ % _ % _ %%
(PAxm ("ProtoPure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3)) %% prf4) =
- Some (equal_elim_axm %> B %> A %% (symmetric_axm % ? A % ? B %% prf1) %%
+ SOME (equal_elim_axm %> B %> A %% (symmetric_axm % ? A % ? B %% prf1) %%
(equal_elim_axm %> D %> B %% (symmetric_axm % ? B % ? D %% prf3) %%
(equal_elim_axm %> C %> D %% prf2 %% prf4)))
| rew' _ ((prf as PAxm ("ProtoPure.combination", _, _) %
- Some ((eq as Const ("==", T)) $ t) % _ % _ % _) %%
+ SOME ((eq as Const ("==", T)) $ t) % _ % _ % _) %%
(PAxm ("ProtoPure.reflexive", _, _) % _)) =
let val (U, V) = (case T of
Type (_, [U, V]) => (U, V) | _ => (dummyT, dummyT))
- in Some (prf %% (ax combination_axm [V, U] %> eq % ? eq % ? t % ? t %%
+ in SOME (prf %% (ax combination_axm [V, U] %> eq % ? eq % ? t % ? t %%
(ax reflexive_axm [T] % ? eq) %% (ax reflexive_axm [U] % ? t)))
end
- | rew' _ _ = None;
+ | rew' _ _ = NONE;
in rew' end;
fun rprocs b = [("Pure/meta_equality", rew b)];
@@ -201,9 +201,9 @@
end;
fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2
- | rew Ts (prf % Some t) = rew Ts prf % Some (rew_term Ts t)
- | rew Ts (Abst (s, Some T, prf)) = Abst (s, Some T, rew (T :: Ts) prf)
- | rew Ts (AbsP (s, Some t, prf)) = AbsP (s, Some (rew_term Ts t), rew Ts prf)
+ | rew Ts (prf % SOME t) = rew Ts prf % SOME (rew_term Ts t)
+ | rew Ts (Abst (s, SOME T, prf)) = Abst (s, SOME T, rew (T :: Ts) prf)
+ | rew Ts (AbsP (s, SOME t, prf)) = AbsP (s, SOME (rew_term Ts t), rew Ts prf)
| rew _ prf = prf
in rew [] end;
@@ -216,12 +216,12 @@
fun insert_refl defs Ts (prf1 %% prf2) =
insert_refl defs Ts prf1 %% insert_refl defs Ts prf2
- | insert_refl defs Ts (Abst (s, Some T, prf)) =
- Abst (s, Some T, insert_refl defs (T :: Ts) prf)
+ | insert_refl defs Ts (Abst (s, SOME T, prf)) =
+ Abst (s, SOME T, insert_refl defs (T :: Ts) prf)
| insert_refl defs Ts (AbsP (s, t, prf)) =
AbsP (s, t, insert_refl defs Ts prf)
| insert_refl defs Ts prf = (case strip_combt prf of
- (PThm ((s, _), _, prop, Some Ts), ts) =>
+ (PThm ((s, _), _, prop, SOME Ts), ts) =>
if s mem defs then
let
val vs = vars_of prop;
@@ -231,7 +231,7 @@
(foldr (fn p => Abs ("", dummyT, abstract_over p)) (vs, rhs)),
map the ts);
in
- change_type (Some [fastype_of1 (Ts, rhs')]) reflexive_axm %> rhs'
+ change_type (SOME [fastype_of1 (Ts, rhs')]) reflexive_axm %> rhs'
end
else prf
| (_, []) => prf
@@ -247,7 +247,7 @@
val cnames = map (fst o dest_Const o fst) defs';
val thms = flat (map (fn (s, ps) =>
if s mem defnames then []
- else map (pair s o Some o fst) (filter_out (fn (p, _) =>
+ else map (pair s o SOME o fst) (filter_out (fn (p, _) =>
null (term_consts p inter cnames)) ps))
(Symtab.dest (thms_of_proof Symtab.empty prf)))
in Reconstruct.expand_proof sign thms end
--- a/src/Pure/Proof/proof_syntax.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Proof/proof_syntax.ML Sun Feb 13 17:15:14 2005 +0100
@@ -133,41 +133,41 @@
fun prf_of [] (Bound i) = PBound i
| prf_of Ts (Const (s, Type ("proof", _))) =
- change_type (if ty then Some Ts else None)
+ change_type (if ty then SOME Ts else NONE)
(case NameSpace.unpack s of
"axm" :: xs =>
let
val name = NameSpace.pack xs;
val prop = (case assoc (axms, name) of
- Some prop => prop
- | None => error ("Unknown axiom " ^ quote name))
- in PAxm (name, prop, None) end
+ SOME prop => prop
+ | NONE => error ("Unknown axiom " ^ quote name))
+ in PAxm (name, prop, NONE) end
| "thm" :: xs =>
let val name = NameSpace.pack xs;
in (case assoc (thms, name) of
- Some thm => fst (strip_combt (Thm.proof_of thm))
- | None => (case Symtab.lookup (tab, name) of
- Some prf => prf
- | None => error ("Unknown theorem " ^ quote name)))
+ SOME thm => fst (strip_combt (Thm.proof_of thm))
+ | NONE => (case Symtab.lookup (tab, name) of
+ SOME prf => prf
+ | NONE => error ("Unknown theorem " ^ quote name)))
end
| _ => error ("Illegal proof constant name: " ^ quote s))
| prf_of Ts (Const ("Hyp", _) $ prop) = Hyp prop
| prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v
| prf_of [] (Const ("Abst", _) $ Abs (s, T, prf)) =
- Abst (s, if ty then Some T else None,
+ Abst (s, if ty then SOME T else NONE,
incr_pboundvars (~1) 0 (prf_of [] prf))
| prf_of [] (Const ("AbsP", _) $ t $ Abs (s, _, prf)) =
AbsP (s, case t of
- Const ("dummy_pattern", _) => None
- | _ $ Const ("dummy_pattern", _) => None
- | _ => Some (mk_term t),
+ Const ("dummy_pattern", _) => NONE
+ | _ $ Const ("dummy_pattern", _) => NONE
+ | _ => SOME (mk_term t),
incr_pboundvars 0 (~1) (prf_of [] prf))
| prf_of [] (Const ("AppP", _) $ prf1 $ prf2) =
prf_of [] prf1 %% prf_of [] prf2
| prf_of Ts (Const ("Appt", _) $ prf $ Const ("TYPE", Type (_, [T]))) =
prf_of (T::Ts) prf
| prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf %
- (case t of Const ("dummy_pattern", _) => None | _ => Some (mk_term t))
+ (case t of Const ("dummy_pattern", _) => NONE | _ => SOME (mk_term t))
| prf_of _ t = error ("Not a proof term:\n" ^
Sign.string_of_term (sign_of thy) t)
@@ -183,12 +183,12 @@
val mk_tyapp = foldl (fn (prf, T) => Const ("Appt",
[proofT, itselfT T] ---> proofT) $ prf $ Logic.mk_type T);
-fun term_of _ (PThm ((name, _), _, _, None)) =
+fun term_of _ (PThm ((name, _), _, _, NONE)) =
Const (add_prefix "thm" name, proofT)
- | term_of _ (PThm ((name, _), _, _, Some Ts)) =
+ | term_of _ (PThm ((name, _), _, _, SOME Ts)) =
mk_tyapp (Const (add_prefix "thm" name, proofT), Ts)
- | term_of _ (PAxm (name, _, None)) = Const (add_prefix "axm" name, proofT)
- | term_of _ (PAxm (name, _, Some Ts)) =
+ | term_of _ (PAxm (name, _, NONE)) = Const (add_prefix "axm" name, proofT)
+ | term_of _ (PAxm (name, _, SOME Ts)) =
mk_tyapp (Const (add_prefix "axm" name, proofT), Ts)
| term_of _ (PBound i) = Bound i
| term_of Ts (Abst (s, opT, prf)) =
--- a/src/Pure/Proof/proofchecker.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Proof/proofchecker.ML Sun Feb 13 17:15:14 2005 +0100
@@ -23,8 +23,8 @@
(flat (map thms_of (thy :: Theory.ancestors_of thy)), Symtab.empty)
in
(fn s => case Symtab.lookup (tab, s) of
- None => error ("Unknown theorem " ^ quote s)
- | Some thm => thm)
+ NONE => error ("Unknown theorem " ^ quote s)
+ | SOME thm => thm)
end;
val beta_eta_convert =
@@ -45,7 +45,7 @@
Thm.instantiate (ctye, []) (forall_intr_vars (forall_intr_frees thm'))
end;
- fun thm_of _ _ (PThm ((name, _), _, prop', Some Ts)) =
+ fun thm_of _ _ (PThm ((name, _), _, prop', SOME Ts)) =
let
val thm = Thm.implies_intr_hyps (lookup name);
val {prop, ...} = rep_thm thm;
@@ -55,12 +55,12 @@
Sign.string_of_term sg prop');
in thm_of_atom thm Ts end
- | thm_of _ _ (PAxm (name, _, Some Ts)) =
+ | thm_of _ _ (PAxm (name, _, SOME Ts)) =
thm_of_atom (get_axiom thy name) Ts
| thm_of _ Hs (PBound i) = nth_elem (i, Hs)
- | thm_of vs Hs (Abst (s, Some T, prf)) =
+ | thm_of vs Hs (Abst (s, SOME T, prf)) =
let
val x = variant (names @ map fst vs) s;
val thm = thm_of ((x, T) :: vs) Hs prf
@@ -68,13 +68,13 @@
Thm.forall_intr (Thm.cterm_of sg (Free (x, T))) thm
end
- | thm_of vs Hs (prf % Some t) =
+ | thm_of vs Hs (prf % SOME t) =
let
val thm = thm_of vs Hs prf
val ct = Thm.cterm_of sg (Term.subst_bounds (map Free vs, t))
in Thm.forall_elim ct thm end
- | thm_of vs Hs (AbsP (s, Some t, prf)) =
+ | thm_of vs Hs (AbsP (s, SOME t, prf)) =
let
val ct = Thm.cterm_of sg (Term.subst_bounds (map Free vs, t));
val thm = thm_of vs (Thm.assume ct :: Hs) prf
--- a/src/Pure/Proof/reconstruct.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Proof/reconstruct.ML Sun Feb 13 17:15:14 2005 +0100
@@ -35,7 +35,7 @@
fun forall_intr_prf (t, prf) =
let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
- in Abst (a, Some T, prf_abstract_over t prf) end;
+ in Abst (a, SOME T, prf_abstract_over t prf) end;
fun forall_intr_vfs_prf prop prf = foldr forall_intr_prf
(vars_of prop @ sort (make_ord atless) (term_frees prop), prf);
@@ -68,13 +68,13 @@
Sign.string_of_typ sg T ^ "\n\n" ^ Sign.string_of_typ sg U);
fun chaseT (env as Envir.Envir {iTs, ...}) (T as TVar (ixn, _)) =
- (case Vartab.lookup (iTs, ixn) of None => T | Some T' => chaseT env T')
+ (case Vartab.lookup (iTs, ixn) of NONE => T | SOME T' => chaseT env T')
| chaseT _ T = T;
fun infer_type sg (env as Envir.Envir {maxidx, asol, iTs}) Ts vTs
(t as Const (s, T)) = if T = dummyT then (case Sign.const_type sg s of
- None => error ("reconstruct_proof: No such constant: " ^ quote s)
- | Some T =>
+ NONE => error ("reconstruct_proof: No such constant: " ^ quote s)
+ | SOME T =>
let val T' = incr_tvar (maxidx + 1) T
in (Const (s, T'), T', vTs,
Envir.Envir {maxidx = maxidx + 1, asol = asol, iTs = iTs})
@@ -82,10 +82,10 @@
else (t, T, vTs, env)
| infer_type sg env Ts vTs (t as Free (s, T)) =
if T = dummyT then (case Symtab.lookup (vTs, s) of
- None =>
+ NONE =>
let val (env', T) = mk_tvar (env, [])
in (Free (s, T), T, Symtab.update_new ((s, T), vTs), env') end
- | Some T => (Free (s, T), T, vTs, env))
+ | SOME T => (Free (s, T), T, vTs, env))
else (t, T, vTs, env)
| infer_type sg env Ts vTs (Var _) = error "reconstruct_proof: internal error"
| infer_type sg env Ts vTs (Abs (s, T, t)) =
@@ -146,13 +146,13 @@
val tfrees = term_tfrees prop;
val (prop', fmap) = Type.varify (prop, []);
val (env', Ts) = (case opTs of
- None => foldl_map mk_tvar (env, map snd tvars @ map snd tfrees)
- | Some Ts => (env, Ts));
+ NONE => foldl_map mk_tvar (env, map snd tvars @ map snd tfrees)
+ | SOME Ts => (env, Ts));
val prop'' = subst_TVars (map fst tvars @ map snd fmap ~~ Ts)
(forall_intr_vfs prop') handle LIST _ =>
error ("Wrong number of type arguments for " ^
quote (fst (get_name_tags [] prop prf)))
- in (prop'', change_type (Some Ts) prf, [], env', vTs) end;
+ in (prop'', change_type (SOME Ts) prf, [], env', vTs) end;
fun head_norm (prop, prf, cnstrts, env, vTs) =
(Envir.head_norm env prop, prf, cnstrts, env, vTs);
@@ -161,23 +161,23 @@
| mk_cnstrts env Ts Hs vTs (Abst (s, opT, cprf)) =
let
val (env', T) = (case opT of
- None => mk_tvar (env, []) | Some T => (env, T));
+ NONE => mk_tvar (env, []) | SOME T => (env, T));
val (t, prf, cnstrts, env'', vTs') =
mk_cnstrts env' (T::Ts) (map (incr_boundvars 1) Hs) vTs cprf;
- in (Const ("all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, Some T, prf),
+ in (Const ("all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, SOME T, prf),
cnstrts, env'', vTs')
end
- | mk_cnstrts env Ts Hs vTs (AbsP (s, Some t, cprf)) =
+ | mk_cnstrts env Ts Hs vTs (AbsP (s, SOME t, cprf)) =
let
val (t', _, vTs', env') = infer_type sg env Ts vTs t;
val (u, prf, cnstrts, env'', vTs'') = mk_cnstrts env' Ts (t'::Hs) vTs' cprf;
- in (Logic.mk_implies (t', u), AbsP (s, Some t', prf), cnstrts, env'', vTs'')
+ in (Logic.mk_implies (t', u), AbsP (s, SOME t', prf), cnstrts, env'', vTs'')
end
- | mk_cnstrts env Ts Hs vTs (AbsP (s, None, cprf)) =
+ | mk_cnstrts env Ts Hs vTs (AbsP (s, NONE, cprf)) =
let
val (env', t) = mk_var env Ts propT;
val (u, prf, cnstrts, env'', vTs') = mk_cnstrts env' Ts (t::Hs) vTs cprf;
- in (Logic.mk_implies (t, u), AbsP (s, Some t, prf), cnstrts, env'', vTs')
+ in (Logic.mk_implies (t, u), AbsP (s, SOME t, prf), cnstrts, env'', vTs')
end
| mk_cnstrts env Ts Hs vTs (cprf1 %% cprf2) =
let val (u, prf2, cnstrts, env', vTs') = mk_cnstrts env Ts Hs vTs cprf2
@@ -191,27 +191,27 @@
env''' vTs'' (t, Logic.mk_implies (u, v))
end)
end
- | mk_cnstrts env Ts Hs vTs (cprf % Some t) =
+ | mk_cnstrts env Ts Hs vTs (cprf % SOME t) =
let val (t', U, vTs1, env1) = infer_type sg env Ts vTs t
in (case head_norm (mk_cnstrts env1 Ts Hs vTs1 cprf) of
(Const ("all", Type ("fun", [Type ("fun", [T, _]), _])) $ f,
prf, cnstrts, env2, vTs2) =>
let val env3 = unifyT sg env2 T U
- in (betapply (f, t'), prf % Some t', cnstrts, env3, vTs2)
+ in (betapply (f, t'), prf % SOME t', cnstrts, env3, vTs2)
end
| (u, prf, cnstrts, env2, vTs2) =>
let val (env3, v) = mk_var env2 Ts (U --> propT);
in
- add_cnstrt Ts (v $ t') (prf % Some t') cnstrts env3 vTs2
+ add_cnstrt Ts (v $ t') (prf % SOME t') cnstrts env3 vTs2
(u, Const ("all", (U --> propT) --> propT) $ v)
end)
end
- | mk_cnstrts env Ts Hs vTs (cprf % None) =
+ | mk_cnstrts env Ts Hs vTs (cprf % NONE) =
(case head_norm (mk_cnstrts env Ts Hs vTs cprf) of
(Const ("all", Type ("fun", [Type ("fun", [T, _]), _])) $ f,
prf, cnstrts, env', vTs') =>
let val (env'', t) = mk_var env' Ts T
- in (betapply (f, t), prf % Some t, cnstrts, env'', vTs')
+ in (betapply (f, t), prf % SOME t, cnstrts, env'', vTs')
end
| (u, prf, cnstrts, env', vTs') =>
let
@@ -219,7 +219,7 @@
val (env2, v) = mk_var env1 Ts (T --> propT);
val (env3, t) = mk_var env2 Ts T
in
- add_cnstrt Ts (v $ t) (prf % Some t) cnstrts env3 vTs'
+ add_cnstrt Ts (v $ t) (prf % SOME t) cnstrts env3 vTs'
(u, Const ("all", (T --> propT) --> propT) $ v)
end)
| mk_cnstrts env _ _ vTs (prf as PThm (_, _, prop, opTs)) =
@@ -294,7 +294,7 @@
fun reconstruct_proof sg prop cprf =
let
- val (cprf' % Some prop', thawf) = freeze_thaw_prf (cprf % Some prop);
+ val (cprf' % SOME prop', thawf) = freeze_thaw_prf (cprf % SOME prop);
val _ = message "Collecting constraints...";
val (t, prf, cs, env, _) = make_constraints_cprf sg
(Envir.empty (maxidx_of_proof cprf)) cprf';
@@ -315,20 +315,20 @@
val head_norm = Envir.head_norm (Envir.empty 0);
fun prop_of0 Hs (PBound i) = nth_elem (i, Hs)
- | prop_of0 Hs (Abst (s, Some T, prf)) =
+ | prop_of0 Hs (Abst (s, SOME T, prf)) =
all T $ (Abs (s, T, prop_of0 Hs prf))
- | prop_of0 Hs (AbsP (s, Some t, prf)) =
+ | prop_of0 Hs (AbsP (s, SOME t, prf)) =
Logic.mk_implies (t, prop_of0 (t :: Hs) prf)
- | prop_of0 Hs (prf % Some t) = (case head_norm (prop_of0 Hs prf) of
+ | prop_of0 Hs (prf % SOME t) = (case head_norm (prop_of0 Hs prf) of
Const ("all", _) $ f => f $ t
| _ => error "prop_of: all expected")
| prop_of0 Hs (prf1 %% prf2) = (case head_norm (prop_of0 Hs prf1) of
Const ("==>", _) $ P $ Q => Q
| _ => error "prop_of: ==> expected")
| prop_of0 Hs (Hyp t) = t
- | prop_of0 Hs (PThm (_, _, prop, Some Ts)) = prop_of_atom prop Ts
- | prop_of0 Hs (PAxm (_, prop, Some Ts)) = prop_of_atom prop Ts
- | prop_of0 Hs (Oracle (_, prop, Some Ts)) = prop_of_atom prop Ts
+ | prop_of0 Hs (PThm (_, _, prop, SOME Ts)) = prop_of_atom prop Ts
+ | prop_of0 Hs (PAxm (_, prop, SOME Ts)) = prop_of_atom prop Ts
+ | prop_of0 Hs (Oracle (_, prop, SOME Ts)) = prop_of_atom prop Ts
| prop_of0 _ _ = error "prop_of: partial proof object";
val prop_of' = Pattern.eta_contract oo (Envir.beta_norm oo prop_of0);
@@ -353,16 +353,16 @@
| expand maxidx prfs (prf % t) =
let val (maxidx', prfs', prf') = expand maxidx prfs prf
in (maxidx', prfs', prf' % t) end
- | expand maxidx prfs (prf as PThm ((a, _), cprf, prop, Some Ts)) =
+ | expand maxidx prfs (prf as PThm ((a, _), cprf, prop, SOME Ts)) =
if not (exists
- (fn (b, None) => a = b
- | (b, Some prop') => a = b andalso prop = prop') thms)
+ (fn (b, NONE) => a = b
+ | (b, SOME prop') => a = b andalso prop = prop') thms)
then (maxidx, prfs, prf) else
let
fun inc i =
map_proof_terms (Logic.incr_indexes ([], i)) (incr_tvar i);
val (maxidx', prf, prfs') = (case assoc (prfs, (a, prop)) of
- None =>
+ NONE =>
let
val _ = message ("Reconstructing proof of " ^ a);
val _ = message (Sign.string_of_term sg prop);
@@ -373,7 +373,7 @@
in (maxidx' + maxidx + 1, inc (maxidx + 1) prf,
((a, prop), (maxidx', prf)) :: prfs')
end
- | Some (maxidx', prf) => (maxidx' + maxidx + 1,
+ | SOME (maxidx', prf) => (maxidx' + maxidx + 1,
inc (maxidx + 1) prf, prfs));
val tfrees = term_tfrees prop;
val tye = map (fn ((s, j), _) => (s, maxidx + 1 + j))
--- a/src/Pure/Syntax/ast.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Syntax/ast.ML Sun Feb 13 17:15:14 2005 +0100
@@ -127,9 +127,9 @@
val lvars = vars_of lhs;
val rvars = vars_of rhs;
in
- if not (unique lvars) then Some "duplicate vars in lhs"
- else if not (rvars subset lvars) then Some "rhs contains extra variables"
- else None
+ if not (unique lvars) then SOME "duplicate vars in lhs"
+ else if not (rvars subset lvars) then SOME "rhs contains extra variables"
+ else NONE
end;
@@ -195,7 +195,7 @@
end
| _ => (ast, []));
in
- Some (mtch head pat Symtab.empty, args) handle NO_MATCH => None
+ SOME (mtch head pat Symtab.empty, args) handle NO_MATCH => NONE
end;
@@ -215,15 +215,15 @@
fun try_rules ((lhs, rhs) :: pats) ast =
(case match ast lhs of
- Some (env, args) =>
- (inc changes; Some (mk_appl (subst env rhs) args))
- | None => (inc failed_matches; try_rules pats ast))
- | try_rules [] _ = None;
+ SOME (env, args) =>
+ (inc changes; SOME (mk_appl (subst env rhs) args))
+ | NONE => (inc failed_matches; try_rules pats ast))
+ | try_rules [] _ = NONE;
val try_headless_rules = try_rules (get_rules "");
fun try ast a =
(case try_rules (get_rules a) ast of
- None => try_headless_rules ast
+ NONE => try_headless_rules ast
| some => some);
fun rewrite (ast as Constant a) = try ast a
@@ -238,8 +238,8 @@
fun norm_root ast =
(case rewrite ast of
- Some new_ast => (rewrote ast new_ast; norm_root new_ast)
- | None => ast);
+ SOME new_ast => (rewrote ast new_ast; norm_root new_ast)
+ | NONE => ast);
fun norm ast =
(case norm_root ast of
--- a/src/Pure/Syntax/lexicon.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Syntax/lexicon.ML Sun Feb 13 17:15:14 2005 +0100
@@ -217,15 +217,15 @@
(* predef_term *)
-fun predef_term "id" = Some (IdentSy "id")
- | predef_term "longid" = Some (LongIdentSy "longid")
- | predef_term "var" = Some (VarSy "var")
- | predef_term "tid" = Some (TFreeSy "tid")
- | predef_term "tvar" = Some (TVarSy "tvar")
- | predef_term "num" = Some (NumSy "num")
- | predef_term "xnum" = Some (XNumSy "xnum")
- | predef_term "xstr" = Some (StrSy "xstr")
- | predef_term _ = None;
+fun predef_term "id" = SOME (IdentSy "id")
+ | predef_term "longid" = SOME (LongIdentSy "longid")
+ | predef_term "var" = SOME (VarSy "var")
+ | predef_term "tid" = SOME (TFreeSy "tid")
+ | predef_term "tvar" = SOME (TVarSy "tvar")
+ | predef_term "num" = SOME (NumSy "num")
+ | predef_term "xnum" = SOME (XNumSy "xnum")
+ | predef_term "xstr" = SOME (StrSy "xstr")
+ | predef_term _ = NONE;
(* xstr tokens *)
@@ -247,7 +247,7 @@
fun explode_xstr str =
(case Scan.read Symbol.stopper scan_str (Symbol.explode str) of
- Some cs => cs
+ SOME cs => cs
| _ => error ("Inner lexical error: literal string expected at " ^ quote str));
@@ -286,10 +286,10 @@
val scan_lit = Scan.literal lex >> (pair Token o implode);
val scan_token =
- scan_comment >> K None ||
- Scan.max (op <= o pairself snd) scan_lit scan_val >> (fn (tk, s) => Some (tk s)) ||
- scan_str >> (Some o StrSy o implode_xstr) ||
- Scan.one Symbol.is_blank >> K None;
+ scan_comment >> K NONE ||
+ Scan.max (op <= o pairself snd) scan_lit scan_val >> (fn (tk, s) => SOME (tk s)) ||
+ scan_str >> (SOME o StrSy o implode_xstr) ||
+ Scan.one Symbol.is_blank >> K NONE;
in
(case Scan.error (Scan.finite Symbol.stopper (Scan.repeat scan_token)) chs of
(toks, []) => mapfilter I toks @ [EndToken]
@@ -330,7 +330,7 @@
fun indexname cs =
(case Scan.read Symbol.stopper scan_indexname (Symbol.explode cs) of
- Some xi => xi
+ SOME xi => xi
| _ => error ("Lexical error in variable name: " ^ quote cs));
@@ -385,9 +385,9 @@
val idents = Scan.repeat1 (blanks |-- scan_id --| blanks) -- junk;
in
(case Scan.read Symbol.stopper idents (Symbol.explode str) of
- Some (ids, []) => ids
- | Some (_, bad) => error ("Bad identifier: " ^ quote (implode bad))
- | None => error ("No identifier found in: " ^ quote str))
+ SOME (ids, []) => ids
+ | SOME (_, bad) => error ("Bad identifier: " ^ quote (implode bad))
+ | NONE => error ("No identifier found in: " ^ quote str))
end;
end;
--- a/src/Pure/Syntax/mixfix.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Syntax/mixfix.ML Sun Feb 13 17:15:14 2005 +0100
@@ -141,13 +141,13 @@
fun mfix_of decl =
let val t = name_of decl in
(case decl of
- (_, _, NoSyn) => None
- | (_, 2, InfixName (sy, p)) => Some (mk_infix sy t (p + 1) (p + 1) p)
- | (_, 2, InfixlName (sy, p)) => Some (mk_infix sy t p (p + 1) p)
- | (_, 2, InfixrName (sy, p)) => Some (mk_infix sy t (p + 1) p p)
- | (sy, 2, Infix p) => Some (mk_infix sy t (p + 1) (p + 1) p)
- | (sy, 2, Infixl p) => Some (mk_infix sy t p (p + 1) p)
- | (sy, 2, Infixr p) => Some (mk_infix sy t (p + 1) p p)
+ (_, _, NoSyn) => NONE
+ | (_, 2, InfixName (sy, p)) => SOME (mk_infix sy t (p + 1) (p + 1) p)
+ | (_, 2, InfixlName (sy, p)) => SOME (mk_infix sy t p (p + 1) p)
+ | (_, 2, InfixrName (sy, p)) => SOME (mk_infix sy t (p + 1) p p)
+ | (sy, 2, Infix p) => SOME (mk_infix sy t (p + 1) (p + 1) p)
+ | (sy, 2, Infixl p) => SOME (mk_infix sy t p (p + 1) p)
+ | (sy, 2, Infixr p) => SOME (mk_infix sy t (p + 1) p p)
| _ => error ("Bad mixfix declaration for type: " ^ quote t))
end;
@@ -186,8 +186,8 @@
[SynExt.Mfix ("(3" ^ sy ^ "_./ _)", binder_typ c ty, sy, [0, p], q)])
end;
- fun binder (c, _, Binder (sy, _, _)) = Some (sy, c)
- | binder _ = None;
+ fun binder (c, _, Binder (sy, _, _)) = SOME (sy, c)
+ | binder _ = NONE;
val mfix = flat (map mfix_of const_decls);
val xconsts = map name_of const_decls;
--- a/src/Pure/Syntax/parser.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Syntax/parser.ML Sun Feb 13 17:15:14 2005 +0100
@@ -63,7 +63,7 @@
(* convert productions to grammar;
N.B. that the chains parameter has the form [(from, [to])];
- prod_count is of type "int option" and is only updated if it is <> None*)
+ prod_count is of type "int option" and is only updated if it is <> NONE*)
fun add_prods _ chains lambdas prod_count [] = (chains, lambdas, prod_count)
| add_prods prods chains lambdas prod_count
((lhs, new_prod as (rhs, name, pri)) :: ps) =
@@ -73,13 +73,13 @@
let (*store chain if it does not already exist*)
fun store_chain from =
let val old_tos = assocs chains from;
- in if lhs mem old_tos then (None, chains)
- else (Some from, overwrite (chains, (from, lhs ins old_tos)))
+ in if lhs mem old_tos then (NONE, chains)
+ else (SOME from, overwrite (chains, (from, lhs ins old_tos)))
end;
in if pri = ~1 then
case rhs of [Nonterminal (id, ~1)] => store_chain id
- | _ => (None, chains)
- else (None, chains)
+ | _ => (NONE, chains)
+ else (NONE, chains)
end;
(*propagate new chain in lookahead and lambda lists;
@@ -118,12 +118,12 @@
(*list optional terminal and all nonterminals on which the lookahead
of a production depends*)
- fun lookahead_dependency _ [] nts = (None, nts)
- | lookahead_dependency _ ((Terminal tk) :: _) nts = (Some tk, nts)
+ fun lookahead_dependency _ [] nts = (NONE, nts)
+ | lookahead_dependency _ ((Terminal tk) :: _) nts = (SOME tk, nts)
| lookahead_dependency lambdas ((Nonterminal (nt, _)) :: symbs) nts =
if nt mem lambdas then
lookahead_dependency lambdas symbs (nt :: nts)
- else (None, nt :: nts);
+ else (NONE, nt :: nts);
(*get all known starting tokens for a nonterminal*)
fun starts_for_nt nt = snd (fst (Array.sub (prods, nt)));
@@ -170,8 +170,8 @@
end;
val nt_prods' =
- let val new_opt_tks = map Some new_tks;
- in copy ((if new_lambda then [None] else []) @
+ let val new_opt_tks = map SOME new_tks;
+ in copy ((if new_lambda then [NONE] else []) @
new_opt_tks) nt_prods
end;
in examine_prods ps (add_lambda orelse new_lambda)
@@ -193,7 +193,7 @@
(*existing productions whose lookahead may depend on l*)
val tk_prods =
assocs nt_prods
- (Some (hd l_starts handle LIST _ => UnknownStart));
+ (SOME (hd l_starts handle LIST _ => UnknownStart));
(*add_lambda is true if an existing production of the nt
produces lambda due to the new lambda NT l*)
@@ -242,9 +242,9 @@
(if is_some start_tk then [the start_tk] else [],
map starts_for_nt start_nts);
- val opt_starts = (if new_lambda then [None]
- else if null start_tks then [Some UnknownStart]
- else []) @ (map Some start_tks);
+ val opt_starts = (if new_lambda then [NONE]
+ else if null start_tks then [SOME UnknownStart]
+ else []) @ (map SOME start_tks);
(*add lhs NT to list of dependent NTs in lookahead*)
fun add_nts [] = ()
@@ -271,7 +271,7 @@
| store (tk :: tks) prods is_new =
let val tk_prods = assocs prods tk;
- (*if prod_count = None then we can assume that
+ (*if prod_count = NONE then we can assume that
grammar does not contain new production already*)
val (tk_prods', is_new') =
if is_some prod_count then
@@ -309,11 +309,11 @@
val key =
case find_first (fn t => not (t mem new_tks))
(starts_for_nt changed_nt) of
- None => Some UnknownStart
+ NONE => SOME UnknownStart
| t => t;
(*copy productions whose lookahead depends on changed_nt;
- if key = Some UnknownToken then tk_prods is used to hold
+ if key = SOME UnknownToken then tk_prods is used to hold
the productions not copied*)
fun update_prods [] result = result
| update_prods ((p as (rhs, _, _)) :: ps)
@@ -335,7 +335,7 @@
fun copy [] nt_prods = nt_prods
| copy (tk :: tks) nt_prods =
let
- val tk_prods = assocs nt_prods (Some tk);
+ val tk_prods = assocs nt_prods (SOME tk);
val tk_prods' =
if not lambda then p :: tk_prods
@@ -343,12 +343,12 @@
(*if production depends on lambda NT we
have to look for duplicates*)
in copy tks
- (overwrite (nt_prods, (Some tk, tk_prods')))
+ (overwrite (nt_prods, (SOME tk, tk_prods')))
end;
val result =
if update then
(tk_prods, copy new_tks nt_prods)
- else if key = Some UnknownStart then
+ else if key = SOME UnknownStart then
(p :: tk_prods, nt_prods)
else (tk_prods, nt_prods);
in update_prods ps result end;
@@ -367,7 +367,7 @@
update_prods tk_prods ([], nt_prods);
val nt_prods' =
- if key = Some UnknownStart then
+ if key = SOME UnknownStart then
overwrite (nt_prods', (key, tk_prods'))
else nt_prods';
@@ -449,8 +449,8 @@
(*Get tag for existing nonterminal or create a new one*)
fun get_tag nt_count tags nt =
case Symtab.lookup (tags, nt) of
- Some tag => (nt_count, tags, tag)
- | None => (nt_count+1, Symtab.update_new ((nt, nt_count), tags),
+ SOME tag => (nt_count, tags, tag)
+ | NONE => (nt_count+1, Symtab.update_new ((nt, nt_count), tags),
nt_count);
(*Convert symbols to the form used by the parser;
@@ -463,10 +463,10 @@
let
val (nt_count', tags', new_symb) =
case predef_term s of
- None =>
+ NONE =>
let val (nt_count', tags', s_tag) = get_tag nt_count tags s;
in (nt_count', tags', Nonterminal (s_tag, p)) end
- | Some tk => (nt_count, tags, Terminal tk);
+ | SOME tk => (nt_count, tags, Terminal tk);
in symb_of ss nt_count' tags' (new_symb :: result) end
| symb_of (_ :: ss) nt_count tags result =
symb_of ss nt_count tags result;
@@ -499,7 +499,7 @@
(*Add new productions to old ones*)
val (fromto_chains', lambdas', _) =
- add_prods prods' fromto_chains lambdas None xprods';
+ add_prods prods' fromto_chains lambdas NONE xprods';
val chains' = inverse_chains fromto_chains' [];
in Gram {nt_count = nt_count', prod_count = prod_count', tags = tags',
@@ -524,8 +524,8 @@
(*get existing tag from grammar1 or create a new one*)
fun get_tag nt_count tags nt =
case Symtab.lookup (tags, nt) of
- Some tag => (nt_count, tags, tag)
- | None => (nt_count+1, Symtab.update_new ((nt, nt_count), tags),
+ SOME tag => (nt_count, tags, tag)
+ | NONE => (nt_count+1, Symtab.update_new ((nt, nt_count), tags),
nt_count)
val ((nt_count1', tags1'), tag_table) =
@@ -591,8 +591,8 @@
val fromto_chains = inverse_chains chains1 [];
- val (fromto_chains', lambdas', Some prod_count1') =
- add_prods prods1' fromto_chains lambdas1 (Some prod_count1) raw_prods;
+ val (fromto_chains', lambdas', SOME prod_count1') =
+ add_prods prods1' fromto_chains lambdas1 (SOME prod_count1) raw_prods;
val chains' = inverse_chains fromto_chains' [];
in Gram {nt_count = nt_count1', prod_count = prod_count1',
@@ -629,10 +629,10 @@
(*Add parse tree to list and eliminate duplicates
saving the maximum precedence*)
-fun conc (t, prec:int) [] = (None, [(t, prec)])
+fun conc (t, prec:int) [] = (NONE, [(t, prec)])
| conc (t, prec) ((t', prec') :: ts) =
if t = t' then
- (Some prec', if prec' >= prec then (t', prec') :: ts
+ (SOME prec', if prec' >= prec then (t', prec') :: ts
else (t, prec) :: ts)
else
let val (n, ts') = conc (t, prec) ts
@@ -697,7 +697,7 @@
(*get all productions of a NT and NTs chained to it which can
be started by specified token*)
fun prods_for prods chains include_none tk nts =
-let (*similar to token_assoc but does not automatically include 'None' key*)
+let (*similar to token_assoc but does not automatically include 'NONE' key*)
fun token_assoc2 (list, key) =
let fun assoc [] result = result
| assoc ((keyi, pi) :: pairs) result =
@@ -725,7 +725,7 @@
let (*predictor operation*)
val (used', new_states) =
(case assoc (used, nt) of
- Some (usedPrec, l) => (*nonterminal has been processed*)
+ SOME (usedPrec, l) => (*nonterminal has been processed*)
if usedPrec <= minPrec then
(*wanted precedence has been processed*)
(used, movedot_lambda S l)
@@ -739,7 +739,7 @@
movedot_lambda S l @ States')
end
- | None => (*nonterminal is parsed for the first time*)
+ | NONE => (*nonterminal is parsed for the first time*)
let val tk_prods = all_prods_for nt;
val States' = mkStates i minPrec nt
(getRHS minPrec tk_prods);
@@ -766,11 +766,11 @@
val (used', O) = update_trees used (A, (tt, prec));
in
case O of
- None =>
+ NONE =>
let val Slist = getS A prec Si;
val States' = map (movedot_nonterm tt) Slist;
in processS used' (States' @ States) (S :: Si, Sii) end
- | Some n =>
+ | SOME n =>
if n >= prec then processS used' States (S :: Si, Sii)
else
let val Slist = getS' A prec n Si;
@@ -825,12 +825,12 @@
fun get_starts [] result = result
| get_starts ((nt, minPrec:int) :: nts) result =
let fun get [] result = result
- | get ((Some tk, prods) :: ps) result =
+ | get ((SOME tk, prods) :: ps) result =
if not (null prods) andalso
exists (reduction tk minPrec [nt]) prods
then get ps (tk :: result)
else get ps result
- | get ((None, _) :: ps) result = get ps result;
+ | get ((NONE, _) :: ps) result = get ps result;
val (_, nt_prods) = Array.sub (prods, nt);
@@ -842,12 +842,12 @@
val nts =
mapfilter (fn (_, _, _, Nonterminal (a, prec) :: _, _, _) =>
- Some (a, prec) | _ => None)
+ SOME (a, prec) | _ => NONE)
(Array.sub (stateset, i-1));
val allowed =
distinct (get_starts nts [] @
- (mapfilter (fn (_, _, _, Terminal a :: _, _, _) => Some a
- | _ => None)
+ (mapfilter (fn (_, _, _, Terminal a :: _, _, _) => SOME a
+ | _ => NONE)
(Array.sub (stateset, i-1))));
in syntax_error (if prev_token = EndToken then indata
else prev_token :: indata) allowed
@@ -863,14 +863,14 @@
end));
-fun get_trees l = mapfilter (fn (_, _, [pt], _, _, _) => Some pt | _ => None)
+fun get_trees l = mapfilter (fn (_, _, [pt], _, _, _) => SOME pt | _ => NONE)
l;
fun earley prods tags chains startsymbol indata =
let
val start_tag = case Symtab.lookup (tags, startsymbol) of
- Some tag => tag
- | None => error ("parse: Unknown startsymbol " ^
+ SOME tag => tag
+ | NONE => error ("parse: Unknown startsymbol " ^
quote startsymbol);
val S0 = [(~1, 0, [], [Nonterminal (start_tag, 0), Terminal EndToken],
"", 0)];
--- a/src/Pure/Syntax/printer.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Syntax/printer.ML Sun Feb 13 17:15:14 2005 +0100
@@ -197,7 +197,7 @@
(* xprod_to_fmt *)
-fun xprod_to_fmt (SynExt.XProd (_, _, "", _)) = None
+fun xprod_to_fmt (SynExt.XProd (_, _, "", _)) = NONE
| xprod_to_fmt (SynExt.XProd (_, xsymbs, const, pri)) =
let
fun cons_str s (String s' :: syms) = String (s ^ s') :: syms
@@ -233,7 +233,7 @@
| nargs [] = 0;
in
(case xsyms_to_syms xsymbs of
- (symbs, []) => Some (const, (symbs, nargs symbs, pri))
+ (symbs, []) => SOME (const, (symbs, nargs symbs, pri))
| _ => sys_error "xprod_to_fmt: unbalanced blocks")
end;
@@ -269,9 +269,9 @@
fun token_trans c [Ast.Variable x] =
(case tokentrf c of
- None => None
- | Some f => Some (f x))
- | token_trans _ _ = None;
+ NONE => NONE
+ | SOME f => SOME (f x))
+ | token_trans _ _ = NONE;
(*default applications: prefix / postfix*)
val appT =
@@ -339,8 +339,8 @@
else prnt (prnps, tbs);
in
(case token_trans a args of (*try token translation function*)
- Some s_len => [Pretty.raw_str s_len]
- | None => (*try print translation functions*)
+ SOME s_len => [Pretty.raw_str s_len]
+ | NONE => (*try print translation functions*)
astT (trans a (trf a) args, p) handle Match => prnt ([], tabs))
end
--- a/src/Pure/Syntax/syn_ext.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Syntax/syn_ext.ML Sun Feb 13 17:15:14 2005 +0100
@@ -140,8 +140,8 @@
fun delims_of xprods =
let
- fun del_of (Delim s) = Some s
- | del_of _ = None;
+ fun del_of (Delim s) = SOME s
+ | del_of _ = NONE;
fun dels_of (XProd (_, xsymbs, _, _)) =
mapfilter del_of xsymbs;
@@ -202,8 +202,8 @@
Scan.repeat1 scan_delim_char >> (Delim o implode);
val scan_symb =
- scan_sym >> Some ||
- $$ "'" -- Scan.one Symbol.is_blank >> K None;
+ scan_sym >> SOME ||
+ $$ "'" -- Scan.one Symbol.is_blank >> K NONE;
val scan_symbs = Scan.repeat scan_symb --| Scan.ahead (Scan.one (not_equal "'"));
val read_symbs = mapfilter I o the o Scan.read Symbol.stopper scan_symbs;
--- a/src/Pure/Syntax/syn_trans.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Syntax/syn_trans.ML Sun Feb 13 17:15:14 2005 +0100
@@ -209,7 +209,7 @@
Lexicon.free (the_struct structs 1)
| struct_tr structs (*"_struct"*) [t as (Const ("_indexnum", _) $ Const (s, _))] =
Lexicon.free (the_struct structs
- (case Lexicon.read_nat s of Some n => n | None => raise TERM ("struct_tr", [t])))
+ (case Lexicon.read_nat s of SOME n => n | NONE => raise TERM ("struct_tr", [t])))
| struct_tr _ (*"_struct"*) ts = raise TERM ("struct_tr", ts);
@@ -430,8 +430,8 @@
fun the_struct' structs s =
[(case Lexicon.read_nat s of
- Some i => Ast.Variable (the_struct structs i handle ERROR_MESSAGE _ => raise Match)
- | None => raise Match)] |> Ast.mk_appl (Ast.Constant "_free");
+ SOME i => Ast.Variable (the_struct structs i handle ERROR_MESSAGE _ => raise Match)
+ | NONE => raise Match)] |> Ast.mk_appl (Ast.Constant "_free");
;
fun struct_ast_tr' structs (*"_struct"*) [Ast.Constant "_indexdefault"] =
@@ -473,8 +473,8 @@
let
fun trans a args =
(case trf a of
- None => Ast.mk_appl (Ast.Constant a) args
- | Some f => f args handle exn => raise TRANSLATION (a, exn));
+ NONE => Ast.mk_appl (Ast.Constant a) args
+ | SOME f => f args handle exn => raise TRANSLATION (a, exn));
(*translate pt bottom-up*)
fun ast_of (Parser.Node (a, pts)) = trans a (map ast_of pts)
@@ -498,8 +498,8 @@
let
fun trans a args =
(case trf a of
- None => Term.list_comb (Lexicon.const a, args)
- | Some f => f args handle exn => raise TRANSLATION (a, exn));
+ NONE => Term.list_comb (Lexicon.const a, args)
+ | SOME f => f args handle exn => raise TRANSLATION (a, exn));
fun term_of (Ast.Constant a) = trans a []
| term_of (Ast.Variable x) = Lexicon.read_var x
--- a/src/Pure/Syntax/syntax.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Syntax/syntax.ML Sun Feb 13 17:15:14 2005 +0100
@@ -335,7 +335,7 @@
val pts = Parser.parse gram root' (Lexicon.tokenize lexicon xids chars);
fun show_pt pt =
- warning (Pretty.string_of (Ast.pretty_ast (hd (SynTrans.pts_to_asts (K None) [pt]))));
+ warning (Pretty.string_of (Ast.pretty_ast (hd (SynTrans.pts_to_asts (K NONE) [pt]))));
in
if length pts > ! ambiguity_level then
if ! ambiguity_is_error then
@@ -389,21 +389,21 @@
| map_trrule f (PrintRule (x, y)) = PrintRule (f x, f y)
| map_trrule f (ParsePrintRule (x, y)) = ParsePrintRule (f x, f y);
-fun parse_rule (ParseRule pats) = Some pats
- | parse_rule (PrintRule _) = None
- | parse_rule (ParsePrintRule pats) = Some pats;
+fun parse_rule (ParseRule pats) = SOME pats
+ | parse_rule (PrintRule _) = NONE
+ | parse_rule (ParsePrintRule pats) = SOME pats;
-fun print_rule (ParseRule _) = None
- | print_rule (PrintRule pats) = Some (swap pats)
- | print_rule (ParsePrintRule pats) = Some (swap pats);
+fun print_rule (ParseRule _) = NONE
+ | print_rule (PrintRule pats) = SOME (swap pats)
+ | print_rule (ParsePrintRule pats) = SOME (swap pats);
fun check_rule (rule as (lhs, rhs)) =
(case Ast.rule_error rule of
- Some msg =>
+ SOME msg =>
error ("Error in syntax translation rule: " ^ msg ^ "\n" ^
Ast.str_of_ast lhs ^ " -> " ^ Ast.str_of_ast rhs)
- | None => rule);
+ | NONE => rule);
fun read_pattern is_logtype syn (root, str) =
--- a/src/Pure/Syntax/type_ext.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Syntax/type_ext.ML Sun Feb 13 17:15:14 2005 +0100
@@ -139,7 +139,7 @@
fun no_brackets () =
Library.find_first (equal bracketsN orf equal no_bracketsN) (! print_mode)
- = Some no_bracketsN;
+ = SOME no_bracketsN;
val type_bracketsN = "type_brackets";
val no_type_bracketsN = "no_type_brackets";
@@ -147,7 +147,7 @@
fun no_type_brackets () =
Library.find_first (equal type_bracketsN orf equal no_type_bracketsN)
(! print_mode)
- <> Some type_bracketsN;
+ <> SOME type_bracketsN;
(* parse ast translations *)
--- a/src/Pure/Thy/html.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Thy/html.ML Sun Feb 13 17:15:14 2005 +0100
@@ -275,11 +275,11 @@
fun href_name s txt = "<a href=" ^ Library.quote s ^ ">" ^ txt ^ "</a>";
fun href_path path txt = href_name (Url.pack path) txt;
-fun href_opt_name None txt = txt
- | href_opt_name (Some s) txt = href_name s txt;
+fun href_opt_name NONE txt = txt
+ | href_opt_name (SOME s) txt = href_name s txt;
-fun href_opt_path None txt = txt
- | href_opt_path (Some p) txt = href_path p txt;
+fun href_opt_path NONE txt = txt
+ | href_opt_path (SOME p) txt = href_path p txt;
fun para txt = "\n<p>" ^ txt ^ "</p>\n";
fun preform txt = "<pre>" ^ txt ^ "</pre>";
@@ -314,8 +314,8 @@
fun begin_index up (index_path, session) opt_readme opt_doc graph =
begin_document ("Index of " ^ session) ^ up_link up ^
para ("View " ^ href_path graph "theory dependencies" ^
- (case opt_readme of None => "" | Some p => "<br>\nView " ^ href_path p "README") ^
- (case opt_doc of None => "" | Some p => "<br>\nView " ^ href_path p "document")) ^
+ (case opt_readme of NONE => "" | SOME p => "<br>\nView " ^ href_path p "README") ^
+ (case opt_doc of NONE => "" | SOME p => "<br>\nView " ^ href_path p "document")) ^
"\n</div>\n<hr>\n<div class=\"theories\">\n<h2>Theories</h2>\n<ul>\n";
val end_index = end_document;
@@ -365,9 +365,9 @@
local
- fun encl None = enclose "[" "]"
- | encl (Some false) = enclose "(" ")"
- | encl (Some true) = I;
+ fun encl NONE = enclose "[" "]"
+ | encl (SOME false) = enclose "(" ")"
+ | encl (SOME true) = I;
fun file (opt_ref, path, loadit) = href_opt_path opt_ref (encl loadit (file_path path));
val files = space_implode " " o map file;
--- a/src/Pure/Thy/present.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Thy/present.ML Sun Feb 13 17:15:14 2005 +0100
@@ -177,8 +177,8 @@
fun change_theory_info name f =
change_browser_info (fn (info as (theories, files, tex_index, html_index, graph)) =>
(case Symtab.lookup (theories, name) of
- None => (warning ("Browser info: cannot access theory document " ^ quote name); info)
- | Some info => (Symtab.update ((name, map_theory_info f info), theories), files,
+ NONE => (warning ("Browser info: cannot access theory document " ^ quote name); info)
+ | SOME info => (Symtab.update ((name, map_theory_info f info), theories), files,
tex_index, html_index, graph)));
@@ -231,9 +231,9 @@
(* state *)
-val session_info = ref (None: session_info option);
+val session_info = ref (NONE: session_info option);
-fun with_session x f = (case ! session_info of None => x | Some info => f info);
+fun with_session x f = (case ! session_info of NONE => x | SOME info => f info);
fun with_context f = f (PureThy.get_name (Context.the_context ()));
@@ -260,8 +260,8 @@
fun update_index dir name =
(case try get_entries dir of
- None => warning ("Browser info: cannot access session index of " ^ quote (Path.pack dir))
- | Some es => (put_entries ((es \ name) @ [name]) dir; create_index dir));
+ NONE => warning ("Browser info: cannot access session index of " ^ quote (Path.pack dir))
+ | SOME es => (put_entries ((es \ name) @ [name]) dir; create_index dir));
(* init session *)
@@ -280,7 +280,7 @@
fun init build info doc doc_graph path name doc_prefix2 (remote_path, first_time) verbose =
if not build andalso not info andalso doc = "" andalso is_none doc_prefix2 then
- (browser_info := empty_browser_info; session_info := None)
+ (browser_info := empty_browser_info; session_info := NONE)
else
let
val parent_name = name_of_session (Library.take (length path - 1, path));
@@ -289,25 +289,25 @@
val html_prefix = Path.append (Path.expand output_path) sess_prefix;
val (doc_prefix1, document_path) =
- if doc = "" then (None, None)
+ if doc = "" then (NONE, NONE)
else if not (File.exists doc_path) then (conditional verbose (fn () =>
- std_error "Warning: missing document directory\n"); (None, None))
- else (Some (Path.append html_prefix doc_path), Some (Path.ext doc doc_path));
+ std_error "Warning: missing document directory\n"); (NONE, NONE))
+ else (SOME (Path.append html_prefix doc_path), SOME (Path.ext doc doc_path));
val parent_index_path = Path.append Path.parent index_path;
val index_up_lnk = if first_time then
Url.append (the remote_path) (Url.File (Path.append sess_prefix parent_index_path))
else Url.File parent_index_path;
val readme =
- if File.exists readme_html_path then Some readme_html_path
- else if File.exists readme_path then Some readme_path
- else None;
+ if File.exists readme_html_path then SOME readme_html_path
+ else if File.exists readme_path then SOME readme_path
+ else NONE;
val index_text = HTML.begin_index (index_up_lnk, parent_name)
(Url.File index_path, session_name) (apsome Url.File readme)
(apsome Url.File document_path) (Url.unpack "medium.html");
in
- session_info := Some (make_session_info (name, parent_name, session_name, path, html_prefix,
+ session_info := SOME (make_session_info (name, parent_name, session_name, path, html_prefix,
info, doc, doc_graph, doc_prefix1, doc_prefix2, remote_path, verbose, readme));
browser_info := init_browser_info remote_path path;
add_html_index index_text
@@ -364,13 +364,13 @@
val opt_graphs =
if doc_graph andalso (is_some doc_prefix1 orelse is_some doc_prefix2) then
- Some (isatool_browser graph)
- else None;
+ SOME (isatool_browser graph)
+ else NONE;
fun prepare_sources path =
(copy_all doc_path path;
copy_files (Path.unpack "~~/lib/texinputs/*.sty") path;
- (case opt_graphs of None => () | Some (pdf, eps) =>
+ (case opt_graphs of NONE => () | SOME (pdf, eps) =>
(File.write (Path.append path graph_pdf_path) pdf;
File.write (Path.append path graph_eps_path) eps));
write_tex_index tex_index path;
@@ -382,7 +382,7 @@
File.write (Path.append html_prefix session_entries_path) "";
create_index html_prefix;
if length path > 1 then update_index parent_html_prefix name else ();
- (case readme of None => () | Some path => File.copy path (Path.append html_prefix path));
+ (case readme of NONE => () | SOME path => File.copy path (Path.append html_prefix path));
write_graph graph (Path.append html_prefix graph_path);
copy_files (Path.unpack "~~/lib/browser/GraphBrowser.jar") html_prefix;
seq (fn (a, txt) => File.write (Path.append html_prefix (Path.basic a)) txt)
@@ -393,18 +393,18 @@
conditional verbose (fn () =>
std_error ("Browser info at " ^ show_path html_prefix ^ "\n"))));
- (case doc_prefix2 of None => () | Some path =>
+ (case doc_prefix2 of NONE => () | SOME path =>
(prepare_sources path;
conditional verbose (fn () => std_error ("Document sources at " ^ show_path path ^ "\n"))));
- (case doc_prefix1 of None => () | Some path =>
+ (case doc_prefix1 of NONE => () | SOME path =>
(prepare_sources path;
isatool_document true doc_format path;
conditional verbose (fn () =>
std_error ("Document at " ^ show_path (Path.ext doc_format path) ^ "\n"))));
browser_info := empty_browser_info;
- session_info := None
+ session_info := NONE
end);
@@ -425,8 +425,8 @@
fun parent_link remote_path curr_session name =
let val {name = _, session, is_local} = get_info (ThyInfo.theory name)
- in (if null session then None else
- Some (if is_some remote_path andalso not is_local then
+ in (if null session then NONE else
+ SOME (if is_some remote_path andalso not is_local then
Url.append (the remote_path) (Url.File
(Path.append (Path.make session) (html_path name)))
else Url.File (Path.append (mk_rel_path curr_session session)
@@ -438,22 +438,22 @@
let
val parents = map (parent_link remote_path path) raw_parents;
val ml_path = ThyLoad.ml_path name;
- val files = map (apsnd Some) orig_files @
- (if is_some (ThyLoad.check_file optpath ml_path) then [(ml_path, None)] else []);
+ val files = map (apsnd SOME) orig_files @
+ (if is_some (ThyLoad.check_file optpath ml_path) then [(ml_path, NONE)] else []);
fun prep_file (raw_path, loadit) =
(case ThyLoad.check_file optpath raw_path of
- Some (path, _) =>
+ SOME (path, _) =>
let
val base = Path.base path;
val base_html = html_ext base;
in
add_file (Path.append html_prefix base_html,
HTML.ml_file (Url.File base) (File.read path));
- (Some (Url.File base_html), Url.File raw_path, loadit)
+ (SOME (Url.File base_html), Url.File raw_path, loadit)
end
- | None => (warning ("Browser info: expected to find ML file" ^ quote (Path.pack raw_path));
- (None, Url.File raw_path, loadit)));
+ | NONE => (warning ("Browser info: expected to find ML file" ^ quote (Path.pack raw_path));
+ (NONE, Url.File raw_path, loadit)));
val files_html = map prep_file files;
--- a/src/Pure/Thy/thm_deps.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Thy/thm_deps.ML Sun Feb 13 17:15:14 2005 +0100
@@ -47,10 +47,10 @@
(case prefx of
(x :: _) =>
(case ThyInfo.lookup_theory x of
- Some thy =>
+ SOME thy =>
let val name = #name (Present.get_info thy)
in if name = "" then [] else [name] end
- | None => [])
+ | NONE => [])
| _ => ["global"]);
in
if name mem parents' then (gra', parents union parents')
--- a/src/Pure/Thy/thy_info.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Thy/thy_info.ML Sun Feb 13 17:15:14 2005 +0100
@@ -104,7 +104,7 @@
fun make_deps present outdated master files =
{present = present, outdated = outdated, master = master, files = files}: deps;
-fun init_deps master files = Some (make_deps false true master (map (rpair None) files));
+fun init_deps master files = SOME (make_deps false true master (map (rpair NONE) files));
type thy = deps option * theory option;
@@ -130,7 +130,7 @@
(* access thy *)
fun lookup_thy name =
- Some (thy_graph Graph.get_node name) handle Graph.UNDEF _ => None;
+ SOME (thy_graph Graph.get_node name) handle Graph.UNDEF _ => NONE;
val known_thy = is_some o lookup_thy;
fun check_known_thy name = known_thy name orelse (warning ("Unknown theory " ^ quote name); false);
@@ -138,8 +138,8 @@
fun get_thy name =
(case lookup_thy name of
- Some thy => thy
- | None => error (loader_msg "nothing known about theory" [name]));
+ SOME thy => thy
+ | NONE => error (loader_msg "nothing known about theory" [name]));
fun change_thy name f = (get_thy name; change_thys (Graph.map_node name f));
@@ -151,13 +151,13 @@
fun change_deps name f = change_thy name (fn (deps, x) => (f deps, x));
fun is_finished name = is_none (get_deps name);
-fun get_files name = (case get_deps name of Some {files, ...} => files | _ => []);
+fun get_files name = (case get_deps name of SOME {files, ...} => files | _ => []);
fun loaded_files name =
(case get_deps name of
- None => []
- | Some {master, files, ...} =>
- (case master of Some m => [#1 (ThyLoad.get_thy m)] | None => []) @
+ NONE => []
+ | SOME {master, files, ...} =>
+ (case master of SOME m => [#1 (ThyLoad.get_thy m)] | NONE => []) @
(mapfilter (apsome #1 o #2) files));
fun get_preds name =
@@ -169,18 +169,18 @@
fun lookup_theory name =
(case lookup_thy name of
- Some (_, Some thy) => Some thy
- | _ => None);
+ SOME (_, SOME thy) => SOME thy
+ | _ => NONE);
fun get_theory name =
(case lookup_theory name of
- (Some theory) => theory
+ (SOME theory) => theory
| _ => error (loader_msg "undefined theory entry for" [name]));
val theory_of_sign = get_theory o Sign.name_of;
val theory_of_thm = theory_of_sign o Thm.sign_of_thm;
-fun put_theory name theory = change_thy name (fn (deps, _) => (deps, Some theory));
+fun put_theory name theory = change_thy name (fn (deps, _) => (deps, SOME theory));
@@ -192,7 +192,7 @@
fun is_outdated name =
(case lookup_deps name of
- Some (Some {outdated, ...}) => outdated
+ SOME (SOME {outdated, ...}) => outdated
| _ => false);
fun outdate_thy name =
@@ -201,8 +201,8 @@
make_deps present true master files)); perform Outdate name);
fun check_unfinished name =
- if is_finished name then (warning (loader_msg "tried to touch finished theory" [name]); None)
- else Some name;
+ if is_finished name then (warning (loader_msg "tried to touch finished theory" [name]); NONE)
+ else SOME name;
in
@@ -228,25 +228,25 @@
(* load_file *)
val opt_path = apsome (Path.dir o fst o ThyLoad.get_thy);
-fun opt_path' (d : deps option) = if_none (apsome (opt_path o #master) d) None;
-fun opt_path'' d = if_none (apsome opt_path' d) None;
+fun opt_path' (d : deps option) = if_none (apsome (opt_path o #master) d) NONE;
+fun opt_path'' d = if_none (apsome opt_path' d) NONE;
local
-fun provide path name info (deps as Some {present, outdated, master, files}) =
+fun provide path name info (deps as SOME {present, outdated, master, files}) =
(if exists (equal path o #1) files then ()
else warning (loader_msg "undeclared dependency of theory" [name] ^
" on file: " ^ quote (Path.pack path));
- Some (make_deps present outdated master (overwrite (files, (path, Some info)))))
- | provide _ _ _ None = None;
+ SOME (make_deps present outdated master (overwrite (files, (path, SOME info)))))
+ | provide _ _ _ NONE = NONE;
fun run_file path =
(case apsome PureThy.get_name (Context.get_context ()) of
- None => (ThyLoad.load_file None path; ())
- | Some name => (case lookup_deps name of
- Some deps => change_deps name (provide path name
+ NONE => (ThyLoad.load_file NONE path; ())
+ | SOME name => (case lookup_deps name of
+ SOME deps => change_deps name (provide path name
(ThyLoad.load_file (opt_path' deps) path))
- | None => (ThyLoad.load_file None path; ())));
+ | NONE => (ThyLoad.load_file NONE path; ())));
in
@@ -283,12 +283,12 @@
else #1 (ThyLoad.deps_thy dir name ml);
val files = get_files name;
- val missing_files = mapfilter (fn (path, None) => Some (Path.pack path) | _ => None) files;
+ val missing_files = mapfilter (fn (path, NONE) => SOME (Path.pack path) | _ => NONE) files;
in
if null missing_files then ()
else warning (loader_msg "unresolved dependencies of theory" [name] ^
" on file(s): " ^ commas_quote missing_files);
- change_deps name (fn _ => Some (make_deps true false (Some master) files));
+ change_deps name (fn _ => SOME (make_deps true false (SOME master) files));
perform Update name
end;
@@ -301,27 +301,27 @@
fun load_deps ml dir name =
let val (master, (parents, files)) = ThyLoad.deps_thy dir name ml
- in (Some (init_deps (Some master) files), parents) end;
+ in (SOME (init_deps (SOME master) files), parents) end;
-fun file_current master (_, None) = false
+fun file_current master (_, NONE) = false
| file_current master (path, info) =
(info = ThyLoad.check_file (opt_path master) path);
fun current_deps ml strict dir name =
(case lookup_deps name of
- None => (false, load_deps ml dir name)
- | Some deps =>
+ NONE => (false, load_deps ml dir name)
+ | SOME deps =>
let
- fun get_name s = (case opt_path'' (lookup_deps s) of None => s
- | Some p => Path.pack (Path.append p (Path.basic s)));
- val same_deps = (None, map get_name (thy_graph Graph.imm_preds name))
+ fun get_name s = (case opt_path'' (lookup_deps s) of NONE => s
+ | SOME p => Path.pack (Path.append p (Path.basic s)));
+ val same_deps = (NONE, map get_name (thy_graph Graph.imm_preds name))
in
(case deps of
- None => (true, same_deps)
- | Some {present, outdated, master, files} =>
+ NONE => (true, same_deps)
+ | SOME {present, outdated, master, files} =>
if present andalso not strict then (true, same_deps)
else
- let val master' = Some (ThyLoad.check_thy dir name ml) in
+ let val master' = SOME (ThyLoad.check_thy dir name ml) in
if master <> master' then (false, load_deps ml dir name)
else (not outdated andalso forall (file_current master') files, same_deps)
end)
@@ -347,13 +347,13 @@
val dir' = if_none (opt_path'' new_deps) dir;
val (visited', parent_results) = foldl_map (req_parent dir') (name :: visited, parents);
- val thy = if not really then Some (get_theory name) else None;
+ val thy = if not really then SOME (get_theory name) else NONE;
val result =
if current andalso forall #1 parent_results then true
else
((case new_deps of
- Some deps => change_thys (update_node name (map base_of_path parents) (deps, thy))
- | None => ());
+ SOME deps => change_thys (update_node name (map base_of_path parents) (deps, thy))
+ | NONE => ());
load_thy really ml (time orelse ! Output.timing) initiators dir name;
false);
in (visited', (result, name)) end
@@ -408,10 +408,10 @@
val theory = PureThy.begin_theory name (map get_theory bparents);
val deps =
if known_thy name then get_deps name
- else (init_deps None (map #1 paths)); (*records additional ML files only!*)
- val use_paths = mapfilter (fn (x, true) => Some x | _ => None) paths;
+ else (init_deps NONE (map #1 paths)); (*records additional ML files only!*)
+ val use_paths = mapfilter (fn (x, true) => SOME x | _ => NONE) paths;
- val _ = change_thys (update_node name bparents (deps, Some (Theory.copy theory)));
+ val _ = change_thys (update_node name bparents (deps, SOME (Theory.copy theory)));
val theory' = theory |> present dir' name bparents paths;
val _ = put_theory name (Theory.copy theory');
val ((), theory'') = Context.pass_theory theory' (seq (load_file false)) use_paths;
@@ -427,7 +427,7 @@
(* finish all theories *)
-fun finish () = change_thys (Graph.map_nodes (fn (_, entry) => (None, entry)));
+fun finish () = change_thys (Graph.map_nodes (fn (_, entry) => (NONE, entry)));
(* register existing theories *)
@@ -443,12 +443,12 @@
val nonfinished = filter_out is_finished parent_names;
fun get_variant (x, y_name) =
- if Theory.eq_thy (x, get_theory y_name) then None
- else Some y_name;
+ if Theory.eq_thy (x, get_theory y_name) then NONE
+ else SOME y_name;
val variants = mapfilter get_variant (parents ~~ parent_names);
fun register G =
- (Graph.new_node (name, (None, Some theory)) G
+ (Graph.new_node (name, (NONE, SOME theory)) G
handle Graph.DUP _ => err "duplicate theory entry" [])
|> add_deps name parent_names;
in
--- a/src/Pure/Thy/thy_load.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Thy/thy_load.ML Sun Feb 13 17:15:14 2005 +0100
@@ -80,8 +80,8 @@
fun find_dir dir =
get_first (find_ext (Path.append dir path)) ml_exts;
- fun add_current ps = (case current of None => ps
- | (Some p) => if Path.is_current p then ps else p :: ps);
+ fun add_current ps = (case current of NONE => ps
+ | (SOME p) => if Path.is_current p then ps else p :: ps);
val paths =
if Path.is_current path then error "Bad file specification"
@@ -94,8 +94,8 @@
fun load_file current raw_path =
(case check_file current raw_path of
- None => error ("Could not find ML file " ^ quote (Path.pack raw_path))
- | Some (path, info) => (File.use path; (path, info)));
+ NONE => error ("Could not find ML file " ^ quote (Path.pack raw_path))
+ | SOME (path, info) => (File.use path; (path, info)));
(* datatype master *)
@@ -110,10 +110,10 @@
local
fun check_thy_aux (name, ml) =
- (case check_file None (thy_path name) of
- None => error ("Could not find theory file for " ^ quote name ^ " in dir(s) " ^
+ (case check_file NONE (thy_path name) of
+ NONE => error ("Could not find theory file for " ^ quote name ^ " in dir(s) " ^
commas_quote (show_path ()))
- | Some thy_info => (thy_info, if ml then check_file None (ml_path name) else None));
+ | SOME thy_info => (thy_info, if ml then check_file NONE (ml_path name) else NONE));
in
--- a/src/Pure/Thy/thy_parse.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Thy/thy_parse.ML Sun Feb 13 17:15:14 2005 +0100
@@ -278,9 +278,9 @@
fun mk_old_type_decl ((ts, n), syn) =
map (fn t => (mk_triple (t, n, syn), false)) ts;
-fun mk_type_decl (((xs, t), None), syn) =
+fun mk_type_decl (((xs, t), NONE), syn) =
[(mk_triple (t, string_of_int (length xs), syn), false)]
- | mk_type_decl (((xs, t), Some rhs), syn) =
+ | mk_type_decl (((xs, t), SOME rhs), syn) =
[(parens (commas [t, mk_list xs, rhs, syn]), true)];
fun mk_type_decls tys =
@@ -296,7 +296,7 @@
empty >> K [];
val type_decl = type_args -- name --
- optional ("=" $$-- typ >> Some) None -- opt_infix >> mk_type_decl;
+ optional ("=" $$-- typ >> SOME) NONE -- opt_infix >> mk_type_decl;
val type_decls =
repeat1 (old_type_decl || type_decl) >> (mk_type_decls o flat);
@@ -401,7 +401,7 @@
val opt_witness =
optional ("(" $$-- list1 axm_or_thm --$$ ")") [] --
- optional (verbatim >> (parens o cat "Some" o parens)) "None"
+ optional (verbatim >> (parens o cat "SOME" o parens)) "NONE"
>> mk_witness;
val instance_decl =
@@ -415,7 +415,7 @@
val locale_decl =
(name --$$ "=") --
- (optional ((ident >> (fn x => parens ("Some" ^ Library.quote x))) --$$ "+") ("None")) --
+ (optional ((ident >> (fn x => parens ("SOME" ^ Library.quote x))) --$$ "+") ("NONE")) --
("fixes" $$--
(repeat (name --$$ "::" -- !! (typ -- opt_mixfix))
>> (mk_big_list o map mk_triple2))) --
@@ -470,8 +470,8 @@
fun sect tab ((Keyword, s, n) :: toks) =
(case Symtab.lookup (tab, s) of
- Some parse => !! parse toks
- | None => syn_err "section" s n)
+ SOME parse => !! parse toks
+ | NONE => syn_err "section" s n)
| sect _ ((_, s, n) :: _) = syn_err "section" s n
| sect _ [] = eof_err ();
@@ -523,7 +523,7 @@
(* standard sections *)
-fun mk_val ax = "val " ^ ax ^ " = PureThy.get_thm thy (" ^ Library.quote ax ^ ", None);";
+fun mk_val ax = "val " ^ ax ^ " = PureThy.get_thm thy (" ^ Library.quote ax ^ ", NONE);";
val mk_vals = cat_lines o map mk_val;
fun mk_axm_sect "" (txt, axs) = (txt, mk_vals axs)
--- a/src/Pure/Thy/thy_scan.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/Thy/thy_scan.ML Sun Feb 13 17:15:14 2005 +0100
@@ -39,8 +39,8 @@
(* diagnostics *)
-fun lex_err None msg = "Outer lexical error: " ^ msg
- | lex_err (Some n) msg = "Outer lexical error on line " ^ string_of_int n ^ ": " ^ msg;
+fun lex_err NONE msg = "Outer lexical error: " ^ msg
+ | lex_err (SOME n) msg = "Outer lexical error on line " ^ string_of_int n ^ ": " ^ msg;
(* line numbering *)
@@ -109,8 +109,8 @@
(* scan token *)
-fun token k None x = (k, x, 0)
- | token k (Some n) x = (k, x, n);
+fun token k NONE x = (k, x, 0)
+ | token k (SOME n) x = (k, x, n);
fun scan_tok lex (n, cs) =
(scan_string >> token String n ||
@@ -129,7 +129,7 @@
fun scan_token lex x =
(case scan_tok lex x of
- ((Keyword, "ML", n), x') => (keep_line scan_rest >> token Verbatim (Some n)) x'
+ ((Keyword, "ML", n), x') => (keep_line scan_rest >> token Verbatim (SOME n)) x'
| tok_x' => tok_x');
@@ -140,7 +140,7 @@
val scan_toks = Scan.repeat (scan_token lex);
val ignore = fn (Ignore, _, _) => true | _ => false;
in
- (case Scan.error (Scan.finite' Symbol.stopper scan_toks) (Some 1, chs) of
+ (case Scan.error (Scan.finite' Symbol.stopper scan_toks) (SOME 1, chs) of
(toks, (n, [])) => filter_out ignore toks @ [token EOF n "end-of-file"]
| (_, (n, cs)) => error (lex_err n ("Bad input " ^ quote (Symbol.beginning 10 cs))))
end;
--- a/src/Pure/axclass.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/axclass.ML Sun Feb 13 17:15:14 2005 +0100
@@ -64,7 +64,7 @@
val is_def = Logic.is_equals o #prop o rep_thm;
fun witnesses thy names thms =
- PureThy.get_thmss thy (map (rpair None) names) @ thms @ filter is_def (map snd (axioms_of thy));
+ PureThy.get_thmss thy (map (rpair NONE) names) @ thms @ filter is_def (map snd (axioms_of thy));
@@ -201,8 +201,8 @@
fun get_axclass_info thy c =
(case lookup_axclass_info_sg (Theory.sign_of thy) c of
- None => error ("Unknown axclass " ^ quote c)
- | Some info => info);
+ NONE => error ("Unknown axclass " ^ quote c)
+ | SOME info => info);
fun put_axclass_info c info thy =
thy |> AxclassesData.put (Symtab.update ((c, info), AxclassesData.get thy));
@@ -391,8 +391,8 @@
prove_arity sign (t, Ss, c) wthms usr_tac);
in add_arity_thms (map prove cs) thy end;
-fun sane_inst_subclass prep sub tac = ext_inst_subclass prep sub [] [] (Some tac);
-fun sane_inst_arity prep arity tac = ext_inst_arity prep arity [] [] (Some tac);
+fun sane_inst_subclass prep sub tac = ext_inst_subclass prep sub [] [] (SOME tac);
+fun sane_inst_arity prep arity tac = ext_inst_arity prep arity [] [] (SOME tac);
val add_inst_subclass_x = ext_inst_subclass read_classrel;
val add_inst_subclass = sane_inst_subclass read_classrel;
--- a/src/Pure/codegen.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/codegen.ML Sun Feb 13 17:15:14 2005 +0100
@@ -111,11 +111,11 @@
{size = Int.max (size1, size2),
iterations = Int.max (iterations1, iterations2),
default_type = case default_type1 of
- None => default_type2
+ NONE => default_type2
| _ => default_type1};
val default_test_params : test_params =
- {size = 10, iterations = 100, default_type = None};
+ {size = 10, iterations = 100, default_type = NONE};
fun set_size size ({iterations, default_type, ...} : test_params) =
{size = size, iterations = iterations, default_type = default_type};
@@ -125,7 +125,7 @@
fun set_default_type s sg ({size, iterations, ...} : test_params) =
{size = size, iterations = iterations,
- default_type = Some (typ_of (read_ctyp sg s))};
+ default_type = SOME (typ_of (read_ctyp sg s))};
(* data kind 'Pure/codegen' *)
@@ -191,20 +191,20 @@
let val {codegens, tycodegens, consts, types, attrs, preprocs, test_params} =
CodegenData.get thy
in (case assoc (codegens, name) of
- None => CodegenData.put {codegens = (name, f) :: codegens,
+ NONE => CodegenData.put {codegens = (name, f) :: codegens,
tycodegens = tycodegens, consts = consts, types = types,
attrs = attrs, preprocs = preprocs, test_params = test_params} thy
- | Some _ => error ("Code generator " ^ name ^ " already declared"))
+ | SOME _ => error ("Code generator " ^ name ^ " already declared"))
end;
fun add_tycodegen name f thy =
let val {codegens, tycodegens, consts, types, attrs, preprocs, test_params} =
CodegenData.get thy
in (case assoc (tycodegens, name) of
- None => CodegenData.put {tycodegens = (name, f) :: tycodegens,
+ NONE => CodegenData.put {tycodegens = (name, f) :: tycodegens,
codegens = codegens, consts = consts, types = types,
attrs = attrs, preprocs = preprocs, test_params = test_params} thy
- | Some _ => error ("Code generator " ^ name ^ " already declared"))
+ | SOME _ => error ("Code generator " ^ name ^ " already declared"))
end;
@@ -214,12 +214,12 @@
let val {codegens, tycodegens, consts, types, attrs, preprocs, test_params} =
CodegenData.get thy
in (case assoc (attrs, name) of
- None => CodegenData.put {tycodegens = tycodegens,
+ NONE => CodegenData.put {tycodegens = tycodegens,
codegens = codegens, consts = consts, types = types,
attrs = if name = "" then attrs @ [(name, att)] else (name, att) :: attrs,
preprocs = preprocs,
test_params = test_params} thy
- | Some _ => error ("Code attribute " ^ name ^ " already declared"))
+ | SOME _ => error ("Code attribute " ^ name ^ " already declared"))
end;
fun mk_parser (a, p) = (if a = "" then Scan.succeed "" else Args.$$$ a) |-- p;
@@ -265,21 +265,21 @@
val cname = Sign.intern_const sg s;
in
(case Sign.const_type sg cname of
- Some T =>
+ SOME T =>
let val T' = (case tyopt of
- None => T
- | Some ty =>
+ NONE => T
+ | SOME ty =>
let val U = prep_type sg ty
in if Sign.typ_instance sg (U, T) then U
else error ("Illegal type constraint for constant " ^ cname)
end)
in (case assoc (consts, (cname, T')) of
- None => CodegenData.put {codegens = codegens,
+ NONE => CodegenData.put {codegens = codegens,
tycodegens = tycodegens,
consts = ((cname, T'), syn) :: consts,
types = types, attrs = attrs, preprocs = preprocs,
test_params = test_params} thy
- | Some _ => error ("Constant " ^ cname ^ " already associated with code"))
+ | SOME _ => error ("Constant " ^ cname ^ " already associated with code"))
end
| _ => error ("Not a constant: " ^ s))
end) (thy, xs);
@@ -296,11 +296,11 @@
val tc = Sign.intern_tycon (sign_of thy) s
in
(case assoc (types, tc) of
- None => CodegenData.put {codegens = codegens,
+ NONE => CodegenData.put {codegens = codegens,
tycodegens = tycodegens, consts = consts,
types = (tc, syn) :: types, attrs = attrs,
preprocs = preprocs, test_params = test_params} thy
- | Some _ => error ("Type " ^ tc ^ " already associated with code"))
+ | SOME _ => error ("Type " ^ tc ^ " already associated with code"))
end) (thy, xs);
fun get_assoc_type thy s = assoc (#types (CodegenData.get thy), s);
@@ -348,7 +348,7 @@
(ts, map (fst o fst) (Drule.vars_of_terms ts));
val reserved = names inter ThmDatabase.ml_reserved;
val (illegal, alt_names) = split_list (mapfilter (fn s =>
- let val s' = mk_id s in if s = s' then None else Some (s, s') end) names)
+ let val s' = mk_id s in if s = s' then NONE else SOME (s, s') end) names)
val ps = (reserved @ illegal) ~~
variantlist (map (suffix "'") reserved @ alt_names, names);
@@ -385,22 +385,22 @@
val (lhs, rhs) = Logic.dest_equals t;
val (c, args) = strip_comb lhs;
val (s', T') = dest_Const c
- in if s = s' then Some (T', (args, rhs)) else None
- end handle TERM _ => None;
+ in if s = s' then SOME (T', (args, rhs)) else NONE
+ end handle TERM _ => NONE;
val defs = mapfilter (fn (name, t) => apsome (pair name) (dest t)) axms;
val i = find_index (is_instance thy T o fst o snd) defs
in
if i >= 0 then
let val (name, (T', (args, _))) = nth_elem (i, defs)
in case dest (prep_def (Thm.get_axiom thy name)) of
- None => None
- | Some (T'', p as (args', rhs)) =>
+ NONE => NONE
+ | SOME (T'', p as (args', rhs)) =>
if T' = T'' andalso args = args' then
- Some (split_last (rename_terms (args @ [rhs])),
- if length defs = 1 then None else Some i)
- else None
+ SOME (split_last (rename_terms (args @ [rhs])),
+ if length defs = 1 then NONE else SOME i)
+ else NONE
end
- else None
+ else NONE
end;
@@ -408,17 +408,17 @@
fun invoke_codegen thy dep brack (gr, t) = (case get_first
(fn (_, f) => f thy gr dep brack t) (#codegens (CodegenData.get thy)) of
- None => error ("Unable to generate code for term:\n" ^
+ NONE => error ("Unable to generate code for term:\n" ^
Sign.string_of_term (sign_of thy) t ^ "\nrequired by:\n" ^
commas (Graph.all_succs gr [dep]))
- | Some x => x);
+ | SOME x => x);
fun invoke_tycodegen thy dep brack (gr, T) = (case get_first
(fn (_, f) => f thy gr dep brack T) (#tycodegens (CodegenData.get thy)) of
- None => error ("Unable to generate code for type:\n" ^
+ NONE => error ("Unable to generate code for type:\n" ^
Sign.string_of_typ (sign_of thy) T ^ "\nrequired by:\n" ^
commas (Graph.all_succs gr [dep]))
- | Some x => x);
+ | SOME x => x);
(**** code generator for mixfix expressions ****)
@@ -468,7 +468,7 @@
let
val (gr', ps) = codegens true (gr, ts);
val (gr'', _) = invoke_tycodegen thy dep false (gr', T)
- in Some (gr'', mk_app brack (Pretty.str (s ^
+ in SOME (gr'', mk_app brack (Pretty.str (s ^
(if i=0 then "" else string_of_int i))) ps)
end
@@ -476,11 +476,11 @@
let
val (gr', ps) = codegens true (gr, ts);
val (gr'', _) = invoke_tycodegen thy dep false (gr', T)
- in Some (gr'', mk_app brack (Pretty.str s) ps) end
+ in SOME (gr'', mk_app brack (Pretty.str s) ps) end
| Const (s, T) =>
(case get_assoc_code thy s T of
- Some ms =>
+ SOME ms =>
let val i = num_args ms
in if length ts < i then
default_codegen thy gr dep brack (eta_expand u ts i)
@@ -491,18 +491,18 @@
val (gr2, ps2) = codegens true (gr1, ts2);
val (gr3, ps3) = codegens false (gr2, quotes_of ms);
in
- Some (gr3, mk_app brack (Pretty.block (pretty_mixfix ms ps1 ps3)) ps2)
+ SOME (gr3, mk_app brack (Pretty.block (pretty_mixfix ms ps1 ps3)) ps2)
end
end
- | None => (case get_defn thy s T of
- None => None
- | Some ((args, rhs), k) =>
+ | NONE => (case get_defn thy s T of
+ NONE => NONE
+ | SOME ((args, rhs), k) =>
let
val id = mk_const_id (sign_of thy) s ^ (case k of
- None => "" | Some i => "_def" ^ string_of_int i);
+ NONE => "" | SOME i => "_def" ^ string_of_int i);
val (gr', ps) = codegens true (gr, ts);
in
- Some (Graph.add_edge (id, dep) gr' handle Graph.UNDEF _ =>
+ SOME (Graph.add_edge (id, dep) gr' handle Graph.UNDEF _ =>
let
val _ = message ("expanding definition of " ^ s);
val (Ts, _) = strip_type T;
@@ -512,10 +512,10 @@
in ([v], betapply (rhs, v)) end;
val (gr1, p) = invoke_codegen thy id false
(Graph.add_edge (id, dep)
- (Graph.new_node (id, (None, "")) gr'), rhs');
+ (Graph.new_node (id, (NONE, "")) gr'), rhs');
val (gr2, xs) = codegens false (gr1, args');
val (gr3, ty) = invoke_tycodegen thy id false (gr2, T);
- in Graph.map_node id (K (None, Pretty.string_of (Pretty.block
+ in Graph.map_node id (K (NONE, Pretty.string_of (Pretty.block
(separate (Pretty.brk 1) (if null args' then
[Pretty.str ("val " ^ id ^ " :"), ty]
else Pretty.str ("fun " ^ id) :: xs) @
@@ -532,26 +532,26 @@
val (gr2, p) = invoke_codegen thy dep false
(gr1, subst_bounds (map Free (rev (bs' ~~ Ts)), t));
in
- Some (gr2, mk_app brack (Pretty.block (Pretty.str "(" :: pretty_fn bs' p @
+ SOME (gr2, mk_app brack (Pretty.block (Pretty.str "(" :: pretty_fn bs' p @
[Pretty.str ")"])) ps)
end
- | _ => None)
+ | _ => NONE)
end;
fun default_tycodegen thy gr dep brack (TVar ((s, i), _)) =
- Some (gr, Pretty.str (s ^ (if i = 0 then "" else string_of_int i)))
- | default_tycodegen thy gr dep brack (TFree (s, _)) = Some (gr, Pretty.str s)
+ SOME (gr, Pretty.str (s ^ (if i = 0 then "" else string_of_int i)))
+ | default_tycodegen thy gr dep brack (TFree (s, _)) = SOME (gr, Pretty.str s)
| default_tycodegen thy gr dep brack (Type (s, Ts)) =
(case assoc (#types (CodegenData.get thy), s) of
- None => None
- | Some ms =>
+ NONE => NONE
+ | SOME ms =>
let
val (gr', ps) = foldl_map
(invoke_tycodegen thy dep false) (gr, fst (args_of ms Ts));
val (gr'', qs) = foldl_map
(invoke_tycodegen thy dep false) (gr', quotes_of ms)
- in Some (gr'', Pretty.block (pretty_mixfix ms ps qs)) end);
+ in SOME (gr'', Pretty.block (pretty_mixfix ms ps qs)) end);
fun output_code gr xs = implode (map (snd o Graph.get_node gr)
@@ -561,7 +561,7 @@
setmp print_mode [] (Pretty.setmp_margin (!margin) (fn xs =>
let
val sg = sign_of thy;
- val gr = Graph.new_node ("<Top>", (None, "")) Graph.empty;
+ val gr = Graph.new_node ("<Top>", (NONE, "")) Graph.empty;
val (gr', ps) = foldl_map (fn (gr, (s, t)) => apsnd (pair s)
(invoke_codegen thy "<Top>" false (gr, t)))
(gr, map (apsnd (prep_term sg)) xs)
@@ -611,7 +611,7 @@
(if s mem xs then "'" else "")) :: map (mk_gen sg true xs a) Ts @
(if s mem xs then [Pretty.str a] else []))));
-val test_fn : (int -> (string * term) list option) ref = ref (fn _ => None);
+val test_fn : (int -> (string * term) list option) ref = ref (fn _ => NONE);
fun test_term thy sz i = setmp print_mode [] (fn t =>
let
@@ -636,9 +636,9 @@
Pretty.brk 1, Pretty.str "in", Pretty.brk 1,
Pretty.block [Pretty.str "if ",
mk_app false (Pretty.str "testf") (map (Pretty.str o mk_id o fst) frees),
- Pretty.brk 1, Pretty.str "then Library.None",
+ Pretty.brk 1, Pretty.str "then NONE",
Pretty.brk 1, Pretty.str "else ",
- Pretty.block [Pretty.str "Library.Some ", Pretty.block (Pretty.str "[" ::
+ Pretty.block [Pretty.str "SOME ", Pretty.block (Pretty.str "[" ::
flat (separate [Pretty.str ",", Pretty.brk 1]
(map (fn (s, T) => [Pretty.block
[Pretty.str ("(" ^ Library.quote (Symbol.escape s) ^ ","), Pretty.brk 1,
@@ -648,14 +648,14 @@
Pretty.brk 1, Pretty.str "end"]), Pretty.str ");"]) ^
"\n\nend;\n";
val _ = use_text Context.ml_output false s;
- fun iter f k = if k > i then None
+ fun iter f k = if k > i then NONE
else (case (f () handle Match =>
- (warning "Exception Match raised in generated code"; None)) of
- None => iter f (k+1) | Some x => Some x);
- fun test k = if k > sz then None
+ (warning "Exception Match raised in generated code"; NONE)) of
+ NONE => iter f (k+1) | SOME x => SOME x);
+ fun test k = if k > sz then NONE
else (priority ("Test data size: " ^ string_of_int k);
case iter (fn () => !test_fn k) 1 of
- None => test (k+1) | Some x => Some x);
+ NONE => test (k+1) | SOME x => SOME x);
in test 0 end);
fun test_goal ({size, iterations, default_type}, tvinsts) i st =
@@ -669,8 +669,8 @@
(map_type_tfree (fn p as (s, _) => if_none (assoc (tvinsts, s))
(if_none default_type (TFree p)))) (subst_bounds (frees, strip gi)));
in case test_term (Toplevel.theory_of st) size iterations gi' of
- None => writeln "No counterexamples found."
- | Some cex => writeln ("Counterexample found:\n" ^
+ NONE => writeln "No counterexamples found."
+ | SOME cex => writeln ("Counterexample found:\n" ^
Pretty.string_of (Pretty.chunks (map (fn (s, t) =>
Pretty.block [Pretty.str (s ^ " ="), Pretty.brk 1,
Sign.pretty_term sg t]) cex)))
@@ -726,8 +726,8 @@
Scan.repeat1 (P.name --| P.$$$ "=" -- P.term) >>
(fn ((opt_fname, mode'), xs) => Toplevel.theory (fn thy =>
((case opt_fname of
- None => use_text Context.ml_output false
- | Some fname => File.write (Path.unpack fname))
+ NONE => use_text Context.ml_output false
+ | SOME fname => File.write (Path.unpack fname))
(setmp mode mode' (generate_code thy) xs); thy))));
val params =
--- a/src/Pure/context.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/context.ML Sun Feb 13 17:15:14 2005 +0100
@@ -36,7 +36,7 @@
(* theory context *)
local
- val current_theory = ref (None: theory option);
+ val current_theory = ref (NONE: theory option);
in
fun get_context () = ! current_theory;
fun set_context opt_thy = current_theory := opt_thy;
@@ -45,19 +45,19 @@
fun the_context () =
(case get_context () of
- Some thy => thy
+ SOME thy => thy
| _ => error "Unknown theory context");
-fun context thy = set_context (Some thy);
-fun reset_context () = set_context None;
+fun context thy = set_context (SOME thy);
+fun reset_context () = set_context NONE;
fun pass opt_thy f x =
setmp opt_thy (fn x => let val y = f x in (y, get_context ()) end) x;
fun pass_theory thy f x =
- (case pass (Some thy) f x of
- (y, Some thy') => (y, thy')
- | (_, None) => error "Lost theory context in ML");
+ (case pass (SOME thy) f x of
+ (y, SOME thy') => (y, thy')
+ | (_, NONE) => error "Lost theory context in ML");
fun save f x = setmp (get_context ()) f x;
@@ -65,7 +65,7 @@
(* map context *)
nonfix >>;
-fun >> f = set_context (Some (f (the_context ())));
+fun >> f = set_context (SOME (f (the_context ())));
(* use ML text *)
--- a/src/Pure/display.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/display.ML Sun Feb 13 17:15:14 2005 +0100
@@ -193,22 +193,22 @@
fun pretty_default S = Pretty.block
[Pretty.str "default sort:", Pretty.brk 1, prt_sort S];
- fun pretty_witness None = Pretty.str "universal non-emptiness witness: -"
- | pretty_witness (Some (T, S)) = Pretty.block
+ fun pretty_witness NONE = Pretty.str "universal non-emptiness witness: -"
+ | pretty_witness (SOME (T, S)) = Pretty.block
[Pretty.str "universal non-emptiness witness:", Pretty.brk 1,
prt_typ_no_tvars T, Pretty.str " ::", Pretty.brk 1, prt_sort S];
val tfrees = map (fn v => TFree (v, []));
fun pretty_type syn (t, Type.LogicalType n) =
- if syn then None
- else Some (prt_typ (Type (t, tfrees (Term.invent_names [] "'a" n))))
+ if syn then NONE
+ else SOME (prt_typ (Type (t, tfrees (Term.invent_names [] "'a" n))))
| pretty_type syn (t, Type.Abbreviation (vs, U, syn')) =
- if syn <> syn' then None
- else Some (Pretty.block
+ if syn <> syn' then NONE
+ else SOME (Pretty.block
[prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U])
| pretty_type syn (t, Type.Nonterminal) =
- if not syn then None
- else Some (prt_typ (Type (t, [])));
+ if not syn then NONE
+ else SOME (prt_typ (Type (t, [])));
val pretty_arities = flat o map (fn (t, ars) => map (prt_arity t) ars);
--- a/src/Pure/drule.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/drule.ML Sun Feb 13 17:15:14 2005 +0100
@@ -223,15 +223,15 @@
(case Symbol.explode a of "'" :: _ => true | _ => false);
val (tvs, vs) = partition is_tv insts;
fun readT (ixn, st) =
- let val S = case rsorts ixn of Some S => S | None => absent ixn;
+ let val S = case rsorts ixn of SOME S => S | NONE => absent ixn;
val T = Sign.read_typ (sign,sorts) st;
in if Sign.typ_instance sign (T, TVar(ixn,S)) then (ixn,T)
else inst_failure ixn
end
val tye = map readT tvs;
fun mkty(ixn,st) = (case rtypes ixn of
- Some T => (ixn,(st,typ_subst_TVars tye T))
- | None => absent ixn);
+ SOME T => (ixn,(st,typ_subst_TVars tye T))
+ | NONE => absent ixn);
val ixnsTs = map mkty vs;
val ixns = map fst ixnsTs
and sTs = map snd ixnsTs
@@ -294,7 +294,7 @@
fun get_kind thm =
(case Library.assoc (#2 (Thm.get_name_tags thm), "kind") of
- Some (k :: _) => k
+ SOME (k :: _) => k
| _ => "unknown");
fun kind_rule k = tag_rule ("kind", [k]) o untag_rule "kind";
@@ -667,7 +667,7 @@
let val p as (ct1, ct2) = Thm.dest_comb ct
in (case pairself term_of p of
(Const ("all", _), Abs (s, _, _)) =>
- let val (v, ct') = Thm.dest_abs (Some "@") ct2;
+ let val (v, ct') = Thm.dest_abs (SOME "@") ct2;
in Thm.combination (Thm.reflexive ct1)
(Thm.abstract_rule s v (forall_conv cv ct'))
end
@@ -677,7 +677,7 @@
(*Use a conversion to transform a theorem*)
fun fconv_rule cv th = equal_elim (cv (cprop_of th)) th;
-(*** Some useful meta-theorems ***)
+(*** SOME useful meta-theorems ***)
(*The rule V/V, obtains assumption solving for eresolve_tac*)
val asm_rl = store_standard_thm_open "asm_rl" (Thm.trivial (read_prop "PROP ?psi"));
@@ -913,8 +913,8 @@
handle TYPE (msg, _, _) => err msg;
fun zip_vars _ [] = []
- | zip_vars (_ :: vs) (None :: opt_ts) = zip_vars vs opt_ts
- | zip_vars (v :: vs) (Some t :: opt_ts) = (v, t) :: zip_vars vs opt_ts
+ | zip_vars (_ :: vs) (NONE :: opt_ts) = zip_vars vs opt_ts
+ | zip_vars (v :: vs) (SOME t :: opt_ts) = (v, t) :: zip_vars vs opt_ts
| zip_vars [] _ = err "more instantiations than variables in thm";
(*instantiate types first!*)
@@ -972,14 +972,14 @@
fun unvarifyT thm =
let
val cT = Thm.ctyp_of (Thm.sign_of_thm thm);
- val tfrees = map (fn ((x, _), S) => Some (cT (TFree (x, S)))) (tvars_of thm);
+ val tfrees = map (fn ((x, _), S) => SOME (cT (TFree (x, S)))) (tvars_of thm);
in instantiate' tfrees [] thm end;
fun unvarify raw_thm =
let
val thm = unvarifyT raw_thm;
val ct = Thm.cterm_of (Thm.sign_of_thm thm);
- val frees = map (fn ((x, _), T) => Some (ct (Free (x, T)))) (vars_of thm);
+ val frees = map (fn ((x, _), T) => SOME (ct (Free (x, T)))) (vars_of thm);
in instantiate' [] frees thm end;
@@ -1014,14 +1014,14 @@
[] => thm
| tvars =>
let val cert = Thm.ctyp_of (Thm.sign_of_thm thm)
- in instantiate' (map (fn ((x, _), S) => Some (cert (TFree (x, S)))) tvars) [] thm end);
+ in instantiate' (map (fn ((x, _), S) => SOME (cert (TFree (x, S)))) tvars) [] thm end);
fun freeze_all_Vars thm =
(case vars_of thm of
[] => thm
| vars =>
let val cert = Thm.cterm_of (Thm.sign_of_thm thm)
- in instantiate' [] (map (fn ((x, _), T) => Some (cert (Free (x, T)))) vars) thm end);
+ in instantiate' [] (map (fn ((x, _), T) => SOME (cert (Free (x, T)))) vars) thm end);
val freeze_all = freeze_all_Vars o freeze_all_TVars;
@@ -1029,7 +1029,7 @@
(* mk_triv_goal *)
(*make an initial proof state, "PROP A ==> (PROP A)" *)
-fun mk_triv_goal ct = instantiate' [] [Some ct] triv_goal;
+fun mk_triv_goal ct = instantiate' [] [SOME ct] triv_goal;
--- a/src/Pure/envir.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/envir.ML Sun Feb 13 17:15:14 2005 +0100
@@ -74,10 +74,10 @@
(*Determine if the least index updated exceeds lim*)
fun above (lim, Envir {asol, iTs, ...}) =
(case (Vartab.min_key asol, Vartab.min_key iTs) of
- (None, None) => true
- | (Some (_, i), None) => i > lim
- | (None, Some (_, i')) => i' > lim
- | (Some (_, i), Some (_, i')) => i > lim andalso i' > lim);
+ (NONE, NONE) => true
+ | (SOME (_, i), NONE) => i > lim
+ | (NONE, SOME (_, i')) => i' > lim
+ | (SOME (_, i), SOME (_, i')) => i > lim andalso i' > lim);
(*Update, checking Var-Var assignments: try to suppress higher indexes*)
fun vupdate((a,t), env) = case t of
@@ -85,8 +85,8 @@
if a=name' then env (*cycle!*)
else if xless(a, name') then
(case lookup(env,name') of (*if already assigned, chase*)
- None => update((name', Var(a,T)), env)
- | Some u => vupdate((a,u), env))
+ NONE => update((name', Var(a,T)), env)
+ | SOME u => vupdate((a,u), env))
else update((a,t), env)
| _ => update((a,t), env);
@@ -105,8 +105,8 @@
fun norm_term1 same (asol,t) : term =
let fun norm (Var (w,T)) =
(case Vartab.lookup (asol, w) of
- Some u => (norm u handle SAME => u)
- | None => raise SAME)
+ SOME u => (norm u handle SAME => u)
+ | NONE => raise SAME)
| norm (Abs(a,T,body)) = Abs(a, T, norm body)
| norm (Abs(_,_,body) $ t) = normh(subst_bound (t, body))
| norm (f $ t) =
@@ -121,8 +121,8 @@
fun normT iTs (Type (a, Ts)) = Type (a, normTs iTs Ts)
| normT iTs (TFree _) = raise SAME
| normT iTs (TVar(v, _)) = (case Vartab.lookup (iTs, v) of
- Some U => normTh iTs U
- | None => raise SAME)
+ SOME U => normTh iTs U
+ | NONE => raise SAME)
and normTh iTs T = ((normT iTs T) handle SAME => T)
and normTs iTs [] = raise SAME
| normTs iTs (T :: Ts) =
@@ -134,8 +134,8 @@
| norm (Free (a, T)) = Free(a, normT iTs T)
| norm (Var (w, T)) =
(case Vartab.lookup (asol, w) of
- Some u => normh u
- | None => Var(w, normT iTs T))
+ SOME u => normh u
+ | NONE => Var(w, normT iTs T))
| norm (Abs (a, T, body)) =
(Abs (a, normT iTs T, normh body) handle SAME => Abs (a, T, norm body))
| norm (Abs(_, _, body) $ t) = normh (subst_bound (t, body))
@@ -170,8 +170,8 @@
fun head_norm env t =
let
fun hnorm (Var (v, T)) = (case lookup (env, v) of
- Some u => head_norm env u
- | None => raise SAME)
+ SOME u => head_norm env u
+ | NONE => raise SAME)
| hnorm (Abs (a, T, body)) = Abs (a, T, hnorm body)
| hnorm (Abs (_, _, body) $ t) =
head_norm env (subst_bound (t, body))
@@ -191,7 +191,7 @@
Type ("fun", [_, T]) => T
| TVar(ixn, _) =>
(case Vartab.lookup (iTs, ixn) of
- Some (Type ("fun", [_, T])) => T
+ SOME (Type ("fun", [_, T])) => T
| _ => raise TERM (funerr, [f $ u]))
| _ => raise TERM (funerr, [f $ u]))
| fast Ts (Const (_, T)) = T
--- a/src/Pure/goals.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/goals.ML Sun Feb 13 17:15:14 2005 +0100
@@ -163,8 +163,8 @@
[Pretty.str (a ^ ":"), Pretty.brk 1, prt_term t];
val anc = case ancestor of
- None => ""
- | Some(loc) => ((Sign.base_name loc) ^ " +")
+ NONE => ""
+ | SOME(loc) => ((Sign.base_name loc) ^ " +")
in
Pretty.big_list (name ^ " = " ^ anc)
[Pretty.big_list "consts:" (map pretty_const consts),
@@ -259,8 +259,8 @@
fun the_locale thy xname =
(case lookup_locale thy xname of
- Some loc => loc
- | None => error ("Unknown locale " ^ quote xname));
+ SOME loc => loc
+ | NONE => error ("Unknown locale " ^ quote xname));
fun open_locale xname thy =
let val loc = the_locale thy xname;
@@ -268,8 +268,8 @@
val cur_sc = get_scope thy;
fun opn lc th = (change_scope (cons lc) th; th)
in case anc of
- None => opn loc thy
- | Some(loc') =>
+ NONE => opn loc thy
+ | SOME(loc') =>
if loc' mem (map fst cur_sc)
then opn loc thy
else (warning ("Opening locale " ^ quote loc' ^ ", required by " ^
@@ -329,8 +329,8 @@
fun get_thmx f get thy name =
(case get_first (get_thm_locale name) (get_scope thy) of
- Some thm => f thm
- | None => get thy (name, None));
+ SOME thm => f thm
+ | NONE => get thy (name, NONE));
val get_thm = get_thmx I PureThy.get_thm;
val get_thms = get_thmx (fn x => [x]) PureThy.get_thms;
@@ -354,7 +354,7 @@
fun read_typ sg (envT, s) =
let
fun def_sort (x, ~1) = assoc (envT, x)
- | def_sort _ = None;
+ | def_sort _ = NONE;
val T = Type.no_tvars (Sign.read_typ (sg, def_sort) s) handle TYPE (msg, _, _) => error msg;
in (Term.add_typ_tfrees (T, envT), T) end;
@@ -373,9 +373,9 @@
fun read_axm sg ((envS, envT, used), (name, s)) =
let
fun def_sort (x, ~1) = assoc (envS, x)
- | def_sort _ = None;
+ | def_sort _ = NONE;
fun def_type (x, ~1) = assoc (envT, x)
- | def_type _ = None;
+ | def_type _ = NONE;
val (_, t) = Theory.read_def_axm (sg, def_type, def_sort) used (name, s);
in
(enter_term t (envS, envT, used), t)
@@ -398,9 +398,9 @@
val envT = flat (map #2 defaults);
val used = flat (map #3 defaults);
fun def_sort (x, ~1) = assoc (envS, x)
- | def_sort _ = None;
+ | def_sort _ = NONE;
fun def_type (x, ~1) = assoc (envT, x)
- | def_type _ = None;
+ | def_type _ = NONE;
in (if (is_open_loc_sg sign)
then (#1 o read_def_cterm (sign, def_type, def_sort) used true)
else Thm.read_cterm sign)
@@ -471,17 +471,17 @@
val (envSb, old_loc_consts, _) =
case bancestor of
- Some(loc) => (get_defaults thy loc)
- | None => ([],[],[]);
+ SOME(loc) => (get_defaults thy loc)
+ | NONE => ([],[],[]);
val old_nosyn = case bancestor of
- Some(loc) => #nosyn(#2(the_locale thy loc))
- | None => [];
+ SOME(loc) => #nosyn(#2(the_locale thy loc))
+ | NONE => [];
(* Get the full name of the ancestor *)
val ancestor = case bancestor of
- Some(loc) => Some(#1(the_locale thy loc))
- | None => None;
+ SOME(loc) => SOME(#1(the_locale thy loc))
+ | NONE => NONE;
(* prepare locale consts *)
@@ -492,7 +492,7 @@
val mx = Syntax.fix_mixfix raw_c raw_mx;
val (envS', T) = prep_typ sign (envS, raw_T) handle ERROR =>
error ("The error(s) above occured in locale constant " ^ quote c);
- val trfun = if mx = Syntax.NoSyn then None else Some (c_syn, mk_loc_tr c);
+ val trfun = if mx = Syntax.NoSyn then NONE else SOME (c_syn, mk_loc_tr c);
in (envS', ((c, T), (c_syn, T, mx), trfun)) end;
val (envS0, loc_consts_syn) = foldl_map prep_const (envSb, raw_consts);
@@ -823,7 +823,7 @@
val tac = EVERY (tacsf prems)
fun statef() =
(case Seq.pull (tac st0) of
- Some(st,_) => st
+ SOME(st,_) => st
| _ => error ("prove_goal: tactic failed"))
in mkresult (check, cond_timeit (!Output.timing) statef) end
handle e => (print_sign_exn_unit (#sign (rep_cterm chorn)) e;
@@ -975,8 +975,8 @@
val As = Drule.strip_imp_prems G;
val B = Drule.strip_imp_concl G;
val asms = map (norm_hhf_rule o assume) As;
- fun check None = error "prove_goal: tactic failed"
- | check (Some (thm, _)) = (case nprems_of thm of
+ fun check NONE = error "prove_goal: tactic failed"
+ | check (SOME (thm, _)) = (case nprems_of thm of
0 => thm
| i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
in
@@ -988,8 +988,8 @@
(*Proof step "by" the given tactic -- apply tactic to the proof state*)
fun by_com tac ((th,ths), pairs) : gstack =
(case Seq.pull(tac th) of
- None => error"by: tactic failed"
- | Some(th2,ths2) =>
+ NONE => error"by: tactic failed"
+ | SOME(th2,ths2) =>
(if eq_thm(th,th2)
then warning "Warning: same as previous level"
else if eq_thm_sg(th,th2) then ()
@@ -1010,8 +1010,8 @@
fun backtrack [] = error"back: no alternatives"
| backtrack ((th,thstr) :: pairs) =
(case Seq.pull thstr of
- None => (writeln"Going back a level..."; backtrack pairs)
- | Some(th2,thstr2) =>
+ NONE => (writeln"Going back a level..."; backtrack pairs)
+ | SOME(th2,thstr2) =>
(if eq_thm(th,th2)
then warning "Warning: same as previous choice at this level"
else if eq_thm_sg(th,th2) then ()
--- a/src/Pure/library.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/library.ML Sun Feb 13 17:15:14 2005 +0100
@@ -36,8 +36,6 @@
val stamp: unit -> stamp
(*options*)
- datatype 'a option = None | Some of 'a
- exception OPTION
val the: 'a option -> 'a
val if_none: 'a option -> 'a -> 'a
val is_some: 'a option -> bool
@@ -80,7 +78,6 @@
(*lists*)
exception LIST of string
- val null: 'a list -> bool
val hd: 'a list -> 'a
val tl: 'a list -> 'a list
val cons: 'a -> 'a list -> 'a list
@@ -315,33 +312,26 @@
(** options **)
-datatype 'a option = None | Some of 'a;
-
-exception OPTION;
-
-fun the (Some x) = x
- | the None = raise OPTION;
+fun the opt = valOf opt;
(*strict!*)
-fun if_none None y = y
- | if_none (Some x) _ = x;
+fun if_none opt a = getOpt(opt,a);
-fun is_some (Some _) = true
- | is_some None = false;
+fun is_some opt = isSome opt;
-fun is_none (Some _) = false
- | is_none None = true;
+fun is_none (SOME _) = false
+ | is_none NONE = true;
-fun apsome f (Some x) = Some (f x)
- | apsome _ None = None;
+fun apsome f (SOME x) = SOME (f x)
+ | apsome _ NONE = NONE;
(* exception handling *)
exception ERROR;
-fun try f x = Some (f x)
- handle Interrupt => raise Interrupt | ERROR => raise ERROR | _ => None;
+fun try f x = SOME (f x)
+ handle Interrupt => raise Interrupt | ERROR => raise ERROR | _ => NONE;
fun can f x = is_some (try f x);
@@ -355,11 +345,11 @@
fun release (Result y) = y
| release (Exn e) = raise e;
-fun get_result (Result x) = Some x
- | get_result _ = None;
+fun get_result (Result x) = SOME x
+ | get_result _ = NONE;
-fun get_exn (Exn exn) = Some exn
- | get_exn _ = None;
+fun get_exn (Exn exn) = SOME exn
+ | get_exn _ = NONE;
@@ -441,9 +431,6 @@
exception LIST of string;
-fun null [] = true
- | null (_ :: _) = false;
-
fun hd [] = raise LIST "hd"
| hd (x :: _) = x;
@@ -456,8 +443,7 @@
fun append xs ys = xs @ ys;
(* tail recursive version *)
-fun rev_append [] ys = ys
- | rev_append (x :: xs) ys = rev_append xs (x :: ys);
+fun rev_append xs ys = List.revAppend(xs,ys);
fun apply [] x = x
| apply (f :: fs) x = apply fs (f x);
@@ -561,19 +547,19 @@
fun find_index_eq x = find_index (equal x);
(*find first element satisfying predicate*)
-fun find_first _ [] = None
+fun find_first _ [] = NONE
| find_first pred (x :: xs) =
- if pred x then Some x else find_first pred xs;
+ if pred x then SOME x else find_first pred xs;
(*get first element by lookup function*)
-fun get_first _ [] = None
+fun get_first _ [] = NONE
| get_first f (x :: xs) =
(case f x of
- None => get_first f xs
+ NONE => get_first f xs
| some => some);
(*flatten a list of lists to a list*)
-fun flat (ls: 'c list list) : 'c list = foldr (op @) (ls, []);
+val flat = List.concat;
fun unflat (xs :: xss) ys =
let val (ps,qs) = splitAt(length xs,ys)
@@ -616,18 +602,15 @@
(* filter *)
(*copy the list preserving elements that satisfy the predicate*)
-fun filter (pred: 'a->bool) : 'a list -> 'a list =
- let fun filt [] = []
- | filt (x :: xs) = if pred x then x :: filt xs else filt xs
- in filt end;
+val filter = List.filter;
fun filter_out f = filter (not o f);
fun mapfilter (f: 'a -> 'b option) ([]: 'a list) = [] : 'b list
| mapfilter f (x :: xs) =
(case f x of
- None => mapfilter f xs
- | Some y => y :: mapfilter f xs);
+ NONE => mapfilter f xs
+ | SOME y => y :: mapfilter f xs);
(* lists of pairs *)
@@ -767,8 +750,8 @@
fun nth_elem_string (i, str) =
(case try String.substring (str, i, 1) of
- Some s => s
- | None => raise LIST "nth_elem_string");
+ SOME s => s
+ | NONE => raise LIST "nth_elem_string");
fun foldl_string f (x0, str) =
let
@@ -1011,40 +994,40 @@
(** association lists **)
(*association list lookup*)
-fun assoc ([], key) = None
+fun assoc ([], key) = NONE
| assoc ((keyi, xi) :: pairs, key) =
- if key = keyi then Some xi else assoc (pairs, key);
+ if key = keyi then SOME xi else assoc (pairs, key);
(*association list lookup, optimized version for ints*)
-fun assoc_int ([], (key:int)) = None
+fun assoc_int ([], (key:int)) = NONE
| assoc_int ((keyi, xi) :: pairs, key) =
- if key = keyi then Some xi else assoc_int (pairs, key);
+ if key = keyi then SOME xi else assoc_int (pairs, key);
(*association list lookup, optimized version for strings*)
-fun assoc_string ([], (key:string)) = None
+fun assoc_string ([], (key:string)) = NONE
| assoc_string ((keyi, xi) :: pairs, key) =
- if key = keyi then Some xi else assoc_string (pairs, key);
+ if key = keyi then SOME xi else assoc_string (pairs, key);
(*association list lookup, optimized version for string*ints*)
-fun assoc_string_int ([], (key:string*int)) = None
+fun assoc_string_int ([], (key:string*int)) = NONE
| assoc_string_int ((keyi, xi) :: pairs, key) =
- if key = keyi then Some xi else assoc_string_int (pairs, key);
+ if key = keyi then SOME xi else assoc_string_int (pairs, key);
fun assocs ps x =
(case assoc (ps, x) of
- None => []
- | Some ys => ys);
+ NONE => []
+ | SOME ys => ys);
(*two-fold association list lookup*)
fun assoc2 (aal, (key1, key2)) =
(case assoc (aal, key1) of
- Some al => assoc (al, key2)
- | None => None);
+ SOME al => assoc (al, key2)
+ | NONE => NONE);
(*generalized association list lookup*)
-fun gen_assoc eq ([], key) = None
+fun gen_assoc eq ([], key) = NONE
| gen_assoc eq ((keyi, xi) :: pairs, key) =
- if eq (key, keyi) then Some xi else gen_assoc eq (pairs, key);
+ if eq (key, keyi) then SOME xi else gen_assoc eq (pairs, key);
(*association list update*)
fun overwrite (al, p as (key, _)) =
--- a/src/Pure/meta_simplifier.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/meta_simplifier.ML Sun Feb 13 17:15:14 2005 +0100
@@ -303,8 +303,8 @@
val basic_mk_rews: mk_rews =
{mk = fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [],
mk_cong = I,
- mk_sym = Some o Drule.symmetric_fun,
- mk_eq_True = K None};
+ mk_sym = SOME o Drule.symmetric_fun,
+ mk_eq_True = K NONE};
in
@@ -445,8 +445,8 @@
fun mk_eq_True (Simpset (_, {mk_rews = {mk_eq_True, ...}, ...})) (thm, name) =
(case mk_eq_True thm of
- None => []
- | Some eq_True =>
+ NONE => []
+ | SOME eq_True =>
let val (_, _, lhs, elhs, _, _) = decomp_simp eq_True
in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end);
@@ -478,8 +478,8 @@
else
let val Simpset (_, {mk_rews = {mk_sym, ...}, ...}) = ss in
(case mk_sym thm of
- None => []
- | Some thm' =>
+ NONE => []
+ | SOME thm' =>
let val (_, _, lhs', elhs', rhs', _) = decomp_simp thm'
in rrule_eq_True (thm', name, lhs', elhs', rhs', ss, thm) end)
end
@@ -516,9 +516,9 @@
(* congs *)
-fun cong_name (Const (a, _)) = Some a
- | cong_name (Free (a, _)) = Some ("Free: " ^ a)
- | cong_name _ = None;
+fun cong_name (Const (a, _)) = SOME a
+ | cong_name (Free (a, _)) = SOME ("Free: " ^ a)
+ | cong_name _ = NONE;
local
@@ -551,7 +551,7 @@
val (lhs, _) = Drule.dest_equals (Drule.strip_imp_concl (Thm.cprop_of thm))
handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", thm);
(*val lhs = Pattern.eta_contract lhs;*)
- val a = the (cong_name (head_of (term_of lhs))) handle Library.OPTION =>
+ val a = the (cong_name (head_of (term_of lhs))) handle Option =>
raise SIMPLIFIER ("Congruence must start with a constant or free variable", thm);
val (alist, weak) = congs;
val alist2 = overwrite_warn (alist, (a, {lhs = lhs, thm = thm}))
@@ -565,12 +565,12 @@
val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) handle TERM _ =>
raise SIMPLIFIER ("Congruence not a meta-equality", thm);
(*val lhs = Pattern.eta_contract lhs;*)
- val a = the (cong_name (head_of lhs)) handle Library.OPTION =>
+ val a = the (cong_name (head_of lhs)) handle Option =>
raise SIMPLIFIER ("Congruence must start with a constant", thm);
val (alist, _) = congs;
val alist2 = filter (fn (x, _) => x <> a) alist;
val weak2 = alist2 |> mapfilter (fn (a, {thm, ...}) =>
- if is_full_cong thm then None else Some a);
+ if is_full_cong thm then NONE else SOME a);
in ((alist2, weak2), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end);
fun mk_cong (Simpset (_, {mk_rews = {mk_cong = f, ...}, ...})) = f;
@@ -710,12 +710,12 @@
let
val thm'' = transitive thm (transitive
(symmetric (Drule.beta_eta_conversion (lhs_of thm'))) thm')
- in if msg then trace_thm "SUCCEEDED" thm' else (); Some thm'' end
+ in if msg then trace_thm "SUCCEEDED" thm' else (); SOME thm'' end
handle THM _ =>
let val {sign, prop = _ $ _ $ prop0, ...} = Thm.rep_thm thm in
trace_thm "Proved wrong thm (Check subgoaler?)" thm';
trace_term false "Should have proved:" sign prop0;
- None
+ NONE
end;
@@ -764,12 +764,12 @@
val Simpset ({depth = depth', ...}, _) = ss';
in
if depth' > ! simp_depth_limit
- then (warning "simp_depth_limit exceeded - giving up"; None)
+ then (warning "simp_depth_limit exceeded - giving up"; NONE)
else
(if depth' mod 10 = 0
then warning ("Simplification depth " ^ string_of_int depth')
else ();
- Some ss')
+ SOME ss')
end;
(*
@@ -803,34 +803,34 @@
in
if perm andalso not (termless (rhs', lhs'))
then (trace_named_thm "Cannot apply permutative rewrite rule" (thm, name);
- trace_thm "Term does not become smaller:" thm'; None)
+ trace_thm "Term does not become smaller:" thm'; NONE)
else (trace_named_thm "Applying instance of rewrite rule" (thm, name);
if unconditional
then
(trace_thm "Rewriting:" thm';
let val lr = Logic.dest_equals prop;
- val Some thm'' = check_conv false eta_thm thm'
- in Some (thm'', uncond_skel (congs, lr)) end)
+ val SOME thm'' = check_conv false eta_thm thm'
+ in SOME (thm'', uncond_skel (congs, lr)) end)
else
(trace_thm "Trying to rewrite:" thm';
case incr_depth ss of
- None => (trace_thm "FAILED - reached depth limit" thm'; None)
- | Some ss' =>
+ NONE => (trace_thm "FAILED - reached depth limit" thm'; NONE)
+ | SOME ss' =>
(case prover ss' thm' of
- None => (trace_thm "FAILED" thm'; None)
- | Some thm2 =>
+ NONE => (trace_thm "FAILED" thm'; NONE)
+ | SOME thm2 =>
(case check_conv true eta_thm thm2 of
- None => None |
- Some thm2' =>
+ NONE => NONE |
+ SOME thm2' =>
let val concl = Logic.strip_imp_concl prop
val lr = Logic.dest_equals concl
- in Some (thm2', cond_skel (congs, lr)) end))))
+ in SOME (thm2', cond_skel (congs, lr)) end))))
end
- fun rews [] = None
+ fun rews [] = NONE
| rews (rrule :: rrules) =
- let val opt = rew rrule handle Pattern.MATCH => None
- in case opt of None => rews rrules | some => some end;
+ let val opt = rew rrule handle Pattern.MATCH => NONE
+ in case opt of NONE => rews rrules | some => some end;
fun sort_rrules rrs = let
fun is_simple({thm, ...}:rrule) = case Thm.prop_of thm of
@@ -842,25 +842,25 @@
else sort rrs (re1,rr::re2)
in sort rrs ([],[]) end
- fun proc_rews [] = None
+ fun proc_rews [] = NONE
| proc_rews (Proc {name, proc, lhs, ...} :: ps) =
if Pattern.matches tsigt (Thm.term_of lhs, Thm.term_of t) then
(debug_term false ("Trying procedure " ^ quote name ^ " on:") signt eta_t;
case transform_failure (curry SIMPROC_FAIL name)
(fn () => proc signt ss eta_t) () of
- None => (debug false "FAILED"; proc_rews ps)
- | Some raw_thm =>
+ NONE => (debug false "FAILED"; proc_rews ps)
+ | SOME raw_thm =>
(trace_thm ("Procedure " ^ quote name ^ " produced rewrite rule:") raw_thm;
(case rews (mk_procrule raw_thm) of
- None => (trace_cterm true ("IGNORED result of simproc " ^ quote name ^
+ NONE => (trace_cterm true ("IGNORED result of simproc " ^ quote name ^
" -- does not match") t; proc_rews ps)
| some => some)))
else proc_rews ps;
in case eta_t of
- Abs _ $ _ => Some (transitive eta_thm
+ Abs _ $ _ => SOME (transitive eta_thm
(beta_conversion false eta_t'), skel0)
| _ => (case rews (sort_rrules (Net.match_term rules eta_t)) of
- None => proc_rews (Net.match_term procs eta_t)
+ NONE => proc_rews (Net.match_term procs eta_t)
| some => some)
end;
@@ -876,67 +876,67 @@
is handled when congc is called *)
val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm);
val unit = trace_thm "Applying congruence rule:" thm';
- fun err (msg, thm) = (trace_thm msg thm; None)
+ fun err (msg, thm) = (trace_thm msg thm; NONE)
in case prover thm' of
- None => err ("Congruence proof failed. Could not prove", thm')
- | Some thm2 => (case check_conv true (Drule.beta_eta_conversion t) thm2 of
- None => err ("Congruence proof failed. Should not have proved", thm2)
- | Some thm2' =>
+ NONE => err ("Congruence proof failed. Could not prove", thm')
+ | SOME thm2 => (case check_conv true (Drule.beta_eta_conversion t) thm2 of
+ NONE => err ("Congruence proof failed. Should not have proved", thm2)
+ | SOME thm2' =>
if op aconv (pairself term_of (dest_equals (cprop_of thm2')))
- then None else Some thm2')
+ then NONE else SOME thm2')
end;
val (cA, (cB, cC)) =
apsnd dest_equals (dest_implies (hd (cprems_of Drule.imp_cong)));
-fun transitive1 None None = None
- | transitive1 (Some thm1) None = Some thm1
- | transitive1 None (Some thm2) = Some thm2
- | transitive1 (Some thm1) (Some thm2) = Some (transitive thm1 thm2)
+fun transitive1 NONE NONE = NONE
+ | transitive1 (SOME thm1) NONE = SOME thm1
+ | transitive1 NONE (SOME thm2) = SOME thm2
+ | transitive1 (SOME thm1) (SOME thm2) = SOME (transitive thm1 thm2)
-fun transitive2 thm = transitive1 (Some thm);
-fun transitive3 thm = transitive1 thm o Some;
+fun transitive2 thm = transitive1 (SOME thm);
+fun transitive3 thm = transitive1 thm o SOME;
fun bottomc ((simprem, useprem, mutsimp), prover, sign, maxidx) =
let
fun botc skel ss t =
- if is_Var skel then None
+ if is_Var skel then NONE
else
(case subc skel ss t of
- some as Some thm1 =>
+ some as SOME thm1 =>
(case rewritec (prover, sign, maxidx) ss (rhs_of thm1) of
- Some (thm2, skel2) =>
+ SOME (thm2, skel2) =>
transitive2 (transitive thm1 thm2)
(botc skel2 ss (rhs_of thm2))
- | None => some)
- | None =>
+ | NONE => some)
+ | NONE =>
(case rewritec (prover, sign, maxidx) ss t of
- Some (thm2, skel2) => transitive2 thm2
+ SOME (thm2, skel2) => transitive2 thm2
(botc skel2 ss (rhs_of thm2))
- | None => None))
+ | NONE => NONE))
and try_botc ss t =
(case botc skel0 ss t of
- Some trec1 => trec1 | None => (reflexive t))
+ SOME trec1 => trec1 | NONE => (reflexive t))
and subc skel (ss as Simpset ({bounds, ...}, {congs, ...})) t0 =
(case term_of t0 of
Abs (a, T, t) =>
let
- val (v, t') = Thm.dest_abs (Some ("." ^ a ^ "." ^ string_of_int bounds)) t0;
+ val (v, t') = Thm.dest_abs (SOME ("." ^ a ^ "." ^ string_of_int bounds)) t0;
val ss' = incr_bounds ss;
val skel' = case skel of Abs (_, _, sk) => sk | _ => skel0;
in case botc skel' ss' t' of
- Some thm => Some (abstract_rule a v thm)
- | None => None
+ SOME thm => SOME (abstract_rule a v thm)
+ | NONE => NONE
end
| t $ _ => (case t of
Const ("==>", _) $ _ => impc t0 ss
| Abs _ =>
let val thm = beta_conversion false t0
in case subc skel0 ss (rhs_of thm) of
- None => Some thm
- | Some thm' => Some (transitive thm thm')
+ NONE => SOME thm
+ | SOME thm' => SOME (transitive thm thm')
end
| _ =>
let fun appc () =
@@ -947,21 +947,21 @@
val (ct, cu) = Thm.dest_comb t0
in
(case botc tskel ss ct of
- Some thm1 =>
+ SOME thm1 =>
(case botc uskel ss cu of
- Some thm2 => Some (combination thm1 thm2)
- | None => Some (combination thm1 (reflexive cu)))
- | None =>
+ SOME thm2 => SOME (combination thm1 thm2)
+ | NONE => SOME (combination thm1 (reflexive cu)))
+ | NONE =>
(case botc uskel ss cu of
- Some thm1 => Some (combination (reflexive ct) thm1)
- | None => None))
+ SOME thm1 => SOME (combination (reflexive ct) thm1)
+ | NONE => NONE))
end
val (h, ts) = strip_comb t
in case cong_name h of
- Some a =>
+ SOME a =>
(case assoc_string (fst congs, a) of
- None => appc ()
- | Some cong =>
+ NONE => appc ()
+ | SOME cong =>
(*post processing: some partial applications h t1 ... tj, j <= length ts,
may be a redex. Example: map (%x. x) = (%xs. xs) wrt map_cong*)
(let
@@ -972,14 +972,14 @@
val skel =
list_comb (h, replicate (length ts) dVar)
in case botc skel ss cl of
- None => thm
- | Some thm' => transitive3 thm
+ NONE => thm
+ | SOME thm' => transitive3 thm
(combination thm' (reflexive cr))
end handle TERM _ => error "congc result"
| Pattern.MATCH => appc ()))
| _ => appc ()
end)
- | _ => None)
+ | _ => NONE)
and impc ct ss =
if mutsimp then mut_impc0 [] ct [] [] ss else nonmut_impc ct ss
@@ -987,10 +987,10 @@
and rules_of_prem ss prem =
if maxidx_of_term (term_of prem) <> ~1
then (trace_cterm true
- "Cannot add premise as rewrite rule because it contains (type) unknowns:" prem; ([], None))
+ "Cannot add premise as rewrite rule because it contains (type) unknowns:" prem; ([], NONE))
else
let val asm = assume prem
- in (extract_safe_rrules (ss, asm), Some asm) end
+ in (extract_safe_rrules (ss, asm), SOME asm) end
and add_rrules (rrss, asms) ss =
foldl (insert_rrule true) (ss, flat rrss) |> add_prems (mapfilter I asms)
@@ -1021,8 +1021,8 @@
Drule.mk_implies (prem, if_none (apsome rhs_of eq) concl);
val dprem = apsome (curry (disch false) prem)
in case rewritec (prover, sign, maxidx) ss' concl' of
- None => rebuild prems concl' rrss asms ss (dprem eq)
- | Some (eq', _) => transitive2 (foldl (disch false o swap)
+ NONE => rebuild prems concl' rrss asms ss (dprem eq)
+ | SOME (eq', _) => transitive2 (foldl (disch false o swap)
(the (transitive3 (dprem eq) eq'), prems))
(mut_impc0 (rev prems) (rhs_of eq') (rev rrss) (rev asms) ss)
end
@@ -1037,7 +1037,7 @@
and mut_impc [] concl [] [] prems' rrss' asms' eqns ss changed k =
transitive1 (foldl (fn (eq2, (eq1, prem)) => transitive1 eq1
- (apsome (curry (disch false) prem) eq2)) (None, eqns ~~ prems'))
+ (apsome (curry (disch false) prem) eq2)) (NONE, eqns ~~ prems'))
(if changed > 0 then
mut_impc (rev prems') concl (rev rrss') (rev asms')
[] [] [] [] ss ~1 changed
@@ -1046,12 +1046,12 @@
| mut_impc (prem :: prems) concl (rrs :: rrss) (asm :: asms)
prems' rrss' asms' eqns ss changed k =
- case (if k = 0 then None else botc skel0 (add_rrules
+ case (if k = 0 then NONE else botc skel0 (add_rrules
(rev rrss' @ rrss, rev asms' @ asms) ss) prem) of
- None => mut_impc prems concl rrss asms (prem :: prems')
- (rrs :: rrss') (asm :: asms') (None :: eqns) ss changed
+ NONE => mut_impc prems concl rrss asms (prem :: prems')
+ (rrs :: rrss') (asm :: asms') (NONE :: eqns) ss changed
(if k = 0 then 0 else k - 1)
- | Some eqn =>
+ | SOME eqn =>
let
val prem' = rhs_of eqn;
val tprems = map term_of prems;
@@ -1059,7 +1059,7 @@
find_index_eq p tprems) (#hyps (rep_thm eqn)));
val (rrs', asm') = rules_of_prem ss prem'
in mut_impc prems concl rrss asms (prem' :: prems')
- (rrs' :: rrss') (asm' :: asms') (Some (foldr (disch true)
+ (rrs' :: rrss') (asm' :: asms') (SOME (foldr (disch true)
(take (i, prems), Drule.imp_cong' eqn (reflexive (Drule.list_implies
(drop (i, prems), concl))))) :: eqns) ss (length prems') ~1
end
@@ -1067,20 +1067,20 @@
(*legacy code - only for backwards compatibility*)
and nonmut_impc ct ss =
let val (prem, conc) = dest_implies ct;
- val thm1 = if simprem then botc skel0 ss prem else None;
+ val thm1 = if simprem then botc skel0 ss prem else NONE;
val prem1 = if_none (apsome rhs_of thm1) prem;
val ss1 = if not useprem then ss else add_rrules
(apsnd single (apfst single (rules_of_prem ss prem1))) ss
in (case botc skel0 ss1 conc of
- None => (case thm1 of
- None => None
- | Some thm1' => Some (Drule.imp_cong' thm1' (reflexive conc)))
- | Some thm2 =>
+ NONE => (case thm1 of
+ NONE => NONE
+ | SOME thm1' => SOME (Drule.imp_cong' thm1' (reflexive conc)))
+ | SOME thm2 =>
let val thm2' = disch false (prem1, thm2)
in (case thm1 of
- None => Some thm2'
- | Some thm1' =>
- Some (transitive (Drule.imp_cong' thm1' (reflexive conc)) thm2'))
+ NONE => SOME thm2'
+ | SOME thm1' =>
+ SOME (transitive (Drule.imp_cong' thm1' (reflexive conc)) thm2'))
end)
end
--- a/src/Pure/pattern.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/pattern.ML Sun Feb 13 17:15:14 2005 +0100
@@ -95,8 +95,8 @@
fun occurs(F,t,env) =
let fun occ(Var(G,_)) = (case Envir.lookup(env,G) of
- Some(t) => occ t
- | None => F=G)
+ SOME(t) => occ t
+ | NONE => F=G)
| occ(t1$t2) = occ t1 orelse occ t2
| occ(Abs(_,_,t)) = occ t
| occ _ = false
@@ -359,9 +359,9 @@
(Var(ixn,T), t) =>
if loose_bvar(t,0) then raise MATCH
else (case assoc_string_int(insts,ixn) of
- None => (typ_match tsig (tyinsts, (T, fastype_of t)),
+ NONE => (typ_match tsig (tyinsts, (T, fastype_of t)),
(ixn,t)::insts)
- | Some u => if t aeconv u then instsp else raise MATCH)
+ | SOME u => if t aeconv u then instsp else raise MATCH)
| (Free (a,T), Free (b,U)) =>
if a=b then (typ_match tsig (tyinsts,(T,U)), insts) else raise MATCH
| (Const (a,T), Const (b,U)) =>
@@ -415,8 +415,8 @@
Var(ixn,_) =>
let val is = ints_of pargs
in case assoc_string_int(itms,ixn) of
- None => (iTs,match_bind(itms,binders,ixn,is,obj))
- | Some u => if obj aeconv (red u is []) then env
+ NONE => (iTs,match_bind(itms,binders,ixn,is,obj))
+ | SOME u => if obj aeconv (red u is []) then env
else raise MATCH
end
| _ =>
@@ -486,49 +486,49 @@
fun match_rew tm (tm1, tm2) =
let val rtm = if_none (Term.rename_abs tm1 tm tm2) tm2
- in Some (subst_vars (match tsig (tm1, tm)) rtm, rtm)
- handle MATCH => None
+ in SOME (subst_vars (match tsig (tm1, tm)) rtm, rtm)
+ handle MATCH => NONE
end;
- fun rew (Abs (_, _, body) $ t) = Some (subst_bound (t, body), skel0)
+ fun rew (Abs (_, _, body) $ t) = SOME (subst_bound (t, body), skel0)
| rew tm = (case get_first (match_rew tm) rules of
- None => apsome (rpair skel0) (get_first (fn p => p tm) procs)
+ NONE => apsome (rpair skel0) (get_first (fn p => p tm) procs)
| x => x);
- fun rew1 (Var _) _ = None
+ fun rew1 (Var _) _ = NONE
| rew1 skel tm = (case rew2 skel tm of
- Some tm1 => (case rew tm1 of
- Some (tm2, skel') => Some (if_none (rew1 skel' tm2) tm2)
- | None => Some tm1)
- | None => (case rew tm of
- Some (tm1, skel') => Some (if_none (rew1 skel' tm1) tm1)
- | None => None))
+ SOME tm1 => (case rew tm1 of
+ SOME (tm2, skel') => SOME (if_none (rew1 skel' tm2) tm2)
+ | NONE => SOME tm1)
+ | NONE => (case rew tm of
+ SOME (tm1, skel') => SOME (if_none (rew1 skel' tm1) tm1)
+ | NONE => NONE))
and rew2 skel (tm1 $ tm2) = (case tm1 of
Abs (_, _, body) =>
let val tm' = subst_bound (tm2, body)
- in Some (if_none (rew2 skel0 tm') tm') end
+ in SOME (if_none (rew2 skel0 tm') tm') end
| _ =>
let val (skel1, skel2) = (case skel of
skel1 $ skel2 => (skel1, skel2)
| _ => (skel0, skel0))
in case rew1 skel1 tm1 of
- Some tm1' => (case rew1 skel2 tm2 of
- Some tm2' => Some (tm1' $ tm2')
- | None => Some (tm1' $ tm2))
- | None => (case rew1 skel2 tm2 of
- Some tm2' => Some (tm1 $ tm2')
- | None => None)
+ SOME tm1' => (case rew1 skel2 tm2 of
+ SOME tm2' => SOME (tm1' $ tm2')
+ | NONE => SOME (tm1' $ tm2))
+ | NONE => (case rew1 skel2 tm2 of
+ SOME tm2' => SOME (tm1 $ tm2')
+ | NONE => NONE)
end)
| rew2 skel (Abs (x, T, tm)) =
let
val (abs, tm') = variant_absfree (x, T, tm);
val skel' = (case skel of Abs (_, _, skel') => skel' | _ => skel0)
in case rew1 skel' tm' of
- Some tm'' => Some (abs tm'')
- | None => None
+ SOME tm'' => SOME (abs tm'')
+ | NONE => NONE
end
- | rew2 _ _ = None
+ | rew2 _ _ = NONE
in if_none (rew1 skel0 tm) tm end;
--- a/src/Pure/proof_general.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/proof_general.ML Sun Feb 13 17:15:14 2005 +0100
@@ -136,15 +136,15 @@
fun free_or_skolem x =
(case try Syntax.dest_skolem x of
- None => tagit free_tag x
- | Some x' => tagit skolem_tag x');
+ NONE => tagit free_tag x
+ | SOME x' => tagit skolem_tag x');
fun var_or_skolem s =
(case Syntax.read_var s of
Var ((x, i), _) =>
(case try Syntax.dest_skolem x of
- None => tagit var_tag s
- | Some x' => tagit skolem_tag (Syntax.string_of_vname (x', i)))
+ NONE => tagit var_tag s
+ | SOME x' => tagit skolem_tag (Syntax.string_of_vname (x', i)))
| _ => tagit var_tag s);
val proof_general_trans =
@@ -161,8 +161,8 @@
(* assembling PGIP packets *)
-val pgip_refseq = ref None : string option ref
-val pgip_refid = ref None : string option ref
+val pgip_refseq = ref NONE : string option ref
+val pgip_refid = ref NONE : string option ref
local
val myseq_no = ref 1; (* PGIP packet counter *)
@@ -370,7 +370,7 @@
(* misc commands for ProofGeneral/isa *)
fun thms_containing ss =
- ProofContext.print_thms_containing (ProofContext.init (the_context ())) None ss;
+ ProofContext.print_thms_containing (ProofContext.init (the_context ())) NONE ss;
val welcome = priority o Session.welcome;
val help = welcome;
@@ -414,14 +414,14 @@
fun which_context () =
(case Context.get_context () of
- Some thy => " Using current (dynamic!) one: " ^
- (case try PureThy.get_name thy of Some name => quote name | None => "<unnamed>")
- | None => "");
+ SOME thy => " Using current (dynamic!) one: " ^
+ (case try PureThy.get_name thy of SOME name => quote name | NONE => "<unnamed>")
+ | NONE => "");
fun try_update_thy_only file =
ThyLoad.cond_add_path (Path.dir (Path.unpack file)) (fn () =>
let val name = thy_name file in
- if is_some (ThyLoad.check_file None (ThyLoad.thy_path name))
+ if is_some (ThyLoad.check_file NONE (ThyLoad.thy_path name))
then update_thy_only name
else warning ("Unkown theory context of ML file." ^ which_context ())
end) ();
@@ -521,8 +521,8 @@
fun nat_option r = (pgipnat,
withdefault (fn () => string_of_int (!r)),
(fn s => (case Syntax.read_nat s of
- None => error ("nat_option: illegal value " ^ s)
- | Some i => r := i)));
+ NONE => error ("nat_option: illegal value " ^ s)
+ | SOME i => r := i)));
fun bool_option r = (pgipbool,
withdefault (fn () => Bool.toString (!r)),
@@ -547,8 +547,8 @@
val print_depth_option = (pgipnat,
withdefault (fn () => string_of_int (!pg_print_depth_val)),
(fn s => (case Syntax.read_nat s of
- None => error ("print_depth_option: illegal value" ^ s)
- | Some i => pg_print_depth i)))
+ NONE => error ("print_depth_option: illegal value" ^ s)
+ | SOME i => pg_print_depth i)))
end
val preferences = ref
@@ -656,8 +656,8 @@
fun xml_idtable ty ctx ids =
let
- fun ctx_attr (Some c)= [("context", c)]
- | ctx_attr None = []
+ fun ctx_attr (SOME c)= [("context", c)]
+ | ctx_attr NONE = []
fun xmlid x = XML.element "identifier" [] [XML.text x];
in
XML.element "idtable" (("objtype", ty)::ctx_attr ctx)
@@ -669,7 +669,7 @@
fun addids t = issue_pgip "addids" [] t
fun delids t = issue_pgip "delids" [] t
-fun delallids ty = setids (xml_idtable ty None [])
+fun delallids ty = setids (xml_idtable ty NONE [])
fun send_thy_idtable ctx thys = setids (xml_idtable objtype_thy ctx thys)
fun clear_thy_idtable () = delallids objtype_thy
@@ -688,24 +688,24 @@
local
val topthy = Toplevel.theory_of o Toplevel.get_state
- fun pthm thy name = print_thm (get_thm thy (name, None))
+ fun pthm thy name = print_thm (get_thm thy (name, NONE))
fun idvalue tp nm = ("idvalue",[("objtype",tp),("name",nm)])
in
-fun askids (None, Some "theory") = send_thy_idtable None (ThyInfo.names())
- | askids (None, None) = send_thy_idtable None (ThyInfo.names())
+fun askids (NONE, SOME "theory") = send_thy_idtable NONE (ThyInfo.names())
+ | askids (NONE, NONE) = send_thy_idtable NONE (ThyInfo.names())
(* only theories known in top context *)
- | askids (Some thy, Some "theory") = send_thy_idtable (Some thy) (ThyInfo.get_preds thy)
- | askids (Some thy, Some "theorem") = send_thm_idtable (Some thy) (ThyInfo.get_theory thy)
- | askids (Some thy, None) = (send_thy_idtable (Some thy) (ThyInfo.get_preds thy);
- send_thm_idtable (Some thy) (ThyInfo.get_theory thy))
- | askids (_, Some ot) = error ("No objects of type "^(quote ot)^" are available here.")
+ | askids (SOME thy, SOME "theory") = send_thy_idtable (SOME thy) (ThyInfo.get_preds thy)
+ | askids (SOME thy, SOME "theorem") = send_thm_idtable (SOME thy) (ThyInfo.get_theory thy)
+ | askids (SOME thy, NONE) = (send_thy_idtable (SOME thy) (ThyInfo.get_preds thy);
+ send_thm_idtable (SOME thy) (ThyInfo.get_theory thy))
+ | askids (_, SOME ot) = error ("No objects of type "^(quote ot)^" are available here.")
fun showid (_, "theory", n) =
with_displaywrap (idvalue "theory" n) (fn ()=>(print_theory (ThyInfo.get_theory n)))
- | showid (Some thy, "theorem", t) =
+ | showid (SOME thy, "theorem", t) =
with_displaywrap (idvalue "theorem" t) (fn ()=>(pthm (ThyInfo.get_theory thy) t))
- | showid (None, "theorem", t) =
+ | showid (NONE, "theorem", t) =
with_displaywrap (idvalue "theorem" t) (fn ()=>pthm (topthy()) t)
| showid (_, ot, _) = error ("Cannot show objects of type "^ot)
@@ -779,15 +779,15 @@
(* FIXME: allow dynamic extensions to language here: we need a hook in
outer_syntax.ML to reset this table. *)
- val keywords_classification_table = ref (None:(string Symtab.table) option)
+ val keywords_classification_table = ref (NONE:(string Symtab.table) option)
fun get_keywords_classification_table () =
case (!keywords_classification_table) of
- Some t => t
- | None => (let
+ SOME t => t
+ | NONE => (let
val tab = foldl (fn (tab,(c,_,k,_))=>Symtab.update((c,k),tab))
(Symtab.empty,OuterSyntax.dest_parsers())
- in (keywords_classification_table := Some tab; tab) end)
+ in (keywords_classification_table := SOME tab; tab) end)
@@ -929,8 +929,8 @@
let
val classification = Symtab.lookup (get_keywords_classification_table(),name)
in case classification of
- Some k => (xmls_of_kind k (name,toks,str))
- | None => (* an uncategorized keyword ignored for undo (maybe wrong) *)
+ SOME k => (xmls_of_kind k (name,toks,str))
+ | NONE => (* an uncategorized keyword ignored for undo (maybe wrong) *)
(parse_warning ("Uncategorized command found: " ^ name ^ " -- using spuriouscmd");
markup_text str "spuriouscmd")
end
@@ -1063,13 +1063,13 @@
fun setpref name value =
(case assoc (allprefs(), name) of
- None => warning ("Unknown pref: " ^ quote name)
- | Some (_, (_, _, set)) => set value);
+ NONE => warning ("Unknown pref: " ^ quote name)
+ | SOME (_, (_, _, set)) => set value);
fun getpref name =
(case assoc (allprefs(), name) of
- None => warning ("Unknown pref: " ^ quote name)
- | Some (_, (_, (_,get), _)) =>
+ NONE => warning ("Unknown pref: " ^ quote name)
+ | SOME (_, (_, (_,get), _)) =>
issue_pgip "prefval" [("name", name)] (get ()));
@@ -1116,8 +1116,8 @@
fun xmlattr attr attrs =
(case xmlattro attr attrs of
- Some value => value
- | None => raise PGIP ("Missing attribute: " ^ attr))
+ SOME value => value
+ | NONE => raise PGIP ("Missing attribute: " ^ attr))
val thyname_attro = xmlattro "thyname"
val thyname_attr = xmlattr "thyname"
@@ -1137,7 +1137,7 @@
val topthy = Toplevel.theory_of o Toplevel.get_state
val topthy_name = PureThy.get_name o topthy
- val currently_open_file = ref (None : string option)
+ val currently_open_file = ref (NONE : string option)
in
fun process_pgip_element pgipxml = (case pgipxml of
@@ -1202,12 +1202,12 @@
| "retracttheory" => isarcmd ("kill_thy " ^ (quote (thyname_attr attrs)))
| "loadfile" => use_thy_or_ml_file (fileurl_of attrs)
| "openfile" => (inform_file_retracted (fileurl_of attrs);
- currently_open_file := Some (fileurl_of attrs))
+ currently_open_file := SOME (fileurl_of attrs))
| "closefile" => (case !currently_open_file of
- Some f => (inform_file_processed f;
- currently_open_file := None)
- | None => raise PGIP ("closefile when no file is open!"))
- | "abortfile" => (currently_open_file := None) (* perhaps error for no file open *)
+ SOME f => (inform_file_processed f;
+ currently_open_file := NONE)
+ | NONE => raise PGIP ("closefile when no file is open!"))
+ | "abortfile" => (currently_open_file := NONE) (* perhaps error for no file open *)
| "changecwd" => ThyLoad.add_path (dirname_attr attrs)
| "systemcmd" => isarscript data
(* unofficial command for debugging *)
@@ -1215,8 +1215,8 @@
| _ => raise PGIP ("Unrecognized PGIP element: " ^ elem)))
fun process_pgip_tree xml =
- (pgip_refseq := None;
- pgip_refid := None;
+ (pgip_refseq := NONE;
+ pgip_refid := NONE;
(case xml of
XML.Elem ("pgip", attrs, pgips) =>
(let
@@ -1254,28 +1254,28 @@
handle e => raise XML_PARSE
in
case pgipo of
- None => ()
- | Some (pgip,src') =>
- ((process_pgip_tree pgip); loop src') handle e => handler (e,Some src')
- end handle e => handler (e,Some src) (* i.e. error in ready/XML parse *)
+ NONE => ()
+ | SOME (pgip,src') =>
+ ((process_pgip_tree pgip); loop src') handle e => handler (e,SOME src')
+ end handle e => handler (e,SOME src) (* i.e. error in ready/XML parse *)
and handler (e,srco) =
case (e,srco) of
- (XML_PARSE,Some src) =>
+ (XML_PARSE,SOME src) =>
(* (!error_fn "Invalid XML input, aborting"; ()) NB: loop OK for tty, but want exit on EOF *)
panic "Invalid XML input, aborting"
| (PGIP_QUIT,_) => ()
- | (ERROR,Some src) => loop src (* message already given *)
- | (Interrupt,Some src) => (!error_fn "Interrupt during PGIP processing"; loop src)
- | (Toplevel.UNDEF,Some src) => (error "No working context defined"; loop src) (* usually *)
- | (e,Some src) => (error (Toplevel.exn_message e); loop src)
- | (_,None) => ()
+ | (ERROR,SOME src) => loop src (* message already given *)
+ | (Interrupt,SOME src) => (!error_fn "Interrupt during PGIP processing"; loop src)
+ | (Toplevel.UNDEF,SOME src) => (error "No working context defined"; loop src) (* usually *)
+ | (e,SOME src) => (error (Toplevel.exn_message e); loop src)
+ | (_,NONE) => ()
in
(* NB: simple tty for now; later might read from other tty-style sources (e.g. sockets). *)
val xmlP = XML.scan_comment_whspc |-- XML.parse_elem >> single
- val tty_src = Source.set_prompt "" (Source.source Symbol.stopper xmlP None Source.tty)
+ val tty_src = Source.set_prompt "" (Source.source Symbol.stopper xmlP NONE Source.tty)
val pgip_toplevel = loop
end
@@ -1388,7 +1388,7 @@
fun make_elisp_commands commands kind =
defconst kind (mapfilter
- (fn (c, _, k, _) => if k = kind then Some c else None)
+ (fn (c, _, k, _) => if k = kind then SOME c else NONE)
commands);
fun make_elisp_syntax (keywords, commands) =
--- a/src/Pure/proofterm.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/proofterm.ML Sun Feb 13 17:15:14 2005 +0100
@@ -136,8 +136,8 @@
| oras_of (tabs, prf1 %% prf2) = oras_of (oras_of (tabs, prf1), prf2)
| oras_of (tabs as (thms, oras), PThm ((name, _), prf, prop, _)) =
(case Symtab.lookup (thms, name) of
- None => oras_of ((Symtab.update ((name, [prop]), thms), oras), prf)
- | Some ps => if prop mem ps then tabs else
+ NONE => oras_of ((Symtab.update ((name, [prop]), thms), oras), prf)
+ | SOME ps => if prop mem ps then tabs else
oras_of ((Symtab.update ((name, prop::ps), thms), oras), prf))
| oras_of ((thms, oras), prf as Oracle _) = (thms, prf ins oras)
| oras_of (tabs, MinProof prfs) = foldl oras_of (tabs, prfs)
@@ -152,8 +152,8 @@
| thms_of_proof tab (prf % _) = thms_of_proof tab prf
| thms_of_proof tab (prf' as PThm ((s, _), prf, prop, _)) =
(case Symtab.lookup (tab, s) of
- None => thms_of_proof (Symtab.update ((s, [(prop, prf')]), tab)) prf
- | Some ps => if exists (equal prop o fst) ps then tab else
+ NONE => thms_of_proof (Symtab.update ((s, [(prop, prf')]), tab)) prf
+ | SOME ps => if exists (equal prop o fst) ps then tab else
thms_of_proof (Symtab.update ((s, (prop, prf')::ps), tab)) prf)
| thms_of_proof tab (MinProof prfs) = foldl (uncurry thms_of_proof) (tab, prfs)
| thms_of_proof tab _ = tab;
@@ -197,7 +197,7 @@
fun infer_derivs' f = infer_derivs (K f) (false, MinProof []);
-fun (prf %> t) = prf % Some t;
+fun (prf %> t) = prf % SOME t;
val proof_combt = foldl (op %>);
val proof_combt' = foldl (op %);
@@ -217,11 +217,11 @@
(PThm (_, prf', _, _), _) => prf'
| _ => prf);
-val mk_Abst = foldr (fn ((s, T:typ), prf) => Abst (s, None, prf));
-fun mk_AbsP (i, prf) = funpow i (fn prf => AbsP ("H", None, prf)) prf;
+val mk_Abst = foldr (fn ((s, T:typ), prf) => Abst (s, NONE, prf));
+fun mk_AbsP (i, prf) = funpow i (fn prf => AbsP ("H", NONE, prf)) prf;
-fun apsome' f None = raise SAME
- | apsome' f (Some x) = Some (f x);
+fun apsome' f NONE = raise SAME
+ | apsome' f (SOME x) = SOME (f x);
fun same f x =
let val x' = f x
@@ -237,25 +237,25 @@
handle SAME => prf % apsome' (same f) t)
| mapp (prf1 %% prf2) = (mapp prf1 %% mapph prf2
handle SAME => prf1 %% mapp prf2)
- | mapp (PThm (a, prf, prop, Some Ts)) =
- PThm (a, prf, prop, Some (same (map g) Ts))
- | mapp (PAxm (a, prop, Some Ts)) =
- PAxm (a, prop, Some (same (map g) Ts))
+ | mapp (PThm (a, prf, prop, SOME Ts)) =
+ PThm (a, prf, prop, SOME (same (map g) Ts))
+ | mapp (PAxm (a, prop, SOME Ts)) =
+ PAxm (a, prop, SOME (same (map g) Ts))
| mapp _ = raise SAME
and mapph prf = (mapp prf handle SAME => prf)
in mapph end;
-fun fold_proof_terms f g (a, Abst (_, Some T, prf)) = fold_proof_terms f g (g (T, a), prf)
- | fold_proof_terms f g (a, Abst (_, None, prf)) = fold_proof_terms f g (a, prf)
- | fold_proof_terms f g (a, AbsP (_, Some t, prf)) = fold_proof_terms f g (f (t, a), prf)
- | fold_proof_terms f g (a, AbsP (_, None, prf)) = fold_proof_terms f g (a, prf)
- | fold_proof_terms f g (a, prf % Some t) = f (t, fold_proof_terms f g (a, prf))
- | fold_proof_terms f g (a, prf % None) = fold_proof_terms f g (a, prf)
+fun fold_proof_terms f g (a, Abst (_, SOME T, prf)) = fold_proof_terms f g (g (T, a), prf)
+ | fold_proof_terms f g (a, Abst (_, NONE, prf)) = fold_proof_terms f g (a, prf)
+ | fold_proof_terms f g (a, AbsP (_, SOME t, prf)) = fold_proof_terms f g (f (t, a), prf)
+ | fold_proof_terms f g (a, AbsP (_, NONE, prf)) = fold_proof_terms f g (a, prf)
+ | fold_proof_terms f g (a, prf % SOME t) = f (t, fold_proof_terms f g (a, prf))
+ | fold_proof_terms f g (a, prf % NONE) = fold_proof_terms f g (a, prf)
| fold_proof_terms f g (a, prf1 %% prf2) = fold_proof_terms f g
(fold_proof_terms f g (a, prf1), prf2)
- | fold_proof_terms _ g (a, PThm (_, _, _, Some Ts)) = foldr g (Ts, a)
- | fold_proof_terms _ g (a, PAxm (_, prop, Some Ts)) = foldr g (Ts, a)
+ | fold_proof_terms _ g (a, PThm (_, _, _, SOME Ts)) = foldr g (Ts, a)
+ | fold_proof_terms _ g (a, PAxm (_, prop, SOME Ts)) = foldr g (Ts, a)
| fold_proof_terms _ _ (a, _) = a;
val add_prf_names = fold_proof_terms add_term_names ((uncurry K) o swap);
@@ -342,10 +342,10 @@
fun prf_loose_bvar1 (prf1 %% prf2) k = prf_loose_bvar1 prf1 k orelse prf_loose_bvar1 prf2 k
- | prf_loose_bvar1 (prf % Some t) k = prf_loose_bvar1 prf k orelse loose_bvar1 (t, k)
- | prf_loose_bvar1 (_ % None) _ = true
- | prf_loose_bvar1 (AbsP (_, Some t, prf)) k = loose_bvar1 (t, k) orelse prf_loose_bvar1 prf k
- | prf_loose_bvar1 (AbsP (_, None, _)) k = true
+ | prf_loose_bvar1 (prf % SOME t) k = prf_loose_bvar1 prf k orelse loose_bvar1 (t, k)
+ | prf_loose_bvar1 (_ % NONE) _ = true
+ | prf_loose_bvar1 (AbsP (_, SOME t, prf)) k = loose_bvar1 (t, k) orelse prf_loose_bvar1 prf k
+ | prf_loose_bvar1 (AbsP (_, NONE, _)) k = true
| prf_loose_bvar1 (Abst (_, _, prf)) k = prf_loose_bvar1 prf (k+1)
| prf_loose_bvar1 _ _ = false;
@@ -363,12 +363,12 @@
(prf_add_loose_bnos plev tlev prf1 p)
| prf_add_loose_bnos plev tlev (prf % opt) (is, js) =
prf_add_loose_bnos plev tlev prf (case opt of
- None => (is, ~1 ins js)
- | Some t => (is, add_loose_bnos (t, tlev, js)))
+ NONE => (is, ~1 ins js)
+ | SOME t => (is, add_loose_bnos (t, tlev, js)))
| prf_add_loose_bnos plev tlev (AbsP (_, opt, prf)) (is, js) =
prf_add_loose_bnos (plev+1) tlev prf (case opt of
- None => (is, ~1 ins js)
- | Some t => (is, add_loose_bnos (t, tlev, js)))
+ NONE => (is, ~1 ins js)
+ | SOME t => (is, add_loose_bnos (t, tlev, js)))
| prf_add_loose_bnos plev tlev (Abst (_, _, prf)) p =
prf_add_loose_bnos plev (tlev+1) prf p
| prf_add_loose_bnos _ _ _ _ = ([], []);
@@ -455,8 +455,8 @@
fun thawT names =
map_type_tfree (fn (s, xs) => case assoc (names, s) of
- None => TFree (s, xs)
- | Some ixn => TVar (ixn, xs));
+ NONE => TFree (s, xs)
+ | SOME ixn => TVar (ixn, xs));
fun freeze names names' (t $ u) =
freeze names names' t $ freeze names names' u
@@ -476,8 +476,8 @@
| thaw names names' (Free (s, T)) =
let val T' = thawT names' T
in case assoc (names, s) of
- None => Free (s, T')
- | Some ixn => Var (ixn, T')
+ NONE => Free (s, T')
+ | SOME ixn => Var (ixn, T')
end
| thaw names names' (Var (ixn, T)) = Var (ixn, thawT names' T)
| thaw names names' t = t;
@@ -517,13 +517,13 @@
| abshyp _ _ = raise SAME
and abshyph i prf = (abshyp i prf handle SAME => prf)
in
- AbsP ("H", None (*h*), abshyph 0 prf)
+ AbsP ("H", NONE (*h*), abshyph 0 prf)
end;
(***** forall introduction *****)
-fun forall_intr_proof x a prf = Abst (a, None, prf_abstract_over x prf);
+fun forall_intr_proof x a prf = Abst (a, NONE, prf_abstract_over x prf);
(***** varify *****)
@@ -535,8 +535,8 @@
val fmap = fs ~~ variantlist (fs, map #1 ixns)
fun thaw (f as (a, S)) =
(case assoc (fmap, a) of
- None => TFree f
- | Some b => TVar ((b, 0), S));
+ NONE => TFree f
+ | SOME b => TVar ((b, 0), S));
in map_proof_terms (map_term_types (map_type_tfree thaw)) (map_type_tfree thaw) prf
end;
@@ -548,8 +548,8 @@
in ((ix, v) :: pairs, v :: used) end;
fun freeze_one alist (ix, sort) = (case assoc (alist, ix) of
- None => TVar (ix, sort)
- | Some name => TFree (name, sort));
+ NONE => TVar (ix, sort)
+ | SOME name => TFree (name, sort));
in
@@ -633,9 +633,9 @@
if b then (i-1, j, prf %% PBound i) else (i, j-1, prf %> Bound j);
fun lift Us bs i j (Const ("==>", _) $ A $ B) =
- AbsP ("H", None (*A*), lift Us (true::bs) (i+1) j B)
+ AbsP ("H", NONE (*A*), lift Us (true::bs) (i+1) j B)
| lift Us bs i j (Const ("all", _) $ Abs (a, T, t)) =
- Abst (a, None (*T*), lift (T::Us) (false::bs) i (j+1) t)
+ Abst (a, NONE (*T*), lift (T::Us) (false::bs) i (j+1) t)
| lift Us bs i j _ = proof_combP (lifth' (rev Us) [] prf,
map (fn k => (#3 (foldr mk_app (bs, (i-1, j-1, PBound k)))))
(i + k - 1 downto i));
@@ -646,8 +646,8 @@
(***** proof by assumption *****)
-fun mk_asm_prf (Const ("==>", _) $ A $ B) i = AbsP ("H", None (*A*), mk_asm_prf B (i+1))
- | mk_asm_prf (Const ("all", _) $ Abs (a, T, t)) i = Abst (a, None (*T*), mk_asm_prf t i)
+fun mk_asm_prf (Const ("==>", _) $ A $ B) i = AbsP ("H", NONE (*A*), mk_asm_prf B (i+1))
+ | mk_asm_prf (Const ("all", _) $ Abs (a, T, t)) i = Abst (a, NONE (*T*), mk_asm_prf t i)
| mk_asm_prf _ i = PBound i;
fun assumption_proof Bs Bi n prf =
@@ -658,9 +658,9 @@
(***** Composition of object rule with proof state *****)
fun flatten_params_proof i j n (Const ("==>", _) $ A $ B, k) =
- AbsP ("H", None (*A*), flatten_params_proof (i+1) j n (B, k))
+ AbsP ("H", NONE (*A*), flatten_params_proof (i+1) j n (B, k))
| flatten_params_proof i j n (Const ("all", _) $ Abs (a, T, t), k) =
- Abst (a, None (*T*), flatten_params_proof i (j+1) n (t, k))
+ Abst (a, NONE (*T*), flatten_params_proof i (j+1) n (t, k))
| flatten_params_proof i j n (_, k) = proof_combP (proof_combt (PBound (k+i),
map Bound (j-1 downto 0)), map PBound (i-1 downto 0 \ i-n));
@@ -706,24 +706,24 @@
val [reflexive_axm, symmetric_axm, transitive_axm, equal_intr_axm,
equal_elim_axm, abstract_rule_axm, combination_axm] =
- map (fn (s, t) => PAxm ("ProtoPure." ^ s, varify t, None)) equality_axms;
+ map (fn (s, t) => PAxm ("ProtoPure." ^ s, varify t, NONE)) equality_axms;
end;
-val reflexive = reflexive_axm % None;
+val reflexive = reflexive_axm % NONE;
fun symmetric (prf as PAxm ("ProtoPure.reflexive", _, _) % _) = prf
- | symmetric prf = symmetric_axm % None % None %% prf;
+ | symmetric prf = symmetric_axm % NONE % NONE %% prf;
fun transitive _ _ (PAxm ("ProtoPure.reflexive", _, _) % _) prf2 = prf2
| transitive _ _ prf1 (PAxm ("ProtoPure.reflexive", _, _) % _) = prf1
| transitive u (Type ("prop", [])) prf1 prf2 =
- transitive_axm % None % Some (remove_types u) % None %% prf1 %% prf2
+ transitive_axm % NONE % SOME (remove_types u) % NONE %% prf1 %% prf2
| transitive u T prf1 prf2 =
- transitive_axm % None % None % None %% prf1 %% prf2;
+ transitive_axm % NONE % NONE % NONE %% prf1 %% prf2;
fun abstract_rule x a prf =
- abstract_rule_axm % None % None %% forall_intr_proof x a prf;
+ abstract_rule_axm % NONE % NONE %% forall_intr_proof x a prf;
fun check_comb (PAxm ("ProtoPure.combination", _, _) % f % g % _ % _ %% prf %% _) =
is_some f orelse check_comb prf
@@ -737,23 +737,23 @@
val f = Envir.beta_norm f;
val g = Envir.beta_norm g;
val prf = if check_comb prf1 then
- combination_axm % None % None
+ combination_axm % NONE % NONE
else (case prf1 of
PAxm ("ProtoPure.reflexive", _, _) % _ =>
- combination_axm %> remove_types f % None
+ combination_axm %> remove_types f % NONE
| _ => combination_axm %> remove_types f %> remove_types g)
in
(case T of
Type ("fun", _) => prf %
(case head_of f of
- Abs _ => Some (remove_types t)
- | Var _ => Some (remove_types t)
- | _ => None) %
+ Abs _ => SOME (remove_types t)
+ | Var _ => SOME (remove_types t)
+ | _ => NONE) %
(case head_of g of
- Abs _ => Some (remove_types u)
- | Var _ => Some (remove_types u)
- | _ => None) %% prf1 %% prf2
- | _ => prf % None % None %% prf1 %% prf2)
+ Abs _ => SOME (remove_types u)
+ | Var _ => SOME (remove_types u)
+ | _ => NONE) %% prf1 %% prf2
+ | _ => prf % NONE % NONE %% prf1 %% prf2)
end;
fun equal_intr A B prf1 prf2 =
@@ -780,7 +780,7 @@
fun add_funvars Ts (vs, t) =
if is_fun (fastype_of1 (Ts, t)) then
vs union mapfilter (fn Var (ixn, T) =>
- if is_fun T then Some ixn else None | _ => None) (vars_of t)
+ if is_fun T then SOME ixn else NONE | _ => NONE) (vars_of t)
else vs;
fun add_npvars q p Ts (vs, Const ("==>", _) $ t $ u) =
@@ -821,10 +821,10 @@
let
val nvs = needed_vars prop;
val args = map (fn (v as Var (ixn, _)) =>
- if ixn mem nvs then Some v else None) (vars_of prop) @
- map Some (sort (make_ord atless) (term_frees prop));
+ if ixn mem nvs then SOME v else NONE) (vars_of prop) @
+ map SOME (sort (make_ord atless) (term_frees prop));
in
- proof_combt' (c (name, prop, None), args)
+ proof_combt' (c (name, prop, NONE), args)
end;
val axm_proof = gen_axm_proof PAxm;
@@ -835,7 +835,7 @@
in (b, is, ch, if ch then Abst (a, T, body') else prf) end
| shrink ls lev (prf as AbsP (a, t, body)) =
let val (b, is, ch, body') = shrink (lev::ls) lev body
- in (b orelse 0 mem is, mapfilter (fn 0 => None | i => Some (i-1)) is,
+ in (b orelse 0 mem is, mapfilter (fn 0 => NONE | i => SOME (i-1)) is,
ch, if ch then AbsP (a, t, body') else prf)
end
| shrink ls lev prf =
@@ -852,9 +852,9 @@
let val (is, ch, (ch', t')::ts', prf') = shrink' ls lev (t::ts) prfs prf1
in (is, ch orelse ch', ts', if ch orelse ch' then prf' % t' else prf) end
| shrink' ls lev ts prfs (prf as PBound i) =
- (if exists (fn Some (Bound j) => lev-j <= nth_elem (i, ls) | _ => true) ts
+ (if exists (fn SOME (Bound j) => lev-j <= nth_elem (i, ls) | _ => true) ts
orelse not (null (duplicates
- (foldl (fn (js, Some (Bound j)) => j :: js | (js, _) => js) ([], ts))))
+ (foldl (fn (js, SOME (Bound j)) => j :: js | (js, _) => js) ([], ts))))
orelse exists #1 prfs then [i] else [], false, map (pair false) ts, prf)
| shrink' ls lev ts prfs (prf as Hyp _) = ([], false, map (pair false) ts, prf)
| shrink' ls lev ts prfs (prf as MinProof _) =
@@ -868,11 +868,11 @@
val insts = take (length ts', map (fst o dest_Var) vs) ~~ ts';
val nvs = foldl (fn (ixns', (ixn, ixns)) =>
ixn ins (case assoc (insts, ixn) of
- Some (Some t) => if is_proj t then ixns union ixns' else ixns'
+ SOME (SOME t) => if is_proj t then ixns union ixns' else ixns'
| _ => ixns union ixns'))
(needed prop ts'' prfs, add_npvars false true [] ([], prop));
val insts' = map
- (fn (ixn, x as Some _) => if ixn mem nvs then (false, x) else (true, None)
+ (fn (ixn, x as SOME _) => if ixn mem nvs then (false, x) else (true, NONE)
| (_, x) => (false, x)) insts
in ([], false, insts' @ map (pair false) ts'', prf) end
and needed (Const ("==>", _) $ t $ u) ts ((b, _, _, _)::prfs) =
@@ -908,9 +908,9 @@
fun match_proof Ts tymatch =
let
- fun optmatch _ inst (None, _) = inst
- | optmatch _ _ (Some _, None) = raise PMatch
- | optmatch mtch inst (Some x, Some y) = mtch inst (x, y)
+ fun optmatch _ inst (NONE, _) = inst
+ | optmatch _ _ (SOME _, NONE) = raise PMatch
+ | optmatch mtch inst (SOME x, SOME y) = mtch inst (x, y)
fun matcht Ts j (pinst, tinst) (t, u) =
(pinst, fomatch Ts tymatch j tinst (t, Envir.beta_norm u));
@@ -958,8 +958,8 @@
val substT = typ_subst_TVars_Vartab tyinsts;
fun subst' lev (t as Var (ixn, _)) = (case assoc (insts, ixn) of
- None => t
- | Some u => incr_boundvars lev u)
+ NONE => t
+ | SOME u => incr_boundvars lev u)
| subst' lev (Const (s, T)) = Const (s, substT T)
| subst' lev (Free (s, T)) = Free (s, substT T)
| subst' lev (Abs (a, T, body)) = Abs (a, substT T, subst' (lev+1) body)
@@ -973,8 +973,8 @@
| subst plev tlev (prf %% prf') = subst plev tlev prf %% subst plev tlev prf'
| subst plev tlev (prf % t) = subst plev tlev prf % apsome (subst' tlev) t
| subst plev tlev (prf as Hyp (Var (ixn, _))) = (case assoc (pinst, ixn) of
- None => prf
- | Some prf' => incr_pboundvars plev tlev prf')
+ NONE => prf
+ | SOME prf' => incr_pboundvars plev tlev prf')
| subst _ _ (PThm (id, prf, prop, Ts)) =
PThm (id, prf, prop, apsome (map substT) Ts)
| subst _ _ (PAxm (id, prop, Ts)) =
@@ -988,7 +988,7 @@
let
fun matchrands (prf1 %% prf2) (prf1' %% prf2') =
could_unify prf2 prf2' andalso matchrands prf1 prf1'
- | matchrands (prf % Some t) (prf' % Some t') =
+ | matchrands (prf % SOME t) (prf' % SOME t') =
Term.could_unify (t, t') andalso matchrands prf prf'
| matchrands (prf % _) (prf' % _) = matchrands prf prf'
| matchrands _ _ = true
@@ -1017,69 +1017,69 @@
fun rewrite_prf tymatch (rules, procs) prf =
let
- fun rew _ (Abst (_, _, body) % Some t) = Some (prf_subst_bounds [t] body, skel0)
- | rew _ (AbsP (_, _, body) %% prf) = Some (prf_subst_pbounds [prf] body, skel0)
+ fun rew _ (Abst (_, _, body) % SOME t) = SOME (prf_subst_bounds [t] body, skel0)
+ | rew _ (AbsP (_, _, body) %% prf) = SOME (prf_subst_pbounds [prf] body, skel0)
| rew Ts prf = (case get_first (fn (_, r) => r Ts prf) procs of
- Some prf' => Some (prf', skel0)
- | None => get_first (fn (prf1, prf2) => Some (prf_subst
+ SOME prf' => SOME (prf', skel0)
+ | NONE => get_first (fn (prf1, prf2) => SOME (prf_subst
(match_proof Ts tymatch ([], (Vartab.empty, [])) (prf1, prf)) prf2, prf2)
- handle PMatch => None) (filter (could_unify prf o fst) rules));
+ handle PMatch => NONE) (filter (could_unify prf o fst) rules));
fun rew0 Ts (prf as AbsP (_, _, prf' %% PBound 0)) =
if prf_loose_Pbvar1 prf' 0 then rew Ts prf
else
let val prf'' = incr_pboundvars (~1) 0 prf'
- in Some (if_none (rew Ts prf'') (prf'', skel0)) end
- | rew0 Ts (prf as Abst (_, _, prf' % Some (Bound 0))) =
+ in SOME (if_none (rew Ts prf'') (prf'', skel0)) end
+ | rew0 Ts (prf as Abst (_, _, prf' % SOME (Bound 0))) =
if prf_loose_bvar1 prf' 0 then rew Ts prf
else
let val prf'' = incr_pboundvars 0 (~1) prf'
- in Some (if_none (rew Ts prf'') (prf'', skel0)) end
+ in SOME (if_none (rew Ts prf'') (prf'', skel0)) end
| rew0 Ts prf = rew Ts prf;
- fun rew1 _ (Hyp (Var _)) _ = None
+ fun rew1 _ (Hyp (Var _)) _ = NONE
| rew1 Ts skel prf = (case rew2 Ts skel prf of
- Some prf1 => (case rew0 Ts prf1 of
- Some (prf2, skel') => Some (if_none (rew1 Ts skel' prf2) prf2)
- | None => Some prf1)
- | None => (case rew0 Ts prf of
- Some (prf1, skel') => Some (if_none (rew1 Ts skel' prf1) prf1)
- | None => None))
+ SOME prf1 => (case rew0 Ts prf1 of
+ SOME (prf2, skel') => SOME (if_none (rew1 Ts skel' prf2) prf2)
+ | NONE => SOME prf1)
+ | NONE => (case rew0 Ts prf of
+ SOME (prf1, skel') => SOME (if_none (rew1 Ts skel' prf1) prf1)
+ | NONE => NONE))
- and rew2 Ts skel (prf % Some t) = (case prf of
+ and rew2 Ts skel (prf % SOME t) = (case prf of
Abst (_, _, body) =>
let val prf' = prf_subst_bounds [t] body
- in Some (if_none (rew2 Ts skel0 prf') prf') end
+ in SOME (if_none (rew2 Ts skel0 prf') prf') end
| _ => (case rew1 Ts (case skel of skel' % _ => skel' | _ => skel0) prf of
- Some prf' => Some (prf' % Some t)
- | None => None))
- | rew2 Ts skel (prf % None) = apsome (fn prf' => prf' % None)
+ SOME prf' => SOME (prf' % SOME t)
+ | NONE => NONE))
+ | rew2 Ts skel (prf % NONE) = apsome (fn prf' => prf' % NONE)
(rew1 Ts (case skel of skel' % _ => skel' | _ => skel0) prf)
| rew2 Ts skel (prf1 %% prf2) = (case prf1 of
AbsP (_, _, body) =>
let val prf' = prf_subst_pbounds [prf2] body
- in Some (if_none (rew2 Ts skel0 prf') prf') end
+ in SOME (if_none (rew2 Ts skel0 prf') prf') end
| _ =>
let val (skel1, skel2) = (case skel of
skel1 %% skel2 => (skel1, skel2)
| _ => (skel0, skel0))
in case rew1 Ts skel1 prf1 of
- Some prf1' => (case rew1 Ts skel2 prf2 of
- Some prf2' => Some (prf1' %% prf2')
- | None => Some (prf1' %% prf2))
- | None => (case rew1 Ts skel2 prf2 of
- Some prf2' => Some (prf1 %% prf2')
- | None => None)
+ SOME prf1' => (case rew1 Ts skel2 prf2 of
+ SOME prf2' => SOME (prf1' %% prf2')
+ | NONE => SOME (prf1' %% prf2))
+ | NONE => (case rew1 Ts skel2 prf2 of
+ SOME prf2' => SOME (prf1 %% prf2')
+ | NONE => NONE)
end)
| rew2 Ts skel (Abst (s, T, prf)) = (case rew1 (if_none T dummyT :: Ts)
(case skel of Abst (_, _, skel') => skel' | _ => skel0) prf of
- Some prf' => Some (Abst (s, T, prf'))
- | None => None)
+ SOME prf' => SOME (Abst (s, T, prf'))
+ | NONE => NONE)
| rew2 Ts skel (AbsP (s, t, prf)) = (case rew1 Ts
(case skel of AbsP (_, _, skel') => skel' | _ => skel0) prf of
- Some prf' => Some (AbsP (s, t, prf'))
- | None => None)
- | rew2 _ _ _ = None
+ SOME prf' => SOME (AbsP (s, t, prf'))
+ | NONE => NONE)
+ | rew2 _ _ _ = NONE
in if_none (rew1 [] skel0 prf) prf end;
@@ -1123,20 +1123,20 @@
val prop = Logic.list_implies (hyps, prop);
val nvs = needed_vars prop;
val args = map (fn (v as Var (ixn, _)) =>
- if ixn mem nvs then Some v else None) (vars_of prop) @
- map Some (sort (make_ord atless) (term_frees prop));
+ if ixn mem nvs then SOME v else NONE) (vars_of prop) @
+ map SOME (sort (make_ord atless) (term_frees prop));
val opt_prf = if ! proofs = 2 then
#4 (shrink [] 0 (rewrite_prf fst (ProofData.get_sg sign)
(foldr (uncurry implies_intr_proof) (hyps, prf))))
else MinProof (mk_min_proof ([], prf));
val head = (case strip_combt (fst (strip_combP prf)) of
- (PThm ((old_name, _), prf', prop', None), args') =>
+ (PThm ((old_name, _), prf', prop', NONE), args') =>
if (old_name="" orelse old_name=name) andalso
prop = prop' andalso args = args' then
- PThm ((name, tags), prf', prop, None)
+ PThm ((name, tags), prf', prop, NONE)
else
- PThm ((name, tags), opt_prf, prop, None)
- | _ => PThm ((name, tags), opt_prf, prop, None))
+ PThm ((name, tags), opt_prf, prop, NONE)
+ | _ => PThm ((name, tags), opt_prf, prop, NONE))
in
proof_combP (proof_combt' (head, args), map Hyp hyps)
end;
--- a/src/Pure/pure_thy.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/pure_thy.ML Sun Feb 13 17:15:14 2005 +0100
@@ -135,14 +135,14 @@
(* selections *)
-fun the_thms _ (Some thms) = thms
- | the_thms name None = error ("Unknown theorem(s) " ^ quote name);
+fun the_thms _ (SOME thms) = thms
+ | the_thms name NONE = error ("Unknown theorem(s) " ^ quote name);
fun single_thm _ [thm] = thm
| single_thm name _ = error ("Single theorem expected " ^ quote name);
-fun select_thm (s, None) xs = xs
- | select_thm (s, Some is) xs = map
+fun select_thm (s, NONE) xs = xs
+ | select_thm (s, SOME is) xs = map
(fn i => (if i < 1 then raise LIST "" else nth_elem (i-1, xs)) handle LIST _ =>
error ("Bad subscript " ^ string_of_int i ^ " for " ^ quote s)) is;
@@ -195,9 +195,9 @@
fun thms_containing thy idx =
let
fun valid (name, ths) =
- (case try (transform_error (get_thms thy)) (name, None) of
- None => false
- | Some ths' => Library.equal_lists Thm.eq_thm (ths, ths'));
+ (case try (transform_error (get_thms thy)) (name, NONE) of
+ NONE => false
+ | SOME ths' => Library.equal_lists Thm.eq_thm (ths, ths'));
in
(thy :: Theory.ancestors_of thy)
|> map (gen_distinct eq_fst o filter valid o FactIndex.find idx o #index o ! o get_theorems)
@@ -214,7 +214,7 @@
(* elim: given a theorem thm,
find a theorem whose major premise eliminates the conclusion of thm *)
-fun top_const t = (case head_of t of Const (c, _) => Some c | _ => None);
+fun top_const t = (case head_of t of Const (c, _) => SOME c | _ => NONE);
(* This is a hack to remove the Trueprop constant that most logics use *)
fun rem_top (_ $ t) = t
@@ -259,8 +259,8 @@
in map (fn (_,_,t) => t) (sort thm_ord (select(named_thms, []))) end;
fun find_matching_thms extract thy prop =
- (case top_const prop of None => []
- | Some c => let val thms = thms_containing_consts thy [c]
+ (case top_const prop of NONE => []
+ | SOME c => let val thms = thms_containing_consts thy [c]
in select_match(c,prop,Theory.sign_of thy,thms,extract) end)
val find_intros =
@@ -322,8 +322,8 @@
val index' = FactIndex.add (K false) (index, (name, named_thms));
in
(case Symtab.lookup (thms_tab, name) of
- None => ()
- | Some thms' =>
+ NONE => ()
+ | SOME thms' =>
if Library.equal_lists Thm.eq_thm (thms', named_thms) then warn_same name
else warn_overwrite name);
r := {space = space', thms_tab = thms_tab', index = index'};
--- a/src/Pure/search.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/search.ML Sun Feb 13 17:15:14 2005 +0100
@@ -58,13 +58,13 @@
Suppresses duplicate solutions to minimize search space.*)
fun DEPTH_FIRST satp tac =
let val tac = tracify trace_DEPTH_FIRST tac
- fun depth used [] = None
+ fun depth used [] = NONE
| depth used (q::qs) =
case Seq.pull q of
- None => depth used qs
- | Some(st,stq) =>
+ NONE => depth used qs
+ | SOME(st,stq) =>
if satp st andalso not (gen_mem eq_thm (st, used))
- then Some(st, Seq.make
+ then SOME(st, Seq.make
(fn()=> depth (st::used) (stq::qs)))
else depth used (tac st :: stq :: qs)
in traced_tac (fn st => depth [] ([Seq.single st])) end;
@@ -156,10 +156,10 @@
implode (map (fn _ => "*") qs))
else ();
Seq.pull q) of
- None => depth (bnd,inc) qs
- | Some(st,stq) =>
+ NONE => depth (bnd,inc) qs
+ | SOME(st,stq) =>
if satp st (*solution!*)
- then Some(st, Seq.make
+ then SOME(st, Seq.make
(fn()=> depth (bnd,inc) ((k,np,rgd,stq)::qs)))
else
@@ -197,8 +197,8 @@
val trace_BEST_FIRST = ref false;
(*For creating output sequence*)
-fun some_of_list [] = None
- | some_of_list (x::l) = Some (x, Seq.make (fn () => some_of_list l));
+fun some_of_list [] = NONE
+ | some_of_list (x::l) = SOME (x, Seq.make (fn () => some_of_list l));
(*Check for and delete duplicate proof states*)
fun deleteAllMin prf heap =
@@ -219,7 +219,7 @@
(map pairsize nonsats, nprf_heap))
| (sats,_) => some_of_list sats)
and next nprf_heap =
- if ThmHeap.is_empty nprf_heap then None
+ if ThmHeap.is_empty nprf_heap then NONE
else
let val (n,prf) = ThmHeap.min nprf_heap
in if !trace_BEST_FIRST
@@ -266,8 +266,8 @@
else (l,m,th)::(l',n,th')::nths;
(*For creating output sequence*)
-fun some_of_list [] = None
- | some_of_list (x::l) = Some (x, Seq.make (fn () => some_of_list l));
+fun some_of_list [] = NONE
+ | some_of_list (x::l) = SOME (x, Seq.make (fn () => some_of_list l));
val trace_ASTAR = ref false;
@@ -280,7 +280,7 @@
=> next (foldr insert_with_level (map cost nonsats, nprfs))
| (sats,_) => some_of_list sats)
end and
- next [] = None
+ next [] = NONE
| next ((level,n,prf)::nprfs) =
(if !trace_ASTAR
then tracing("level = " ^ string_of_int level ^
--- a/src/Pure/sign.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/sign.ML Sun Feb 13 17:15:14 2005 +0100
@@ -237,7 +237,7 @@
(* id and self *)
fun check_stale (sg as Sg ({id, ...},
- {self = SgRef (Some (ref (Sg ({id = id', ...}, _)))), ...})) =
+ {self = SgRef (SOME (ref (Sg ({id = id', ...}, _)))), ...})) =
if id = id' then sg
else raise TERM ("Stale signature: " ^ str_of_sg sg, [])
| check_stale _ = sys_error "Sign.check_stale";
@@ -246,8 +246,8 @@
fun self_ref (sg as Sg (_, {self, ...})) = (check_stale sg; self);
-fun deref (SgRef (Some (ref sg))) = sg
- | deref (SgRef None) = sys_error "Sign.deref";
+fun deref (SgRef (SOME (ref sg))) = sg
+ | deref (SgRef NONE) = sys_error "Sign.deref";
fun name_of (sg as Sg ({id = ref name, ...}, _)) =
if name = "" orelse ord name = ord "#" then
@@ -378,8 +378,8 @@
fun entry data kind =
(case gen_assoc Object.eq_kind (data, kind) of
- None => []
- | Some x => [(kind, x)]);
+ NONE => []
+ | SOME x => [(kind, x)]);
fun merge_entries [(kind, (e, mths as (_, ext, _, _)))] =
(kind, (ext e handle exn => err_method "prep_ext" (Object.name_of_kind kind) exn, mths))
@@ -411,10 +411,10 @@
fun lookup_data sg tab kind =
let val name = Object.name_of_kind kind in
(case Symtab.lookup (tab, name) of
- Some (k, x) =>
+ SOME (k, x) =>
if Object.eq_kind (kind, k) then x
else err_access sg name
- | None => err_uninit sg name)
+ | NONE => err_uninit sg name)
end;
fun get_data kind f (sg as Sg (_, {data = Data tab, ...})) =
@@ -447,7 +447,7 @@
make_sign (id, self, tsig, ctab, syn, path, spaces, data, ext_stamps stamps id);
in
(case self of
- SgRef (Some r) => r := sign
+ SgRef (SOME r) => r := sign
| _ => sys_error "Sign.create_sign");
sign
end;
@@ -458,7 +458,7 @@
val _ = check_stale sg;
val (self', data') =
if is_draft sg andalso keep then (self, data)
- else (SgRef (Some (ref sg)), prep_ext_data data);
+ else (SgRef (SOME (ref sg)), prep_ext_data data);
in
create_sign self' stamps name
(extfun sg (syn, tsig, consts, (path, spaces), data') decls)
@@ -493,8 +493,8 @@
(*make full names*)
fun full _ "" = error "Attempt to declare empty name \"\""
- | full None bname = bname
- | full (Some path) bname =
+ | full NONE bname = bname
+ | full (SOME path) bname =
if NameSpace.is_qualified bname then
error ("Attempt to declare qualified name " ^ quote bname)
else
@@ -513,7 +513,7 @@
(*prepare mapping of names*)
fun mapping f add_xs t =
let
- fun f' x = let val y = f x in if x = y then None else Some (x, y) end;
+ fun f' x = let val y = f x in if x = y then NONE else SOME (x, y) end;
val table = mapfilter f' (add_xs (t, []));
fun lookup x = if_none (assoc (table, x)) x;
in lookup end;
@@ -568,7 +568,7 @@
val full_name = full o #path o rep_sg;
fun full_name_path sg elems =
- full (Some (if_none (#path (rep_sg sg)) [] @ NameSpace.unpack elems));
+ full (SOME (if_none (#path (rep_sg sg)) [] @ NameSpace.unpack elems));
end;
@@ -579,12 +579,12 @@
val pure_syn =
Syn (Syntax.pure_syn, (Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty));
-val dummy_sg = make_sign (ref "", SgRef None, Type.empty_tsig,
- Symtab.empty, pure_syn, Some [], [], empty_data, []);
+val dummy_sg = make_sign (ref "", SgRef NONE, Type.empty_tsig,
+ Symtab.empty, pure_syn, SOME [], [], empty_data, []);
val pre_pure =
- create_sign (SgRef (Some (ref dummy_sg))) [] "#"
- (pure_syn, Type.empty_tsig, Symtab.empty, (Some [], []), empty_data);
+ create_sign (SgRef (SOME (ref dummy_sg))) [] "#"
+ (pure_syn, Type.empty_tsig, Symtab.empty, (SOME [], []), empty_data);
val PureN = "Pure";
val CPureN = "CPure";
@@ -734,18 +734,18 @@
fun nodup_tvars (env, Type (_, Ts)) = nodup_tvars_list (env, Ts)
| nodup_tvars (env as (tfrees, tvars), T as TFree (v as (a, S))) =
(case assoc_string (tfrees, a) of
- Some S' =>
+ SOME S' =>
if S = S' then env
else raise TYPE ("Type variable " ^ quote a ^
" has two distinct sorts", [TFree (a, S'), T], [])
- | None => (v :: tfrees, tvars))
+ | NONE => (v :: tfrees, tvars))
| nodup_tvars (env as (tfrees, tvars), T as TVar (v as (a, S))) =
(case assoc_string_int (tvars, a) of
- Some S' =>
+ SOME S' =>
if S = S' then env
else raise TYPE ("Type variable " ^ quote (Syntax.string_of_vname a) ^
" has two distinct sorts", [TVar (a, S'), T], [])
- | None => (tfrees, v :: tvars))
+ | NONE => (tfrees, v :: tvars))
(*equivalent to foldl nodup_tvars_list, but 3X faster under Poly/ML*)
and nodup_tvars_list (env, []) = env
| nodup_tvars_list (env, T :: Ts) = nodup_tvars_list (nodup_tvars (env, T), Ts);
@@ -760,18 +760,18 @@
Const (c, T) => (env, nodup_tvars (envT, T))
| Free (v as (a, T)) =>
(case assoc_string (frees, a) of
- Some T' =>
+ SOME T' =>
if T = T' then (env, nodup_tvars (envT, T))
else raise TYPE ("Variable " ^ quote a ^
" has two distinct types", [T', T], [])
- | None => ((v :: frees, vars), nodup_tvars (envT, T)))
+ | NONE => ((v :: frees, vars), nodup_tvars (envT, T)))
| Var (v as (ixn, T)) =>
(case assoc_string_int (vars, ixn) of
- Some T' =>
+ SOME T' =>
if T = T' then (env, nodup_tvars (envT, T))
else raise TYPE ("Variable " ^ quote (Syntax.string_of_vname ixn) ^
" has two distinct types", [T', T], [])
- | None => ((frees, v :: vars), nodup_tvars (envT, T)))
+ | NONE => ((frees, v :: vars), nodup_tvars (envT, T)))
| Bound _ => envs
| Abs (_, T, t) => nodups (env, nodup_tvars (envT, T)) t
| s $ t => nodups (nodups envs s) t)
@@ -792,8 +792,8 @@
| check_atoms (Abs (_, _, t)) = check_atoms t
| check_atoms (Const (a, T)) =
(case const_type sg a of
- None => err ("Undeclared constant " ^ show_const a T)
- | Some U =>
+ NONE => err ("Undeclared constant " ^ show_const a T)
+ | SOME U =>
if typ_instance sg (T, U) then ()
else err ("Illegal type for constant " ^ show_const a T))
| check_atoms (Var ((x, i), _)) =
@@ -884,7 +884,7 @@
read_def_terms' (pp sign) (is_logtype sign) (syn_of sign) (sign, types, sorts);
fun simple_read_term sign T s =
- (read_def_terms (sign, K None, K None) [] true [(s, T)]
+ (read_def_terms (sign, K NONE, K NONE) [] true [(s, T)]
handle ERROR => error ("The error(s) above occurred for " ^ s)) |> #1 |> hd;
@@ -922,7 +922,7 @@
(* add type abbreviations *)
fun read_abbr sg syn tsig spaces (t, vs, rhs_src) =
- (t, vs, rd_raw_typ (make_syntax sg syn) tsig spaces (K None) rhs_src)
+ (t, vs, rd_raw_typ (make_syntax sg syn) tsig spaces (K NONE) rhs_src)
handle ERROR => error ("in type abbreviation " ^ t);
fun ext_abbrs rd_abbr sg (syn, tsig, ctab, (path, spaces), data) abbrs =
@@ -978,7 +978,7 @@
fun read_const sg syn tsig (path, spaces) (c, ty_src, mx) =
- (c, rd_raw_typ (make_syntax sg syn) tsig spaces (K None) ty_src, mx)
+ (c, rd_raw_typ (make_syntax sg syn) tsig spaces (K NONE) ty_src, mx)
handle ERROR => err_in_const (const_name path c mx);
fun ext_cnsts rd_const syn_only prmode sg (syn, tsig, ctab, (path, spaces), data)
@@ -1086,11 +1086,11 @@
fun ext_path _ (syn, tsig, ctab, (path, spaces), data) elems =
let
val path' =
- if elems = "//" then None
- else if elems = "/" then Some []
- else if elems = ".." andalso is_some path andalso path <> Some [] then
- Some (fst (split_last (the path)))
- else Some (if_none path [] @ NameSpace.unpack elems);
+ if elems = "//" then NONE
+ else if elems = "/" then SOME []
+ else if elems = ".." andalso is_some path andalso path <> SOME [] then
+ SOME (fst (split_last (the path)))
+ else SOME (if_none path [] @ NameSpace.unpack elems);
in
(syn, tsig, ctab, (path', spaces), data)
end;
@@ -1123,7 +1123,7 @@
fun copy (sg as Sg ({id = _, stamps, syn}, {self, tsig, consts, path, spaces, data})) =
let
val _ = check_stale sg;
- val self' = SgRef (Some (ref sg));
+ val self' = SgRef (SOME (ref sg));
val Data tab = data;
val data' = Data (Symtab.map copy_data tab);
in create_sign self' stamps "#" (syn, tsig, consts, (path, spaces), data') end;
@@ -1182,8 +1182,8 @@
(* trivial merge *)
-fun merge_refs (sgr1 as SgRef (Some (ref (sg1 as Sg ({stamps = s1, ...}, _)))),
- sgr2 as SgRef (Some (ref (sg2 as Sg ({stamps = s2, ...}, _))))) =
+fun merge_refs (sgr1 as SgRef (SOME (ref (sg1 as Sg ({stamps = s1, ...}, _)))),
+ sgr2 as SgRef (SOME (ref (sg2 as Sg ({stamps = s2, ...}, _))))) =
if fast_subsig (sg2, sg1) then sgr1
else if fast_subsig (sg1, sg2) then sgr2
else if subsig (sg2, sg1) then sgr1
@@ -1220,7 +1220,7 @@
val consts = Symtab.merge eq_snd (consts1, consts2)
handle Symtab.DUPS cs => err_dup_consts cs;
- val path = Some [];
+ val path = SOME [];
val kinds = distinct (map fst (spaces1 @ spaces2));
val spaces =
kinds ~~
@@ -1229,12 +1229,12 @@
val data = merge_data (data1, data2);
- val pre_sign = make_sign (ref "", SgRef (Some (ref dummy_sg)),
+ val pre_sign = make_sign (ref "", SgRef (SOME (ref dummy_sg)),
tsig1, consts, Syn (syntax, trfuns), path, spaces, data, ref "#" :: stamps1);
val tsig = Type.merge_tsigs (pp pre_sign) (tsig1, tsig2);
val self_ref = ref dummy_sg;
- val self = SgRef (Some self_ref);
+ val self = SgRef (SOME self_ref);
val sign = make_sign (ref "", self, tsig, consts, Syn (syntax, trfuns),
path, spaces, data, stamps);
in self_ref := sign; syn_of sign; sign end;
--- a/src/Pure/sorts.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/sorts.ML Sun Feb 13 17:15:14 2005 +0100
@@ -176,8 +176,8 @@
let
fun mg_dom c =
(case Library.assoc_string (Symtab.lookup_multi (arities, a), c) of
- None => raise DOMAIN (a, c)
- | Some Ss => Ss);
+ NONE => raise DOMAIN (a, c)
+ | SOME Ss => Ss);
val doms = map mg_dom S;
in foldl (ListPair.map (inter_sort classes)) (hd doms, tl doms) end;
@@ -205,27 +205,27 @@
let
val top_witn = (propT, []);
fun le S1 S2 = sort_le classes (S1, S2);
- fun get_solved S2 (T, S1) = if le S1 S2 then Some (T, S2) else None;
- fun get_hyp S2 S1 = if le S1 S2 then Some (TFree ("'hyp", S1), S2) else None;
- fun mg_dom t S = Some (mg_domain (classes, arities) t S) handle DOMAIN _ => None;
+ fun get_solved S2 (T, S1) = if le S1 S2 then SOME (T, S2) else NONE;
+ fun get_hyp S2 S1 = if le S1 S2 then SOME (TFree ("'hyp", S1), S2) else NONE;
+ fun mg_dom t S = SOME (mg_domain (classes, arities) t S) handle DOMAIN _ => NONE;
- fun witn_sort _ (solved_failed, []) = (solved_failed, Some top_witn)
+ fun witn_sort _ (solved_failed, []) = (solved_failed, SOME top_witn)
| witn_sort path ((solved, failed), S) =
- if exists (le S) failed then ((solved, failed), None)
+ if exists (le S) failed then ((solved, failed), NONE)
else
(case get_first (get_solved S) solved of
- Some w => ((solved, failed), Some w)
- | None =>
+ SOME w => ((solved, failed), SOME w)
+ | NONE =>
(case get_first (get_hyp S) hyps of
- Some w => ((w :: solved, failed), Some w)
- | None => witn_types path log_types ((solved, failed), S)))
+ SOME w => ((w :: solved, failed), SOME w)
+ | NONE => witn_types path log_types ((solved, failed), S)))
and witn_sorts path x = foldl_map (witn_sort path) x
- and witn_types _ [] ((solved, failed), S) = ((solved, S :: failed), None)
+ and witn_types _ [] ((solved, failed), S) = ((solved, S :: failed), NONE)
| witn_types path (t :: ts) (solved_failed, S) =
(case mg_dom t S of
- Some SS =>
+ SOME SS =>
(*do not descend into stronger args (achieving termination)*)
if exists (fn D => le D S orelse exists (le D) path) SS then
witn_types path ts (solved_failed, S)
@@ -233,10 +233,10 @@
let val ((solved', failed'), ws) = witn_sorts (S :: path) (solved_failed, SS) in
if forall is_some ws then
let val w = (Type (t, map (#1 o the) ws), S)
- in ((w :: solved', failed'), Some w) end
+ in ((w :: solved', failed'), SOME w) end
else witn_types path ts ((solved', failed'), S)
end
- | None => witn_types path ts (solved_failed, S));
+ | NONE => witn_types path ts (solved_failed, S));
in witn_sorts [] (([], []), sorts) end;
@@ -248,9 +248,9 @@
fun witness_sorts (classes, arities) log_types hyps sorts =
let
(*double check result of witness search*)
- fun check_result None = None
- | check_result (Some (T, S)) =
- if of_sort (classes, arities) (T, S) then Some (T, S)
+ fun check_result NONE = NONE
+ | check_result (SOME (T, S)) =
+ if of_sort (classes, arities) (T, S) then SOME (T, S)
else sys_error ("Sorts.witness_sorts: bad witness for sort " ^ str_of_sort S);
in mapfilter check_result (#2 (witness_aux (classes, arities) log_types hyps sorts)) end;
--- a/src/Pure/tactic.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/tactic.ML Sun Feb 13 17:15:14 2005 +0100
@@ -124,7 +124,7 @@
(*Discover which goal is chosen: SOMEGOAL(trace_goalno_tac tac) *)
fun trace_goalno_tac tac i st =
case Seq.pull(tac i st) of
- None => Seq.empty
+ NONE => Seq.empty
| seqcell => (tracing ("Subgoal " ^ string_of_int i ^ " selected");
Seq.make(fn()=> seqcell));
@@ -132,8 +132,8 @@
fun rule_by_tactic tac rl =
let val (st, thaw) = freeze_thaw (zero_var_indexes rl)
in case Seq.pull (tac st) of
- None => raise THM("rule_by_tactic", 0, [rl])
- | Some(st',_) => Thm.varifyT (thaw st')
+ NONE => raise THM("rule_by_tactic", 0, [rl])
+ | SOME(st',_) => Thm.varifyT (thaw st')
end;
(*** Basic tactics ***)
@@ -227,7 +227,7 @@
let val {sign,...} = rep_thm st
and params = params_of_state st i
and rts = types_sorts rule and (types,sorts) = types_sorts st
- fun types'(a,~1) = (case assoc(params,a) of None => types(a,~1) | sm => sm)
+ fun types'(a,~1) = (case assoc(params,a) of NONE => types(a,~1) | sm => sm)
| types'(ixn) = types ixn;
val used = add_term_tvarnames
(prop_of st $ prop_of rule,[])
@@ -420,8 +420,8 @@
fun insert_tagged_brl (kbrl as (k, (eres, th)), (inet, enet)) =
if eres then
(case try Thm.major_prem_of th of
- Some prem => (inet, Net.insert_term ((prem, kbrl), enet, K false))
- | None => error "insert_tagged_brl: elimination rule with no premises")
+ SOME prem => (inet, Net.insert_term ((prem, kbrl), enet, K false))
+ | NONE => error "insert_tagged_brl: elimination rule with no premises")
else (Net.insert_term ((concl_of th, kbrl), inet, K false), enet);
(*build a pair of nets for biresolution*)
@@ -435,8 +435,8 @@
fun delete_tagged_brl (brl as (eres, th), (inet, enet)) =
(if eres then
(case try Thm.major_prem_of th of
- Some prem => (inet, Net.delete_term ((prem, ((), brl)), enet, eq_kbrl))
- | None => (inet, enet)) (*no major premise: ignore*)
+ SOME prem => (inet, Net.delete_term ((prem, ((), brl)), enet, eq_kbrl))
+ | NONE => (inet, enet)) (*no major premise: ignore*)
else (Net.delete_term ((Thm.concl_of th, ((), brl)), inet, eq_kbrl), enet))
handle Net.DELETE => (inet,enet);
end;
@@ -575,8 +575,8 @@
it affects nothing but the names of bound variables!*)
fun rename_params_tac xs i =
case Library.find_first (not o Syntax.is_identifier) xs of
- Some x => error ("Not an identifier: " ^ x)
- | None =>
+ SOME x => error ("Not an identifier: " ^ x)
+ | NONE =>
(if !Logic.auto_rename
then (warning "Resetting Logic.auto_rename";
Logic.auto_rename := false)
@@ -613,15 +613,15 @@
(* remove premises that do not satisfy p; fails if all prems satisfy p *)
fun filter_prems_tac p =
- let fun Then None tac = Some tac
- | Then (Some tac) tac' = Some(tac THEN' tac');
+ let fun Then NONE tac = SOME tac
+ | Then (SOME tac) tac' = SOME(tac THEN' tac');
fun thins ((tac,n),H) =
if p H then (tac,n+1)
else (Then tac (rotate_tac n THEN' etac thin_rl),0);
in SUBGOAL(fn (subg,n) =>
let val Hs = Logic.strip_assums_hyp subg
- in case fst(foldl thins ((None,0),Hs)) of
- None => no_tac | Some tac => tac n
+ in case fst(foldl thins ((NONE,0),Hs)) of
+ NONE => no_tac | SOME tac => tac n
end)
end;
@@ -663,7 +663,7 @@
val goal = Drule.mk_triv_goal (cert_safe prop);
val goal' =
- (case Seq.pull (tac prems goal) of Some (goal', _) => goal' | _ => err "Tactic failed.");
+ (case Seq.pull (tac prems goal) of SOME (goal', _) => goal' | _ => err "Tactic failed.");
val ngoals = Thm.nprems_of goal';
val raw_result = goal' RS Drule.rev_triv_goal;
val prop' = prop_of raw_result;
--- a/src/Pure/tctical.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/tctical.ML Sun Feb 13 17:15:14 2005 +0100
@@ -94,7 +94,7 @@
Does not backtrack to tac2 if tac1 was initially chosen. *)
fun (tac1 ORELSE tac2) st =
case Seq.pull(tac1 st) of
- None => tac2 st
+ NONE => tac2 st
| sequencecell => Seq.make(fn()=> sequencecell);
@@ -116,7 +116,7 @@
*)
fun (tac THEN_ELSE (tac1, tac2)) st =
case Seq.pull(tac st) of
- None => tac2 st (*failed; try tactic 2*)
+ NONE => tac2 st (*failed; try tactic 2*)
| seqcell => Seq.flat (*succeeded; use tactic 1*)
(Seq.map tac1 (Seq.make(fn()=> seqcell)));
@@ -152,17 +152,17 @@
(*This version of EVERY avoids backtracking over repeated states*)
fun EVY (trail, []) st =
- Seq.make (fn()=> Some(st,
+ Seq.make (fn()=> SOME(st,
Seq.make (fn()=> Seq.pull (evyBack trail))))
| EVY (trail, tac::tacs) st =
case Seq.pull(tac st) of
- None => evyBack trail (*failed: backtrack*)
- | Some(st',q) => EVY ((st',q,tacs)::trail, tacs) st'
+ NONE => evyBack trail (*failed: backtrack*)
+ | SOME(st',q) => EVY ((st',q,tacs)::trail, tacs) st'
and evyBack [] = Seq.empty (*no alternatives*)
| evyBack ((st',q,tacs)::trail) =
case Seq.pull q of
- None => evyBack trail
- | Some(st,q') => if eq_thm (st',st)
+ NONE => evyBack trail
+ | SOME(st,q') => if eq_thm (st',st)
then evyBack ((st',q',tacs)::trail)
else EVY ((st,q',tacs)::trail, tacs) st
in
@@ -246,17 +246,17 @@
fun traced_tac seqf st =
(suppress_tracing := false;
Seq.make (fn()=> seqf st
- handle TRACE_EXIT st' => Some(st', Seq.empty)));
+ handle TRACE_EXIT st' => SOME(st', Seq.empty)));
(*Deterministic DO..UNTIL: only retains the first outcome; tail recursive.
Forces repitition until predicate on state is fulfilled.*)
fun DETERM_UNTIL p tac =
let val tac = tracify trace_REPEAT tac
- fun drep st = if p st then Some (st, Seq.empty)
+ fun drep st = if p st then SOME (st, Seq.empty)
else (case Seq.pull(tac st) of
- None => None
- | Some(st',_) => drep st')
+ NONE => NONE
+ | SOME(st',_) => drep st')
in traced_tac drep end;
(*Deterministic REPEAT: only retains the first outcome;
@@ -264,11 +264,11 @@
If non-negative, n bounds the number of repetitions.*)
fun REPEAT_DETERM_N n tac =
let val tac = tracify trace_REPEAT tac
- fun drep 0 st = Some(st, Seq.empty)
+ fun drep 0 st = SOME(st, Seq.empty)
| drep n st =
(case Seq.pull(tac st) of
- None => Some(st, Seq.empty)
- | Some(st',_) => drep (n-1) st')
+ NONE => SOME(st, Seq.empty)
+ | SOME(st',_) => drep (n-1) st')
in traced_tac (drep n) end;
(*Allows any number of repetitions*)
@@ -279,12 +279,12 @@
let val tac = tracify trace_REPEAT tac
fun rep qs st =
case Seq.pull(tac st) of
- None => Some(st, Seq.make(fn()=> repq qs))
- | Some(st',q) => rep (q::qs) st'
- and repq [] = None
+ NONE => SOME(st, Seq.make(fn()=> repq qs))
+ | SOME(st',q) => rep (q::qs) st'
+ and repq [] = NONE
| repq(q::qs) = case Seq.pull q of
- None => repq qs
- | Some(st,q) => rep (q::qs) st
+ NONE => repq qs
+ | SOME(st,q) => rep (q::qs) st
in traced_tac (rep []) end;
(*Repeat 1 or more times*)
--- a/src/Pure/term.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/term.ML Sun Feb 13 17:15:14 2005 +0100
@@ -468,7 +468,7 @@
Abs (if_none (assoc_string (ren, x)) x, T, ren_abs b)
| ren_abs (f $ t) = ren_abs f $ ren_abs t
| ren_abs t = t
- in if null ren then None else Some (ren_abs t) end;
+ in if null ren then NONE else SOME (ren_abs t) end;
(*Accumulate all 'loose' bound vars referring to level 'lev' or beyond.
(Bound 0) is loose at level 0 *)
@@ -596,8 +596,8 @@
| subst_free pairs =
let fun substf u =
case gen_assoc (op aconv) (pairs, u) of
- Some u' => u'
- | None => (case u of Abs(a,T,t) => Abs(a, T, substf t)
+ SOME u' => u'
+ | NONE => (case u of Abs(a,T,t) => Abs(a, T, substf t)
| t$u' => substf t $ substf u'
| _ => u)
in substf end;
@@ -671,8 +671,8 @@
let fun subst(Const(a,T)) = Const(a,typ_subst_TVars iTs T)
| subst(Free(a,T)) = Free(a,typ_subst_TVars iTs T)
| subst(v as Var(ixn,T)) = (case assoc(itms,ixn) of
- None => Var(ixn,typ_subst_TVars iTs T)
- | Some t => t)
+ NONE => Var(ixn,typ_subst_TVars iTs T)
+ | SOME t => t)
| subst(b as Bound _) = b
| subst(Abs(a,T,t)) = Abs(a,typ_subst_TVars iTs T,subst t)
| subst(f$t) = subst f $ subst t
@@ -1011,7 +1011,7 @@
let fun subst(Type(a, Ts)) = Type(a, map subst Ts)
| subst(T as TFree _) = T
| subst(T as TVar(ixn, _)) =
- (case Vartab.lookup (iTs, ixn) of None => T | Some(U) => U)
+ (case Vartab.lookup (iTs, ixn) of NONE => T | SOME(U) => U)
in subst T end;
(*Substitute for type Vars in a term, version using Vartab*)
@@ -1029,8 +1029,8 @@
fun compress_type T =
(case Typtab.lookup (! memo_types, T) of
- Some T' => T'
- | None =>
+ SOME T' => T'
+ | NONE =>
let val T' = (case T of Type (a, Ts) => Type (a, map compress_type Ts) | _ => T)
in memo_types := Typtab.update ((T', T'), ! memo_types); T' end);
@@ -1043,8 +1043,8 @@
| share_term (Abs (a, T, u)) = Abs (a, T, share_term u)
| share_term t =
(case Termtab.lookup (! memo_terms, t) of
- Some t' => t'
- | None => (memo_terms := Termtab.update ((t, t), ! memo_terms); t));
+ SOME t' => t'
+ | NONE => (memo_terms := Termtab.update ((t, t), ! memo_terms); t));
val compress_term = share_term o map_term_types compress_type;
--- a/src/Pure/theory.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/theory.ML Sun Feb 13 17:15:14 2005 +0100
@@ -276,11 +276,11 @@
in cert_axm sg (name, t) end
handle ERROR => err_in_axm name;
-fun read_axm sg name_str = read_def_axm (sg, K None, K None) [] name_str;
+fun read_axm sg name_str = read_def_axm (sg, K NONE, K NONE) [] name_str;
fun inferT_axm sg (name, pre_tm) =
let val (t, _) = Sign.infer_types (Sign.pp sg) sg
- (K None) (K None) [] true ([pre_tm], propT);
+ (K NONE) (K NONE) [] true ([pre_tm], propT);
in (name, no_vars sg t) end
handle ERROR => err_in_axm name;
@@ -329,8 +329,8 @@
fun clash_defn c_ty (name, tm) =
let val (c, ty') = dest_Const (head_of (fst (Logic.dest_equals (Logic.strip_imp_concl tm)))) in
- if clash_consts c_ty (c, ty') then Some (name, ty') else None
- end handle TERM _ => None;
+ if clash_consts c_ty (c, ty') then SOME (name, ty') else NONE
+ end handle TERM _ => NONE;
fun clash_defns c_ty axms =
distinct (mapfilter (clash_defn c_ty) axms);
@@ -402,9 +402,9 @@
let
fun is_final (c, ty) =
case Symtab.lookup(final_consts,c) of
- Some [] => true
- | Some tys => exists (curry (Sign.typ_instance sg) ty) tys
- | None => false;
+ SOME [] => true
+ | SOME tys => exists (curry (Sign.typ_instance sg) ty) tys
+ | NONE => false;
fun pretty_const (c, ty) = [Pretty.str c, Pretty.str " ::", Pretty.brk 1,
Pretty.quote (Sign.pretty_typ sg (#1 (Type.freeze_thaw_type ty)))];
@@ -419,8 +419,8 @@
val (c_ty as (c, ty), rhs) = dest_defn tm
handle TERM (msg, _) => err_in_defn sg name msg;
val c_decl =
- (case Sign.const_type sg c of Some T => T
- | None => err_in_defn sg name ("Undeclared constant " ^ quote c));
+ (case Sign.const_type sg c of SOME T => T
+ | NONE => err_in_defn sg name ("Undeclared constant " ^ quote c));
val clashed = clash_defns c_ty axms;
in
@@ -479,8 +479,8 @@
val (c,ty) = dest_Const tm
handle TERM _ => err "Can only make constants final";
val c_decl =
- (case Sign.const_type sign c of Some T => T
- | None => err ("Undeclared constant " ^ quote c));
+ (case Sign.const_type sign c of SOME T => T
+ | NONE => err ("Undeclared constant " ^ quote c));
val simple =
case overloading sign c_decl ty of
NoOverloading => true
@@ -491,15 +491,15 @@
in
if simple then
(case Symtab.lookup(finals,c) of
- Some [] => err "Constant already final"
+ SOME [] => err "Constant already final"
| _ => Symtab.update((c,[]),finals))
else
(case Symtab.lookup(finals,c) of
- Some [] => err "Constant already completely final"
- | Some tys => if exists (curry typ_instance ty) tys
+ SOME [] => err "Constant already completely final"
+ | SOME tys => if exists (curry typ_instance ty) tys
then err "Instance of constant is already final"
else Symtab.update((c,ty::gen_rem typ_instance (tys,ty)),finals)
- | None => Symtab.update((c,[ty]),finals))
+ | NONE => Symtab.update((c,[ty]),finals))
end;
val consts = map (prep_term sign) raw_terms;
val final_consts' = foldl mk_final (final_consts,consts);
@@ -525,7 +525,7 @@
| merge (_,[]) = []
| merge input = foldl merge_single input;
in
- Some o merge
+ SOME o merge
end;
--- a/src/Pure/thm.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/thm.ML Sun Feb 13 17:15:14 2005 +0100
@@ -152,7 +152,7 @@
Ctyp {sign_ref = Sign.self_ref sign, T = Sign.certify_typ sign T};
fun read_ctyp sign s =
- Ctyp {sign_ref = Sign.self_ref sign, T = Sign.read_typ (sign, K None) s};
+ Ctyp {sign_ref = Sign.self_ref sign, T = Sign.read_typ (sign, K NONE) s};
fun dest_ctyp (Ctyp {sign_ref, T = Type (s, Ts)}) =
map (fn T => Ctyp {sign_ref = sign_ref, T = T}) Ts
@@ -279,7 +279,7 @@
let val ([ct],tye) = read_def_cterms args used freeze [aT]
in (ct,tye) end;
-fun read_cterm sign = #1 o read_def_cterm (sign, K None, K None) [] true;
+fun read_cterm sign = #1 o read_def_cterm (sign, K NONE, K NONE) [] true;
(*tags provide additional comment, apart from the axiom/theorem name*)
@@ -479,12 +479,12 @@
let
val name = Sign.intern (Theory.sign_of theory) Theory.axiomK raw_name;
- fun get_ax [] = None
+ fun get_ax [] = NONE
| get_ax (thy :: thys) =
let val {sign, axioms, ...} = Theory.rep_theory thy in
(case Symtab.lookup (axioms, name) of
- Some t =>
- Some (fix_shyps [] []
+ SOME t =>
+ SOME (fix_shyps [] []
(Thm {sign_ref = Sign.self_ref sign,
der = Pt.infer_derivs' I
(false, Pt.axm_proof name t),
@@ -493,12 +493,12 @@
hyps = [],
tpairs = [],
prop = t}))
- | None => get_ax thys)
+ | NONE => get_ax thys)
end;
in
(case get_ax (theory :: Theory.ancestors_of theory) of
- Some thm => thm
- | None => raise THEORY ("No axiom " ^ quote name, [theory]))
+ SOME thm => thm
+ | NONE => raise THEORY ("No axiom " ^ quote name, [theory]))
end;
fun def_name name = name ^ "_def";
@@ -669,7 +669,7 @@
raise THM("forall_elim: type mismatch", 0, [th])
else let val thm = fix_shyps [th] []
(Thm{sign_ref= Sign.merge_refs(sign_ref,sign_reft),
- der = Pt.infer_derivs' (Pt.% o rpair (Some t)) der,
+ der = Pt.infer_derivs' (Pt.% o rpair (SOME t)) der,
maxidx = Int.max(maxidx, maxt),
shyps = [],
hyps = hyps,
@@ -1012,7 +1012,7 @@
raise THM("trivial: the term must have type prop", 0, [])
else fix_shyps [] []
(Thm{sign_ref = sign_ref,
- der = Pt.infer_derivs' I (false, Pt.AbsP ("H", None, Pt.PBound 0)),
+ der = Pt.infer_derivs' I (false, Pt.AbsP ("H", NONE, Pt.PBound 0)),
maxidx = maxidx,
shyps = [],
hyps = [],
@@ -1029,7 +1029,7 @@
fix_shyps [] []
(Thm {sign_ref = sign_ref,
der = Pt.infer_derivs' I
- (false, Pt.PAxm ("ProtoPure.class_triv:" ^ c, t, Some [])),
+ (false, Pt.PAxm ("ProtoPure.class_triv:" ^ c, t, SOME [])),
maxidx = maxidx,
shyps = [],
hyps = [],
@@ -1256,8 +1256,8 @@
fun rename_boundvars pat obj (thm as Thm {sign_ref, der, maxidx, hyps, shyps, tpairs, prop}) =
(case Term.rename_abs pat obj prop of
- None => thm
- | Some prop' => Thm
+ NONE => thm
+ | SOME prop' => Thm
{sign_ref = sign_ref,
der = der,
maxidx = maxidx,
@@ -1288,9 +1288,9 @@
val vids = map (#1 o #1 o dest_Var) vars;
fun rename(t as Var((x,i),T)) =
(case assoc(al,x) of
- Some(y) => if x mem_string vids orelse y mem_string vids then t
+ SOME(y) => if x mem_string vids orelse y mem_string vids then t
else Var((y,i),T)
- | None=> t)
+ | NONE=> t)
| rename(Abs(x,T,t)) =
Abs(if_none(assoc_string(al,x)) x, T, rename t)
| rename(f$t) = rename f $ rename t
@@ -1403,17 +1403,17 @@
fun tryasms (_, _, _, []) = Seq.empty
| tryasms (A, As, n, (t,u)::apairs) =
(case Seq.pull(Unify.unifiers(sign, env, (t,u)::dpairs)) of
- None => tryasms (A, As, n+1, apairs)
- | cell as Some((_,tpairs),_) =>
+ NONE => tryasms (A, As, n+1, apairs)
+ | cell as SOME((_,tpairs),_) =>
Seq.it_right (addth A (newAs(As, n, [BBi,(u,t)], tpairs)))
(Seq.make(fn()=> cell),
Seq.make(fn()=> Seq.pull (tryasms(A, As, n+1, apairs)))))
fun eres [] = raise THM("bicompose: no premises", 0, [orule,state])
- | eres (A1::As) = tryasms(Some A1, As, 1, Logic.assum_pairs(nlift+1,A1))
+ | eres (A1::As) = tryasms(SOME A1, As, 1, Logic.assum_pairs(nlift+1,A1))
(*ordinary resolution*)
- fun res(None) = Seq.empty
- | res(cell as Some((_,tpairs),_)) =
- Seq.it_right (addth None (newAs(rev rAs, 0, [BBi], tpairs)))
+ fun res(NONE) = Seq.empty
+ | res(cell as SOME((_,tpairs),_)) =
+ Seq.it_right (addth NONE (newAs(rev rAs, 0, [BBi], tpairs)))
(Seq.make (fn()=> cell), Seq.empty)
in if eres_flg then eres(rev rAs)
else res(Seq.pull(Unify.unifiers(sign, env, dpairs)))
@@ -1446,7 +1446,7 @@
if !Pattern.trace_unify_fail orelse
could_bires (Hs, B, eres_flg, rule)
then Seq.make (*delay processing remainder till needed*)
- (fn()=> Some(comp (eres_flg, lift rule, nprems_of rule),
+ (fn()=> SOME(comp (eres_flg, lift rule, nprems_of rule),
res brules))
else res brules
in Seq.flat (res brules) end;
@@ -1460,8 +1460,8 @@
val name = Sign.intern sg Theory.oracleK raw_name;
val oracle =
(case Symtab.lookup (oracles, name) of
- None => raise THM ("Unknown oracle: " ^ name, 0, [])
- | Some (f, _) => f);
+ NONE => raise THM ("Unknown oracle: " ^ name, 0, [])
+ | SOME (f, _) => f);
in
fn (sign, exn) =>
let
--- a/src/Pure/type.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/type.ML Sun Feb 13 17:15:14 2005 +0100
@@ -112,7 +112,7 @@
|> Library.sort (Library.int_ord o pairself #2) |> map #1;
val witness =
(case Sorts.witness_sorts (classes, arities) log_types [] [Graph.keys classes] of
- [w] => Some w | _ => None);
+ [w] => SOME w | _ => NONE);
in make_tsig (classes, default, types, arities, log_types, witness) end;
fun change_tsig f (TSig {classes, default, types, arities, log_types = _, witness = _}) =
@@ -165,13 +165,13 @@
fun nargs n = if length Ts <> n then err (bad_nargs c) else ();
in
(case Symtab.lookup (types, c) of
- Some (LogicalType n, _) => (nargs n; Type (c, Ts'))
- | Some (Abbreviation (vs, U, syn), _) => (nargs (length vs);
+ SOME (LogicalType n, _) => (nargs n; Type (c, Ts'))
+ | SOME (Abbreviation (vs, U, syn), _) => (nargs (length vs);
if syn then check_syntax c else ();
if normalize then inst_typ (vs ~~ Ts') U
else Type (c, Ts'))
- | Some (Nonterminal, _) => (nargs 0; check_syntax c; T)
- | None => err (undecl_type c))
+ | SOME (Nonterminal, _) => (nargs 0; check_syntax c; T)
+ | NONE => err (undecl_type c))
end
| cert (TFree (x, S)) = TFree (x, Sorts.certify_sort classes S)
| cert (TVar (xi as (_, i), S)) =
@@ -224,8 +224,8 @@
val fmap = fs ~~ map (rpair 0) (variantlist (fs, map #1 ixns))
fun thaw (f as (a, S)) =
(case assoc (fmap, a) of
- None => TFree f
- | Some b => TVar (b, S));
+ NONE => TFree f
+ | SOME b => TVar (b, S));
in (map_term_types (map_type_tfree thaw) t, fmap) end;
@@ -239,11 +239,11 @@
fun freeze_one alist (ix,sort) =
TFree (the (assoc (alist, ix)), sort)
- handle OPTION =>
+ handle Option =>
raise TYPE ("Failure during freezing of ?" ^ string_of_indexname ix, [], []);
fun thaw_one alist (a,sort) = TVar (the (assoc (alist,a)), sort)
- handle OPTION => TFree(a, sort);
+ handle Option => TFree(a, sort);
in
@@ -279,10 +279,10 @@
let
fun inst (var as (v, S)) =
(case assoc_string_int (tye, v) of
- Some U =>
+ SOME U =>
if of_sort tsig (U, S) then U
else raise TYPE ("Type not of sort " ^ Pretty.string_of_sort pp S, [U], [])
- | None => TVar var);
+ | NONE => TVar var);
in map_type_tvar inst end;
fun inst_term_tvars _ _ [] t = t
@@ -297,10 +297,10 @@
let
fun match (subs, (TVar (v, S), T)) =
(case Vartab.lookup (subs, v) of
- None =>
+ NONE =>
if of_sort tsig (T, S) then Vartab.update ((v, T), subs)
else raise TYPE_MATCH
- | Some U => if U = T then subs else raise TYPE_MATCH)
+ | SOME U => if U = T then subs else raise TYPE_MATCH)
| match (subs, (Type (a, Ts), Type (b, Us))) =
if a <> b then raise TYPE_MATCH
else foldl match (subs, Ts ~~ Us)
@@ -325,15 +325,15 @@
| occ (TVar (w, _)) =
eq_ix (v, w) orelse
(case Vartab.lookup (tye, w) of
- None => false
- | Some U => occ U);
+ NONE => false
+ | SOME U => occ U);
in occ end;
(*chase variable assignments; if devar returns a type var then it must be unassigned*)
fun devar (T as TVar (v, _), tye) =
(case Vartab.lookup (tye, v) of
- Some U => devar (U, tye)
- | None => T)
+ SOME U => devar (U, tye)
+ | NONE => T)
| devar (T, tye) = T;
fun unify (tsig as TSig {classes, arities, ...}) (tyenv, maxidx) TU =
@@ -397,8 +397,8 @@
fun err_decl t decl = error ("Illegal " ^ str_of_decl decl ^ ": " ^ quote t);
-fun for_classes _ None = ""
- | for_classes pp (Some (c1, c2)) =
+fun for_classes _ NONE = ""
+ | for_classes pp (SOME (c1, c2)) =
" for classes " ^ Pretty.string_of_classrel pp [c1, c2];
fun err_conflict pp t cc (c, Ss) (c', Ss') =
@@ -410,24 +410,24 @@
let
fun conflict (c', Ss') =
if Sorts.class_le C (c, c') andalso not (Sorts.sorts_le C (Ss, Ss')) then
- Some ((c, c'), (c', Ss'))
+ SOME ((c, c'), (c', Ss'))
else if Sorts.class_le C (c', c) andalso not (Sorts.sorts_le C (Ss', Ss)) then
- Some ((c', c), (c', Ss'))
- else None;
+ SOME ((c', c), (c', Ss'))
+ else NONE;
in
(case Library.get_first conflict ars of
- Some ((c1, c2), (c', Ss')) => err_conflict pp t (Some (c1, c2)) (c, Ss) (c', Ss')
- | None => (c, Ss) :: ars)
+ SOME ((c1, c2), (c', Ss')) => err_conflict pp t (SOME (c1, c2)) (c, Ss) (c', Ss')
+ | NONE => (c, Ss) :: ars)
end;
fun insert pp C t ((c, Ss), ars) =
(case assoc_string (ars, c) of
- None => coregular pp C t (c, Ss) ars
- | Some Ss' =>
+ NONE => coregular pp C t (c, Ss) ars
+ | SOME Ss' =>
if Sorts.sorts_le C (Ss, Ss') then ars
else if Sorts.sorts_le C (Ss', Ss)
then coregular pp C t (c, Ss) (ars \ (c, Ss'))
- else err_conflict pp t None (c, Ss) (c, Ss'));
+ else err_conflict pp t NONE (c, Ss) (c, Ss'));
fun complete C (c, Ss) = map (rpair Ss) (Graph.all_succs C [c]);
@@ -446,13 +446,13 @@
let
fun prep (t, Ss, S) =
(case Symtab.lookup (types, t) of
- Some (LogicalType n, _) =>
+ SOME (LogicalType n, _) =>
if length Ss = n then
(t, map (cert_sort tsig) Ss, cert_sort tsig S)
handle TYPE (msg, _, _) => error msg
else error (bad_nargs t)
- | Some (decl, _) => err_decl t decl
- | None => error (undecl_type t));
+ | SOME (decl, _) => err_decl t decl
+ | NONE => error (undecl_type t));
val ars = decls |> map ((fn (t, Ss, S) => (t, map (fn c => (c, Ss)) S)) o prep);
val arities' = foldl (insert_arities pp classes) (arities, ars);
@@ -530,8 +530,8 @@
fun new_decl (c, decl) types =
(case Symtab.lookup (types, c) of
- Some (decl', _) => err_in_decls c decl decl'
- | None => Symtab.update ((c, (decl, stamp ())), types));
+ SOME (decl', _) => err_in_decls c decl decl'
+ | NONE => Symtab.update ((c, (decl, stamp ())), types));
fun the_decl types c = fst (the (Symtab.lookup (types, c)));
@@ -539,7 +539,7 @@
(classes, default, f types, arities));
fun syntactic types (Type (c, Ts)) =
- (case Symtab.lookup (types, c) of Some (Nonterminal, _) => true | _ => false)
+ (case Symtab.lookup (types, c) of SOME (Nonterminal, _) => true | _ => false)
orelse exists (syntactic types) Ts
| syntactic _ _ = false;
--- a/src/Pure/type_infer.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/type_infer.ML Sun Feb 13 17:15:14 2005 +0100
@@ -120,8 +120,8 @@
fun pre_of (TVar (v as (xi, _))) =
(case assoc (params', xi) of
- None => PTVar v
- | Some p => p)
+ NONE => PTVar v
+ | SOME p => p)
| pre_of (TFree ("'_dummy_", S)) = mk_param S
| pre_of (TFree v) = PTFree v
| pre_of (T as Type (a, Ts)) =
@@ -162,8 +162,8 @@
fun pre_of (ps, Const (c, T)) =
(case const_type c of
- Some U => constrain (ps, PConst (c, snd (pretyp_of (K true) ([], U)))) T
- | None => raise TYPE ("No such constant: " ^ quote c, [], []))
+ SOME U => constrain (ps, PConst (c, snd (pretyp_of (K true) ([], U)))) T
+ | NONE => raise TYPE ("No such constant: " ^ quote c, [], []))
| pre_of (ps, Var (xi, Type ("_polymorphic_", [T]))) =
(ps, PVar (xi, snd (pretyp_of (K true) ([], T))))
| pre_of (ps, Var (xi, T)) = constrain (ps, PVar (xi, var_param xi)) T
@@ -408,8 +408,8 @@
val _ = seq (fn t => (infer pp classes arities t; ())) tTs';
(*collect result unifier*)
- fun ch_var (xi, Link (r as ref (Param S))) = (r := PTVar (xi, S); None)
- | ch_var xi_T = Some xi_T;
+ fun ch_var (xi, Link (r as ref (Param S))) = (r := PTVar (xi, S); NONE)
+ | ch_var xi_T = SOME xi_T;
val env = mapfilter ch_var Tps;
(*convert back to terms/typs*)
@@ -458,10 +458,10 @@
fun get xi =
(case (assoc (env, xi), def_sort xi) of
- (None, None) => Type.defaultS tsig
- | (None, Some S) => S
- | (Some S, None) => S
- | (Some S, Some S') =>
+ (NONE, NONE) => Type.defaultS tsig
+ | (NONE, SOME S) => S
+ | (SOME S, NONE) => S
+ | (SOME S, SOME S') =>
if Type.eq_sort tsig (S, S') then S'
else error ("Sort constraint inconsistent with default for type variable " ^
quote (Syntax.string_of_vname' xi)));
--- a/src/Pure/unify.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/Pure/unify.ML Sun Feb 13 17:15:14 2005 +0100
@@ -50,14 +50,14 @@
fun body_type(Envir.Envir{iTs,...}) =
let fun bT(Type("fun",[_,T])) = bT T
| bT(T as TVar(ixn,_)) = (case Vartab.lookup (iTs,ixn) of
- None => T | Some(T') => bT T')
+ NONE => T | SOME(T') => bT T')
| bT T = T
in bT end;
fun binder_types(Envir.Envir{iTs,...}) =
let fun bTs(Type("fun",[T,U])) = T :: bTs U
| bTs(T as TVar(ixn,_)) = (case Vartab.lookup (iTs,ixn) of
- None => [] | Some(T') => bTs T')
+ NONE => [] | SOME(T') => bTs T')
| bTs _ = []
in bTs end;
@@ -72,7 +72,7 @@
Abs("", T, etif(U, incr_boundvars 1 t $ Bound 0))
| etif (TVar(ixn,_),t) =
(case Vartab.lookup (iTs,ixn) of
- None => t | Some(T) => etif(T,t))
+ NONE => t | SOME(T) => etif(T,t))
| etif (_,t) = t;
fun eta_nm (rbinder, Abs(a,T,body)) =
Abs(a, T, eta_nm ((a,T)::rbinder, body))
@@ -98,8 +98,8 @@
(*no need to lookup: v has no assignment*)
else (seen := w:: !seen;
case Envir.lookup(env,w) of
- None => false
- | Some t => occur t)
+ NONE => false
+ | SOME t => occur t)
| occur (Abs(_,_,body)) = occur body
| occur (f$t) = occur t orelse occur f
in occurs ts end;
@@ -110,7 +110,7 @@
fun head_of_in (env,t) : term = case t of
f$_ => head_of_in(env,f)
| Var (v,_) => (case Envir.lookup(env,v) of
- Some u => head_of_in(env,u) | None => t)
+ SOME u => head_of_in(env,u) | NONE => t)
| _ => t;
@@ -160,8 +160,8 @@
else if eq_ix(v,w) then Rigid
else (seen := w:: !seen;
case Envir.lookup(env,w) of
- None => NoOcc
- | Some t => occur t)
+ NONE => NoOcc
+ | SOME t => occur t)
| occur (Abs(_,_,body)) = occur body
| occur (t as f$_) = (*switch to nonrigid search?*)
(case head_of_in (env,f) of
@@ -283,7 +283,7 @@
(* changed(env,t) checks whether the head of t is a variable assigned in env*)
fun changed (env, f$_) = changed (env,f)
| changed (env, Var (v,_)) =
- (case Envir.lookup(env,v) of None=>false | _ => true)
+ (case Envir.lookup(env,v) of NONE=>false | _ => true)
| changed _ = false;
@@ -358,7 +358,7 @@
val dp = (rbinder, list_comb(targ, map plugin args), u);
val (env2,frigid,fflex) = SIMPL (env', dp::dpairs)
(*may raise exception CANTUNIFY*)
- in Some ((list_comb(head,args), (env2, frigid@fflex)),
+ in SOME ((list_comb(head,args), (env2, frigid@fflex)),
tail)
end handle CANTUNIFY => Seq.pull tail)
end handle CANTUNIFY => tail;
@@ -400,8 +400,8 @@
fun new_dset (u', (env',dpairs')) =
(*if v was updated to s, must unify s with u' *)
case Envir.lookup(env',v) of
- None => (Envir.update ((v, types_abs(Ts, u')), env'), dpairs')
- | Some s => (env', ([], s, types_abs(Ts, u'))::dpairs')
+ NONE => (Envir.update ((v, types_abs(Ts, u')), env'), dpairs')
+ | SOME s => (env', ([], s, types_abs(Ts, u'))::dpairs')
in Seq.map new_dset
(matchcopy (#1 v) (rbinder, targs, u, (env,dpairs)))
end;
@@ -575,7 +575,7 @@
then print_dpairs "Enter SIMPL" (env,dpairs) else ();
SIMPL (env,dpairs))
in case flexrigid of
- [] => Some (foldr add_ffpair (flexflex, (env',[])), reseq)
+ [] => SOME (foldr add_ffpair (flexflex, (env',[])), reseq)
| dp::frigid' =>
if tdepth > !search_bound then
(warning "Unification bound exceeded"; Seq.pull reseq)
--- a/src/ZF/Datatype.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/ZF/Datatype.ML Sun Feb 13 17:15:14 2005 +0100
@@ -74,9 +74,9 @@
val lname = #1 (dest_Const lhead) handle TERM _ => raise Match;
val rname = #1 (dest_Const rhead) handle TERM _ => raise Match;
val lcon_info = the (Symtab.lookup (ConstructorsData.get_sg sg, lname))
- handle Library.OPTION => raise Match;
+ handle Option => raise Match;
val rcon_info = the (Symtab.lookup (ConstructorsData.get_sg sg, rname))
- handle Library.OPTION => raise Match;
+ handle Option => raise Match;
val new =
if #big_rec_name lcon_info = #big_rec_name rcon_info
andalso not (null (#free_iffs lcon_info)) then
@@ -92,8 +92,8 @@
handle ERROR_MESSAGE msg =>
(warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Sign.string_of_term sg goal);
raise Match)
- in Some thm end
- handle Match => None;
+ in SOME thm end
+ handle Match => NONE;
val conv =
--- a/src/ZF/Tools/cartprod.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/ZF/Tools/cartprod.ML Sun Feb 13 17:15:14 2005 +0100
@@ -61,7 +61,7 @@
functor CartProd_Fun (Pr: PR) : CARTPROD =
struct
-(* Some of these functions expect "pseudo-types" containing products,
+(* SOME of these functions expect "pseudo-types" containing products,
as in HOL; the true ZF types would just be "i" *)
fun mk_prod (T1,T2) = Type("*", [T1,T2]);
--- a/src/ZF/Tools/datatype_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/ZF/Tools/datatype_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -296,9 +296,9 @@
(*** Prove the recursor theorems ***)
val recursor_eqns = case try (get_def thy1) recursor_base_name of
- None => (writeln " [ No recursion operator ]";
+ NONE => (writeln " [ No recursion operator ]";
[])
- | Some recursor_def =>
+ | SOME recursor_def =>
let
(*Replace subterms rec`x (where rec is a Free var) by recursor_tm(x) *)
fun subst_rec (Const("op `",_) $ Free _ $ arg) = recursor_tm $ arg
--- a/src/ZF/Tools/ind_cases.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/ZF/Tools/ind_cases.ML Sun Feb 13 17:15:14 2005 +0100
@@ -42,8 +42,8 @@
(Logic.strip_imp_concl A)))))) handle TERM _ => err ();
in
(case Symtab.lookup (IndCasesData.get thy, c) of
- None => error ("Unknown inductive cases rule for set " ^ quote c)
- | Some f => f ss (Thm.cterm_of (Theory.sign_of thy) A))
+ NONE => error ("Unknown inductive cases rule for set " ^ quote c)
+ | SOME f => f ss (Thm.cterm_of (Theory.sign_of thy) A))
end;
--- a/src/ZF/Tools/induct_tacs.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/ZF/Tools/induct_tacs.ML Sun Feb 13 17:15:14 2005 +0100
@@ -78,8 +78,8 @@
fun datatype_info_sg sign name =
(case Symtab.lookup (DatatypesData.get_sg sign, name) of
- Some info => info
- | None => error ("Unknown datatype " ^ quote name));
+ SOME info => info
+ | NONE => error ("Unknown datatype " ^ quote name));
(*Given a variable, find the inductive set associated it in the assumptions*)
@@ -92,8 +92,8 @@
val pairs = mapfilter (try (mk_pair o FOLogic.dest_Trueprop))
(#2 (strip_context Bi))
in case assoc (pairs, var) of
- None => raise Find_tname ("Cannot determine datatype of " ^ quote var)
- | Some t => t
+ NONE => raise Find_tname ("Cannot determine datatype of " ^ quote var)
+ | SOME t => t
end;
(** generic exhaustion and induction tactic for datatypes
--- a/src/ZF/Tools/inductive_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/ZF/Tools/inductive_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -301,8 +301,8 @@
fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $
(Const("op :",_)$t$X), iprems) =
(case gen_assoc (op aconv) (ind_alist, X) of
- Some pred => prem :: FOLogic.mk_Trueprop (pred $ t) :: iprems
- | None => (*possibly membership in M(rec_tm), for M monotone*)
+ SOME pred => prem :: FOLogic.mk_Trueprop (pred $ t) :: iprems
+ | NONE => (*possibly membership in M(rec_tm), for M monotone*)
let fun mk_sb (rec_tm,pred) =
(rec_tm, Ind_Syntax.Collect_const$rec_tm$pred)
in subst_free (map mk_sb ind_alist) prem :: iprems end)
@@ -314,7 +314,7 @@
val iprems = foldr (add_induct_prem ind_alist)
(Logic.strip_imp_prems intr,[])
val (t,X) = Ind_Syntax.rule_concl intr
- val (Some pred) = gen_assoc (op aconv) (ind_alist, X)
+ val (SOME pred) = gen_assoc (op aconv) (ind_alist, X)
val concl = FOLogic.mk_Trueprop (pred $ t)
in list_all_free (quantfrees, Logic.list_implies (iprems,concl)) end
handle Bind => error"Recursion term not found in conclusion";
--- a/src/ZF/Tools/primrec_package.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/ZF/Tools/primrec_package.ML Sun Feb 13 17:15:14 2005 +0100
@@ -57,7 +57,7 @@
val (cname, _) = dest_Const constr
handle TERM _ => raise RecError "ill-formed constructor";
val con_info = the (Symtab.lookup (ConstructorsData.get thy, cname))
- handle Library.OPTION =>
+ handle Option =>
raise RecError "cannot determine datatype associated with function"
val (ls, cargs, rs) = (map dest_Free ls_frees,
@@ -78,8 +78,8 @@
else if length middle > 1 then
raise RecError "more than one non-variable in pattern"
else case rec_fn_opt of
- None => Some (fname, ftype, ls, rs, con_info, [new_eqn])
- | Some (fname', _, ls', rs', con_info': constructor_info, eqns) =>
+ NONE => SOME (fname, ftype, ls, rs, con_info, [new_eqn])
+ | SOME (fname', _, ls', rs', con_info': constructor_info, eqns) =>
if is_some (assoc (eqns, cname)) then
raise RecError "constructor already occurred as pattern"
else if (ls <> ls') orelse (rs <> rs') then
@@ -89,7 +89,7 @@
else if fname <> fname' then
raise RecError ("inconsistent functions for datatype " ^
#big_rec_name con_info)
- else Some (fname, ftype, ls, rs, con_info, new_eqn::eqns)
+ else SOME (fname, ftype, ls, rs, con_info, new_eqn::eqns)
end
handle RecError s => primrec_eq_err (sign_of thy) s eq;
@@ -121,11 +121,11 @@
fun add_case ((cname, recursor_pair), cases) =
let val (rhs, recursor_rhs, eq) =
case assoc (eqns, cname) of
- None => (warning ("no equation for constructor " ^ cname ^
+ NONE => (warning ("no equation for constructor " ^ cname ^
"\nin definition of function " ^ fname);
(Const ("0", Ind_Syntax.iT),
#2 recursor_pair, Const ("0", Ind_Syntax.iT)))
- | Some (rhs, cargs', eq) =>
+ | SOME (rhs, cargs', eq) =>
(rhs, inst_recursor (recursor_pair, cargs'), eq)
val allowed_terms = map use_fabs (#2 (strip_comb recursor_rhs))
val abs = foldr absterm (allowed_terms, rhs)
@@ -171,8 +171,8 @@
fun add_primrec_i args thy =
let
val ((eqn_names, eqn_terms), eqn_atts) = apfst split_list (split_list args);
- val Some (fname, ftype, ls, rs, con_info, eqns) =
- foldr (process_eqn thy) (eqn_terms, None);
+ val SOME (fname, ftype, ls, rs, con_info, eqns) =
+ foldr (process_eqn thy) (eqn_terms, NONE);
val def = process_fun thy (fname, ftype, ls, rs, con_info, eqns);
val (thy1, [def_thm]) = thy
--- a/src/ZF/arith_data.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/ZF/arith_data.ML Sun Feb 13 17:15:14 2005 +0100
@@ -53,8 +53,8 @@
| dest_sum tm = [tm];
(*Apply the given rewrite (if present) just once*)
-fun gen_trans_tac th2 None = all_tac
- | gen_trans_tac th2 (Some th) = ALLGOALS (rtac (th RS th2));
+fun gen_trans_tac th2 NONE = all_tac
+ | gen_trans_tac th2 (SOME th) = ALLGOALS (rtac (th RS th2));
(*Use <-> or = depending on the type of t*)
fun mk_eq_iff(t,u) =
@@ -69,14 +69,14 @@
fun add_chyps chyps ct = Drule.list_implies (map cprop_of chyps, ct);
fun prove_conv name tacs sg hyps xs (t,u) =
- if t aconv u then None
+ if t aconv u then NONE
else
let val hyps' = filter (not o is_eq_thm) hyps
val goal = Logic.list_implies (map (#prop o Thm.rep_thm) hyps',
FOLogic.mk_Trueprop (mk_eq_iff (t, u)));
- in Some (hyps' MRS Tactic.prove sg xs [] goal (K (EVERY tacs)))
+ in SOME (hyps' MRS Tactic.prove sg xs [] goal (K (EVERY tacs)))
handle ERROR_MESSAGE msg =>
- (warning (msg ^ "\nCancellation failed: no typing information? (" ^ name ^ ")"); None)
+ (warning (msg ^ "\nCancellation failed: no typing information? (" ^ name ^ ")"); NONE)
end;
fun prep_simproc (name, pats, proc) =
--- a/src/ZF/simpdata.ML Fri Feb 11 18:51:00 2005 +0100
+++ b/src/ZF/simpdata.ML Sun Feb 13 17:15:14 2005 +0100
@@ -16,9 +16,9 @@
case head_of t of
Const(a,_) =>
(case assoc(pairs,a) of
- Some rls => flat (map (atomize (conn_pairs, mem_pairs))
+ SOME rls => flat (map (atomize (conn_pairs, mem_pairs))
([th] RL rls))
- | None => [th])
+ | NONE => [th])
| _ => [th]
in case concl_of th of
Const("Trueprop",_) $ P =>