--- a/src/HOL/Import/replay.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/HOL/Import/replay.ML Thu May 31 01:25:24 2007 +0200
@@ -291,10 +291,10 @@
let
fun get_facts facts =
case TextIO.inputLine is of
- "" => (case facts of
+ NONE => (case facts of
i::facts => (valOf (Int.fromString i),map P.protect_factname (rev facts))
| _ => raise ERR "replay_thm" "Bad facts.lst file")
- | fact => get_facts ((String.substring(fact,0,String.size fact -1 ))::facts)
+ | SOME fact => get_facts ((String.substring(fact,0,String.size fact -1 ))::facts)
in
get_facts []
end
--- a/src/HOL/Matrix/cplex/Cplex_tools.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/HOL/Matrix/cplex/Cplex_tools.ML Thu May 31 01:25:24 2007 +0200
@@ -320,8 +320,9 @@
)
end
else
- let val s = TextIO.inputLine f in
- if s = "" then NONE else
+ (case TextIO.inputLine f of
+ NONE => NONE
+ | SOME s =>
let val t = tokenize s in
if (length t >= 2 andalso
snd(hd (tl t)) = ":")
@@ -334,8 +335,7 @@
else
rest := t;
readToken_helper ()
- end
- end
+ end)
fun readToken_helper2 () =
let val c = readToken_helper () in
@@ -872,9 +872,9 @@
)
end
else
- let val s = TextIO.inputLine f in
- if s = "" then NONE else (rest := tokenize s; readToken_helper())
- end
+ (case TextIO.inputLine f of
+ NONE => NONE
+ | SOME s => (rest := tokenize s; readToken_helper()))
fun is_tt tok ty = (tok <> NONE andalso (fst (the tok)) = ty)
@@ -1007,11 +1007,9 @@
)
end
else
- let
- val s = TextIO.inputLine f
- in
- if s = "" then NONE else (rest := tokenize s; readToken_helper())
- end
+ (case TextIO.inputLine f of
+ NONE => NONE
+ | SOME s => (rest := tokenize s; readToken_helper()))
fun is_tt tok ty = (tok <> NONE andalso (fst (the tok)) = ty)
--- a/src/HOL/Tools/ATP/watcher.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/HOL/Tools/ATP/watcher.ML Thu May 31 01:25:24 2007 +0200
@@ -23,7 +23,7 @@
(* Start a watcher and set up signal handlers*)
val createWatcher :
- Proof.context * thm * string Vector.vector list ->
+ Proof.context * thm * string Vector.vector list ->
TextIO.instream * TextIO.outstream * Posix.Process.pid
val killWatcher : Posix.Process.pid -> unit
val killChild : ('a, 'b) Unix.proc -> OS.Process.status
@@ -45,7 +45,7 @@
val trace_path = Path.basic "watcher_trace";
-fun trace s = if !Output.debugging then File.append (File.tmp_path trace_path) s
+fun trace s = if !Output.debugging then File.append (File.tmp_path trace_path) s
else ();
(*Representation of a watcher process*)
@@ -56,7 +56,7 @@
(*Representation of a child (ATP) process*)
type cmdproc = {
prover: string, (* Name of the resolution prover used, e.g. "spass"*)
- file: string, (* The file containing the goal for the ATP to prove *)
+ file: string, (* The file containing the goal for the ATP to prove *)
proc_handle : (TextIO.instream,TextIO.outstream) Unix.proc,
instr : TextIO.instream, (*Output of the child process *)
outstr : TextIO.outstream}; (*Input to the child process *)
@@ -69,7 +69,7 @@
Posix.IO.mkTextWriter {
appendMode = false,
initBlkMode = true,
- name = name,
+ name = name,
chunkSize=4096,
fd = fd};
@@ -83,14 +83,14 @@
TextIO.StreamIO.mkInstream (
fdReader (name, fd), ""));
-
+
(* Send request to Watcher for a vampire to be called for filename in arg*)
-fun callResProver (toWatcherStr, arg) =
- (TextIO.output (toWatcherStr, arg^"\n");
+fun callResProver (toWatcherStr, arg) =
+ (TextIO.output (toWatcherStr, arg^"\n");
TextIO.flushOut toWatcherStr)
(* Send request to Watcher for multiple provers to be called*)
-fun callResProvers (toWatcherStr, []) =
+fun callResProvers (toWatcherStr, []) =
(TextIO.output (toWatcherStr, "End of calls\n"); TextIO.flushOut toWatcherStr)
| callResProvers (toWatcherStr,
(prover,proverCmd,settings,probfile) :: args) =
@@ -101,45 +101,45 @@
inc goals_being_watched;
TextIO.flushOut toWatcherStr;
callResProvers (toWatcherStr,args))
-
+
(*Send message to watcher to kill currently running vampires. NOT USED and possibly
buggy. Note that killWatcher kills the entire process group anyway.*)
-fun callSlayer toWatcherStr = (TextIO.output (toWatcherStr, "Kill children\n");
+fun callSlayer toWatcherStr = (TextIO.output (toWatcherStr, "Kill children\n");
TextIO.flushOut toWatcherStr)
-
+
(* Get commands from Isabelle*)
-fun getCmds (toParentStr, fromParentStr, cmdList) =
- let val thisLine = TextIO.inputLine fromParentStr
+fun getCmds (toParentStr, fromParentStr, cmdList) =
+ let val thisLine = the_default "" (TextIO.inputLine fromParentStr)
in
trace("\nGot command from parent: " ^ thisLine);
if thisLine = "End of calls\n" orelse thisLine = "" then cmdList
else if thisLine = "Kill children\n"
- then (TextIO.output (toParentStr,thisLine);
+ then (TextIO.output (toParentStr,thisLine);
TextIO.flushOut toParentStr;
[("","Kill children",[],"")])
else
- let val [prover,proverCmd,settingstr,probfile,_] =
+ let val [prover,proverCmd,settingstr,probfile,_] =
String.tokens (fn c => c = command_sep) thisLine
val settings = String.tokens (fn c => c = setting_sep) settingstr
in
trace ("\nprover: " ^ prover ^ " prover path: " ^ proverCmd ^
"\n problem file: " ^ probfile);
- getCmds (toParentStr, fromParentStr,
- (prover, proverCmd, settings, probfile)::cmdList)
+ getCmds (toParentStr, fromParentStr,
+ (prover, proverCmd, settings, probfile)::cmdList)
end
- handle Bind =>
+ handle Bind =>
(trace "\ngetCmds: command parsing failed!";
getCmds (toParentStr, fromParentStr, cmdList))
end
-
-
+
+
(*Get Io-descriptor for polling of an input stream*)
-fun getInIoDesc someInstr =
+fun getInIoDesc someInstr =
let val (rd, buf) = TextIO.StreamIO.getReader(TextIO.getInstream someInstr)
val _ = TextIO.output (TextIO.stdOut, buf)
- val ioDesc =
+ val ioDesc =
case rd
of TextPrimIO.RD{ioDesc = SOME iod, ...} => SOME iod
| _ => NONE
@@ -149,9 +149,9 @@
ioDesc
end
-fun pollChild fromStr =
+fun pollChild fromStr =
case getInIoDesc fromStr of
- SOME iod =>
+ SOME iod =>
(case OS.IO.pollDesc iod of
SOME pd =>
let val pd' = OS.IO.pollIn pd in
@@ -179,7 +179,7 @@
Posix.Process.exit 0w0);
(* take an instream and poll its underlying reader for input *)
-fun pollParentInput (fromParentIOD, fromParentStr, toParentStr) =
+fun pollParentInput (fromParentIOD, fromParentStr, toParentStr) =
case OS.IO.pollDesc fromParentIOD of
SOME pd =>
(case OS.IO.poll ([OS.IO.pollIn pd], SOME (Time.fromSeconds 2)) of
@@ -192,7 +192,7 @@
(*get the number of the subgoal from the filename: the last digit string*)
fun number_from_filename s =
case String.tokens (not o Char.isDigit) s of
- [] => (trace ("\nWatcher could not read subgoal nunber! " ^ s);
+ [] => (trace ("\nWatcher could not read subgoal nunber! " ^ s);
error "")
| numbers => valOf (Int.fromString (List.last numbers));
@@ -204,7 +204,7 @@
| execCmds ((prover,proverCmd,settings,file)::cmds) procList =
let val _ = trace ("\nAbout to execute command: " ^ proverCmd ^ " " ^ file)
val settings' = List.concat (map (String.tokens Char.isSpace) settings)
- val childhandle:(TextIO.instream,TextIO.outstream) Unix.proc =
+ val childhandle:(TextIO.instream,TextIO.outstream) Unix.proc =
Unix.execute(proverCmd, settings' @ [file])
val (instr, outstr) = Unix.streamsOf childhandle
val newProcList = {prover=prover, file=file, proc_handle=childhandle,
@@ -213,42 +213,42 @@
Date.toString(Date.fromTimeLocal(Time.now())))
in execCmds cmds newProcList end
-fun checkChildren (ctxt, th, thm_names_list) toParentStr children =
+fun checkChildren (ctxt, th, thm_names_list) toParentStr children =
let fun check [] = [] (* no children to check *)
- | check (child::children) =
+ | check (child::children) =
let val {prover, file, proc_handle, instr=childIn, ...} : cmdproc = child
val _ = trace ("\nprobfile = " ^ file)
val sgno = number_from_filename file (*subgoal number*)
val thm_names = List.nth(thm_names_list, sgno-1);
val ppid = Posix.ProcEnv.getppid()
- in
+ in
if pollChild childIn
then (* check here for prover label on child*)
let val _ = trace ("\nInput available from child: " ^ file)
val childDone = (case prover of
"vampire" => ResReconstruct.checkVampProofFound
- (childIn, toParentStr, ppid, file, ctxt, th, sgno, thm_names)
+ (childIn, toParentStr, ppid, file, ctxt, th, sgno, thm_names)
| "E" => ResReconstruct.checkEProofFound
- (childIn, toParentStr, ppid, file, ctxt, th, sgno, thm_names)
+ (childIn, toParentStr, ppid, file, ctxt, th, sgno, thm_names)
| "spass" => ResReconstruct.checkSpassProofFound
- (childIn, toParentStr, ppid, file, ctxt, th, sgno, thm_names)
+ (childIn, toParentStr, ppid, file, ctxt, th, sgno, thm_names)
| _ => (trace ("\nBad prover! " ^ prover); true) )
in
if childDone (*ATP has reported back (with success OR failure)*)
- then (Unix.reap proc_handle;
+ then (Unix.reap proc_handle;
if !Output.debugging then () else OS.FileSys.remove file;
check children)
else child :: check children
end
else (trace "\nNo child output"; child :: check children)
end
- in
+ in
trace ("\nIn checkChildren, length of queue: " ^ Int.toString(length children));
- check children
+ check children
end;
-fun setupWatcher (ctxt, th, thm_names_list) =
+fun setupWatcher (ctxt, th, thm_names_list) =
let
val checkc = checkChildren (ctxt, th, thm_names_list)
val p1 = Posix.IO.pipe() (*pipes for communication between parent and watcher*)
@@ -257,9 +257,9 @@
fun startWatcher procList =
case Posix.Process.fork() of
SOME pid => pid (* parent - i.e. main Isabelle process *)
- | NONE =>
+ | NONE =>
let (* child - i.e. watcher *)
- val oldchildin = #infd p1
+ val oldchildin = #infd p1
val fromParent = Posix.FileSys.wordToFD 0w0
val oldchildout = #outfd p2
val toParent = Posix.FileSys.wordToFD 0w1
@@ -272,27 +272,27 @@
val () = trace "\nsubgoals forked to startWatcher"
val limit = ref 200; (*don't let watcher run forever*)
(*Watcher Loop : Check running ATP processes for output*)
- fun keepWatching procList =
- (trace ("\npollParentInput. Limit = " ^ Int.toString (!limit) ^
+ fun keepWatching procList =
+ (trace ("\npollParentInput. Limit = " ^ Int.toString (!limit) ^
" length(procList) = " ^ Int.toString(length procList));
OS.Process.sleep (Time.fromMilliseconds 100); limit := !limit - 1;
- if !limit < 0 then killWatcher (toParentStr, procList)
- else
+ if !limit < 0 then killWatcher (toParentStr, procList)
+ else
case pollParentInput(fromParentIOD, fromParentStr, toParentStr) of
- SOME [(_,"Kill children",_,_)] =>
- (trace "\nReceived Kill command";
+ SOME [(_,"Kill children",_,_)] =>
+ (trace "\nReceived Kill command";
killChildren procList; keepWatching [])
| SOME cmds => (* deal with commands from Isabelle process *)
- if length procList < 40 then (* Execute locally *)
- let val _ = trace("\nCommands from parent: " ^
+ if length procList < 40 then (* Execute locally *)
+ let val _ = trace("\nCommands from parent: " ^
Int.toString(length cmds))
- val newProcList' = checkc toParentStr (execCmds cmds procList)
+ val newProcList' = checkc toParentStr (execCmds cmds procList)
in trace "\nCommands executed"; keepWatching newProcList' end
else (* Execute remotely [FIXME: NOT REALLY] *)
- let val newProcList' = checkc toParentStr (execCmds cmds procList)
+ let val newProcList' = checkc toParentStr (execCmds cmds procList)
in keepWatching newProcList' end
| NONE => (* No new input from Isabelle *)
- (trace "\nNothing from parent...";
+ (trace "\nNothing from parent...";
keepWatching(checkc toParentStr procList)))
handle exn => (*FIXME: exn handler is too general!*)
(trace ("\nkeepWatching exception handler: " ^ Toplevel.exn_message exn);
@@ -305,7 +305,7 @@
Posix.IO.dup2{old = oldchildout, new = toParent};
Posix.IO.close oldchildout;
keepWatching (procList)
- end;
+ end;
val _ = TextIO.flushOut TextIO.stdOut
val pid = startWatcher []
@@ -328,15 +328,15 @@
(**********************************************************)
(*Signal handler to tidy away zombie processes*)
-fun reapAll() =
+fun reapAll() =
(case Posix.Process.waitpid_nh(Posix.Process.W_ANY_CHILD, []) of
- SOME _ => reapAll() | NONE => ())
+ SOME _ => reapAll() | NONE => ())
handle OS.SysErr _ => ()
(*FIXME: does the main process need something like this?
IsaSignal.signal (IsaSignal.chld, IsaSignal.SIG_HANDLE reap)??*)
-fun killWatcher pid =
+fun killWatcher pid =
(goals_being_watched := 0;
Posix.Process.kill(Posix.Process.K_GROUP pid, Posix.Signal.kill);
reapAll());
@@ -357,19 +357,19 @@
fun decr_watched() =
(goals_being_watched := !goals_being_watched - 1;
if !goals_being_watched = 0
- then
+ then
(Output.debug (fn () => ("\nReaping a watcher, childpid = " ^ string_of_pid childpid));
killWatcher childpid (*???; reapWatcher (childpid, childin, childout)*) )
else ())
- fun proofHandler _ =
+ fun proofHandler _ =
let val _ = trace("\nIn signal handler for pid " ^ string_of_pid childpid)
- val outcome = ResReconstruct.restore_newlines (TextIO.inputLine childin)
- val probfile = TextIO.inputLine childin
+ val outcome = ResReconstruct.restore_newlines (the_default "" (TextIO.inputLine childin))
+ val probfile = the_default "" (TextIO.inputLine childin)
val sgno = number_from_filename probfile
val text = string_of_subgoal th sgno
fun report s = priority ("Subgoal " ^ Int.toString sgno ^ ": " ^ s);
- in
- Output.debug (fn () => ("In signal handler. outcome = \"" ^ outcome ^
+ in
+ Output.debug (fn () => ("In signal handler. outcome = \"" ^ outcome ^
"\"\nprobfile = " ^ probfile ^
"\nGoals being watched: "^ Int.toString (!goals_being_watched)));
if String.isPrefix "Invalid" outcome
@@ -377,7 +377,7 @@
decr_watched())
else if String.isPrefix "Failure" outcome
then (report ("Proof attempt failed:\n" ^ text);
- decr_watched())
+ decr_watched())
(* print out a list of rules used from clasimpset*)
else if String.isPrefix "Success" outcome
then (report (outcome ^ text);
--- a/src/HOL/Tools/res_atp_provers.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/HOL/Tools/res_atp_provers.ML Thu May 31 01:25:24 2007 +0200
@@ -4,20 +4,17 @@
Functions used for ATP Oracle.
*)
-
structure ResAtpProvers =
-
struct
fun seek_line s instr =
- let val thisLine = TextIO.inputLine instr
- in thisLine <> "" andalso
- (thisLine = s orelse seek_line s instr)
- end;
+ (case TextIO.inputLine instr of
+ NONE => false
+ | SOME thisLine => thisLine = s orelse seek_line s instr);
fun location s = warning ("ATP input at: " ^ s);
-fun call_vampire (file:string, time: int) =
+fun call_vampire (file:string, time: int) =
let val _ = location file
val runtime = "-t " ^ (string_of_int time)
val vampire = ResAtp.helper_path "VAMPIRE_HOME" "vampire"
@@ -29,35 +26,33 @@
let val _ = location file
val eprover = ResAtp.helper_path "E_HOME" "eprover"
val runtime = "--cpu-limit=" ^ (string_of_int time)
- val (instr,outstr) = Unix.streamsOf (Unix.execute(eprover,
+ val (instr,outstr) = Unix.streamsOf (Unix.execute(eprover,
[runtime,"--tstp-in","-tAutoDev","-xAutoDev",file]))
in seek_line "# Proof found!\n" instr
- end;
+ end;
fun call_spass (file:string, time:int) =
let val _ = location file
val runtime = "-TimeLimit=" ^ (string_of_int time)
val spass = ResAtp.helper_path "SPASS_HOME" "SPASS"
- val (instr,outstr) = Unix.streamsOf (Unix.execute(spass,
+ val (instr,outstr) = Unix.streamsOf (Unix.execute(spass,
[runtime,"-Splits=0", "-FullRed=0", "-SOS=1",file]))
in seek_line "SPASS beiseite: Proof found.\n" instr
end;
-fun vampire_o _ (file,time) =
- if call_vampire (file,time)
- then (warning file; ResAtp.cond_rm_tmp file; HOLogic.mk_Trueprop HOLogic.false_const)
+fun vampire_o _ (file,time) =
+ if call_vampire (file,time)
+ then (warning file; ResAtp.cond_rm_tmp file; HOLogic.mk_Trueprop HOLogic.false_const)
else (ResAtp.cond_rm_tmp file; raise Fail ("vampire oracle failed"));
-fun eprover_o _ (file,time) =
- if call_eprover (file,time)
+fun eprover_o _ (file,time) =
+ if call_eprover (file,time)
then (warning file; ResAtp.cond_rm_tmp file; HOLogic.mk_Trueprop HOLogic.false_const)
else (ResAtp.cond_rm_tmp file; raise Fail ("eprover oracle failed"));
fun spass_o _ (file,time) =
- if call_spass (file,time)
- then (warning file; ResAtp.cond_rm_tmp file; HOLogic.mk_Trueprop HOLogic.false_const)
- else (ResAtp.cond_rm_tmp file; raise Fail ("SPASS oracle failed"));
-
+ if call_spass (file,time)
+ then (warning file; ResAtp.cond_rm_tmp file; HOLogic.mk_Trueprop HOLogic.false_const)
+ else (ResAtp.cond_rm_tmp file; raise Fail ("SPASS oracle failed"));
end;
-
--- a/src/HOL/Tools/res_reconstruct.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/HOL/Tools/res_reconstruct.ML Thu May 31 01:25:24 2007 +0200
@@ -8,16 +8,16 @@
(***************************************************************************)
signature RES_RECONSTRUCT =
sig
- val checkEProofFound:
- TextIO.instream * TextIO.outstream * Posix.Process.pid *
+ val checkEProofFound:
+ TextIO.instream * TextIO.outstream * Posix.Process.pid *
string * Proof.context * thm * int * string Vector.vector -> bool
- val checkVampProofFound:
- TextIO.instream * TextIO.outstream * Posix.Process.pid *
+ val checkVampProofFound:
+ TextIO.instream * TextIO.outstream * Posix.Process.pid *
string * Proof.context * thm * int * string Vector.vector -> bool
- val checkSpassProofFound:
- TextIO.instream * TextIO.outstream * Posix.Process.pid *
+ val checkSpassProofFound:
+ TextIO.instream * TextIO.outstream * Posix.Process.pid *
string * Proof.context * thm * int * string Vector.vector -> bool
- val signal_parent:
+ val signal_parent:
TextIO.outstream * Posix.Process.pid * string * string -> unit
end;
@@ -27,7 +27,7 @@
val trace_path = Path.basic "atp_trace";
-fun trace s = if !Output.debugging then File.append (File.tmp_path trace_path) s
+fun trace s = if !Output.debugging then File.append (File.tmp_path trace_path) s
else ();
(*Full proof reconstruction wanted*)
@@ -74,7 +74,7 @@
| syn_equal (t1, SOME (SOME _, t2)) = negate (equate (t1,t2));
(*Literals can involve negation, = and !=.*)
-val literal = $$"~" |-- term >> negate ||
+val literal = $$"~" |-- term >> negate ||
(term -- Scan.option (Scan.option ($$"!") --| $$"=" -- term) >> syn_equal) ;
val literals = literal -- Scan.repeat ($$"|" |-- literal) >> op:: ;
@@ -86,8 +86,8 @@
(*<cnf_annotated> ::=Ęcnf(<name>,<formula_role>,<cnf_formula><annotations>).
The <name> could be an identifier, but we assume integers.*)
-val tstp_line = (Scan.this_string "cnf" -- $$"(") |--
- integer --| $$"," -- Symbol.scan_id --| $$"," --
+val tstp_line = (Scan.this_string "cnf" -- $$"(") |--
+ integer --| $$"," -- Symbol.scan_id --| $$"," --
clause -- Scan.option annotations --| $$ ")";
@@ -121,7 +121,7 @@
else if check_valid_int sti
then (Char.chr (cList2int sti) :: s)
else (sti @ (#"_" :: s))
- in normalise_s s' cs false []
+ in normalise_s s' cs false []
end
else normalise_s s cs true []
| normalise_s s (c::cs) true sti =
@@ -147,7 +147,7 @@
exception STREE of stree;
(*If string s has the prefix s1, return the result of deleting it.*)
-fun strip_prefix s1 s =
+fun strip_prefix s1 s =
if String.isPrefix s1 s then SOME (normalise_string (String.extract (s, size s1, NONE)))
else NONE;
@@ -168,17 +168,17 @@
fun type_of_stree t =
case t of
Int _ => raise STREE t
- | Br (a,ts) =>
+ | Br (a,ts) =>
let val Ts = map type_of_stree ts
- in
+ in
case strip_prefix ResClause.tconst_prefix a of
SOME b => Type(invert_type_const b, Ts)
- | NONE =>
+ | NONE =>
if not (null ts) then raise STREE t (*only tconsts have type arguments*)
- else
+ else
case strip_prefix ResClause.tfree_prefix a of
SOME b => TFree("'" ^ b, HOLogic.typeS)
- | NONE =>
+ | NONE =>
case strip_prefix ResClause.tvar_prefix a of
SOME b => make_tvar b
| NONE => make_tvar a (*Variable from the ATP, say X1*)
@@ -186,7 +186,7 @@
(*Invert the table of translations between Isabelle and ATPs*)
val const_trans_table_inv =
- Symtab.update ("fequal", "op =")
+ Symtab.update ("fequal", "op =")
(Symtab.make (map swap (Symtab.dest ResClause.const_trans_table)));
fun invert_const c =
@@ -207,11 +207,11 @@
Int _ => raise STREE t
| Br ("hBOOL",[t]) => term_of_stree [] thy t (*ignore hBOOL*)
| Br ("hAPP",[t,u]) => term_of_stree (u::args) thy t
- | Br (a,ts) =>
+ | Br (a,ts) =>
case strip_prefix ResClause.const_prefix a of
- SOME "equal" =>
+ SOME "equal" =>
list_comb(Const ("op =", HOLogic.typeT), List.map (term_of_stree [] thy) ts)
- | SOME b =>
+ | SOME b =>
let val c = invert_const b
val nterms = length ts - num_typargs thy c
val us = List.map (term_of_stree [] thy) (List.take(ts,nterms) @ args)
@@ -224,17 +224,17 @@
val opr = (*a Free variable is typically a Skolem function*)
case strip_prefix ResClause.fixed_var_prefix a of
SOME b => Free(b,T)
- | NONE =>
+ | NONE =>
case strip_prefix ResClause.schematic_var_prefix a of
SOME b => make_var (b,T)
| NONE => make_var (a,T) (*Variable from the ATP, say X1*)
in list_comb (opr, List.map (term_of_stree [] thy) (args@ts)) end;
-(*Type class literal applied to a type. Returns triple of polarity, class, type.*)
+(*Type class literal applied to a type. Returns triple of polarity, class, type.*)
fun constraint_of_stree pol (Br("c_Not",[t])) = constraint_of_stree (not pol) t
| constraint_of_stree pol t = case t of
Int _ => raise STREE t
- | Br (a,ts) =>
+ | Br (a,ts) =>
(case (strip_prefix ResClause.class_prefix a, map type_of_stree ts) of
(SOME b, [T]) => (pol, b, T)
| _ => raise STREE t);
@@ -252,16 +252,16 @@
(*Final treatment of the list of "real" literals from a clause.*)
fun finish [] = HOLogic.true_const (*No "real" literals means only type information*)
- | finish lits =
+ | finish lits =
case nofalses lits of
[] => HOLogic.false_const (*The empty clause, since we started with real literals*)
| xs => foldr1 HOLogic.mk_disj (rev xs);
(*Accumulate sort constraints in vt, with "real" literals in lits.*)
fun lits_of_strees ctxt (vt, lits) [] = (vt, finish lits)
- | lits_of_strees ctxt (vt, lits) (t::ts) =
+ | lits_of_strees ctxt (vt, lits) (t::ts) =
lits_of_strees ctxt (add_constraint (constraint_of_stree true t, vt), lits) ts
- handle STREE _ =>
+ handle STREE _ =>
lits_of_strees ctxt (vt, term_of_stree [] (ProofContext.theory_of ctxt) t :: lits) ts;
(*Update TVars/TFrees with detected sort constraints.*)
@@ -279,7 +279,7 @@
(*Interpret a list of syntax trees as a clause, given by "real" literals and sort constraints.
vt0 holds the initial sort constraints, from the conjecture clauses.*)
-fun clause_of_strees_aux ctxt vt0 ts =
+fun clause_of_strees_aux ctxt vt0 ts =
let val (vt, dt) = lits_of_strees ctxt (vt0,[]) ts in
singleton (ProofContext.infer_types ctxt) (TypeInfer.constrain (fix_sorts vt dt) HOLogic.boolT)
end;
@@ -315,7 +315,7 @@
| add_tfree_constraint (_, vt) = vt;
fun tfree_constraints_of_clauses vt [] = vt
- | tfree_constraints_of_clauses vt ([lit]::tss) =
+ | tfree_constraints_of_clauses vt ([lit]::tss) =
(tfree_constraints_of_clauses (add_tfree_constraint (constraint_of_stree true lit, vt)) tss
handle STREE _ => (*not a positive type constraint: ignore*)
tfree_constraints_of_clauses vt tss)
@@ -341,13 +341,13 @@
| smash_types (f$t) = smash_types f $ smash_types t
| smash_types t = t;
-val sort_lits = sort Term.fast_term_ord o dest_disj o
+val sort_lits = sort Term.fast_term_ord o dest_disj o
smash_types o HOLogic.dest_Trueprop o strip_all_body;
fun permuted_clause t =
let val lits = sort_lits t
fun perm [] = NONE
- | perm (ctm::ctms) =
+ | perm (ctm::ctms) =
if forall (op aconv) (ListPair.zip (lits, sort_lits ctm)) then SOME ctm
else perm ctms
in perm end;
@@ -364,7 +364,7 @@
SOME u => "assume " ^ lname ^ ": \"" ^ string_of u ^ "\"\n"
| NONE => "assume? " ^ lname ^ ": \"" ^ string_of t ^ "\"\n") (*no match!!*)
| doline have (lname, t, deps) =
- have_or_show have lname ^ string_of t ^
+ have_or_show have lname ^ string_of t ^
"\"\n by (metis " ^ space_implode " " deps ^ ")\n"
fun dolines [(lname, t, deps)] = [doline "show " (lname, t, deps)]
| dolines ((lname, t, deps)::lines) = doline "have " (lname, t, deps) :: dolines lines
@@ -377,14 +377,14 @@
fun replace_dep (old:int, new) dep = if dep=old then new else [dep];
-fun replace_deps (old:int, new) (lno, t, deps) =
+fun replace_deps (old:int, new) (lno, t, deps) =
(lno, t, foldl (op union_int) [] (map (replace_dep (old, new)) deps));
(*Discard axioms; consolidate adjacent lines that prove the same clause, since they differ
only in type information.*)
fun add_prfline ((lno, "axiom", t, []), lines) = (*axioms are not proof lines*)
if eq_types t (*must be clsrel/clsarity: type information, so delete refs to it*)
- then map (replace_deps (lno, [])) lines
+ then map (replace_deps (lno, [])) lines
else
(case take_prefix (notequal t) lines of
(_,[]) => lines (*no repetition of proof line*)
@@ -398,22 +398,22 @@
else (*FIXME: Doesn't this code risk conflating proofs involving different types??*)
case take_prefix (notequal t) lines of
(_,[]) => (lno, t, deps) :: lines (*no repetition of proof line*)
- | (pre, (lno',t',deps')::post) =>
+ | (pre, (lno',t',deps')::post) =>
(lno, t', deps) :: (*repetition: replace later line by earlier one*)
(pre @ map (replace_deps (lno', [lno])) post);
(*Recursively delete empty lines (type information) from the proof.*)
fun add_nonnull_prfline ((lno, t, []), lines) = (*no dependencies, so a conjecture clause*)
if eq_types t (*must be type information, tfree_tcs, clsrel, clsarity: delete refs to it*)
- then delete_dep lno lines
- else (lno, t, []) :: lines
+ then delete_dep lno lines
+ else (lno, t, []) :: lines
| add_nonnull_prfline ((lno, t, deps), lines) = (lno, t, deps) :: lines
and delete_dep lno lines = foldr add_nonnull_prfline [] (map (replace_deps (lno, [])) lines);
fun bad_free (Free (a,_)) = String.isPrefix "llabs_" a orelse String.isPrefix "sko_" a
| bad_free _ = false;
-(*TVars are forbidden in goals. Also, we don't want lines with <2 dependencies.
+(*TVars are forbidden in goals. Also, we don't want lines with <2 dependencies.
To further compress proofs, setting modulus:=n deletes every nth line, and nlines
counts the number of proof lines processed so far.
Deleted lines are replaced by their own dependencies. Note that the "add_nonnull_prfline"
@@ -423,7 +423,7 @@
| add_wanted_prfline (line, (nlines, [])) = (nlines, [line]) (*final line must be kept*)
| add_wanted_prfline ((lno, t, deps), (nlines, lines)) =
if eq_types t orelse not (null (term_tvars t)) orelse
- length deps < 2 orelse nlines mod !modulus <> 0 orelse
+ length deps < 2 orelse nlines mod !modulus <> 0 orelse
exists bad_free (term_frees t)
then (nlines+1, map (replace_deps (lno, deps)) lines) (*Delete line*)
else (nlines+1, (lno, t, deps) :: lines);
@@ -432,9 +432,9 @@
fun stringify_deps thm_names deps_map [] = []
| stringify_deps thm_names deps_map ((lno, t, deps) :: lines) =
if lno <= Vector.length thm_names (*axiom*)
- then (Vector.sub(thm_names,lno-1), t, []) :: stringify_deps thm_names deps_map lines
+ then (Vector.sub(thm_names,lno-1), t, []) :: stringify_deps thm_names deps_map lines
else let val lname = Int.toString (length deps_map)
- fun fix lno = if lno <= Vector.length thm_names
+ fun fix lno = if lno <= Vector.length thm_names
then SOME(Vector.sub(thm_names,lno-1))
else AList.lookup op= deps_map lno;
in (lname, t, List.mapPartial fix (distinct (op=) deps)) ::
@@ -448,15 +448,15 @@
fun decode_tstp_file cnfs ctxt th sgno thm_names =
let val tuples = map (dest_tstp o tstp_line o explode) cnfs
- val nonnull_lines =
- foldr add_nonnull_prfline []
+ val nonnull_lines =
+ foldr add_nonnull_prfline []
(foldr add_prfline [] (decode_tstp_list ctxt tuples))
val (_,lines) = foldr add_wanted_prfline (0,[]) nonnull_lines
val (ccls,fixes) = ResAxioms.neg_conjecture_clauses th sgno
val ccls = map forall_intr_vars ccls
- in
+ in
app (fn th => Output.debug (fn () => string_of_thm th)) ccls;
- isar_header (map #1 fixes) ^
+ isar_header (map #1 fixes) ^
String.concat (isar_lines ctxt (map prop_of ccls) (stringify_deps thm_names [] lines))
end;
@@ -469,17 +469,17 @@
val encode_newlines = String.translate (fn c => if c = #"\n" then "\t" else str c);
val restore_newlines = String.translate (fn c => if c = #"\t" then "\n" else str c);
-fun signal_success probfile toParent ppid msg =
+fun signal_success probfile toParent ppid msg =
(trace ("\nReporting Success for " ^ probfile ^ "\n" ^ msg);
TextIO.output (toParent, "Success. " ^ encode_newlines msg ^ "\n");
TextIO.output (toParent, probfile ^ "\n");
TextIO.flushOut toParent;
Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2));
-fun tstp_extract proofextract probfile toParent ppid ctxt th sgno thm_names =
+fun tstp_extract proofextract probfile toParent ppid ctxt th sgno thm_names =
let val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract))
- in
- signal_success probfile toParent ppid
+ in
+ signal_success probfile toParent ppid
(decode_tstp_file cnfs ctxt th sgno thm_names)
end;
@@ -487,8 +487,8 @@
(**** retrieve the axioms that were used in the proof ****)
(*Get names of axioms used. Axioms are indexed from 1, while the vector is indexed from 0*)
-fun get_axiom_names (thm_names: string vector) step_nums =
- let fun is_axiom n = n <= Vector.length thm_names
+fun get_axiom_names (thm_names: string vector) step_nums =
+ let fun is_axiom n = n <= Vector.length thm_names
fun index i = Vector.sub(thm_names, i-1)
val axnums = List.filter is_axiom step_nums
val axnames = sort_distinct string_ord (map index axnums)
@@ -497,10 +497,10 @@
else axnames
end
- (*String contains multiple lines. We want those of the form
+ (*String contains multiple lines. We want those of the form
"253[0:Inp] et cetera..."
A list consisting of the first number in each line is returned. *)
-fun get_spass_linenums proofstr =
+fun get_spass_linenums proofstr =
let val toks = String.tokens (not o Char.isAlphaNum)
fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok
| inputno _ = NONE
@@ -509,7 +509,7 @@
fun get_axiom_names_spass proofstr thm_names =
get_axiom_names thm_names (get_spass_linenums proofstr);
-
+
fun not_comma c = c <> #",";
(*A valid TSTP axiom line has the form cnf(NNN,axiom,...) where NNN is a positive integer.*)
@@ -520,18 +520,18 @@
(*We only allow negated_conjecture because the line number will be removed in
get_axiom_names above, while suppressing the UNSOUND warning*)
val ints = if Substring.string rolef mem_string ["axiom","negated_conjecture"]
- then Substring.string intf
- else "error"
+ then Substring.string intf
+ else "error"
in Int.fromString ints end
- handle Fail _ => NONE;
+ handle Fail _ => NONE;
fun get_axiom_names_tstp proofstr thm_names =
get_axiom_names thm_names (List.mapPartial parse_tstp_line (split_lines proofstr));
-
- (*String contains multiple lines. We want those of the form
+
+ (*String contains multiple lines. We want those of the form
"*********** [448, input] ***********".
A list consisting of the first number in each line is returned. *)
-fun get_vamp_linenums proofstr =
+fun get_vamp_linenums proofstr =
let val toks = String.tokens (not o Char.isAlphaNum)
fun inputno [ntok,"input"] = Int.fromString ntok
| inputno _ = NONE
@@ -540,21 +540,21 @@
fun get_axiom_names_vamp proofstr thm_names =
get_axiom_names thm_names (get_vamp_linenums proofstr);
-
+
fun rules_to_metis [] = "metis"
| rules_to_metis xs = "(metis " ^ space_implode " " xs ^ ")"
(*The signal handler in watcher.ML must be able to read the output of this.*)
-fun prover_lemma_list_aux getax proofstr probfile toParent ppid thm_names =
+fun prover_lemma_list_aux getax proofstr probfile toParent ppid thm_names =
(trace ("\n\nGetting lemma names. proofstr is " ^ proofstr ^
" num of clauses is " ^ string_of_int (Vector.length thm_names));
- signal_success probfile toParent ppid
+ signal_success probfile toParent ppid
("Try this command: \n apply " ^ rules_to_metis (getax proofstr thm_names))
)
handle e => (*FIXME: exn handler is too general!*)
(trace ("\nprover_lemma_list_aux: In exception handler: " ^ Toplevel.exn_message e);
- TextIO.output (toParent, "Translation failed for the proof: " ^
+ TextIO.output (toParent, "Translation failed for the proof: " ^
String.toString proofstr ^ "\n");
TextIO.output (toParent, probfile);
TextIO.flushOut toParent;
@@ -570,9 +570,9 @@
(**** Extracting proofs from an ATP's output ****)
(*Return everything in s that comes before the string t*)
-fun cut_before t s =
+fun cut_before t s =
let val (s1,s2) = Substring.position t (Substring.full s)
- in if Substring.size s2 = 0 then error "cut_before: string not found"
+ in if Substring.size s2 = 0 then error "cut_before: string not found"
else Substring.string s2
end;
@@ -593,29 +593,28 @@
return value is currently never used!*)
fun startTransfer endS (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names) =
let fun transferInput currentString =
- let val thisLine = TextIO.inputLine fromChild
- in
- if thisLine = "" (*end of file?*)
- then (trace ("\n extraction_failed. End bracket: " ^ endS ^
- "\naccumulated text: " ^ currentString);
- false)
- else if String.isPrefix endS thisLine
+ (case TextIO.inputLine fromChild of
+ NONE => (*end of file?*)
+ (trace ("\n extraction_failed. End bracket: " ^ endS ^
+ "\naccumulated text: " ^ currentString);
+ false)
+ | SOME thisLine =>
+ if String.isPrefix endS thisLine
then let val proofextract = currentString ^ cut_before endS thisLine
val lemma_list = if endS = end_V8 then vamp_lemma_list
else if endS = end_SPASS then spass_lemma_list
else e_lemma_list
in
- trace ("\nExtracted proof:\n" ^ proofextract);
+ trace ("\nExtracted proof:\n" ^ proofextract);
if !full andalso String.isPrefix "cnf(" proofextract
then tstp_extract proofextract probfile toParent ppid ctxt th sgno thm_names
else lemma_list proofextract probfile toParent ppid thm_names;
true
end
- else transferInput (currentString^thisLine)
- end
+ else transferInput (currentString^thisLine))
in
transferInput ""
- end
+ end
(*The signal handler in watcher.ML must be able to read the output of this.*)
@@ -632,27 +631,23 @@
(*Called from watcher. Returns true if the Vampire process has returned a verdict.*)
fun checkVampProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =
- let val thisLine = TextIO.inputLine fromChild
- in
- trace thisLine;
- if thisLine = ""
- then (trace "\nNo proof output seen"; false)
- else if String.isPrefix start_V8 thisLine
+ (case TextIO.inputLine fromChild of
+ NONE => (trace "\nNo proof output seen"; false)
+ | SOME thisLine =>
+ if String.isPrefix start_V8 thisLine
then startTransfer end_V8 arg
else if (String.isPrefix "Satisfiability detected" thisLine) orelse
(String.isPrefix "Refutation not found" thisLine)
then (signal_parent (toParent, ppid, "Failure\n", probfile);
true)
- else checkVampProofFound arg
- end
+ else checkVampProofFound arg);
(*Called from watcher. Returns true if the E process has returned a verdict.*)
-fun checkEProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =
- let val thisLine = TextIO.inputLine fromChild
- in
- trace thisLine;
- if thisLine = "" then (trace "\nNo proof output seen"; false)
- else if String.isPrefix start_E thisLine
+fun checkEProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =
+ (case TextIO.inputLine fromChild of
+ NONE => (trace "\nNo proof output seen"; false)
+ | SOME thisLine =>
+ if String.isPrefix start_E thisLine
then startTransfer end_E arg
else if String.isPrefix "# Problem is satisfiable" thisLine
then (signal_parent (toParent, ppid, "Invalid\n", probfile);
@@ -660,16 +655,14 @@
else if String.isPrefix "# Cannot determine problem status within resource limit" thisLine
then (signal_parent (toParent, ppid, "Failure\n", probfile);
true)
- else checkEProofFound arg
- end;
+ else checkEProofFound arg);
(*Called from watcher. Returns true if the SPASS process has returned a verdict.*)
-fun checkSpassProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =
- let val thisLine = TextIO.inputLine fromChild
- in
- trace thisLine;
- if thisLine = "" then (trace "\nNo proof output seen"; false)
- else if String.isPrefix "Here is a proof" thisLine
+fun checkSpassProofFound (arg as (fromChild, toParent, ppid, probfile, ctxt, th, sgno, thm_names)) =
+ (case TextIO.inputLine fromChild of
+ NONE => (trace "\nNo proof output seen"; false)
+ | SOME thisLine =>
+ if String.isPrefix "Here is a proof" thisLine
then startTransfer end_SPASS arg
else if thisLine = "SPASS beiseite: Completion found.\n"
then (signal_parent (toParent, ppid, "Invalid\n", probfile);
@@ -678,7 +671,6 @@
thisLine = "SPASS beiseite: Maximal number of loops exceeded.\n"
then (signal_parent (toParent, ppid, "Failure\n", probfile);
true)
- else checkSpassProofFound arg
- end
+ else checkSpassProofFound arg);
end;
--- a/src/Pure/General/source.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/Pure/General/source.ML Thu May 31 01:25:24 2007 +0200
@@ -124,7 +124,9 @@
else
(TextIO.output (outstream, prompt);
TextIO.flushOut outstream;
- (input @ explode (TextIO.inputLine instream), ()))
+ (case TextIO.inputLine instream of
+ SOME line => (input @ explode line, ())
+ | NONE => (input, ())))
end;
fun of_stream instream outstream =
--- a/src/Pure/ML-Systems/alice.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/Pure/ML-Systems/alice.ML Thu May 31 01:25:24 2007 +0200
@@ -104,11 +104,7 @@
structure TextIO =
struct
open TextIO;
-
- fun inputLine is =
- (case TextIO.inputLine is of
- SOME str => str
- | NONE => "")
+ fun inputLine is = TextIO.inputLine is
handle IO.Io _ => raise Interrupt;
end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/polyml-4.1.3.ML Thu May 31 01:25:24 2007 +0200
@@ -0,0 +1,8 @@
+(* Title: Pure/ML-Systems/polyml-4.1.3.ML
+ ID: $Id$
+
+Compatibility wrapper for Poly/ML 4.1.3.
+*)
+
+use "ML-Systems/polyml-old-basis.ML";
+use "ML-Systems/polyml.ML";
--- a/src/Pure/ML-Systems/polyml-4.1.4-patch.ML Wed May 30 23:32:54 2007 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,28 +0,0 @@
-(* Title: Pure/ML-Systems/polyml-4.1.4-patch.ML
- ID: $Id$
-
-Basis library fixes for Poly/ML 4.2.0 or later.
-*)
-
-structure Posix =
-struct
- open Posix;
- structure IO =
- struct
- open IO;
- val mkReader = mkTextReader;
- val mkWriter = mkTextWriter;
- end;
-end;
-
-structure TextIO =
-struct
- open TextIO;
- fun inputLine is = Option.getOpt (TextIO.inputLine is, "");
-end;
-
-structure Substring =
-struct
- open Substring;
- val all = full;
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/polyml-4.1.4.ML Thu May 31 01:25:24 2007 +0200
@@ -0,0 +1,8 @@
+(* Title: Pure/ML-Systems/polyml-4.1.4.ML
+ ID: $Id$
+
+Compatibility wrapper for Poly/ML 4.1.4.
+*)
+
+use "ML-Systems/polyml-old-basis.ML";
+use "ML-Systems/polyml.ML";
--- a/src/Pure/ML-Systems/polyml-4.2.0.ML Wed May 30 23:32:54 2007 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(* Title: Pure/ML-Systems/polyml-4.2.0.ML
- ID: $Id$
-
-Compatibility wrapper for Poly/ML 4.2.0.
-*)
-
-use "ML-Systems/polyml-4.1.4-patch.ML";
-use "ML-Systems/polyml.ML";
--- a/src/Pure/ML-Systems/polyml-5.0.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/Pure/ML-Systems/polyml-5.0.ML Thu May 31 01:25:24 2007 +0200
@@ -4,7 +4,6 @@
Compatibility wrapper for Poly/ML 5.0.
*)
-use "ML-Systems/polyml-4.1.4-patch.ML";
use "ML-Systems/polyml.ML";
val pointer_eq = PolyML.pointerEq;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/polyml-old-basis.ML Thu May 31 01:25:24 2007 +0200
@@ -0,0 +1,30 @@
+(* Title: Pure/ML-Systems/polyml-old-basis.ML
+ ID: $Id$
+
+Fixes for the old SML basis library (before Poly/ML 4.2.0).
+*)
+
+structure String =
+struct
+ open String;
+ fun isSuffix s1 s2 =
+ let val n1 = size s1 and n2 = size s2
+ in if n1 = n2 then s1 = s2 else n1 <= n2 andalso String.substring (s2, n2 - n1, n1) = s1 end;
+end;
+
+structure Substring =
+struct
+ open Substring;
+ val full = all;
+end;
+
+structure Posix =
+struct
+ open Posix;
+ structure IO =
+ struct
+ open IO;
+ val mkTextReader = mkReader;
+ val mkTextWriter = mkWriter;
+ end;
+end;
--- a/src/Pure/ML-Systems/polyml-posix.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/Pure/ML-Systems/polyml-posix.ML Thu May 31 01:25:24 2007 +0200
@@ -4,20 +4,6 @@
Posix patches for Poly/ML.
*)
-structure OriginalPosix = Posix;
-structure OriginalIO = Posix.IO;
-
-structure Posix =
-struct
- open OriginalPosix
- structure IO =
- struct
- open OriginalIO
- val mkTextReader = mkReader
- val mkTextWriter = mkWriter
- end;
-end;
-
(*This extension of the Poly/ML Signal structure is only necessary
because in SML/NJ, types Posix.Signal.signal and Signals.signal differ.*)
structure IsaSignal =
--- a/src/Pure/ML-Systems/polyml.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/Pure/ML-Systems/polyml.ML Thu May 31 01:25:24 2007 +0200
@@ -1,28 +1,13 @@
(* Title: Pure/ML-Systems/polyml.ML
ID: $Id$
-Compatibility file for Poly/ML (version 4.0, 4.1, 4.1.x).
+Compatibility file for Poly/ML (version 4.1.x and 4.2.0).
*)
(** ML system and platform related **)
(* String compatibility *)
-structure String =
-struct
- fun isSuffix s1 s2 =
- let val n1 = size s1 and n2 = size s2
- in if n1 = n2 then s1 = s2 else n1 <= n2 andalso String.substring (s2, n2 - n1, n1) = s1 end;
- open String;
-end;
-
-structure Substring =
-struct
- open Substring;
- val full = all;
-end;
-
-
(*low-level pointer equality*)
val pointer_eq = Address.wordEq;
@@ -189,7 +174,7 @@
(*Convert a process ID to a decimal string (chiefly for tracing)*)
fun string_of_pid pid =
- Word.fmt StringCvt.DEC (Word.fromLargeWord (Posix.Process.pidToWord pid));
+ Word.fmt StringCvt.DEC (Word.fromLargeWord (Posix.Process.pidToWord pid));
(* getenv *)
--- a/src/Pure/ML-Systems/smlnj.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/Pure/ML-Systems/smlnj.ML Thu May 31 01:25:24 2007 +0200
@@ -159,11 +159,7 @@
structure TextIO =
struct
open TextIO;
-
- fun inputLine is =
- (case TextIO.inputLine is of
- SOME str => str
- | NONE => "")
+ fun inputLine is = TextIO.inputLine is
handle IO.Io _ => raise Interrupt;
end;
--- a/src/Pure/tctical.ML Wed May 30 23:32:54 2007 +0200
+++ b/src/Pure/tctical.ML Thu May 31 01:25:24 2007 +0200
@@ -203,7 +203,7 @@
(*Pause until a line is typed -- if non-empty then fail. *)
fun pause_tac st =
(tracing "** Press RETURN to continue:";
- if TextIO.inputLine TextIO.stdIn = "\n" then Seq.single st
+ if TextIO.inputLine TextIO.stdIn = SOME "\n" then Seq.single st
else (tracing "Goodbye"; Seq.empty));
exception TRACE_EXIT of thm
@@ -215,13 +215,13 @@
(*Handle all tracing commands for current state and tactic *)
fun exec_trace_command flag (tac, st) =
- case TextIO.inputLine(TextIO.stdIn) of
- "\n" => tac st
- | "f\n" => Seq.empty
- | "o\n" => (flag:=false; tac st)
- | "s\n" => (suppress_tracing:=true; tac st)
- | "x\n" => (tracing "Exiting now"; raise (TRACE_EXIT st))
- | "quit\n" => raise TRACE_QUIT
+ case TextIO.inputLine TextIO.stdIn of
+ SOME "\n" => tac st
+ | SOME "f\n" => Seq.empty
+ | SOME "o\n" => (flag:=false; tac st)
+ | SOME "s\n" => (suppress_tracing:=true; tac st)
+ | SOME "x\n" => (tracing "Exiting now"; raise (TRACE_EXIT st))
+ | SOME "quit\n" => raise TRACE_QUIT
| _ => (tracing
"Type RETURN to continue or...\n\
\ f - to fail here\n\