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;