merged
authorbulwahn
Thu, 16 Sep 2010 16:20:20 +0200
changeset 39469 d58b91a9b8b1
parent 39468 3cb3b1668c5d (current diff)
parent 39458 13c8577e1783 (diff)
child 39470 d7caf48c4676
merged
--- a/src/HOL/IsaMakefile	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/IsaMakefile	Thu Sep 16 16:20:20 2010 +0200
@@ -268,6 +268,7 @@
   $(SRC)/Tools/Metis/metis.ML \
   Tools/async_manager.ML \
   Tools/ATP/atp_problem.ML \
+  Tools/ATP/atp_proof.ML \
   Tools/ATP/atp_systems.ML \
   Tools/choice_specification.ML \
   Tools/int_arith.ML \
--- a/src/HOL/Mirabelle/Tools/mirabelle.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Mirabelle/Tools/mirabelle.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -159,9 +159,9 @@
     val {context = ctxt, facts, goal} = Proof.goal st
     val full_tac = HEADGOAL (Method.insert_tac facts THEN' tac ctxt)
   in
-    (case TimeLimit.timeLimit time (Seq.pull o full_tac) goal of
-      SOME _ => true
-    | NONE => false)
+    (case try (TimeLimit.timeLimit time (Seq.pull o full_tac)) goal of
+      SOME (SOME _) => true
+    | _ => false)
   end
 
 local
--- a/src/HOL/Sledgehammer.thy	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Sledgehammer.thy	Thu Sep 16 16:20:20 2010 +0200
@@ -11,6 +11,7 @@
 imports Plain Hilbert_Choice
 uses
   ("Tools/ATP/atp_problem.ML")
+  ("Tools/ATP/atp_proof.ML")
   ("Tools/ATP/atp_systems.ML")
   ("~~/src/Tools/Metis/metis.ML")
   ("Tools/Sledgehammer/clausifier.ML")
@@ -92,6 +93,7 @@
 done
 
 use "Tools/ATP/atp_problem.ML"
+use "Tools/ATP/atp_proof.ML"
 use "Tools/ATP/atp_systems.ML"
 setup ATP_Systems.setup
 
--- a/src/HOL/Tools/ATP/atp_problem.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_problem.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -2,7 +2,7 @@
     Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
     Author:     Jasmin Blanchette, TU Muenchen
 
-TPTP syntax.
+Abstract representation of ATP problems and TPTP syntax.
 *)
 
 signature ATP_PROBLEM =
@@ -14,16 +14,17 @@
     AQuant of quantifier * 'a list * ('a, 'b) formula |
     AConn of connective * ('a, 'b) formula list |
     AAtom of 'b
+  type 'a uniform_formula = ('a, 'a fo_term) formula
 
   datatype kind = Axiom | Hypothesis | Conjecture
   datatype 'a problem_line = Fof of string * kind * ('a, 'a fo_term) formula
   type 'a problem = (string * 'a problem_line list) list
 
   val timestamp : unit -> string
-  val is_tptp_variable : string -> bool
-  val strings_for_tptp_problem :
+  val is_atp_variable : string -> bool
+  val tptp_strings_for_atp_problem :
     bool -> (string * string problem_line list) list -> string list
-  val nice_tptp_problem :
+  val nice_atp_problem :
     bool -> ('a * (string * string) problem_line list) list
     -> ('a * string problem_line list) list
        * (string Symtab.table * string Symtab.table) option
@@ -41,6 +42,7 @@
   AQuant of quantifier * 'a list * ('a, 'b) formula |
   AConn of connective * ('a, 'b) formula list |
   AAtom of 'b
+type 'a uniform_formula = ('a, 'a fo_term) formula
 
 datatype kind = Axiom | Hypothesis | Conjecture
 datatype 'a problem_line = Fof of string * kind * ('a, 'a fo_term) formula
@@ -90,7 +92,7 @@
     "fof(" ^ ident ^ ", " ^ string_for_kind kind ^ ",\n    (" ^
     string_for_formula phi ^ ")).\n"
   end
-fun strings_for_tptp_problem use_conjecture_for_hypotheses problem =
+fun tptp_strings_for_atp_problem use_conjecture_for_hypotheses problem =
   "% This file was generated by Isabelle (most likely Sledgehammer)\n\
   \% " ^ timestamp () ^ "\n" ::
   maps (fn (_, []) => []
@@ -99,7 +101,7 @@
            map (string_for_problem_line use_conjecture_for_hypotheses) lines)
        problem
 
-fun is_tptp_variable s = Char.isUpper (String.sub (s, 0))
+fun is_atp_variable s = Char.isUpper (String.sub (s, 0))
 
 
 (** Nice names **)
@@ -163,7 +165,7 @@
 fun nice_problem problem =
   pool_map (fn (heading, lines) =>
                pool_map nice_problem_line lines #>> pair heading) problem
-fun nice_tptp_problem readable_names problem =
+fun nice_atp_problem readable_names problem =
   nice_problem problem (empty_name_pool readable_names)
 
 end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/ATP/atp_proof.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -0,0 +1,259 @@
+(*  Title:      HOL/Tools/ATP/atp_proof.ML
+    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
+    Author:     Claire Quigley, Cambridge University Computer Laboratory
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Abstract representation of ATP proofs and TSTP/Vampire/SPASS syntax.
+*)
+
+signature ATP_PROOF =
+sig
+  type 'a fo_term = 'a ATP_Problem.fo_term
+  type 'a uniform_formula = 'a ATP_Problem.uniform_formula
+
+  type step_name = string * string option
+
+  datatype 'a step =
+    Definition of step_name * 'a * 'a |
+    Inference of step_name * 'a * step_name list
+
+  type 'a proof = 'a uniform_formula step list
+
+  val strip_spaces : (char -> bool) -> string -> string
+  val is_same_step : step_name * step_name -> bool
+  val atp_proof_from_tstplike_string : string -> string proof
+  val map_term_names_in_atp_proof :
+    (string -> string) -> string proof -> string proof
+  val nasty_atp_proof : string Symtab.table -> string proof -> string proof
+end;
+
+structure ATP_Proof : ATP_PROOF =
+struct
+
+fun strip_spaces_in_list _ [] = []
+  | strip_spaces_in_list _ [c1] = if Char.isSpace c1 then [] else [str c1]
+  | strip_spaces_in_list is_evil [c1, c2] =
+    strip_spaces_in_list is_evil [c1] @ strip_spaces_in_list is_evil [c2]
+  | strip_spaces_in_list is_evil (c1 :: c2 :: c3 :: cs) =
+    if Char.isSpace c1 then
+      strip_spaces_in_list is_evil (c2 :: c3 :: cs)
+    else if Char.isSpace c2 then
+      if Char.isSpace c3 then
+        strip_spaces_in_list is_evil (c1 :: c3 :: cs)
+      else
+        str c1 :: (if forall is_evil [c1, c3] then [" "] else []) @
+        strip_spaces_in_list is_evil (c3 :: cs)
+    else
+      str c1 :: strip_spaces_in_list is_evil (c2 :: c3 :: cs)
+fun strip_spaces is_evil =
+  implode o strip_spaces_in_list is_evil o String.explode
+
+fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
+val strip_spaces_except_between_ident_chars = strip_spaces is_ident_char
+
+open ATP_Problem
+
+fun mk_anot (AConn (ANot, [phi])) = phi
+  | mk_anot phi = AConn (ANot, [phi])
+fun mk_aconn c (phi1, phi2) = AConn (c, [phi1, phi2])
+
+type step_name = string * string option
+
+fun is_same_step p = p |> pairself fst |> op =
+
+fun step_name_ord p =
+  let val q = pairself fst p in
+    (* The "unprefix" part is to cope with remote Vampire's output. The proper
+       solution would be to perform a topological sort, e.g. using the nice
+       "Graph" functor. *)
+    case pairself (Int.fromString o perhaps (try (unprefix "f"))) q of
+      (NONE, NONE) => string_ord q
+    | (NONE, SOME _) => LESS
+    | (SOME _, NONE) => GREATER
+    | (SOME i, SOME j) => int_ord (i, j)
+  end
+
+datatype 'a step =
+  Definition of step_name * 'a * 'a |
+  Inference of step_name * 'a * step_name list
+
+type 'a proof = 'a uniform_formula step list
+
+fun step_name (Definition (name, _, _)) = name
+  | step_name (Inference (name, _, _)) = name
+
+(**** PARSING OF TSTP FORMAT ****)
+
+(*Strings enclosed in single quotes, e.g. filenames*)
+val scan_general_id =
+  $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode
+  || Scan.repeat ($$ "$") -- Scan.many1 Symbol.is_letdig
+     >> (fn (ss1, ss2) => implode ss1 ^ implode ss2)
+
+(* Generalized first-order terms, which include file names, numbers, etc. *)
+fun parse_annotation strict x =
+  ((scan_general_id ::: Scan.repeat ($$ " " |-- scan_general_id)
+      >> (strict ? filter (is_some o Int.fromString)))
+   -- Scan.optional (parse_annotation strict) [] >> op @
+   || $$ "(" |-- parse_annotations strict --| $$ ")"
+   || $$ "[" |-- parse_annotations strict --| $$ "]") x
+and parse_annotations strict x =
+  (Scan.optional (parse_annotation strict
+                  ::: Scan.repeat ($$ "," |-- parse_annotation strict)) []
+   >> flat) x
+
+(* Vampire proof lines sometimes contain needless information such as "(0:3)",
+   which can be hard to disambiguate from function application in an LL(1)
+   parser. As a workaround, we extend the TPTP term syntax with such detritus
+   and ignore it. *)
+fun parse_vampire_detritus x =
+  (scan_general_id |-- $$ ":" --| scan_general_id >> K []) x
+
+fun parse_term x =
+  (scan_general_id
+    -- Scan.optional ($$ "(" |-- (parse_vampire_detritus || parse_terms)
+                      --| $$ ")") []
+    --| Scan.optional ($$ "(" |-- parse_vampire_detritus --| $$ ")") []
+   >> ATerm) x
+and parse_terms x = (parse_term ::: Scan.repeat ($$ "," |-- parse_term)) x
+
+val parse_atom =
+  parse_term -- Scan.option (Scan.option ($$ "!") --| $$ "=" -- parse_term)
+  >> (fn (u1, NONE) => AAtom u1
+       | (u1, SOME (NONE, u2)) => AAtom (ATerm ("c_equal", [u1, u2]))
+       | (u1, SOME (SOME _, u2)) =>
+         mk_anot (AAtom (ATerm ("c_equal", [u1, u2]))))
+
+fun fo_term_head (ATerm (s, _)) = s
+
+(* TPTP formulas are fully parenthesized, so we don't need to worry about
+   operator precedence. *)
+fun parse_formula x =
+  (($$ "(" |-- parse_formula --| $$ ")"
+    || ($$ "!" >> K AForall || $$ "?" >> K AExists)
+       --| $$ "[" -- parse_terms --| $$ "]" --| $$ ":" -- parse_formula
+       >> (fn ((q, ts), phi) => AQuant (q, map fo_term_head ts, phi))
+    || $$ "~" |-- parse_formula >> mk_anot
+    || parse_atom)
+   -- Scan.option ((Scan.this_string "=>" >> K AImplies
+                    || Scan.this_string "<=>" >> K AIff
+                    || Scan.this_string "<~>" >> K ANotIff
+                    || Scan.this_string "<=" >> K AIf
+                    || $$ "|" >> K AOr || $$ "&" >> K AAnd)
+                   -- parse_formula)
+   >> (fn (phi1, NONE) => phi1
+        | (phi1, SOME (c, phi2)) => mk_aconn c (phi1, phi2))) x
+
+val parse_tstp_extra_arguments =
+  Scan.optional ($$ "," |-- parse_annotation false
+                 --| Scan.option ($$ "," |-- parse_annotations false)) []
+
+(* Syntax: (fof|cnf)\(<num>, <formula_role>, <formula> <extra_arguments>\).
+   The <num> could be an identifier, but we assume integers. *)
+val parse_tstp_line =
+  ((Scan.this_string "fof" || Scan.this_string "cnf") -- $$ "(")
+    |-- scan_general_id --| $$ "," -- Symbol.scan_id --| $$ ","
+    -- parse_formula -- parse_tstp_extra_arguments --| $$ ")" --| $$ "."
+   >> (fn (((num, role), phi), deps) =>
+          let
+            val (name, deps) =
+              case deps of
+                ["file", _, s] => ((num, SOME s), [])
+              | _ => ((num, NONE), deps)
+          in
+            case role of
+              "definition" =>
+              (case phi of
+                 AConn (AIff, [phi1 as AAtom _, phi2]) =>
+                 Definition (name, phi1, phi2)
+               | AAtom (ATerm ("c_equal", _)) =>
+                 (* Vampire's equality proxy axiom *)
+                 Inference (name, phi, map (rpair NONE) deps)
+               | _ => raise Fail "malformed definition")
+            | _ => Inference (name, phi, map (rpair NONE) deps)
+          end)
+
+(**** PARSING OF VAMPIRE OUTPUT ****)
+
+(* Syntax: <num>. <formula> <annotation> *)
+val parse_vampire_line =
+  scan_general_id --| $$ "." -- parse_formula -- parse_annotation true
+  >> (fn ((num, phi), deps) =>
+         Inference ((num, NONE), phi, map (rpair NONE) deps))
+
+(**** PARSING OF SPASS OUTPUT ****)
+
+(* SPASS returns clause references of the form "x.y". We ignore "y", whose role
+   is not clear anyway. *)
+val parse_dot_name = scan_general_id --| $$ "." --| scan_general_id
+
+val parse_spass_annotations =
+  Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name
+                                         --| Scan.option ($$ ","))) []
+
+(* It is not clear why some literals are followed by sequences of stars and/or
+   pluses. We ignore them. *)
+val parse_decorated_atom =
+  parse_atom --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")
+
+fun mk_horn ([], []) = AAtom (ATerm ("c_False", []))
+  | mk_horn ([], pos_lits) = foldr1 (mk_aconn AOr) pos_lits
+  | mk_horn (neg_lits, []) = mk_anot (foldr1 (mk_aconn AAnd) neg_lits)
+  | mk_horn (neg_lits, pos_lits) =
+    mk_aconn AImplies (foldr1 (mk_aconn AAnd) neg_lits,
+                       foldr1 (mk_aconn AOr) pos_lits)
+
+val parse_horn_clause =
+  Scan.repeat parse_decorated_atom --| $$ "|" --| $$ "|"
+    -- Scan.repeat parse_decorated_atom --| $$ "-" --| $$ ">"
+    -- Scan.repeat parse_decorated_atom
+  >> (mk_horn o apfst (op @))
+
+(* Syntax: <num>[0:<inference><annotations>]
+   <atoms> || <atoms> -> <atoms>. *)
+val parse_spass_line =
+  scan_general_id --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id
+    -- parse_spass_annotations --| $$ "]" -- parse_horn_clause --| $$ "."
+  >> (fn ((num, deps), u) => Inference ((num, NONE), u, map (rpair NONE) deps))
+
+val parse_line = parse_tstp_line || parse_vampire_line || parse_spass_line
+val parse_proof =
+  fst o Scan.finite Symbol.stopper
+            (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output")
+                            (Scan.repeat1 parse_line)))
+  o explode o strip_spaces_except_between_ident_chars
+
+fun clean_up_dependency seen dep = find_first (curry is_same_step dep) seen
+fun clean_up_dependencies _ [] = []
+  | clean_up_dependencies seen ((step as Definition (name, _, _)) :: steps) =
+    step :: clean_up_dependencies (name :: seen) steps
+  | clean_up_dependencies seen (Inference (name, u, deps) :: steps) =
+    Inference (name, u, map_filter (clean_up_dependency seen) deps) ::
+    clean_up_dependencies (name :: seen) steps
+
+val atp_proof_from_tstplike_string =
+  suffix "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
+  #> parse_proof
+  #> sort (step_name_ord o pairself step_name)
+  #> clean_up_dependencies []
+
+fun map_term_names_in_term f (ATerm (s, ts)) =
+  ATerm (f s, map (map_term_names_in_term f) ts)
+fun map_term_names_in_formula f (AQuant (q, xs, phi)) =
+    AQuant (q, xs, map_term_names_in_formula f phi)
+  | map_term_names_in_formula f (AConn (c, phis)) =
+    AConn (c, map (map_term_names_in_formula f) phis)
+  | map_term_names_in_formula f (AAtom t) = AAtom (map_term_names_in_term f t)
+fun map_term_names_in_step f (Definition (name, phi1, phi2)) =
+    Definition (name, map_term_names_in_formula f phi1,
+                map_term_names_in_formula f phi2)
+  | map_term_names_in_step f (Inference (name, phi, deps)) =
+    Inference (name, map_term_names_in_formula f phi, deps)
+fun map_term_names_in_atp_proof f = map (map_term_names_in_step f)
+
+fun nasty_name pool s = s |> Symtab.lookup pool |> the_default s
+fun nasty_atp_proof pool =
+  if Symtab.is_empty pool then I
+  else map_term_names_in_atp_proof (nasty_name pool)
+
+end;
--- a/src/HOL/Tools/Nitpick/kodkod.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Tools/Nitpick/kodkod.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -2,7 +2,7 @@
     Author:     Jasmin Blanchette, TU Muenchen
     Copyright   2008, 2009, 2010
 
-ML interface on top of Kodkod.
+ML interface for Kodkod.
 *)
 
 signature KODKOD =
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -1019,7 +1019,7 @@
 fun kodkod_formula_from_nut ofs
         (kk as {kk_all, kk_exist, kk_formula_let, kk_formula_if, kk_or, kk_not,
                 kk_iff, kk_implies, kk_and, kk_subset, kk_rel_eq, kk_no,
-                kk_lone, kk_one, kk_some, kk_rel_let, kk_rel_if, kk_union,
+                kk_lone, kk_some, kk_rel_let, kk_rel_if, kk_union,
                 kk_difference, kk_intersect, kk_product, kk_join, kk_closure,
                 kk_comprehension, kk_project, kk_project_seq, kk_not3,
                 kk_nat_less, kk_int_less, ...}) u =
--- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -736,8 +736,24 @@
           in lmap |> fold (add_thm false) helper_ths end
   in (mode, add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap) end
 
+val clause_params =
+  {ordering = Metis_KnuthBendixOrder.default,
+   orderLiterals = Metis_Clause.UnsignedLiteralOrder,
+   orderTerms = true}
+val active_params =
+  {clause = clause_params,
+   prefactor = #prefactor Metis_Active.default,
+   postfactor = #postfactor Metis_Active.default}
+val waiting_params =
+  {symbolsWeight = 1.0,
+   variablesWeight = 0.0,
+   literalsWeight = 0.0,
+   models = []}
+val refute_params = {active = active_params, waiting = waiting_params}
+
 fun refute cls =
-    Metis_Resolution.loop (Metis_Resolution.new Metis_Resolution.default {axioms = cls, conjecture = []});
+  Metis_Resolution.new refute_params {axioms = cls, conjecture = []}
+  |> Metis_Resolution.loop
 
 fun is_false t = t aconv (HOLogic.mk_Trueprop HOLogic.false_const);
 
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -40,7 +40,7 @@
      used_thm_names: (string * locality) list,
      atp_run_time_in_msecs: int,
      output: string,
-     proof: string,
+     tstplike_proof: string,
      axiom_names: (string * locality) list vector,
      conjecture_shape: int list list}
   type prover = params -> minimize_command -> problem -> prover_result
@@ -111,7 +111,7 @@
    used_thm_names: (string * locality) list,
    atp_run_time_in_msecs: int,
    output: string,
-   proof: string,
+   tstplike_proof: string,
    axiom_names: (string * locality) list vector,
    conjecture_shape: int list list}
 
@@ -153,23 +153,23 @@
            |> space_implode "\n " |> quote
 
 (* Splits by the first possible of a list of delimiters. *)
-fun extract_proof delims output =
+fun extract_tstplike_proof delims output =
   case pairself (find_first (fn s => String.isSubstring s output))
                 (ListPair.unzip delims) of
     (SOME begin_delim, SOME end_delim) =>
     extract_delimited (begin_delim, end_delim) output
   | _ => ""
 
-fun extract_proof_and_outcome complete res_code proof_delims known_failures
-                              output =
+fun extract_tstplike_proof_and_outcome complete res_code proof_delims
+                                       known_failures output =
   case known_failure_in_output output known_failures of
-    NONE => (case extract_proof proof_delims output of
+    NONE => (case extract_tstplike_proof proof_delims output of
              "" => ("", SOME (if res_code = 0 andalso output = "" then
                                 Interrupted
                               else
                                  UnknownError))
-           | proof => if res_code = 0 then (proof, NONE)
-                      else ("", SOME UnknownError))
+           | tstplike_proof => if res_code = 0 then (tstplike_proof, NONE)
+                               else ("", SOME UnknownError))
   | SOME failure =>
     ("", SOME (if failure = IncompleteUnprovable andalso complete then
                  Unprovable
@@ -297,16 +297,16 @@
                        else
                          I)
                   |>> (if measure_run_time then split_time else rpair 0)
-                val (proof, outcome) =
-                  extract_proof_and_outcome complete res_code proof_delims
-                                            known_failures output
-              in (output, msecs, proof, outcome) end
+                val (tstplike_proof, outcome) =
+                  extract_tstplike_proof_and_outcome complete res_code
+                      proof_delims known_failures output
+              in (output, msecs, tstplike_proof, outcome) end
             val readable_names = debug andalso overlord
             val (problem, pool, conjecture_offset, axiom_names) =
               prepare_problem ctxt readable_names explicit_forall full_types
                               explicit_apply hyp_ts concl_t axioms
-            val ss = strings_for_tptp_problem use_conjecture_for_hypotheses
-                                              problem
+            val ss = tptp_strings_for_atp_problem use_conjecture_for_hypotheses
+                                                  problem
             val _ = File.write_list probfile ss
             val conjecture_shape =
               conjecture_offset + 1 upto conjecture_offset + length hyp_ts + 1
@@ -322,8 +322,8 @@
               |> run_twice
                  ? (fn (_, msecs0, _, SOME _) =>
                        run true (Time.- (timeout, Timer.checkRealTimer timer))
-                       |> (fn (output, msecs, proof, outcome) =>
-                              (output, msecs0 + msecs, proof, outcome))
+                       |> (fn (output, msecs, tstplike_proof, outcome) =>
+                              (output, msecs0 + msecs, tstplike_proof, outcome))
                      | result => result)
           in ((pool, conjecture_shape, axiom_names), result) end
         else
@@ -339,7 +339,7 @@
       else
         File.write (Path.explode (Path.implode probfile ^ "_proof")) output
     val ((pool, conjecture_shape, axiom_names),
-         (output, msecs, proof, outcome)) =
+         (output, msecs, tstplike_proof, outcome)) =
       with_path cleanup export run_on problem_path_name
     val (conjecture_shape, axiom_names) =
       repair_conjecture_shape_and_theorem_names output conjecture_shape
@@ -352,8 +352,8 @@
         NONE =>
         proof_text isar_proof
             (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
-            (banner, full_types, minimize_command, proof, axiom_names, goal,
-             subgoal)
+            (banner, full_types, minimize_command, tstplike_proof, axiom_names,
+             goal, subgoal)
         |>> (fn message =>
                 message ^ (if verbose then
                              "\nATP real CPU time: " ^ string_of_int msecs ^
@@ -369,14 +369,14 @@
   in
     {outcome = outcome, message = message, pool = pool,
      used_thm_names = used_thm_names, atp_run_time_in_msecs = msecs,
-     output = output, proof = proof, axiom_names = axiom_names,
-     conjecture_shape = conjecture_shape}
+     output = output, tstplike_proof = tstplike_proof,
+     axiom_names = axiom_names, conjecture_shape = conjecture_shape}
   end
 
 fun get_prover_fun thy name = prover_fun false name (get_prover thy name)
 
-fun run_prover (params as {blocking, verbose, max_relevant, timeout, expect,
-                           ...})
+fun run_prover (params as {blocking, debug, verbose, max_relevant, timeout,
+                           expect, ...})
                auto i n minimize_command (problem as {state, goal, axioms, ...})
                (prover as {default_max_relevant, ...}, atp_name) =
   let
@@ -398,14 +398,19 @@
          "\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)))
     fun go () =
       let
-        val (outcome_code, message) =
+        fun really_go () =
           prover_fun auto atp_name prover params (minimize_command atp_name)
                      problem
           |> (fn {outcome, message, ...} =>
                  (if is_some outcome then "none" else "some", message))
-          handle ERROR message => ("unknown", "Error: " ^ message ^ "\n")
-               | exn => ("unknown", "Internal error:\n" ^
-                                    ML_Compiler.exn_message exn ^ "\n")
+        val (outcome_code, message) =
+          if debug then
+            really_go ()
+          else
+            (really_go ()
+             handle ERROR message => ("unknown", "Error: " ^ message ^ "\n")
+                  | exn => ("unknown", "Internal error:\n" ^
+                                       ML_Compiler.exn_message exn ^ "\n"))
         val _ =
           if expect = "" orelse outcome_code = expect then
             ()
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -117,7 +117,7 @@
          val new_timeout =
            Int.min (msecs, Time.toMilliseconds time + fudge_msecs)
            |> Time.fromMilliseconds
-         val (min_thms, {proof, axiom_names, ...}) =
+         val (min_thms, {tstplike_proof, axiom_names, ...}) =
            sublinear_minimize (do_test new_timeout)
                (filter_used_facts used_thm_names axioms) ([], result)
          val n = length min_thms
@@ -130,8 +130,8 @@
          (SOME min_thms,
           proof_text isar_proof
               (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
-              ("Minimized command", full_types, K "", proof, axiom_names, goal,
-               i) |> fst)
+              ("Minimized command", full_types, K "", tstplike_proof,
+               axiom_names, goal, i) |> fst)
        end
      | {outcome = SOME TimedOut, ...} =>
        (NONE, "Timeout: You can increase the time limit using the \"timeout\" \
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -26,6 +26,7 @@
 struct
 
 open ATP_Problem
+open ATP_Proof
 open Metis_Clauses
 open Sledgehammer_Util
 open Sledgehammer_Filter
@@ -39,6 +40,76 @@
   string Symtab.table * bool * int * Proof.context * int list list
 type text_result = string * (string * locality) list
 
+
+(** Soft-core proof reconstruction: Metis one-liner **)
+
+fun string_for_label (s, num) = s ^ string_of_int num
+
+fun metis_using [] = ""
+  | metis_using ls =
+    "using " ^ space_implode " " (map string_for_label ls) ^ " "
+fun metis_apply _ 1 = "by "
+  | metis_apply 1 _ = "apply "
+  | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply "
+fun metis_name full_types = if full_types then "metisFT" else "metis"
+fun metis_call full_types [] = metis_name full_types
+  | metis_call full_types ss =
+    "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")"
+fun metis_command full_types i n (ls, ss) =
+  metis_using ls ^ metis_apply i n ^ metis_call full_types ss
+fun metis_line banner full_types i n ss =
+  banner ^ ": " ^
+  Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "."
+fun minimize_line _ [] = ""
+  | minimize_line minimize_command ss =
+    case minimize_command ss of
+      "" => ""
+    | command =>
+      "\nTo minimize the number of lemmas, try this: " ^
+      Markup.markup Markup.sendback command ^ "."
+
+fun resolve_axiom axiom_names ((_, SOME s)) =
+    (case strip_prefix_and_unascii axiom_prefix s of
+       SOME s' => (case find_first_in_list_vector axiom_names s' of
+                     SOME x => [(s', x)]
+                   | NONE => [])
+     | NONE => [])
+  | resolve_axiom axiom_names (num, NONE) =
+    case Int.fromString num of
+      SOME j =>
+      if j > 0 andalso j <= Vector.length axiom_names then
+        Vector.sub (axiom_names, j - 1)
+      else
+        []
+    | NONE => []
+
+fun add_fact axiom_names (Inference (name, _, [])) =
+    append (resolve_axiom axiom_names name)
+  | add_fact _ _ = I
+
+fun used_facts_in_tstplike_proof axiom_names =
+  atp_proof_from_tstplike_string #> rpair [] #-> fold (add_fact axiom_names)
+
+fun used_facts axiom_names =
+  used_facts_in_tstplike_proof axiom_names
+  #> List.partition (curry (op =) Chained o snd)
+  #> pairself (sort_distinct (string_ord o pairself fst))
+
+fun metis_proof_text (banner, full_types, minimize_command,
+                      tstplike_proof, axiom_names, goal, i) =
+  let
+    val (chained_lemmas, other_lemmas) =
+      used_facts axiom_names tstplike_proof
+    val n = Logic.count_prems (prop_of goal)
+  in
+    (metis_line banner full_types i n (map fst other_lemmas) ^
+     minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
+     other_lemmas @ chained_lemmas)
+  end
+
+
+(** Hard-core proof reconstruction: structured Isar proofs **)
+
 (* Simple simplifications to ensure that sort annotations don't leave a trail of
    spurious "True"s. *)
 fun s_not @{const False} = @{const True}
@@ -58,61 +129,9 @@
   | s_iff (t1, @{const True}) = t1
   | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
 
-fun mk_anot (AConn (ANot, [phi])) = phi
-  | mk_anot phi = AConn (ANot, [phi])
-fun mk_aconn c (phi1, phi2) = AConn (c, [phi1, phi2])
-
 fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t
 fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda v t
 
-datatype raw_step_name = Str of string * string | Num of string
-
-fun raw_step_num (Str (num, _)) = num
-  | raw_step_num (Num num) = num
-
-fun same_raw_step s t = raw_step_num s = raw_step_num t
-
-fun raw_step_name_ord p =
-  let val q = pairself raw_step_num p in
-    (* The "unprefix" part is to cope with remote Vampire's output. The proper
-       solution would be to perform a topological sort, e.g. using the nice
-       "Graph" functor. *)
-    case pairself (Int.fromString o perhaps (try (unprefix "f"))) q of
-      (NONE, NONE) => string_ord q
-    | (NONE, SOME _) => LESS
-    | (SOME _, NONE) => GREATER
-    | (SOME i, SOME j) => int_ord (i, j)
-  end
-
-fun index_in_shape x = find_index (exists (curry (op =) x))
-fun resolve_axiom axiom_names (Str (_, s)) =
-    (case strip_prefix_and_unascii axiom_prefix s of
-       SOME s' => (case find_first_in_list_vector axiom_names s' of
-                     SOME x => [(s', x)]
-                   | NONE => [])
-     | NONE => [])
-  | resolve_axiom axiom_names (Num num) =
-    case Int.fromString num of
-      SOME j =>
-      if j > 0 andalso j <= Vector.length axiom_names then
-        Vector.sub (axiom_names, j - 1)
-      else
-        []
-    | NONE => []
-val is_axiom = not o null oo resolve_axiom
-
-fun resolve_conjecture conjecture_shape (Str (num, s)) =
-    let
-      val k = case try (unprefix conjecture_prefix) s of
-                SOME s => Int.fromString s |> the_default ~1
-              | NONE => case Int.fromString num of
-                          SOME j => index_in_shape j conjecture_shape
-                        | NONE => ~1
-    in if k >= 0 then [k] else [] end
-  | resolve_conjecture conjecture_shape (Num num) =
-    resolve_conjecture conjecture_shape (Str (num, "")) (* HACK *)
-val is_conjecture = not o null oo resolve_conjecture
-
 fun negate_term (Const (@{const_name All}, T) $ Abs (s, T', t')) =
     Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t')
   | negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
@@ -126,174 +145,32 @@
   | negate_term (@{const Not} $ t) = t
   | negate_term t = @{const Not} $ t
 
-datatype ('a, 'b, 'c) raw_step =
-  Definition of raw_step_name * 'a * 'b |
-  Inference of raw_step_name * 'c * raw_step_name list
-
-(**** PARSING OF TSTP FORMAT ****)
-
-(*Strings enclosed in single quotes, e.g. filenames*)
-val scan_general_id =
-  $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode
-  || Scan.repeat ($$ "$") -- Scan.many1 Symbol.is_letdig
-     >> (fn (ss1, ss2) => implode ss1 ^ implode ss2)
-
-fun repair_name _ "$true" = "c_True"
-  | repair_name _ "$false" = "c_False"
-  | repair_name _ "$$e" = "c_equal" (* seen in Vampire proofs *)
-  | repair_name _ "equal" = "c_equal" (* needed by SPASS? *)
-  | repair_name pool s =
-    case Symtab.lookup pool s of
-      SOME s' => s'
-    | NONE =>
-      if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then
-        "c_equal" (* seen in Vampire proofs *)
-      else
-        s
-(* Generalized first-order terms, which include file names, numbers, etc. *)
-fun parse_annotation strict x =
-  ((scan_general_id ::: Scan.repeat ($$ " " |-- scan_general_id)
-      >> (strict ? filter (is_some o Int.fromString)))
-   -- Scan.optional (parse_annotation strict) [] >> op @
-   || $$ "(" |-- parse_annotations strict --| $$ ")"
-   || $$ "[" |-- parse_annotations strict --| $$ "]") x
-and parse_annotations strict x =
-  (Scan.optional (parse_annotation strict
-                  ::: Scan.repeat ($$ "," |-- parse_annotation strict)) []
-   >> flat) x
+val indent_size = 2
+val no_label = ("", ~1)
 
-(* Vampire proof lines sometimes contain needless information such as "(0:3)",
-   which can be hard to disambiguate from function application in an LL(1)
-   parser. As a workaround, we extend the TPTP term syntax with such detritus
-   and ignore it. *)
-fun parse_vampire_detritus x =
-  (scan_general_id |-- $$ ":" --| scan_general_id >> K []) x
-
-fun parse_term pool x =
-  ((scan_general_id >> repair_name pool)
-    -- Scan.optional ($$ "(" |-- (parse_vampire_detritus || parse_terms pool)
-                      --| $$ ")") []
-    --| Scan.optional ($$ "(" |-- parse_vampire_detritus --| $$ ")") []
-   >> ATerm) x
-and parse_terms pool x =
-  (parse_term pool ::: Scan.repeat ($$ "," |-- parse_term pool)) x
-
-fun parse_atom pool =
-  parse_term pool -- Scan.option (Scan.option ($$ "!") --| $$ "="
-                                  -- parse_term pool)
-  >> (fn (u1, NONE) => AAtom u1
-       | (u1, SOME (NONE, u2)) => AAtom (ATerm ("c_equal", [u1, u2]))
-       | (u1, SOME (SOME _, u2)) =>
-         mk_anot (AAtom (ATerm ("c_equal", [u1, u2]))))
-
-fun fo_term_head (ATerm (s, _)) = s
-
-(* TPTP formulas are fully parenthesized, so we don't need to worry about
-   operator precedence. *)
-fun parse_formula pool x =
-  (($$ "(" |-- parse_formula pool --| $$ ")"
-    || ($$ "!" >> K AForall || $$ "?" >> K AExists)
-       --| $$ "[" -- parse_terms pool --| $$ "]" --| $$ ":"
-       -- parse_formula pool
-       >> (fn ((q, ts), phi) => AQuant (q, map fo_term_head ts, phi))
-    || $$ "~" |-- parse_formula pool >> mk_anot
-    || parse_atom pool)
-   -- Scan.option ((Scan.this_string "=>" >> K AImplies
-                    || Scan.this_string "<=>" >> K AIff
-                    || Scan.this_string "<~>" >> K ANotIff
-                    || Scan.this_string "<=" >> K AIf
-                    || $$ "|" >> K AOr || $$ "&" >> K AAnd)
-                   -- parse_formula pool)
-   >> (fn (phi1, NONE) => phi1
-        | (phi1, SOME (c, phi2)) => mk_aconn c (phi1, phi2))) x
+val raw_prefix = "X"
+val assum_prefix = "A"
+val fact_prefix = "F"
 
-val parse_tstp_extra_arguments =
-  Scan.optional ($$ "," |-- parse_annotation false
-                 --| Scan.option ($$ "," |-- parse_annotations false)) []
-
-(* Syntax: (fof|cnf)\(<num>, <formula_role>, <formula> <extra_arguments>\).
-   The <num> could be an identifier, but we assume integers. *)
- fun parse_tstp_line pool =
-   ((Scan.this_string "fof" || Scan.this_string "cnf") -- $$ "(")
-     |-- scan_general_id --| $$ "," -- Symbol.scan_id --| $$ ","
-     -- parse_formula pool -- parse_tstp_extra_arguments --| $$ ")" --| $$ "."
-    >> (fn (((num, role), phi), deps) =>
-           let
-             val (name, deps) =
-               case deps of
-                 ["file", _, s] => (Str (num, s), [])
-               | _ => (Num num, deps)
-           in
-             case role of
-               "definition" =>
-               (case phi of
-                  AConn (AIff, [phi1 as AAtom _, phi2]) =>
-                  Definition (name, phi1, phi2)
-                | AAtom (ATerm ("c_equal", _)) =>
-                  (* Vampire's equality proxy axiom *)
-                  Inference (name, phi, map Num deps)
-                | _ => raise Fail "malformed definition")
-             | _ => Inference (name, phi, map Num deps)
-           end)
-
-(**** PARSING OF VAMPIRE OUTPUT ****)
-
-(* Syntax: <num>. <formula> <annotation> *)
-fun parse_vampire_line pool =
-  scan_general_id --| $$ "." -- parse_formula pool -- parse_annotation true
-  >> (fn ((num, phi), deps) => Inference (Num num, phi, map Num deps))
-
-(**** PARSING OF SPASS OUTPUT ****)
-
-(* SPASS returns clause references of the form "x.y". We ignore "y", whose role
-   is not clear anyway. *)
-val parse_dot_name = scan_general_id --| $$ "." --| scan_general_id
+fun resolve_conjecture conjecture_shape (num, s_opt) =
+  let
+    val k = case try (unprefix conjecture_prefix) (the_default "" s_opt) of
+              SOME s => Int.fromString s |> the_default ~1
+            | NONE => case Int.fromString num of
+                        SOME j => find_index (exists (curry (op =) j))
+                                             conjecture_shape
+                      | NONE => ~1
+  in if k >= 0 then [k] else [] end
 
-val parse_spass_annotations =
-  Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name
-                                         --| Scan.option ($$ ","))) []
-
-(* It is not clear why some literals are followed by sequences of stars and/or
-   pluses. We ignore them. *)
-fun parse_decorated_atom pool =
-  parse_atom pool --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")
-
-fun mk_horn ([], []) = AAtom (ATerm ("c_False", []))
-  | mk_horn ([], pos_lits) = foldr1 (mk_aconn AOr) pos_lits
-  | mk_horn (neg_lits, []) = mk_anot (foldr1 (mk_aconn AAnd) neg_lits)
-  | mk_horn (neg_lits, pos_lits) =
-    mk_aconn AImplies (foldr1 (mk_aconn AAnd) neg_lits,
-                       foldr1 (mk_aconn AOr) pos_lits)
-
-fun parse_horn_clause pool =
-  Scan.repeat (parse_decorated_atom pool) --| $$ "|" --| $$ "|"
-    -- Scan.repeat (parse_decorated_atom pool) --| $$ "-" --| $$ ">"
-    -- Scan.repeat (parse_decorated_atom pool)
-  >> (mk_horn o apfst (op @))
+val is_axiom = not o null oo resolve_axiom
+val is_conjecture = not o null oo resolve_conjecture
 
-(* Syntax: <num>[0:<inference><annotations>]
-   <atoms> || <atoms> -> <atoms>. *)
-fun parse_spass_line pool =
-  scan_general_id --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id
-    -- parse_spass_annotations --| $$ "]" -- parse_horn_clause pool --| $$ "."
-  >> (fn ((num, deps), u) => Inference (Num num, u, map Num deps))
-
-fun parse_line pool =
-  parse_tstp_line pool || parse_vampire_line pool || parse_spass_line pool
-fun parse_lines pool = Scan.repeat1 (parse_line pool)
-fun parse_proof pool =
-  fst o Scan.finite Symbol.stopper
-            (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output")
-                            (parse_lines pool)))
-  o explode o strip_spaces_except_between_ident_chars
-
-fun clean_up_dependencies _ [] = []
-  | clean_up_dependencies seen ((step as Definition (name, _, _)) :: steps) =
-    step :: clean_up_dependencies (name :: seen) steps
-  | clean_up_dependencies seen (Inference (name, u, deps) :: steps) =
-    Inference (name, u,
-             map_filter (fn dep => find_first (same_raw_step dep) seen) deps) ::
-    clean_up_dependencies (name :: seen) steps
+fun raw_label_for_name conjecture_shape name =
+  case resolve_conjecture conjecture_shape name of
+    [j] => (conjecture_prefix, j)
+  | _ => case Int.fromString (fst name) of
+           SOME j => (raw_prefix, j)
+         | NONE => (raw_prefix ^ fst name, 0)
 
 (**** INTERPRETATION OF TSTP SYNTAX TREES ****)
 
@@ -411,7 +288,7 @@
                 case strip_prefix_and_unascii schematic_var_prefix a of
                   SOME b => Var ((b, 0), T)
                 | NONE =>
-                  if is_tptp_variable a then
+                  if is_atp_variable a then
                     Var ((repair_atp_variable_name Char.toLower a, 0), T)
                   else
                     (* Skolem constants? *)
@@ -537,7 +414,7 @@
 val is_only_type_information = curry (op aconv) HOLogic.true_const
 
 fun replace_one_dependency (old, new) dep =
-  if raw_step_num dep = raw_step_num old then new else [dep]
+  if is_same_step (dep, old) then new else [dep]
 fun replace_dependencies_in_line _ (line as Definition _) = line
   | replace_dependencies_in_line p (Inference (name, t, deps)) =
     Inference (name, t, fold (union (op =) o replace_one_dependency p) deps [])
@@ -607,104 +484,6 @@
      else
        map (replace_dependencies_in_line (name, deps)) lines)  (* drop line *)
 
-(** EXTRACTING LEMMAS **)
-
-(* Like "split_line" but splits on ".\n" (for TSTP and SPASS) or "]\n" (for
-   Vampire). *)
-val split_proof_lines =
-  let
-    fun aux [] [] = []
-      | aux line [] = [implode (rev line)]
-      | aux line ("." :: "\n" :: rest) = aux line [] @ aux [] rest
-      | aux line ("]" :: "\n" :: rest) = aux line [] @ aux [] rest
-      | aux line (s :: rest) = aux (s :: line) rest
-  in aux [] o explode end
-
-(* A list consisting of the first number in each line is returned. For TSTP,
-   interesting lines have the form "fof(108, axiom, ...)", where the number
-   (108) is extracted. For SPASS, lines have the form "108[0:Inp] ...", where
-   the first number (108) is extracted. For Vampire, we look for
-   "108. ... [input]". *)
-fun used_facts_in_atp_proof axiom_names atp_proof =
-  let
-    val tokens_of =
-      String.tokens (fn c => not (Char.isAlphaNum c) andalso c <> #"_")
-    fun do_line (tag :: num :: "axiom" :: (rest as _ :: _)) =
-        if tag = "cnf" orelse tag = "fof" then
-          (case strip_prefix_and_unascii axiom_prefix (List.last rest) of
-             SOME name =>
-             if member (op =) rest "file" then
-               ([(name, name |> find_first_in_list_vector axiom_names |> the)]
-                handle Option.Option =>
-                       error ("No such fact: " ^ quote name ^ "."))
-             else
-               resolve_axiom axiom_names (Num num)
-           | NONE => resolve_axiom axiom_names (Num num))
-        else
-          []
-      | do_line (num :: "0" :: "Inp" :: _) = resolve_axiom axiom_names (Num num)
-      | do_line (num :: rest) =
-        (case List.last rest of
-           "input" => resolve_axiom axiom_names (Num num)
-         | _ => [])
-      | do_line _ = []
-  in atp_proof |> split_proof_lines |> maps (do_line o tokens_of) end
-
-val indent_size = 2
-val no_label = ("", ~1)
-
-val raw_prefix = "X"
-val assum_prefix = "A"
-val fact_prefix = "F"
-
-fun string_for_label (s, num) = s ^ string_of_int num
-
-fun raw_label_for_name conjecture_shape name =
-  case resolve_conjecture conjecture_shape name of
-    [j] => (conjecture_prefix, j)
-  | _ => case Int.fromString (raw_step_num name) of
-           SOME j => (raw_prefix, j)
-         | NONE => (raw_prefix ^ raw_step_num name, 0)
-
-fun metis_using [] = ""
-  | metis_using ls =
-    "using " ^ space_implode " " (map string_for_label ls) ^ " "
-fun metis_apply _ 1 = "by "
-  | metis_apply 1 _ = "apply "
-  | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply "
-fun metis_name full_types = if full_types then "metisFT" else "metis"
-fun metis_call full_types [] = metis_name full_types
-  | metis_call full_types ss =
-    "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")"
-fun metis_command full_types i n (ls, ss) =
-  metis_using ls ^ metis_apply i n ^ metis_call full_types ss
-fun metis_line banner full_types i n ss =
-  banner ^ ": " ^
-  Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "."
-fun minimize_line _ [] = ""
-  | minimize_line minimize_command ss =
-    case minimize_command ss of
-      "" => ""
-    | command =>
-      "\nTo minimize the number of lemmas, try this: " ^
-      Markup.markup Markup.sendback command ^ "."
-
-fun used_facts axiom_names =
-  used_facts_in_atp_proof axiom_names
-  #> List.partition (curry (op =) Chained o snd)
-  #> pairself (sort_distinct (string_ord o pairself fst))
-
-fun metis_proof_text (banner, full_types, minimize_command, atp_proof,
-                      axiom_names, goal, i) =
-  let
-    val (chained_lemmas, other_lemmas) = used_facts axiom_names atp_proof
-    val n = Logic.count_prems (prop_of goal)
-  in
-    (metis_line banner full_types i n (map fst other_lemmas) ^
-     minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
-     other_lemmas @ chained_lemmas)
-  end
-
 (** Isar proof construction and manipulation **)
 
 fun merge_fact_sets (ls1, ss1) (ls2, ss2) =
@@ -713,16 +492,16 @@
 type label = string * int
 type facts = label list * string list
 
-datatype qualifier = Show | Then | Moreover | Ultimately
+datatype isar_qualifier = Show | Then | Moreover | Ultimately
 
-datatype step =
+datatype isar_step =
   Fix of (string * typ) list |
   Let of term * term |
   Assume of label * term |
-  Have of qualifier list * label * term * byline
+  Have of isar_qualifier list * label * term * byline
 and byline =
   ByMetis of facts |
-  CaseSplit of step list list * facts
+  CaseSplit of isar_step list list * facts
 
 fun smart_case_split [] facts = ByMetis facts
   | smart_case_split proofs facts = CaseSplit (proofs, facts)
@@ -743,17 +522,24 @@
           ByMetis (fold (add_fact_from_dependency conjecture_shape axiom_names)
                         deps ([], [])))
 
-fun raw_step_name (Definition (name, _, _)) = name
-  | raw_step_name (Inference (name, _, _)) = name
+fun repair_name "$true" = "c_True"
+  | repair_name "$false" = "c_False"
+  | repair_name "$$e" = "c_equal" (* seen in Vampire proofs *)
+  | repair_name "equal" = "c_equal" (* needed by SPASS? *)
+  | repair_name s =
+    if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then
+      "c_equal" (* seen in Vampire proofs *)
+    else
+      s
 
-fun proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
-                         atp_proof conjecture_shape axiom_names params frees =
+fun isar_proof_from_tstplike_proof pool ctxt full_types tfrees isar_shrink_factor
+        tstplike_proof conjecture_shape axiom_names params frees =
   let
     val lines =
-      atp_proof ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
-      |> parse_proof pool
-      |> sort (raw_step_name_ord o pairself raw_step_name)
-      |> clean_up_dependencies []
+      tstplike_proof
+      |> atp_proof_from_tstplike_string
+      |> nasty_atp_proof pool
+      |> map_term_names_in_atp_proof repair_name
       |> decode_lines ctxt full_types tfrees
       |> rpair [] |-> fold_rev (add_line conjecture_shape axiom_names)
       |> rpair [] |-> fold_rev add_nontrivial_line
@@ -1049,14 +835,13 @@
     and do_proof [Have (_, _, _, ByMetis _)] = ""
       | do_proof proof =
         (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
-        do_indent 0 ^ "proof -\n" ^
-        do_steps "" "" 1 proof ^
-        do_indent 0 ^ (if n <> 1 then "next" else "qed")
+        do_indent 0 ^ "proof -\n" ^ do_steps "" "" 1 proof ^ do_indent 0 ^
+        (if n <> 1 then "next" else "qed")
   in do_proof end
 
 fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
-                    (other_params as (_, full_types, _, atp_proof, axiom_names,
-                                      goal, i)) =
+                    (other_params as (_, full_types, _, tstplike_proof,
+                                      axiom_names, goal, i)) =
   let
     val (params, hyp_ts, concl_t) = strip_subgoal goal i
     val frees = fold Term.add_frees (concl_t :: hyp_ts) []
@@ -1064,9 +849,9 @@
     val n = Logic.count_prems (prop_of goal)
     val (one_line_proof, lemma_names) = metis_proof_text other_params
     fun isar_proof_for () =
-      case proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
-                                atp_proof conjecture_shape axiom_names params
-                                frees
+      case isar_proof_from_tstplike_proof pool ctxt full_types tfrees
+               isar_shrink_factor tstplike_proof conjecture_shape axiom_names
+               params frees
            |> redirect_proof hyp_ts concl_t
            |> kill_duplicate_assumptions_in_proof
            |> then_chain_proof
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -376,7 +376,7 @@
 type const_info = {min_arity: int, max_arity: int, sub_level: bool}
 
 fun consider_term top_level (ATerm ((s, _), ts)) =
-  (if is_tptp_variable s then
+  (if is_atp_variable s then
      I
    else
      let val n = length ts in
@@ -468,7 +468,7 @@
 fun close_universally phi =
   let
     fun term_vars bounds (ATerm (name as (s, _), tms)) =
-        (is_tptp_variable s andalso not (member (op =) bounds name))
+        (is_atp_variable s andalso not (member (op =) bounds name))
           ? insert (op =) name
         #> fold (term_vars bounds) tms
     fun formula_vars bounds (AQuant (_, xs, phi)) =
@@ -524,7 +524,7 @@
        ("Conjectures", conjecture_lines),
        ("Type variables", tfree_lines)]
       |> repair_problem thy explicit_forall full_types explicit_apply
-    val (problem, pool) = nice_tptp_problem readable_names problem
+    val (problem, pool) = nice_atp_problem readable_names problem
     val conjecture_offset =
       length axiom_lines + length class_rel_lines + length arity_lines
       + length helper_lines
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -10,7 +10,6 @@
   val plural_s : int -> string
   val serial_commas : string -> string list -> string list
   val simplify_spaces : string -> string
-  val strip_spaces_except_between_ident_chars : string -> string
   val parse_bool_option : bool -> string -> string -> bool option
   val parse_time_option : string -> string -> Time.time option
   val scan_integer : string list -> int * string list
@@ -41,28 +40,7 @@
   | serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3]
   | serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss
 
-fun strip_spaces_in_list _ [] = []
-  | strip_spaces_in_list _ [c1] = if Char.isSpace c1 then [] else [str c1]
-  | strip_spaces_in_list is_evil [c1, c2] =
-    strip_spaces_in_list is_evil [c1] @ strip_spaces_in_list is_evil [c2]
-  | strip_spaces_in_list is_evil (c1 :: c2 :: c3 :: cs) =
-    if Char.isSpace c1 then
-      strip_spaces_in_list is_evil (c2 :: c3 :: cs)
-    else if Char.isSpace c2 then
-      if Char.isSpace c3 then
-        strip_spaces_in_list is_evil (c1 :: c3 :: cs)
-      else
-        str c1 :: (if forall is_evil [c1, c3] then [" "] else []) @
-        strip_spaces_in_list is_evil (c3 :: cs)
-    else
-      str c1 :: strip_spaces_in_list is_evil (c2 :: c3 :: cs)
-fun strip_spaces is_evil =
-  implode o strip_spaces_in_list is_evil o String.explode
-
-val simplify_spaces = strip_spaces (K true)
-
-fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
-val strip_spaces_except_between_ident_chars = strip_spaces is_ident_char
+val simplify_spaces = ATP_Proof.strip_spaces (K true)
 
 fun parse_bool_option option name s =
   (case s of
--- a/src/Tools/Metis/Makefile	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/Makefile	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 ###############################################################################
 # METIS MAKEFILE
-# Copyright (c) 2001 Joe Hurd, distributed under the GNU GPL version 2
+# Copyright (c) 2001 Joe Hurd, distributed under the BSD License
 ###############################################################################
 
 .SUFFIXES:
--- a/src/Tools/Metis/Makefile.FILES	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/Makefile.FILES	Thu Sep 16 16:20:20 2010 +0200
@@ -4,7 +4,7 @@
 	echo > $@
 refresh_FILES:
 	echo $(POLYML_SRC) | \
-	sed "s/src\///g" | \
-	sed "s/ Tptp\.s[a-z][a-z]//g" | \
-	sed "s/ Options\.s[a-z][a-z]//g" \
+	sed "s/src\/PortablePolyml/PortableIsabelle/g" | \
+	sed "s/ src\/Tptp\.s[a-z][a-z]//g" | \
+	sed "s/ src\/Options\.s[a-z][a-z]//g" \
 	> FILES
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Metis/PortableIsabelle.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -0,0 +1,26 @@
+(*  Title:      Tools/Metis/PortableIsabelle.sml
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2010
+
+Isabelle-specific setup for Metis. Based on "src/PortablePolyml.sml".
+*)
+
+structure Portable :> Portable =
+struct
+
+val ml = "isabelle"
+
+fun pointerEqual (x : 'a, y : 'a) = pointer_eq (x, y)
+
+fun critical e () = NAMED_CRITICAL "metis" e
+
+val randomWord = Random.nextWord
+val randomBool = Random.nextBool
+val randomInt = Random.nextInt
+val randomReal = Random.nextReal
+
+fun time f x = f x (* dummy *)
+
+end
+
+datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a
--- a/src/Tools/Metis/README	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/README	Thu Sep 16 16:20:20 2010 +0200
@@ -5,57 +5,43 @@
 unfortunately somewhat involved and frustrating, and hopefully
 temporary.
 
- 1. The file "Makefile" and the directory "src/" and "script/" were
-    initially copied from Joe Hurd's Metis package. (His "script/"
-    directory has many files in it, but we only need "mlpp".) The
-    package that was used when these notes where written was an
-    unnumbered version of Metis, more recent than 2.2 and dated 19
-    July 2010.
+ 1. The files "Makefile" and "script/mlpp" and the directory "src/"
+    were initially copied from Joe Hurd's official Metis package. The
+    package that was used when these notes where written was Metis 2.3
+    (15 Sept. 2010).
 
  2. The license in each source file will probably not be something we
-    can use in Isabelle. Lawrence C. Paulson's command
-
-        perl -p -i~ -w -e 's/MIT license/BSD License/g' *sig *sml
-
-    run in the "src/" directory should do the trick. In a 13 Sept.
-    2010 email to Gerwin Klein, Joe Hurd, the sole copyright holder of
-    Metis, wrote:
+    can use in Isabelle. The "fix_metis_license" script can be run to
+    replace all occurrences of "MIT license" with "BSD License". In a
+    13 Sept. 2010 email to Gerwin Klein, Joe Hurd, the sole copyright
+    holder of Metis, wrote:
 
         I hereby give permission to the Isabelle team to release Metis
         as part of Isabelle, with the Metis code covered under the
         Isabelle BSD license.
 
  3. Some modifications have to be done manually to the source files.
-    Some of these are necessary so that the sources compile at all
-    with Isabelle, some are necessary to avoid race conditions in a
-    multithreaded environment, and some affect the defaults of Metis's
-    heuristics and are needed for backward compatibility with older
-    Isabelle proofs and for performance reasons. These changes should
-    be identified by a comment
-
-        (* MODIFIED by ?Joe ?Blow *)
-
-    but the ultimate way to track them down is to use Mercurial. The
+    The ultimate way to track them down is to use Mercurial. The
     command
 
-        hg diff -r2d0a4361c3ef: src
+        hg diff -rbeabb8443ee4: src
 
-    should do the trick. (You might need to specify a different
+    should do the trick. You might need to specify a different
     revision number if somebody updated the Metis sources without
-    updating these instructions.)
+    updating these instructions; consult the history in case of doubt.
 
- 4. Isabelle itself only cares about "metis.ML", which is generated
+ 4. Isabelle itself cares only about "metis.ML", which is generated
     from the files in "src/" by the script "make_metis". The script
-    relies on "Makefile", "src/", and "scripts/", as well as a special
-    file "Makefile.FILES" used to determine all the files needed to be
-    included in "metis.ML".
+    relies on "Makefile", "scripts/mlpp", and "src/", as well as
+    the pseudo-makefile "Makefile.FILES" used to determine all the
+    files needed to be included in "metis.ML".
 
  5. The output of "make_metis" should then work as is with Isabelle,
-    but there are of course no guarantees. The script "make_metis" has
-    a few evil regex hacks that could go wrong. It also produces a
-    few temporary files ("FILES", "Makefile.dev", and
-    "bin/mosml/Makefile.src") as side-effects; you can safely ignore
-    them or delete them.
+    but there are of course no guarantees. The script "make_metis" and
+    the pseudo-makefile "Makefile.FILES" have a few regexes that could
+    go wrong. They also produce a few temporary files ("FILES",
+    "Makefile.dev", and "bin/mosml/Makefile.src") as side-effects; you
+    can safely ignore them or delete them.
 
  6. Once you have successfully regenerated "metis.ML", build all of
     Isabelle and keep an eye on the time taken by Metis.
@@ -69,4 +55,4 @@
 Good luck!
 
     Jasmin Blanchette
-    15 Sept. 2010
+    16 Sept. 2010
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Metis/fix_metis_license	Thu Sep 16 16:20:20 2010 +0200
@@ -0,0 +1,4 @@
+#!/usr/bin/env bash
+THIS=$(cd "$(dirname "$0")"; echo $PWD)
+(cd $THIS;
+ perl -p -i~ -w -e 's/MIT license/BSD License/g' Makefile src/*.s* scripts/mlpp)
--- a/src/Tools/Metis/make_metis	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/make_metis	Thu Sep 16 16:20:20 2010 +0200
@@ -7,8 +7,8 @@
 # compile within Isabelle on Poly/ML and SML/NJ.
 
 THIS=$(cd "$(dirname "$0")"; echo $PWD)
-
 make -f Makefile.FILES refresh_FILES
+FILES=$(cat "$THIS/FILES")
 
 (
   cat <<EOF
@@ -32,19 +32,17 @@
 
 EOF
 
-  for FILE in $(cat "$THIS/FILES")
+  for FILE in $FILES
   do
     echo
     echo "(**** Original file: $FILE ****)"
     echo
     echo -e "$FILE" >&2
-    "$THIS/scripts/mlpp" -c polyml "src/$FILE" | \
+    "$THIS/scripts/mlpp" -c polyml "$FILE" | \
     perl -p -e \
 's/type name$/type name = string/;'\
 's/\bref\b/Unsynchronized.ref/g;'\
-'s/\bPolyML.pointerEq\b/pointer_eq/g;'\
-'s/\bRL\b/Metis_RL/g;'\
-"`grep "^\(signature\|structure\|functor\)" src/*.{sig,sml} | \
+"`grep "^\(signature\|structure\|functor\)" $FILES | \
   sed "s/[^:]*:[a-z]* \([A-Za-z0-9_]*\).*/\1/" | \
   tr " " "\n" | \
   sort | \
--- a/src/Tools/Metis/metis.ML	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/metis.ML	Thu Sep 16 16:20:20 2010 +0200
@@ -17,7 +17,7 @@
 print_depth 0;
 
 
-(**** Original file: Random.sig ****)
+(**** Original file: src/Random.sig ****)
 
 (*  Title:      Tools/random_word.ML
     Author:     Makarius
@@ -29,6 +29,7 @@
 
 signature Metis_Random =
 sig
+
   val nextWord : unit -> word
 
   val nextBool : unit -> bool
@@ -39,7 +40,7 @@
 
 end;
 
-(**** Original file: Random.sml ****)
+(**** Original file: src/Random.sml ****)
 
 (*  Title:      Tools/random_word.ML
     Author:     Makarius
@@ -90,11 +91,11 @@
 
 end;
 
-(**** Original file: Portable.sig ****)
-
-(* ========================================================================= *)
-(* ML SPECIFIC FUNCTIONS                                                     *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
+(**** Original file: src/Portable.sig ****)
+
+(* ========================================================================= *)
+(* ML COMPILER SPECIFIC FUNCTIONS                                            *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Portable =
@@ -113,17 +114,10 @@
 val pointerEqual : 'a * 'a -> bool
 
 (* ------------------------------------------------------------------------- *)
-(* Timing function applications.                                             *)
-(* ------------------------------------------------------------------------- *)
-
-val time : ('a -> 'b) -> 'a -> 'b
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing)                                 *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-val CRITICAL: (unit -> 'a) -> 'a
+(* Marking critical sections of code.                                        *)
+(* ------------------------------------------------------------------------- *)
+
+val critical : (unit -> 'a) -> unit -> 'a
 
 (* ------------------------------------------------------------------------- *)
 (* Generating random values.                                                 *)
@@ -137,97 +131,44 @@
 
 val randomWord : unit -> Word.word
 
-end
-
-(**** Original file: PortablePolyml.sml ****)
-
-(* ========================================================================= *)
-(* POLY/ML SPECIFIC FUNCTIONS                                                *)
-(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License            *)
-(* ========================================================================= *)
+(* ------------------------------------------------------------------------- *)
+(* Timing function applications.                                             *)
+(* ------------------------------------------------------------------------- *)
+
+val time : ('a -> 'b) -> 'a -> 'b
+
+end
+
+(**** Original file: PortableIsabelle.sml ****)
+
+(*  Title:      Tools/Metis/PortableIsabelle.sml
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2010
+
+Isabelle-specific setup for Metis. Based on "src/PortablePolyml.sml".
+*)
 
 structure Metis_Portable :> Metis_Portable =
 struct
 
-(* ------------------------------------------------------------------------- *)
-(* The ML implementation.                                                    *)
-(* ------------------------------------------------------------------------- *)
-
-val ml = "polyml";
-
-(* ------------------------------------------------------------------------- *)
-(* Pointer equality using the run-time system.                               *)
-(* ------------------------------------------------------------------------- *)
-
-fun pointerEqual (x : 'a, y : 'a) = pointer_eq(x,y);
-
-(* ------------------------------------------------------------------------- *)
-(* Timing function applications.                                             *)
-(* ------------------------------------------------------------------------- *)
-
-fun time f x =
-    let
-      fun p t =
-          let
-            val s = Time.fmt 3 t
-          in
-            case size (List.last (String.fields (fn x => x = #".") s)) of
-              3 => s
-            | 2 => s ^ "0"
-            | 1 => s ^ "00"
-            | _ => raise Fail "Metis_Portable.time"
-          end
-
-      val c = Timer.startCPUTimer ()
-
-      val r = Timer.startRealTimer ()
-
-      fun pt () =
-          let
-            val {usr,sys} = Timer.checkCPUTimer c
-            val real = Timer.checkRealTimer r
-          in
-            TextIO.print (* MODIFIED by Jasmin Blanchette *)
-              ("User: " ^ p usr ^ "  System: " ^ p sys ^
-               "  Real: " ^ p real ^ "\n")
-          end
-
-      val y = f x handle e => (pt (); raise e)
-
-      val () = pt ()
-    in
-      y
-    end;
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing)                                 *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-fun CRITICAL e = NAMED_CRITICAL "metis" e;
-
-
-(* ------------------------------------------------------------------------- *)
-(* Generating random values.                                                 *)
-(* ------------------------------------------------------------------------- *)
-
-val randomWord = Metis_Random.nextWord;
-
-val randomBool = Metis_Random.nextBool;
-
-val randomInt = Metis_Random.nextInt;
-
-val randomReal = Metis_Random.nextReal;
-
-end
-
-(* ------------------------------------------------------------------------- *)
-(* Quotations a la Moscow ML.                                                *)
-(* ------------------------------------------------------------------------- *)
-
-datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a;
-
-(**** Original file: Useful.sig ****)
+val ml = "isabelle"
+
+fun pointerEqual (x : 'a, y : 'a) = pointer_eq (x, y)
+
+fun critical e () = NAMED_CRITICAL "metis" e
+
+val randomWord = Metis_Random.nextWord
+val randomBool = Metis_Random.nextBool
+val randomInt = Metis_Random.nextInt
+val randomReal = Metis_Random.nextReal
+
+fun time f x = f x (* dummy *)
+
+end
+
+datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a
+
+(**** Original file: src/Useful.sig ****)
 
 (* ========================================================================= *)
 (* ML UTILITY FUNCTIONS                                                      *)
@@ -253,8 +194,6 @@
 (* Tracing.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-val print : string -> unit (* MODIFIED by Jasmin Blanchette *)
-
 val tracePrint : (string -> unit) Unsynchronized.ref
 
 val trace : string -> unit
@@ -340,10 +279,6 @@
 (* Lists: note we count elements from 0.                                     *)
 (* ------------------------------------------------------------------------- *)
 
-val foldl : ('a * 'b -> 'b) -> 'b -> 'a list -> 'b (* MODIFIED by Jasmin Blanchette *)
-
-val foldr : ('a * 'b -> 'b) -> 'b -> 'a list -> 'b (* MODIFIED by Jasmin Blanchette *)
-
 val cons : 'a -> 'a list -> 'a list
 
 val hdTl : 'a list -> 'a * 'a list
@@ -448,10 +383,6 @@
 (* Strings.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-val implode : char list -> string (* MODIFIED by Jasmin Blanchette *)
-
-val explode : string -> char list (* MODIFIED by Jasmin Blanchette *)
-
 val rot : int -> char -> char
 
 val charToInt : char -> int option
@@ -554,7 +485,9 @@
 
 val try : ('a -> 'b) -> 'a -> 'b
 
-val chat : string -> unit
+val chat : string -> unit  (* stdout *)
+
+val chide : string -> unit  (* stderr *)
 
 val warn : string -> unit
 
@@ -568,7 +501,7 @@
 
 end
 
-(**** Original file: Useful.sml ****)
+(**** Original file: src/Useful.sml ****)
 
 (* ========================================================================= *)
 (* ML UTILITY FUNCTIONS                                                      *)
@@ -622,9 +555,7 @@
 (* Tracing.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-val print = TextIO.print; (* MODIFIED by Jasmin Blanchette *)
-
-val tracePrint = Unsynchronized.ref print;
+val tracePrint = Unsynchronized.ref TextIO.print;
 
 fun trace mesg = !tracePrint mesg;
 
@@ -744,10 +675,6 @@
 (* Lists.                                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-val foldl = List.foldl; (* MODIFIED by Jasmin Blanchette *)
-
-val foldr = List.foldr; (* MODIFIED by Jasmin Blanchette *)
-
 fun cons x y = x :: y;
 
 fun hdTl l = (hd l, tl l);
@@ -783,19 +710,22 @@
 
 fun zip xs ys = zipWith pair xs ys;
 
-fun unzip ab =
-    foldl (fn ((x, y), (xs, ys)) => (x :: xs, y :: ys)) ([], []) (rev ab);
+local
+  fun inc ((x,y),(xs,ys)) = (x :: xs, y :: ys);
+in
+  fun unzip ab = List.foldl inc ([],[]) (rev ab);
+end;
 
 fun cartwith f =
-  let
-    fun aux _ res _ [] = res
-      | aux xsCopy res [] (y :: yt) = aux xsCopy res xsCopy yt
-      | aux xsCopy res (x :: xt) (ys as y :: _) =
-        aux xsCopy (f x y :: res) xt ys
-  in
-    fn xs => fn ys =>
-       let val xs' = rev xs in aux xs' [] xs' (rev ys) end
-  end;
+    let
+      fun aux _ res _ [] = res
+        | aux xsCopy res [] (y :: yt) = aux xsCopy res xsCopy yt
+        | aux xsCopy res (x :: xt) (ys as y :: _) =
+          aux xsCopy (f x y :: res) xt ys
+    in
+      fn xs => fn ys =>
+         let val xs' = rev xs in aux xs' [] xs' (rev ys) end
+    end;
 
 fun cart xs ys = cartwith pair xs ys;
 
@@ -914,15 +844,32 @@
 
 fun delete x s = List.filter (not o equal x) s;
 
-fun setify s = rev (foldl (fn (v,x) => if mem v x then x else v :: x) [] s);
-
-fun union s t = foldl (fn (v,x) => if mem v t then x else v :: x) t (rev s);
+local
+  fun inc (v,x) = if mem v x then x else v :: x;
+in
+  fun setify s = rev (List.foldl inc [] s);
+end;
+
+fun union s t =
+    let
+      fun inc (v,x) = if mem v t then x else v :: x
+    in
+      List.foldl inc t (rev s)
+    end;
 
 fun intersect s t =
-    foldl (fn (v,x) => if mem v t then v :: x else x) [] (rev s);
+    let
+      fun inc (v,x) = if mem v t then v :: x else x
+    in
+      List.foldl inc [] (rev s)
+    end;
 
 fun difference s t =
-    foldl (fn (v,x) => if mem v t then x else v :: x) [] (rev s);
+    let
+      fun inc (v,x) = if mem v t then x else v :: x
+    in
+      List.foldl inc [] (rev s)
+    end;
 
 fun subset s t = List.all (fn x => mem x t) s;
 
@@ -1032,8 +979,7 @@
 
   val primesList = Unsynchronized.ref [2];
 in
-  (* MODIFIED by Jasmin Blanchette *)
-  fun primes n = CRITICAL (fn () =>
+  fun primes n = Metis_Portable.critical (fn () =>
       let
         val Unsynchronized.ref ps = primesList
 
@@ -1048,11 +994,10 @@
           in
             ps
           end
-      end);
-end;
-
-(* MODIFIED by Jasmin Blanchette *)
-fun primesUpTo n = CRITICAL (fn () =>
+      end) ();
+end;
+
+fun primesUpTo n = Metis_Portable.critical (fn () =>
     let
       fun f k =
           let
@@ -1064,22 +1009,18 @@
           end
     in
       f 8
-    end);
+    end) ();
 
 (* ------------------------------------------------------------------------- *)
 (* Strings.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-val implode = String.implode (* MODIFIED by Jasmin Blanchette *)
-
-val explode = String.explode (* MODIFIED by Jasmin Blanchette *)
-
 local
   fun len l = (length l, l)
 
-  val upper = len (explode "ABCDEFGHIJKLMNOPQRSTUVWXYZ");
-
-  val lower = len (explode "abcdefghijklmnopqrstuvwxyz");
+  val upper = len (String.explode "ABCDEFGHIJKLMNOPQRSTUVWXYZ");
+
+  val lower = len (String.explode "abcdefghijklmnopqrstuvwxyz");
 
   fun rotate (n,l) c k =
       List.nth (l, (k + Option.valOf (index (equal c) l)) mod n);
@@ -1118,7 +1059,7 @@
     let
       fun dup 0 l = l | dup n l = dup (n - 1) (x :: l)
     in
-      fn n => implode (dup n [])
+      fn n => String.implode (dup n [])
     end;
 
 fun chomp s =
@@ -1130,14 +1071,15 @@
     end;
 
 local
-  fun chop [] = []
-    | chop (l as (h :: t)) = if Char.isSpace h then chop t else l;
-in
-  val trim = implode o chop o rev o chop o rev o explode;
-end;
-
-fun join _ [] = ""
-  | join s (h :: t) = foldl (fn (x,y) => y ^ s ^ x) h t;
+  fun chop l =
+      case l of
+        [] => []
+      | h :: t => if Char.isSpace h then chop t else l;
+in
+  val trim = String.implode o chop o rev o chop o rev o String.explode;
+end;
+
+val join = String.concatWith;
 
 local
   fun match [] l = SOME l
@@ -1145,18 +1087,19 @@
     | match (x :: xs) (y :: ys) = if x = y then match xs ys else NONE;
 
   fun stringify acc [] = acc
-    | stringify acc (h :: t) = stringify (implode h :: acc) t;
+    | stringify acc (h :: t) = stringify (String.implode h :: acc) t;
 in
   fun split sep =
       let
         val pat = String.explode sep
+
         fun div1 prev recent [] = stringify [] (rev recent :: prev)
           | div1 prev recent (l as h :: t) =
             case match pat l of
               NONE => div1 prev (h :: recent) t
             | SOME rest => div1 (rev recent :: prev) [] rest
       in
-        fn s => div1 [] [] (explode s)
+        fn s => div1 [] [] (String.explode s)
       end;
 end;
 
@@ -1286,24 +1229,22 @@
 local
   val generator = Unsynchronized.ref 0
 in
-  (* MODIFIED by Jasmin Blanchette *)
-  fun newInt () = CRITICAL (fn () =>
+  fun newInt () = Metis_Portable.critical (fn () =>
       let
         val n = !generator
         val () = generator := n + 1
       in
         n
-      end);
+      end) ();
 
   fun newInts 0 = []
-      (* MODIFIED by Jasmin Blanchette *)
-    | newInts k = CRITICAL (fn () =>
+    | newInts k = Metis_Portable.critical (fn () =>
       let
         val n = !generator
         val () = generator := n + k
       in
         interval n k
-      end);
+      end) ();
 end;
 
 fun withRef (r,new) f x =
@@ -1383,10 +1324,12 @@
 (* Profiling and error reporting.                                            *)
 (* ------------------------------------------------------------------------- *)
 
-fun chat s = TextIO.output (TextIO.stdErr, s ^ "\n");
-
-local
-  fun err x s = chat (x ^ ": " ^ s);
+fun chat s = TextIO.output (TextIO.stdOut, s ^ "\n");
+
+fun chide s = TextIO.output (TextIO.stdErr, s ^ "\n");
+
+local
+  fun err x s = chide (x ^ ": " ^ s);
 in
   fun try f x = f x
       handle e as Error _ => (err "try" (errorToString e); raise e)
@@ -1431,7 +1374,7 @@
 
 end
 
-(**** Original file: Lazy.sig ****)
+(**** Original file: src/Lazy.sig ****)
 
 (* ========================================================================= *)
 (* SUPPORT FOR LAZY EVALUATION                                               *)
@@ -1453,7 +1396,7 @@
 
 end
 
-(**** Original file: Lazy.sml ****)
+(**** Original file: src/Lazy.sml ****)
 
 (* ========================================================================= *)
 (* SUPPORT FOR LAZY EVALUATION                                               *)
@@ -1494,11 +1437,11 @@
 
 end
 
-(**** Original file: Ordered.sig ****)
+(**** Original file: src/Ordered.sig ****)
 
 (* ========================================================================= *)
 (* ORDERED TYPES                                                             *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Ordered =
@@ -1530,11 +1473,11 @@
 
 end
 
-(**** Original file: Ordered.sml ****)
+(**** Original file: src/Ordered.sml ****)
 
 (* ========================================================================= *)
 (* ORDERED TYPES                                                             *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_IntOrdered =
@@ -1556,7 +1499,7 @@
 structure Metis_StringOrdered =
 struct type t = string val compare = String.compare end;
 
-(**** Original file: Map.sig ****)
+(**** Original file: src/Map.sig ****)
 
 (* ========================================================================= *)
 (* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES                      *)
@@ -1749,7 +1692,7 @@
 
 end
 
-(**** Original file: Map.sml ****)
+(**** Original file: src/Map.sml ****)
 
 (* ========================================================================= *)
 (* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES                      *)
@@ -2632,38 +2575,40 @@
 (* ------------------------------------------------------------------------- *)
 
 datatype ('key,'value) iterator =
-    LR of ('key * 'value) * ('key,'value) tree * ('key,'value) node list
-  | Metis_RL of ('key * 'value) * ('key,'value) tree * ('key,'value) node list;
-
-fun fromSpineLR nodes =
+    LeftToRightIterator of
+      ('key * 'value) * ('key,'value) tree * ('key,'value) node list
+  | RightToLeftIterator of
+      ('key * 'value) * ('key,'value) tree * ('key,'value) node list;
+
+fun fromSpineLeftToRightIterator nodes =
     case nodes of
       [] => NONE
     | Node {key,value,right,...} :: nodes =>
-      SOME (LR ((key,value),right,nodes));
-
-fun fromSpineRL nodes =
+      SOME (LeftToRightIterator ((key,value),right,nodes));
+
+fun fromSpineRightToLeftIterator nodes =
     case nodes of
       [] => NONE
     | Node {key,value,left,...} :: nodes =>
-      SOME (Metis_RL ((key,value),left,nodes));
-
-fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
-
-fun addRL nodes tree = fromSpineRL (treeRightSpine nodes tree);
-
-fun treeMkIterator tree = addLR [] tree;
-
-fun treeMkRevIterator tree = addRL [] tree;
+      SOME (RightToLeftIterator ((key,value),left,nodes));
+
+fun addLeftToRightIterator nodes tree = fromSpineLeftToRightIterator (treeLeftSpine nodes tree);
+
+fun addRightToLeftIterator nodes tree = fromSpineRightToLeftIterator (treeRightSpine nodes tree);
+
+fun treeMkIterator tree = addLeftToRightIterator [] tree;
+
+fun treeMkRevIterator tree = addRightToLeftIterator [] tree;
 
 fun readIterator iter =
     case iter of
-      LR (key_value,_,_) => key_value
-    | Metis_RL (key_value,_,_) => key_value;
+      LeftToRightIterator (key_value,_,_) => key_value
+    | RightToLeftIterator (key_value,_,_) => key_value;
 
 fun advanceIterator iter =
     case iter of
-      LR (_,tree,nodes) => addLR nodes tree
-    | Metis_RL (_,tree,nodes) => addRL nodes tree;
+      LeftToRightIterator (_,tree,nodes) => addLeftToRightIterator nodes tree
+    | RightToLeftIterator (_,tree,nodes) => addRightToLeftIterator nodes tree;
 
 fun foldIterator f acc io =
     case io of
@@ -3177,7 +3122,7 @@
 
 end
 
-(**** Original file: KeyMap.sig ****)
+(**** Original file: src/KeyMap.sig ****)
 
 (* ========================================================================= *)
 (* FINITE MAPS WITH A FIXED KEY TYPE                                         *)
@@ -3376,7 +3321,7 @@
 
 end
 
-(**** Original file: KeyMap.sml ****)
+(**** Original file: src/KeyMap.sml ****)
 
 (* ========================================================================= *)
 (* FINITE MAPS WITH A FIXED KEY TYPE                                         *)
@@ -4267,38 +4212,40 @@
 (* ------------------------------------------------------------------------- *)
 
 datatype 'value iterator =
-    LR of (key * 'value) * 'value tree * 'value node list
-  | Metis_RL of (key * 'value) * 'value tree * 'value node list;
-
-fun fromSpineLR nodes =
+    LeftToRightIterator of
+      (key * 'value) * 'value tree * 'value node list
+  | RightToLeftIterator of
+      (key * 'value) * 'value tree * 'value node list;
+
+fun fromSpineLeftToRightIterator nodes =
     case nodes of
       [] => NONE
     | Node {key,value,right,...} :: nodes =>
-      SOME (LR ((key,value),right,nodes));
-
-fun fromSpineRL nodes =
+      SOME (LeftToRightIterator ((key,value),right,nodes));
+
+fun fromSpineRightToLeftIterator nodes =
     case nodes of
       [] => NONE
     | Node {key,value,left,...} :: nodes =>
-      SOME (Metis_RL ((key,value),left,nodes));
-
-fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
-
-fun addRL nodes tree = fromSpineRL (treeRightSpine nodes tree);
-
-fun treeMkIterator tree = addLR [] tree;
-
-fun treeMkRevIterator tree = addRL [] tree;
+      SOME (RightToLeftIterator ((key,value),left,nodes));
+
+fun addLeftToRightIterator nodes tree = fromSpineLeftToRightIterator (treeLeftSpine nodes tree);
+
+fun addRightToLeftIterator nodes tree = fromSpineRightToLeftIterator (treeRightSpine nodes tree);
+
+fun treeMkIterator tree = addLeftToRightIterator [] tree;
+
+fun treeMkRevIterator tree = addRightToLeftIterator [] tree;
 
 fun readIterator iter =
     case iter of
-      LR (key_value,_,_) => key_value
-    | Metis_RL (key_value,_,_) => key_value;
+      LeftToRightIterator (key_value,_,_) => key_value
+    | RightToLeftIterator (key_value,_,_) => key_value;
 
 fun advanceIterator iter =
     case iter of
-      LR (_,tree,nodes) => addLR nodes tree
-    | Metis_RL (_,tree,nodes) => addRL nodes tree;
+      LeftToRightIterator (_,tree,nodes) => addLeftToRightIterator nodes tree
+    | RightToLeftIterator (_,tree,nodes) => addRightToLeftIterator nodes tree;
 
 fun foldIterator f acc io =
     case io of
@@ -4821,7 +4768,7 @@
 structure Metis_StringMap =
 Metis_KeyMap (Metis_StringOrdered);
 
-(**** Original file: Set.sig ****)
+(**** Original file: src/Set.sig ****)
 
 (* ========================================================================= *)
 (* FINITE SETS IMPLEMENTED WITH RANDOMLY BALANCED TREES                      *)
@@ -4993,7 +4940,7 @@
 
 end
 
-(**** Original file: Set.sml ****)
+(**** Original file: src/Set.sml ****)
 
 (* ========================================================================= *)
 (* FINITE SETS IMPLEMENTED WITH RANDOMLY BALANCED TREES                      *)
@@ -5327,7 +5274,7 @@
 
 end
 
-(**** Original file: ElementSet.sig ****)
+(**** Original file: src/ElementSet.sig ****)
 
 (* ========================================================================= *)
 (* FINITE SETS WITH A FIXED ELEMENT TYPE                                     *)
@@ -5505,7 +5452,7 @@
 
 end
 
-(**** Original file: ElementSet.sml ****)
+(**** Original file: src/ElementSet.sml ****)
 
 (* ========================================================================= *)
 (* FINITE SETS WITH A FIXED ELEMENT TYPE                                     *)
@@ -5855,11 +5802,11 @@
 structure Metis_StringSet =
 Metis_ElementSet (Metis_StringMap);
 
-(**** Original file: Sharing.sig ****)
+(**** Original file: src/Sharing.sig ****)
 
 (* ========================================================================= *)
 (* PRESERVING SHARING OF ML VALUES                                           *)
-(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2005 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Sharing =
@@ -5903,11 +5850,11 @@
 
 end
 
-(**** Original file: Sharing.sml ****)
+(**** Original file: src/Sharing.sml ****)
 
 (* ========================================================================= *)
 (* PRESERVING SHARING OF ML VALUES                                           *)
-(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2005 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Sharing :> Metis_Sharing =
@@ -6064,11 +6011,11 @@
 
 end
 
-(**** Original file: Stream.sig ****)
+(**** Original file: src/Stream.sig ****)
 
 (* ========================================================================= *)
 (* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Stream =
@@ -6176,18 +6123,16 @@
 
 end
 
-(**** Original file: Stream.sml ****)
+(**** Original file: src/Stream.sml ****)
 
 (* ========================================================================= *)
 (* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Stream :> Metis_Stream =
 struct
 
-open Metis_Useful; (* MODIFIED by Jasmin Blanchette *)
-
 val K = Metis_Useful.K;
 
 val pair = Metis_Useful.pair;
@@ -6381,9 +6326,9 @@
 
 fun listConcat s = concat (map fromList s);
 
-fun toString s = implode (toList s);
-
-fun fromString s = fromList (explode s);
+fun toString s = String.implode (toList s);
+
+fun fromString s = fromList (String.explode s);
 
 fun toTextFile {filename = f} s =
     let
@@ -6415,11 +6360,11 @@
 
 end
 
-(**** Original file: Heap.sig ****)
+(**** Original file: src/Heap.sig ****)
 
 (* ========================================================================= *)
 (* A HEAP DATATYPE FOR ML                                                    *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Heap =
@@ -6449,11 +6394,11 @@
 
 end
 
-(**** Original file: Heap.sml ****)
+(**** Original file: src/Heap.sml ****)
 
 (* ========================================================================= *)
 (* A HEAP DATATYPE FOR ML                                                    *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Heap :> Metis_Heap =
@@ -6529,32 +6474,44 @@
 
 end
 
-(**** Original file: Print.sig ****)
+(**** Original file: src/Print.sig ****)
 
 (* ========================================================================= *)
 (* PRETTY-PRINTING                                                           *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Print =
 sig
 
 (* ------------------------------------------------------------------------- *)
+(* Escaping strings.                                                         *)
+(* ------------------------------------------------------------------------- *)
+
+val escapeString : {escape : char list} -> string -> string
+
+(* ------------------------------------------------------------------------- *)
+(* A type of strings annotated with their size.                              *)
+(* ------------------------------------------------------------------------- *)
+
+type stringSize = string * int
+
+val mkStringSize : string -> stringSize
+
+(* ------------------------------------------------------------------------- *)
 (* A type of pretty-printers.                                                *)
 (* ------------------------------------------------------------------------- *)
 
 datatype breakStyle = Consistent | Inconsistent
 
-datatype ppStep =
+datatype step =
     BeginBlock of breakStyle * int
   | EndBlock
-  | AddString of string
+  | AddString of stringSize
   | AddBreak of int
   | AddNewline
 
-type ppstream = ppStep Metis_Stream.stream
-
-type 'a pp = 'a -> ppstream
+type ppstream = step Metis_Stream.stream
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty-printer primitives.                                                *)
@@ -6564,7 +6521,7 @@
 
 val endBlock : ppstream
 
-val addString : string -> ppstream
+val addString : stringSize -> ppstream
 
 val addBreak : int -> ppstream
 
@@ -6584,18 +6541,24 @@
 
 val blockProgram : breakStyle -> int -> ppstream list -> ppstream
 
-val bracket : string -> string -> ppstream -> ppstream
-
-val field : string -> ppstream -> ppstream
-
-val record : (string * ppstream) list -> ppstream
+(* ------------------------------------------------------------------------- *)
+(* Executing pretty-printers to generate lines.                              *)
+(* ------------------------------------------------------------------------- *)
+
+val execute :
+    {lineLength : int} -> ppstream ->
+    {indent : int, line : string} Metis_Stream.stream
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty-printer combinators.                                               *)
 (* ------------------------------------------------------------------------- *)
 
+type 'a pp = 'a -> ppstream
+
 val ppMap : ('a -> 'b) -> 'b pp -> 'a pp
 
+val ppString : string pp
+
 val ppBracket : string -> string -> 'a pp -> 'a pp
 
 val ppOp : string -> ppstream
@@ -6614,8 +6577,6 @@
 
 val ppChar : char pp
 
-val ppString : string pp
-
 val ppEscapeString : {escape : char list} -> string pp
 
 val ppUnit : unit pp
@@ -6644,7 +6605,7 @@
 
 val ppBreakStyle : breakStyle pp
 
-val ppPpStep : ppStep pp
+val ppStep : step pp
 
 val ppPpstream : ppstream pp
 
@@ -6652,28 +6613,27 @@
 (* Pretty-printing infix operators.                                          *)
 (* ------------------------------------------------------------------------- *)
 
+type token = string
+
+datatype assoc =
+    LeftAssoc
+  | NonAssoc
+  | RightAssoc
+
 datatype infixes =
     Infixes of
-      {token : string,
+      {token : token,
        precedence : int,
-       leftAssoc : bool} list
+       assoc : assoc} list
 
 val tokensInfixes : infixes -> Metis_StringSet.set
 
-val layerInfixes :
-    infixes ->
-    {tokens : {leftSpaces : int, token : string, rightSpaces : int} list,
-     leftAssoc : bool} list
+val layerInfixes : infixes -> {tokens : Metis_StringSet.set, assoc : assoc} list
 
 val ppInfixes :
-    infixes -> ('a -> (string * 'a * 'a) option) -> ('a * bool) pp ->
-    ('a * bool) pp
-
-(* ------------------------------------------------------------------------- *)
-(* Executing pretty-printers to generate lines.                              *)
-(* ------------------------------------------------------------------------- *)
-
-val execute : {lineLength : int} -> ppstream -> string Metis_Stream.stream
+    infixes ->
+    ('a -> (token * 'a * 'a) option) -> ('a * token) pp ->
+    ('a * bool) pp -> ('a * bool) pp
 
 (* ------------------------------------------------------------------------- *)
 (* Executing pretty-printers with a global line length.                      *)
@@ -6689,11 +6649,11 @@
 
 end
 
-(**** Original file: Print.sml ****)
+(**** Original file: src/Print.sml ****)
 
 (* ========================================================================= *)
 (* PRETTY-PRINTING                                                           *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Print :> Metis_Print =
@@ -6722,47 +6682,69 @@
     | Metis_Stream.Cons (h,t) => revAppend h (revConcat o t);
 
 local
-  fun join current prev = (prev ^ "\n", current);
-in
-  fun joinNewline strm =
-      case strm of
-        Metis_Stream.Nil => Metis_Stream.Nil
-      | Metis_Stream.Cons (h,t) => Metis_Stream.maps join Metis_Stream.singleton h (t ());
-end;
-
-local
   fun calcSpaces n = nChars #" " n;
 
-  val cachedSpaces = Vector.tabulate (initialLineLength,calcSpaces);
+  val cacheSize = 2 * initialLineLength;
+
+  val cachedSpaces = Vector.tabulate (cacheSize,calcSpaces);
 in
   fun nSpaces n =
-      if n < initialLineLength then Vector.sub (cachedSpaces,n)
+      if n < cacheSize then Vector.sub (cachedSpaces,n)
       else calcSpaces n;
 end;
 
 (* ------------------------------------------------------------------------- *)
+(* Escaping strings.                                                         *)
+(* ------------------------------------------------------------------------- *)
+
+fun escapeString {escape} =
+    let
+      val escapeMap = map (fn c => (c, "\\" ^ str c)) escape
+
+      fun escapeChar c =
+          case c of
+            #"\\" => "\\\\"
+          | #"\n" => "\\n"
+          | #"\t" => "\\t"
+          | _ =>
+            case List.find (equal c o fst) escapeMap of
+              SOME (_,s) => s
+            | NONE => str c
+    in
+      String.translate escapeChar
+    end;
+
+(* ------------------------------------------------------------------------- *)
+(* A type of strings annotated with their size.                              *)
+(* ------------------------------------------------------------------------- *)
+
+type stringSize = string * int;
+
+fun mkStringSize s = (s, size s);
+
+val emptyStringSize = mkStringSize "";
+
+(* ------------------------------------------------------------------------- *)
 (* A type of pretty-printers.                                                *)
 (* ------------------------------------------------------------------------- *)
 
 datatype breakStyle = Consistent | Inconsistent;
 
-datatype ppStep =
+datatype step =
     BeginBlock of breakStyle * int
   | EndBlock
-  | AddString of string
+  | AddString of stringSize
   | AddBreak of int
   | AddNewline;
 
-type ppstream = ppStep Metis_Stream.stream;
-
-type 'a pp = 'a -> ppstream;
+type ppstream = step Metis_Stream.stream;
 
 fun breakStyleToString style =
     case style of
       Consistent => "Consistent"
     | Inconsistent => "Inconsistent";
 
-fun ppStepToString step =
+fun stepToString step =
     case step of
       BeginBlock _ => "BeginBlock"
     | EndBlock => "EndBlock"
@@ -6802,330 +6784,6 @@
 
 fun blockProgram style indent pps = block style indent (program pps);
 
-fun bracket l r pp =
-    blockProgram Inconsistent (size l)
-      [addString l,
-       pp,
-       addString r];
-
-fun field f pp =
-    blockProgram Inconsistent 2
-      [addString (f ^ " ="),
-       addBreak 1,
-       pp];
-
-val record =
-    let
-      val sep = sequence (addString ",") (addBreak 1)
-
-      fun recordField (f,pp) = field f pp
-
-      fun sepField f = sequence sep (recordField f)
-
-      fun fields [] = []
-        | fields (f :: fs) = recordField f :: map sepField fs
-    in
-      bracket "{" "}" o blockProgram Consistent 0 o fields
-    end;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printer combinators.                                               *)
-(* ------------------------------------------------------------------------- *)
-
-fun ppMap f ppB a : ppstream = ppB (f a);
-
-fun ppBracket l r ppA a = bracket l r (ppA a);
-
-fun ppOp s = sequence (if s = "" then skip else addString s) (addBreak 1);
-
-fun ppOp2 ab ppA ppB (a,b) =
-    blockProgram Inconsistent 0
-      [ppA a,
-       ppOp ab,
-       ppB b];
-
-fun ppOp3 ab bc ppA ppB ppC (a,b,c) =
-    blockProgram Inconsistent 0
-      [ppA a,
-       ppOp ab,
-       ppB b,
-       ppOp bc,
-       ppC c];
-
-fun ppOpList s ppA =
-    let
-      fun ppOpA a = sequence (ppOp s) (ppA a)
-    in
-      fn [] => skip
-       | h :: t => blockProgram Inconsistent 0 (ppA h :: map ppOpA t)
-    end;
-
-fun ppOpStream s ppA =
-    let
-      fun ppOpA a = sequence (ppOp s) (ppA a)
-    in
-      fn Metis_Stream.Nil => skip
-       | Metis_Stream.Cons (h,t) =>
-         blockProgram Inconsistent 0
-           [ppA h,
-            Metis_Stream.concat (Metis_Stream.map ppOpA (t ()))]
-    end;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printers for common types.                                         *)
-(* ------------------------------------------------------------------------- *)
-
-fun ppChar c = addString (str c);
-
-val ppString = addString;
-
-fun ppEscapeString {escape} =
-    let
-      val escapeMap = map (fn c => (c, "\\" ^ str c)) escape
-
-      fun escapeChar c =
-          case c of
-            #"\\" => "\\\\"
-          | #"\n" => "\\n"
-          | #"\t" => "\\t"
-          | _ =>
-            case List.find (equal c o fst) escapeMap of
-              SOME (_,s) => s
-            | NONE => str c
-    in
-      fn s => addString (String.translate escapeChar s)
-    end;
-
-val ppUnit : unit pp = K (addString "()");
-
-fun ppBool b = addString (if b then "true" else "false");
-
-fun ppInt i = addString (Int.toString i);
-
-local
-  val ppNeg = addString "~"
-  and ppSep = addString ","
-  and ppZero = addString "0"
-  and ppZeroZero = addString "00";
-
-  fun ppIntBlock i =
-      if i < 10 then sequence ppZeroZero (ppInt i)
-      else if i < 100 then sequence ppZero (ppInt i)
-      else ppInt i;
-
-  fun ppIntBlocks i =
-      if i < 1000 then ppInt i
-      else sequence (ppIntBlocks (i div 1000))
-             (sequence ppSep (ppIntBlock (i mod 1000)));
-in
-  fun ppPrettyInt i =
-      if i < 0 then sequence ppNeg (ppIntBlocks (~i))
-      else ppIntBlocks i;
-end;
-
-fun ppReal r = addString (Real.toString r);
-
-fun ppPercent p = addString (percentToString p);
-
-fun ppOrder x =
-    addString
-      (case x of
-         LESS => "Less"
-       | EQUAL => "Equal"
-       | GREATER => "Greater");
-
-fun ppList ppA = ppBracket "[" "]" (ppOpList "," ppA);
-
-fun ppStream ppA = ppBracket "[" "]" (ppOpStream "," ppA);
-
-fun ppOption ppA ao =
-    case ao of
-      SOME a => ppA a
-    | NONE => addString "-";
-
-fun ppPair ppA ppB = ppBracket "(" ")" (ppOp2 "," ppA ppB);
-
-fun ppTriple ppA ppB ppC = ppBracket "(" ")" (ppOp3 "," "," ppA ppB ppC);
-
-fun ppBreakStyle style = addString (breakStyleToString style);
-
-fun ppPpStep step =
-    let
-      val cmd = ppStepToString step
-    in
-      blockProgram Inconsistent 2
-        (addString cmd ::
-         (case step of
-            BeginBlock style_indent =>
-              [addBreak 1,
-               ppPair ppBreakStyle ppInt style_indent]
-          | EndBlock => []
-          | AddString s =>
-              [addBreak 1,
-               addString ("\"" ^ s ^ "\"")]
-          | AddBreak n =>
-              [addBreak 1,
-               ppInt n]
-          | AddNewline => []))
-    end;
-
-val ppPpstream = ppStream ppPpStep;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printing infix operators.                                          *)
-(* ------------------------------------------------------------------------- *)
-
-datatype infixes =
-    Infixes of
-      {token : string,
-       precedence : int,
-       leftAssoc : bool} list;
-
-local
-  fun chop l =
-      case l of
-        #" " :: l => let val (n,l) = chop l in (n + 1, l) end
-      | _ => (0,l);
-in
-  fun opSpaces tok =
-      let
-        val tok = explode tok
-        val (r,tok) = chop (rev tok)
-        val (l,tok) = chop (rev tok)
-        val tok = implode tok
-      in
-        {leftSpaces = l, token = tok, rightSpaces = r}
-      end;
-end;
-
-fun ppOpSpaces {leftSpaces,token,rightSpaces} =
-    let
-      val leftSpacesToken =
-          if leftSpaces = 0 then token else nSpaces leftSpaces ^ token
-    in
-      sequence
-        (addString leftSpacesToken)
-        (addBreak rightSpaces)
-    end;
-
-local
-  fun new t l acc = {tokens = [opSpaces t], leftAssoc = l} :: acc;
-
-  fun add t l acc =
-      case acc of
-        [] => raise Bug "Metis_Print.layerInfixOps.layer"
-      | {tokens = ts, leftAssoc = l'} :: acc =>
-        if l = l' then {tokens = opSpaces t :: ts, leftAssoc = l} :: acc
-        else raise Bug "Metis_Print.layerInfixOps: mixed assocs";
-
-  fun layer ({token = t, precedence = p, leftAssoc = l}, (acc,p')) =
-      let
-        val acc = if p = p' then add t l acc else new t l acc
-      in
-        (acc,p)
-      end;
-in
-  fun layerInfixes (Infixes i) =
-      case sortMap #precedence Int.compare i of
-        [] => []
-      | {token = t, precedence = p, leftAssoc = l} :: i =>
-        let
-          val acc = new t l []
-
-          val (acc,_) = List.foldl layer (acc,p) i
-        in
-          acc
-        end;
-end;
-
-val tokensLayeredInfixes =
-    let
-      fun addToken ({leftSpaces = _, token = t, rightSpaces = _}, s) =
-          Metis_StringSet.add s t
-
-      fun addTokens ({tokens = t, leftAssoc = _}, s) =
-          List.foldl addToken s t
-    in
-      List.foldl addTokens Metis_StringSet.empty
-    end;
-
-val tokensInfixes = tokensLayeredInfixes o layerInfixes;
-
-local
-  val mkTokenMap =
-      let
-        fun add (token,m) =
-            let
-              val {leftSpaces = _, token = t, rightSpaces = _} = token
-            in
-              Metis_StringMap.insert m (t, ppOpSpaces token)
-            end
-      in
-        List.foldl add (Metis_StringMap.new ())
-      end;
-in
-  fun ppGenInfix {tokens,leftAssoc} =
-      let
-        val tokenMap = mkTokenMap tokens
-      in
-        fn dest => fn ppSub =>
-           let
-             fun dest' tm =
-                 case dest tm of
-                   NONE => NONE
-                 | SOME (t,a,b) =>
-                   case Metis_StringMap.peek tokenMap t of
-                     NONE => NONE
-                   | SOME p => SOME (p,a,b)
-
-             fun ppGo (tmr as (tm,r)) =
-                 case dest' tm of
-                   NONE => ppSub tmr
-                 | SOME (p,a,b) =>
-                   program
-                     [(if leftAssoc then ppGo else ppSub) (a,true),
-                      p,
-                      (if leftAssoc then ppSub else ppGo) (b,r)]
-           in
-             fn tmr as (tm,_) =>
-                if Option.isSome (dest' tm) then
-                  block Inconsistent 0 (ppGo tmr)
-                else
-                  ppSub tmr
-           end
-      end;
-end
-
-fun ppInfixes ops =
-    let
-      val layeredOps = layerInfixes ops
-
-      val toks = tokensLayeredInfixes layeredOps
-
-      val iprinters = List.map ppGenInfix layeredOps
-    in
-      fn dest => fn ppSub =>
-         let
-           fun printer sub = foldl (fn (ip,p) => ip dest p) sub iprinters
-
-           fun isOp t =
-               case dest t of
-                 SOME (x,_,_) => Metis_StringSet.member x toks
-               | NONE => false
-
-           fun subpr (tmr as (tm,_)) =
-               if isOp tm then
-                 blockProgram Inconsistent 1
-                   [addString "(",
-                    printer subpr (tm,false),
-                    addString ")"]
-               else
-                 ppSub tmr
-         in
-           fn tmr => block Inconsistent 0 (printer subpr tmr)
-         end
-    end;
-
 (* ------------------------------------------------------------------------- *)
 (* Executing pretty-printers to generate lines.                              *)
 (* ------------------------------------------------------------------------- *)
@@ -7189,18 +6847,22 @@
 val sizeChunks = List.foldl (fn (c,z) => sizeChunk c + z) 0;
 
 local
-  fun render acc [] = acc
-    | render acc (chunk :: chunks) =
-      case chunk of
-        StringChunk {string = s, ...} => render (acc ^ s) chunks
-      | BreakChunk n => render (acc ^ nSpaces n) chunks
-      | BlockChunk (Block {chunks = l, ...}) =>
-        render acc (List.revAppend (l,chunks));
-in
-  fun renderChunks indent chunks = render (nSpaces indent) (rev chunks);
-
-  fun renderChunk indent chunk = renderChunks indent [chunk];
-end;
+  fun flatten acc chunks =
+      case chunks of
+        [] => rev acc
+      | chunk :: chunks =>
+        case chunk of
+          StringChunk {string = s, ...} => flatten (s :: acc) chunks
+        | BreakChunk n => flatten (nSpaces n :: acc) chunks
+        | BlockChunk (Block {chunks = l, ...}) =>
+          flatten acc (List.revAppend (l,chunks));
+in
+  fun renderChunks indent chunks =
+      {indent = indent,
+       line = String.concat (flatten [] (rev chunks))};
+end;
+
+fun renderChunk indent chunk = renderChunks indent [chunk];
 
 fun isEmptyBlock block =
     let
@@ -7216,6 +6878,7 @@
       empty
     end;
 
+(*BasicDebug
 fun checkBlock ind block =
     let
       val Block {indent, style = _, size, chunks} = block
@@ -7251,6 +6914,7 @@
     in
       check initialIndent o rev
     end;
+*)
 
 val initialBlock =
     let
@@ -7662,12 +7326,10 @@
         (lines,state)
       end;
 
-  fun executeAddString lineLength s lines state =
+  fun executeAddString lineLength (s,n) lines state =
       let
         val State {blocks,lineIndent,lineSize} = state
 
-        val n = size s
-
         val blocks =
             case blocks of
               [] => raise Bug "Metis_Print.executeAddString: no block"
@@ -7754,10 +7416,13 @@
 
   fun executeAddNewline lineLength lines state =
       let
-        val (lines,state) = executeAddString lineLength "" lines state
-        val (lines,state) = executeBigBreak lineLength lines state
-      in
-        executeAddString lineLength "" lines state
+        val (lines,state) =
+            executeAddString lineLength emptyStringSize lines state
+
+        val (lines,state) =
+            executeBigBreak lineLength lines state
+      in
+        executeAddString lineLength emptyStringSize lines state
       end;
 
   fun final lineLength lines state =
@@ -7801,7 +7466,7 @@
               (lines,state)
             end
             handle Bug bug =>
-              raise Bug ("Metis_Print.advance: after " ^ ppStepToString step ^
+              raise Bug ("Metis_Print.advance: after " ^ stepToString step ^
                          " command:\n" ^ bug)
 *)
       in
@@ -7810,26 +7475,377 @@
 end;
 
 (* ------------------------------------------------------------------------- *)
+(* Pretty-printer combinators.                                               *)
+(* ------------------------------------------------------------------------- *)
+
+type 'a pp = 'a -> ppstream;
+
+fun ppMap f ppB a : ppstream = ppB (f a);
+
+fun ppBracket' l r ppA a =
+    let
+      val (_,n) = l
+    in
+      blockProgram Inconsistent n
+        [addString l,
+         ppA a,
+         addString r]
+    end;
+
+fun ppOp' s = sequence (addString s) (addBreak 1);
+
+fun ppOp2' ab ppA ppB (a,b) =
+    blockProgram Inconsistent 0
+      [ppA a,
+       ppOp' ab,
+       ppB b];
+
+fun ppOp3' ab bc ppA ppB ppC (a,b,c) =
+    blockProgram Inconsistent 0
+      [ppA a,
+       ppOp' ab,
+       ppB b,
+       ppOp' bc,
+       ppC c];
+
+fun ppOpList' s ppA =
+    let
+      fun ppOpA a = sequence (ppOp' s) (ppA a)
+    in
+      fn [] => skip
+       | h :: t => blockProgram Inconsistent 0 (ppA h :: map ppOpA t)
+    end;
+
+fun ppOpStream' s ppA =
+    let
+      fun ppOpA a = sequence (ppOp' s) (ppA a)
+    in
+      fn Metis_Stream.Nil => skip
+       | Metis_Stream.Cons (h,t) =>
+         blockProgram Inconsistent 0
+           [ppA h,
+            Metis_Stream.concat (Metis_Stream.map ppOpA (t ()))]
+    end;
+
+fun ppBracket l r = ppBracket' (mkStringSize l) (mkStringSize r);
+
+fun ppOp s = ppOp' (mkStringSize s);
+
+fun ppOp2 ab = ppOp2' (mkStringSize ab);
+
+fun ppOp3 ab bc = ppOp3' (mkStringSize ab) (mkStringSize bc);
+
+fun ppOpList s = ppOpList' (mkStringSize s);
+
+fun ppOpStream s = ppOpStream' (mkStringSize s);
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printers for common types.                                         *)
+(* ------------------------------------------------------------------------- *)
+
+fun ppChar c = addString (str c, 1);
+
+fun ppString s = addString (mkStringSize s);
+
+fun ppEscapeString escape = ppMap (escapeString escape) ppString;
+
+local
+  val pp = ppString "()";
+in
+  fun ppUnit () = pp;
+end;
+
+local
+  val ppTrue = ppString "true"
+  and ppFalse = ppString "false";
+in
+  fun ppBool b = if b then ppTrue else ppFalse;
+end;
+
+val ppInt = ppMap Int.toString ppString;
+
+local
+  val ppNeg = ppString "~"
+  and ppSep = ppString ","
+  and ppZero = ppString "0"
+  and ppZeroZero = ppString "00";
+
+  fun ppIntBlock i =
+      if i < 10 then sequence ppZeroZero (ppInt i)
+      else if i < 100 then sequence ppZero (ppInt i)
+      else ppInt i;
+
+  fun ppIntBlocks i =
+      if i < 1000 then ppInt i
+      else sequence (ppIntBlocks (i div 1000))
+             (sequence ppSep (ppIntBlock (i mod 1000)));
+in
+  fun ppPrettyInt i =
+      if i < 0 then sequence ppNeg (ppIntBlocks (~i))
+      else ppIntBlocks i;
+end;
+
+val ppReal = ppMap Real.toString ppString;
+
+val ppPercent = ppMap percentToString ppString;
+
+local
+  val ppLess = ppString "Less"
+  and ppEqual = ppString "Equal"
+  and ppGreater = ppString "Greater";
+in
+  fun ppOrder ord =
+      case ord of
+        LESS => ppLess
+      | EQUAL => ppEqual
+      | GREATER => ppGreater;
+end;
+
+local
+  val left = mkStringSize "["
+  and right = mkStringSize "]"
+  and sep = mkStringSize ",";
+in
+  fun ppList ppX xs = ppBracket' left right (ppOpList' sep ppX) xs;
+
+  fun ppStream ppX xs = ppBracket' left right (ppOpStream' sep ppX) xs;
+end;
+
+local
+  val ppNone = ppString "-";
+in
+  fun ppOption ppX xo =
+      case xo of
+        SOME x => ppX x
+      | NONE => ppNone;
+end;
+
+local
+  val left = mkStringSize "("
+  and right = mkStringSize ")"
+  and sep = mkStringSize ",";
+in
+  fun ppPair ppA ppB =
+      ppBracket' left right (ppOp2' sep ppA ppB);
+
+  fun ppTriple ppA ppB ppC =
+      ppBracket' left right (ppOp3' sep sep ppA ppB ppC);
+end;
+
+val ppBreakStyle = ppMap breakStyleToString ppString;
+
+val ppStep = ppMap stepToString ppString;
+
+val ppStringSize =
+    let
+      val left = mkStringSize "\""
+      and right = mkStringSize "\""
+      and escape = {escape = [#"\""]}
+
+      val pp = ppBracket' left right (ppEscapeString escape)
+    in
+      fn (s,n) => if size s = n then pp s else ppPair pp ppInt (s,n)
+    end;
+
+fun ppStep step =
+    blockProgram Inconsistent 2
+      (ppStep step ::
+       (case step of
+          BeginBlock style_indent =>
+            [addBreak 1,
+             ppPair ppBreakStyle ppInt style_indent]
+        | EndBlock => []
+        | AddString s =>
+            [addBreak 1,
+             ppStringSize s]
+        | AddBreak n =>
+            [addBreak 1,
+             ppInt n]
+        | AddNewline => []));
+
+val ppPpstream = ppStream ppStep;
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing infix operators.                                          *)
+(* ------------------------------------------------------------------------- *)
+
+type token = string;
+
+datatype assoc =
+    LeftAssoc
+  | NonAssoc
+  | RightAssoc;
+
+datatype infixes =
+    Infixes of
+      {token : token,
+       precedence : int,
+       assoc : assoc} list;
+
+local
+  fun comparePrecedence (io1,io2) =
+      let
+        val {token = _, precedence = p1, assoc = _} = io1
+        and {token = _, precedence = p2, assoc = _} = io2
+      in
+        Int.compare (p2,p1)
+      end;
+
+  fun equalAssoc a a' =
+      case a of
+        LeftAssoc => (case a' of LeftAssoc => true | _ => false)
+      | NonAssoc => (case a' of NonAssoc => true | _ => false)
+      | RightAssoc => (case a' of RightAssoc => true | _ => false);
+
+  fun new t a acc = {tokens = Metis_StringSet.singleton t, assoc = a} :: acc;
+
+  fun add t a' acc =
+      case acc of
+        [] => raise Bug "Metis_Print.layerInfixes: null"
+      | {tokens = ts, assoc = a} :: acc =>
+        if equalAssoc a a' then {tokens = Metis_StringSet.add ts t, assoc = a} :: acc
+        else raise Bug "Metis_Print.layerInfixes: mixed assocs";
+
+  fun layer ({token = t, precedence = p, assoc = a}, (acc,p')) =
+      let
+        val acc = if p = p' then add t a acc else new t a acc
+      in
+        (acc,p)
+      end;
+in
+  fun layerInfixes (Infixes ios) =
+      case sort comparePrecedence ios of
+        [] => []
+      | {token = t, precedence = p, assoc = a} :: ios =>
+        let
+          val acc = new t a []
+
+          val (acc,_) = List.foldl layer (acc,p) ios
+        in
+          acc
+        end;
+end;
+
+local
+  fun add ({tokens = ts, assoc = _}, tokens) = Metis_StringSet.union ts tokens;
+in
+  fun tokensLayeredInfixes l = List.foldl add Metis_StringSet.empty l;
+end;
+
+fun tokensInfixes ios = tokensLayeredInfixes (layerInfixes ios);
+
+fun destInfixOp dest tokens tm =
+    case dest tm of
+      NONE => NONE
+    | s as SOME (t,a,b) => if Metis_StringSet.member t tokens then s else NONE;
+
+fun ppLayeredInfixes dest ppTok {tokens,assoc} ppLower ppSub =
+    let
+      fun isLayer t = Metis_StringSet.member t tokens
+
+      fun ppTerm aligned (tm,r) =
+          case dest tm of
+            NONE => ppSub (tm,r)
+          | SOME (t,a,b) =>
+            if aligned andalso isLayer t then ppLayer (tm,t,a,b,r)
+            else ppLower (tm,t,a,b,r)
+
+      and ppLeft tm_r =
+          let
+            val aligned = case assoc of LeftAssoc => true | _ => false
+          in
+            ppTerm aligned tm_r
+          end
+
+      and ppLayer (tm,t,a,b,r) =
+          program
+            [ppLeft (a,true),
+             ppTok (tm,t),
+             ppRight (b,r)]
+
+      and ppRight tm_r =
+          let
+            val aligned = case assoc of RightAssoc => true | _ => false
+          in
+            ppTerm aligned tm_r
+          end
+    in
+      fn tm_t_a_b_r as (_,t,_,_,_) =>
+         if isLayer t then block Inconsistent 0 (ppLayer tm_t_a_b_r)
+         else ppLower tm_t_a_b_r
+    end;
+
+local
+  val leftBrack = mkStringSize "("
+  and rightBrack = mkStringSize ")";
+in
+  fun ppInfixes ops =
+      let
+        val layers = layerInfixes ops
+
+        val toks = tokensLayeredInfixes layers
+      in
+        fn dest => fn ppTok => fn ppSub =>
+           let
+             fun destOp tm = destInfixOp dest toks tm
+
+             fun ppInfix tm_t_a_b_r = ppLayers layers tm_t_a_b_r
+
+             and ppLayers ls (tm,t,a,b,r) =
+                 case ls of
+                   [] =>
+                   ppBracket' leftBrack rightBrack ppInfix (tm,t,a,b,false)
+                 | l :: ls =>
+                   let
+                     val ppLower = ppLayers ls
+                   in
+                     ppLayeredInfixes destOp ppTok l ppLower ppSub (tm,t,a,b,r)
+                   end
+           in
+             fn (tm,r) =>
+                case destOp tm of
+                  SOME (t,a,b) => ppInfix (tm,t,a,b,r)
+                | NONE => ppSub (tm,r)
+           end
+      end;
+end;
+
+(* ------------------------------------------------------------------------- *)
 (* Executing pretty-printers with a global line length.                      *)
 (* ------------------------------------------------------------------------- *)
 
 val lineLength = Unsynchronized.ref initialLineLength;
 
 fun toStream ppA a =
-    Metis_Stream.map (fn s => s ^ "\n")
+    Metis_Stream.map (fn {indent,line} => nSpaces indent ^ line ^ "\n")
       (execute {lineLength = !lineLength} (ppA a));
 
-fun toString ppA a =
-    case execute {lineLength = !lineLength} (ppA a) of
-      Metis_Stream.Nil => ""
-    | Metis_Stream.Cons (h,t) => Metis_Stream.foldl (fn (s,z) => z ^ "\n" ^ s) h (t ());
-
-fun trace ppX nameX x =
-    Metis_Useful.trace (toString (ppOp2 " =" ppString ppX) (nameX,x) ^ "\n");
-
-end
-
-(**** Original file: Parse.sig ****)
+local
+  fun inc {indent,line} acc = line :: nSpaces indent :: acc;
+
+  fun incn (indent_line,acc) = inc indent_line ("\n" :: acc);
+in
+  fun toString ppA a =
+      case execute {lineLength = !lineLength} (ppA a) of
+        Metis_Stream.Nil => ""
+      | Metis_Stream.Cons (h,t) =>
+        let
+          val lines = Metis_Stream.foldl incn (inc h []) (t ())
+        in
+          String.concat (rev lines)
+        end;
+end;
+
+local
+  val sep = mkStringSize " =";
+in
+  fun trace ppX nameX x =
+      Metis_Useful.trace (toString (ppOp2' sep ppString ppX) (nameX,x) ^ "\n");
+end;
+
+end
+
+(**** Original file: src/Parse.sig ****)
 
 (* ========================================================================= *)
 (* PARSING                                                                   *)
@@ -7932,8 +7948,9 @@
 (* ------------------------------------------------------------------------- *)
 
 val parseInfixes :
-    Metis_Print.infixes -> (string * 'a * 'a -> 'a) -> (string,'a) parser ->
-    (string,'a) parser
+    Metis_Print.infixes ->
+    (Metis_Print.token * 'a * 'a -> 'a) -> ('b,Metis_Print.token) parser ->
+    ('b,'a) parser -> ('b,'a) parser
 
 (* ------------------------------------------------------------------------- *)
 (* Quotations.                                                               *)
@@ -7945,7 +7962,7 @@
 
 end
 
-(**** Original file: Parse.sml ****)
+(**** Original file: src/Parse.sml ****)
 
 (* ========================================================================= *)
 (* PARSING                                                                   *)
@@ -8083,7 +8100,7 @@
                   val Unsynchronized.ref (n,_,l2,l3) = lastLine
                   val () = lastLine := (n + 1, l2, l3, line)
                 in
-                  explode line
+                  String.explode line
                 end
           in
             Metis_Stream.memoize (Metis_Stream.map saveLast lines)
@@ -8109,7 +8126,7 @@
       [] => nothing
     | c :: cs => (exactChar c ++ exactCharList cs) >> snd;
 
-fun exactString s = exactCharList (explode s);
+fun exactString s = exactCharList (String.explode s);
 
 fun escapeString {escape} =
     let
@@ -8132,7 +8149,7 @@
           ((exactChar #"\\" ++ escapeParser) >> snd) ||
           some isNormal
     in
-      many charParser >> implode
+      many charParser >> String.implode
     end;
 
 local
@@ -8145,46 +8162,51 @@
   val atLeastOneSpace = atLeastOne space >> K ();
 end;
 
-fun fromString parser s = fromList parser (explode s);
+fun fromString parser s = fromList parser (String.explode s);
 
 (* ------------------------------------------------------------------------- *)
 (* Infix operators.                                                          *)
 (* ------------------------------------------------------------------------- *)
 
-fun parseGenInfix update sof toks parse inp =
-    let
-      val (e,rest) = parse inp
-
-      val continue =
-          case rest of
-            Metis_Stream.Nil => NONE
-          | Metis_Stream.Cons (h_t as (h,_)) =>
-            if Metis_StringSet.member h toks then SOME h_t else NONE
-    in
-      case continue of
-        NONE => (sof e, rest)
-      | SOME (h,t) => parseGenInfix update (update sof h e) toks parse (t ())
-    end;
-
-local
-  fun add ({leftSpaces = _, token = t, rightSpaces = _}, s) = Metis_StringSet.add s t;
-
-  fun parse leftAssoc toks con =
-      let
-        val update =
-            if leftAssoc then (fn f => fn t => fn a => fn b => con (t, f a, b))
-            else (fn f => fn t => fn a => fn b => f (con (t, a, b)))
-      in
-        parseGenInfix update I toks
-      end;
-in
-  fun parseLayeredInfixes {tokens,leftAssoc} =
-      let
-        val toks = List.foldl add Metis_StringSet.empty tokens
-      in
-        parse leftAssoc toks
-      end;
-end;
+fun parseLayeredInfixes {tokens,assoc} mk tokParser subParser =
+    let
+      fun layerTokParser inp =
+          let
+            val tok_rest as (tok,_) = tokParser inp
+          in
+            if Metis_StringSet.member tok tokens then tok_rest
+            else raise NoParse
+          end
+
+      fun layerMk (x,txs) =
+          case assoc of
+            Metis_Print.LeftAssoc =>
+            let
+              fun inc ((t,y),z) = mk (t,z,y)
+            in
+              List.foldl inc x txs
+            end
+          | Metis_Print.NonAssoc =>
+            (case txs of
+               [] => x
+             | [(t,y)] => mk (t,x,y)
+             | _ => raise NoParse)
+          | Metis_Print.RightAssoc =>
+            (case rev txs of
+               [] => x
+             | tx :: txs =>
+               let
+                 fun inc ((t,y),(u,z)) = (t, mk (u,y,z))
+
+                 val (t,y) = List.foldl inc tx txs
+               in
+                 mk (t,x,y)
+               end)
+
+      val layerParser = subParser ++ many (layerTokParser ++ subParser)
+    in
+      layerParser >> layerMk
+    end;
 
 fun parseInfixes ops =
     let
@@ -8192,7 +8214,8 @@
 
       val iparsers = List.map parseLayeredInfixes layeredOps
     in
-      fn con => fn subparser => foldl (fn (p,sp) => p con sp) subparser iparsers
+      fn mk => fn tokParser => fn subParser =>
+         List.foldr (fn (p,sp) => p mk tokParser sp) subParser iparsers
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -8206,14 +8229,14 @@
     fun expand (QUOTE q, s) = s ^ q
       | expand (ANTIQUOTE a, s) = s ^ printer a
 
-    val string = foldl expand "" quote
+    val string = List.foldl expand "" quote
   in
     parser string
   end;
 
 end
 
-(**** Original file: Name.sig ****)
+(**** Original file: src/Name.sig ****)
 
 (* ========================================================================= *)
 (* NAMES                                                                     *)
@@ -8261,7 +8284,7 @@
 
 end
 
-(**** Original file: Name.sml ****)
+(**** Original file: src/Name.sml ****)
 
 (* ========================================================================= *)
 (* NAMES                                                                     *)
@@ -8349,24 +8372,24 @@
 structure Metis_NameSet =
 struct
 
-  local
-    structure S = Metis_ElementSet (Metis_NameMap);
-  in
-    open S;
-  end;
-
-  val pp =
-      Metis_Print.ppMap
-        toList
-        (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_Name.pp));
-
-end
-
-(**** Original file: NameArity.sig ****)
+local
+  structure S = Metis_ElementSet (Metis_NameMap);
+in
+  open S;
+end;
+
+val pp =
+    Metis_Print.ppMap
+      toList
+      (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_Name.pp));
+
+end
+
+(**** Original file: src/NameArity.sig ****)
 
 (* ========================================================================= *)
 (* NAME/ARITY PAIRS                                                          *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_NameArity =
@@ -8412,11 +8435,11 @@
 
 end
 
-(**** Original file: NameArity.sml ****)
+(**** Original file: src/NameArity.sml ****)
 
 (* ========================================================================= *)
 (* NAME/ARITY PAIRS                                                          *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_NameArity :> Metis_NameArity =
@@ -8462,7 +8485,7 @@
 fun pp (n,i) =
     Metis_Print.blockProgram Metis_Print.Inconsistent 0
       [Metis_Name.pp n,
-       Metis_Print.addString "/",
+       Metis_Print.ppString "/",
        Metis_Print.ppInt i];
 
 end
@@ -8473,40 +8496,41 @@
 structure Metis_NameArityMap =
 struct
 
-  local
-    structure S = Metis_KeyMap (Metis_NameArityOrdered);
-  in
-    open S;
-  end;
-
-  fun compose m1 m2 =
-      let
-        fun pk ((_,a),n) = peek m2 (n,a)
-      in
-        mapPartial pk m1
-      end;
+local
+  structure S = Metis_KeyMap (Metis_NameArityOrdered);
+in
+  open S;
+end;
+
+fun compose m1 m2 =
+    let
+      fun pk ((_,a),n) = peek m2 (n,a)
+    in
+      mapPartial pk m1
+    end;
 
 end
 
 structure Metis_NameAritySet =
 struct
 
-  local
-    structure S = Metis_ElementSet (Metis_NameArityMap);
-  in
-    open S;
-  end;
-
-  val allNullary = all Metis_NameArity.nullary;
-
-  val pp =
-      Metis_Print.ppMap
-        toList
-        (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_NameArity.pp));
-
-end
-
-(**** Original file: Term.sig ****)
+local
+  structure S = Metis_ElementSet (Metis_NameArityMap);
+in
+  open S;
+end;
+
+val allNullary = all Metis_NameArity.nullary;
+
+val pp =
+    Metis_Print.ppMap
+      toList
+      (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_NameArity.pp));
+
+
+end
+
+(**** Original file: src/Term.sig ****)
 
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC TERMS                                                   *)
@@ -8696,7 +8720,7 @@
 
 end
 
-(**** Original file: Term.sml ****)
+(**** Original file: src/Term.sml ****)
 
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC TERMS                                                   *)
@@ -9048,7 +9072,7 @@
 
 val isApp = can destApp;
 
-fun listMkApp (f,l) = foldl mkApp f l;
+fun listMkApp (f,l) = List.foldl mkApp f l;
 
 local
   fun strip tms tm =
@@ -9068,38 +9092,38 @@
 val infixes =
     (Unsynchronized.ref o Metis_Print.Infixes)
       [(* ML symbols *)
-       {token = " / ", precedence = 7, leftAssoc = true},
-       {token = " div ", precedence = 7, leftAssoc = true},
-       {token = " mod ", precedence = 7, leftAssoc = true},
-       {token = " * ", precedence = 7, leftAssoc = true},
-       {token = " + ", precedence = 6, leftAssoc = true},
-       {token = " - ", precedence = 6, leftAssoc = true},
-       {token = " ^ ", precedence = 6, leftAssoc = true},
-       {token = " @ ", precedence = 5, leftAssoc = false},
-       {token = " :: ", precedence = 5, leftAssoc = false},
-       {token = " = ", precedence = 4, leftAssoc = true},
-       {token = " <> ", precedence = 4, leftAssoc = true},
-       {token = " <= ", precedence = 4, leftAssoc = true},
-       {token = " < ", precedence = 4, leftAssoc = true},
-       {token = " >= ", precedence = 4, leftAssoc = true},
-       {token = " > ", precedence = 4, leftAssoc = true},
-       {token = " o ", precedence = 3, leftAssoc = true},
-       {token = " -> ", precedence = 2, leftAssoc = false},  (* inferred prec *)
-       {token = " : ", precedence = 1, leftAssoc = false},  (* inferred prec *)
-       {token = ", ", precedence = 0, leftAssoc = false},  (* inferred prec *)
+       {token = "/", precedence = 7, assoc = Metis_Print.LeftAssoc},
+       {token = "div", precedence = 7, assoc = Metis_Print.LeftAssoc},
+       {token = "mod", precedence = 7, assoc = Metis_Print.LeftAssoc},
+       {token = "*", precedence = 7, assoc = Metis_Print.LeftAssoc},
+       {token = "+", precedence = 6, assoc = Metis_Print.LeftAssoc},
+       {token = "-", precedence = 6, assoc = Metis_Print.LeftAssoc},
+       {token = "^", precedence = 6, assoc = Metis_Print.LeftAssoc},
+       {token = "@", precedence = 5, assoc = Metis_Print.RightAssoc},
+       {token = "::", precedence = 5, assoc = Metis_Print.RightAssoc},
+       {token = "=", precedence = 4, assoc = Metis_Print.NonAssoc},
+       {token = "<>", precedence = 4, assoc = Metis_Print.NonAssoc},
+       {token = "<=", precedence = 4, assoc = Metis_Print.NonAssoc},
+       {token = "<", precedence = 4, assoc = Metis_Print.NonAssoc},
+       {token = ">=", precedence = 4, assoc = Metis_Print.NonAssoc},
+       {token = ">", precedence = 4, assoc = Metis_Print.NonAssoc},
+       {token = "o", precedence = 3, assoc = Metis_Print.LeftAssoc},
+       {token = "->", precedence = 2, assoc = Metis_Print.RightAssoc},
+       {token = ":", precedence = 1, assoc = Metis_Print.NonAssoc},
+       {token = ",", precedence = 0, assoc = Metis_Print.RightAssoc},
 
        (* Logical connectives *)
-       {token = " /\\ ", precedence = ~1, leftAssoc = false},
-       {token = " \\/ ", precedence = ~2, leftAssoc = false},
-       {token = " ==> ", precedence = ~3, leftAssoc = false},
-       {token = " <=> ", precedence = ~4, leftAssoc = false},
+       {token = "/\\", precedence = ~1, assoc = Metis_Print.RightAssoc},
+       {token = "\\/", precedence = ~2, assoc = Metis_Print.RightAssoc},
+       {token = "==>", precedence = ~3, assoc = Metis_Print.RightAssoc},
+       {token = "<=>", precedence = ~4, assoc = Metis_Print.RightAssoc},
 
        (* Other symbols *)
-       {token = " . ", precedence = 9, leftAssoc = true},  (* function app *)
-       {token = " ** ", precedence = 8, leftAssoc = true},
-       {token = " ++ ", precedence = 6, leftAssoc = true},
-       {token = " -- ", precedence = 6, leftAssoc = true},
-       {token = " == ", precedence = 4, leftAssoc = true}];
+       {token = ".", precedence = 9, assoc = Metis_Print.LeftAssoc},
+       {token = "**", precedence = 8, assoc = Metis_Print.LeftAssoc},
+       {token = "++", precedence = 6, assoc = Metis_Print.LeftAssoc},
+       {token = "--", precedence = 6, assoc = Metis_Print.LeftAssoc},
+       {token = "==", precedence = 4, assoc = Metis_Print.NonAssoc}];
 
 (* The negation symbol *)
 
@@ -9122,9 +9146,14 @@
       and neg = !negation
       and bracks = !brackets
 
-      val bracks = map (fn (b1,b2) => (b1 ^ b2, b1, b2)) bracks
-
-      val bTokens = map #2 bracks @ map #3 bracks
+      val bMap =
+          let
+            fun f (b1,b2) = (b1 ^ b2, b1, b2)
+          in
+            List.map f bracks
+          end
+
+      val bTokens = op@ (unzip bracks)
 
       val iTokens = Metis_Print.tokensInfixes iOps
 
@@ -9138,7 +9167,15 @@
             end
           | _ => NONE
 
-      val iPrinter = Metis_Print.ppInfixes iOps destI
+      fun isI tm = Option.isSome (destI tm)
+
+      fun iToken (_,tok) =
+          Metis_Print.program
+            [(if tok = "," then Metis_Print.skip else Metis_Print.ppString " "),
+             Metis_Print.ppString tok,
+             Metis_Print.addBreak 1];
+
+      val iPrinter = Metis_Print.ppInfixes iOps destI iToken
 
       val specialTokens =
           Metis_StringSet.addList iTokens (neg :: quants @ ["$","(",")"] @ bTokens)
@@ -9166,8 +9203,6 @@
 
       fun functionName bv = Metis_Print.ppMap (checkFunctionName bv) Metis_Print.ppString
 
-      fun isI tm = Option.isSome (destI tm)
-
       fun stripNeg tm =
           case tm of
             Fn (f,[a]) =>
@@ -9194,7 +9229,7 @@
           let
             val s = Metis_Name.toString b
           in
-            case List.find (fn (n,_,_) => n = s) bracks of
+            case List.find (fn (n,_,_) => n = s) bMap of
               NONE => NONE
             | SOME (_,b1,b2) => SOME (b1,tm,b2)
           end
@@ -9227,11 +9262,11 @@
               val bv = Metis_StringSet.addList bv (map Metis_Name.toString (v :: vs))
             in
               Metis_Print.program
-                [Metis_Print.addString q,
+                [Metis_Print.ppString q,
                  varName bv v,
                  Metis_Print.program
                    (map (Metis_Print.sequence (Metis_Print.addBreak 1) o varName bv) vs),
-                 Metis_Print.addString ".",
+                 Metis_Print.ppString ".",
                  Metis_Print.addBreak 1,
                  innerQuant bv tm]
             end
@@ -9245,7 +9280,7 @@
             val (n,tm) = stripNeg tm
           in
             Metis_Print.blockProgram Metis_Print.Inconsistent n
-              [Metis_Print.duplicate n (Metis_Print.addString neg),
+              [Metis_Print.duplicate n (Metis_Print.ppString neg),
                if isI tm orelse (r andalso isQuant tm) then bracket bv tm
                else quantifier bv tm]
           end
@@ -9271,31 +9306,32 @@
 
   val isAlphaNum =
       let
-        val alphaNumChars = explode "_'"
+        val alphaNumChars = String.explode "_'"
       in
         fn c => mem c alphaNumChars orelse Char.isAlphaNum c
       end;
 
   local
-    val alphaNumToken = atLeastOne (some isAlphaNum) >> implode;
+    val alphaNumToken = atLeastOne (some isAlphaNum) >> String.implode;
 
     val symbolToken =
         let
           fun isNeg c = str c = !negation
 
-          val symbolChars = explode "<>=-*+/\\?@|!$%&#^:;~"
+          val symbolChars = String.explode "<>=-*+/\\?@|!$%&#^:;~"
 
           fun isSymbol c = mem c symbolChars
 
           fun isNonNegSymbol c = not (isNeg c) andalso isSymbol c
         in
           some isNeg >> str ||
-          (some isNonNegSymbol ++ many (some isSymbol)) >> (implode o op::)
+          (some isNonNegSymbol ++ many (some isSymbol)) >>
+          (String.implode o op::)
         end;
 
     val punctToken =
         let
-          val punctChars = explode "()[]{}.,"
+          val punctChars = String.explode "()[]{}.,"
 
           fun isPunct c = mem c punctChars
         in
@@ -9327,8 +9363,9 @@
 
         val iTokens = Metis_Print.tokensInfixes iOps
 
-        val iParser =
-            parseInfixes iOps (fn (f,a,b) => Fn (Metis_Name.fromString f, [a,b]))
+        fun iMk (f,a,b) = Fn (Metis_Name.fromString f, [a,b])
+
+        val iParser = parseInfixes iOps iMk any
 
         val specialTokens =
             Metis_StringSet.addList iTokens (neg :: quants @ ["$"] @ bTokens)
@@ -9367,7 +9404,7 @@
                      some (Metis_Useful.equal ".")) >>++
                     (fn (_,(vs,_)) =>
                         term (Metis_StringSet.addList bv vs) >>
-                        (fn body => foldr bind body vs))
+                        (fn body => List.foldr bind body vs))
                   end
             in
               var ||
@@ -9396,7 +9433,7 @@
 in
   fun fromString input =
       let
-        val chars = Metis_Stream.fromList (explode input)
+        val chars = Metis_Stream.fromList (String.explode input)
 
         val tokens = everything (lexer >> singleton) chars
 
@@ -9409,7 +9446,8 @@
 end;
 
 local
-  val antiquotedTermToString = Metis_Print.toString (Metis_Print.ppBracket "(" ")" pp);
+  val antiquotedTermToString =
+      Metis_Print.toString (Metis_Print.ppBracket "(" ")" pp);
 in
   val parse = Metis_Parse.parseQuotation antiquotedTermToString fromString;
 end;
@@ -9423,11 +9461,11 @@
 
 structure Metis_TermSet = Metis_ElementSet (Metis_TermMap);
 
-(**** Original file: Subst.sig ****)
+(**** Original file: src/Subst.sig ****)
 
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC SUBSTITUTIONS                                           *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Subst =
@@ -9545,11 +9583,11 @@
 
 end
 
-(**** Original file: Subst.sml ****)
+(**** Original file: src/Subst.sml ****)
 
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC SUBSTITUTIONS                                           *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Subst :> Metis_Subst =
@@ -9797,11 +9835,11 @@
 
 end
 
-(**** Original file: Atom.sig ****)
+(**** Original file: src/Atom.sig ****)
 
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC ATOMS                                                   *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Atom =
@@ -9941,11 +9979,11 @@
 
 end
 
-(**** Original file: Atom.sml ****)
+(**** Original file: src/Atom.sml ****)
 
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC ATOMS                                                   *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Atom :> Metis_Atom =
@@ -9979,14 +10017,14 @@
     let
       fun f (tm,acc) = Metis_NameAritySet.union (Metis_Term.functions tm) acc
     in
-      fn atm => foldl f Metis_NameAritySet.empty (arguments atm)
+      fn atm => List.foldl f Metis_NameAritySet.empty (arguments atm)
     end;
 
 val functionNames =
     let
       fun f (tm,acc) = Metis_NameSet.union (Metis_Term.functionNames tm) acc
     in
-      fn atm => foldl f Metis_NameSet.empty (arguments atm)
+      fn atm => List.foldl f Metis_NameSet.empty (arguments atm)
     end;
 
 (* Binary relations *)
@@ -10003,7 +10041,8 @@
 (* The size of an atom in symbols.                                           *)
 (* ------------------------------------------------------------------------- *)
 
-fun symbols atm = foldl (fn (tm,z) => Metis_Term.symbols tm + z) 1 (arguments atm);
+fun symbols atm =
+    List.foldl (fn (tm,z) => Metis_Term.symbols tm + z) 1 (arguments atm);
 
 (* ------------------------------------------------------------------------- *)
 (* A total comparison function for atoms.                                    *)
@@ -10030,7 +10069,7 @@
     let
       fun f ((n,tm),l) = map (fn (p,s) => (n :: p, s)) (Metis_Term.subterms tm) @ l
     in
-      foldl f [] (enumerate tms)
+      List.foldl f [] (enumerate tms)
     end;
 
 fun replace _ ([],_) = raise Bug "Metis_Atom.replace: empty path"
@@ -10065,7 +10104,7 @@
     let
       fun f (tm,acc) = Metis_NameSet.union (Metis_Term.freeVars tm) acc
     in
-      fn atm => foldl f Metis_NameSet.empty (arguments atm)
+      fn atm => List.foldl f Metis_NameSet.empty (arguments atm)
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -10091,7 +10130,7 @@
         val _ = (Metis_Name.equal p1 p2 andalso length tms1 = length tms2) orelse
                 raise Error "Metis_Atom.match"
       in
-        foldl matchArg sub (zip tms1 tms2)
+        List.foldl matchArg sub (zip tms1 tms2)
       end;
 end;
 
@@ -10107,7 +10146,7 @@
         val _ = (Metis_Name.equal p1 p2 andalso length tms1 = length tms2) orelse
                 raise Error "Metis_Atom.unify"
       in
-        foldl unifyArg sub (zip tms1 tms2)
+        List.foldl unifyArg sub (zip tms1 tms2)
       end;
 end;
 
@@ -10156,7 +10195,7 @@
 (* ------------------------------------------------------------------------- *)
 
 fun typedSymbols ((_,tms) : atom) =
-    foldl (fn (tm,z) => Metis_Term.typedSymbols tm + z) 1 tms;
+    List.foldl (fn (tm,z) => Metis_Term.typedSymbols tm + z) 1 tms;
 
 fun nonVarTypedSubterms (_,tms) =
     let
@@ -10164,10 +10203,10 @@
           let
             fun addTm ((path,tm),acc) = (n :: path, tm) :: acc
           in
-            foldl addTm acc (Metis_Term.nonVarTypedSubterms arg)
-          end
-    in
-      foldl addArg [] (enumerate tms)
+            List.foldl addTm acc (Metis_Term.nonVarTypedSubterms arg)
+          end
+    in
+      List.foldl addArg [] (enumerate tms)
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -10191,11 +10230,11 @@
 
 structure Metis_AtomSet = Metis_ElementSet (Metis_AtomMap);
 
-(**** Original file: Formula.sig ****)
+(**** Original file: src/Formula.sig ****)
 
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC FORMULAS                                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Formula =
@@ -10391,11 +10430,11 @@
 
 end
 
-(**** Original file: Formula.sml ****)
+(**** Original file: src/Formula.sml ****)
 
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC FORMULAS                                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Formula :> Metis_Formula =
@@ -10540,7 +10579,7 @@
 (* Conjunctions *)
 
 fun listMkConj fms =
-    case rev fms of [] => True | fm :: fms => foldl And fm fms;
+    case rev fms of [] => True | fm :: fms => List.foldl And fm fms;
 
 local
   fun strip cs (And (p,q)) = strip (p :: cs) q
@@ -10563,7 +10602,7 @@
 (* Disjunctions *)
 
 fun listMkDisj fms =
-    case rev fms of [] => False | fm :: fms => foldl Or fm fms;
+    case rev fms of [] => False | fm :: fms => List.foldl Or fm fms;
 
 local
   fun strip cs (Or (p,q)) = strip (p :: cs) q
@@ -10586,7 +10625,7 @@
 (* Equivalences *)
 
 fun listMkEquiv fms =
-    case rev fms of [] => True | fm :: fms => foldl Iff fm fms;
+    case rev fms of [] => True | fm :: fms => List.foldl Iff fm fms;
 
 local
   fun strip cs (Iff (p,q)) = strip (p :: cs) q
@@ -10995,11 +11034,11 @@
 
 structure Metis_FormulaSet = Metis_ElementSet (Metis_FormulaMap);
 
-(**** Original file: Literal.sig ****)
+(**** Original file: src/Literal.sig ****)
 
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC LITERALS                                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Literal =
@@ -11163,11 +11202,11 @@
 
 end
 
-(**** Original file: Literal.sml ****)
+(**** Original file: src/Literal.sml ****)
 
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC LITERALS                                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Literal :> Metis_Literal =
@@ -11459,7 +11498,7 @@
 
 structure Metis_LiteralSetSet = Metis_ElementSet (Metis_LiteralSetMap);
 
-(**** Original file: Thm.sig ****)
+(**** Original file: src/Thm.sig ****)
 
 (* ========================================================================= *)
 (* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS                         *)
@@ -11610,7 +11649,7 @@
 
 end
 
-(**** Original file: Thm.sml ****)
+(**** Original file: src/Thm.sml ****)
 
 (* ========================================================================= *)
 (* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS                         *)
@@ -11725,7 +11764,7 @@
 in
   fun pp th =
       Metis_Print.blockProgram Metis_Print.Inconsistent 3
-        [Metis_Print.addString "|- ",
+        [Metis_Print.ppString "|- ",
          Metis_Formula.pp (toFormula th)];
 end;
 
@@ -11827,11 +11866,11 @@
 
 end
 
-(**** Original file: Proof.sig ****)
+(**** Original file: src/Proof.sig ****)
 
 (* ========================================================================= *)
 (* PROOFS IN FIRST ORDER LOGIC                                               *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Proof =
@@ -11891,11 +11930,11 @@
 
 end
 
-(**** Original file: Proof.sml ****)
+(**** Original file: src/Proof.sml ****)
 
 (* ========================================================================= *)
 (* PROOFS IN FIRST ORDER LOGIC                                               *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Proof :> Metis_Proof =
@@ -11934,40 +11973,40 @@
   fun ppSubst ppThm (sub,thm) =
       Metis_Print.sequence (Metis_Print.addBreak 1)
         (Metis_Print.blockProgram Metis_Print.Inconsistent 1
-           [Metis_Print.addString "{",
+           [Metis_Print.ppString "{",
             Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Subst.pp ("sub",sub),
-            Metis_Print.addString ",",
+            Metis_Print.ppString ",",
             Metis_Print.addBreak 1,
             Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("thm",thm),
-            Metis_Print.addString "}"]);
+            Metis_Print.ppString "}"]);
 
   fun ppResolve ppThm (res,pos,neg) =
       Metis_Print.sequence (Metis_Print.addBreak 1)
         (Metis_Print.blockProgram Metis_Print.Inconsistent 1
-           [Metis_Print.addString "{",
+           [Metis_Print.ppString "{",
             Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Atom.pp ("res",res),
-            Metis_Print.addString ",",
+            Metis_Print.ppString ",",
             Metis_Print.addBreak 1,
             Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("pos",pos),
-            Metis_Print.addString ",",
+            Metis_Print.ppString ",",
             Metis_Print.addBreak 1,
             Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("neg",neg),
-            Metis_Print.addString "}"]);
+            Metis_Print.ppString "}"]);
 
   fun ppRefl tm = Metis_Print.sequence (Metis_Print.addBreak 1) (Metis_Term.pp tm);
 
   fun ppEquality (lit,path,res) =
       Metis_Print.sequence (Metis_Print.addBreak 1)
         (Metis_Print.blockProgram Metis_Print.Inconsistent 1
-           [Metis_Print.addString "{",
+           [Metis_Print.ppString "{",
             Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Literal.pp ("lit",lit),
-            Metis_Print.addString ",",
+            Metis_Print.ppString ",",
             Metis_Print.addBreak 1,
             Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Term.ppPath ("path",path),
-            Metis_Print.addString ",",
+            Metis_Print.ppString ",",
             Metis_Print.addBreak 1,
             Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Term.pp ("res",res),
-            Metis_Print.addString "}"]);
+            Metis_Print.ppString "}"]);
 
   fun ppInf ppAxiom ppThm inf =
       let
@@ -11975,7 +12014,7 @@
       in
         Metis_Print.block Metis_Print.Inconsistent 2
           (Metis_Print.sequence
-             (Metis_Print.addString infString)
+             (Metis_Print.ppString infString)
              (case inf of
                 Axiom cl => ppAxiom cl
               | Assume x => ppAssume x
@@ -12001,7 +12040,7 @@
         val prf = enumerate prf
 
         fun ppThm th =
-            Metis_Print.addString
+            Metis_Print.ppString
             let
               val cl = Metis_Thm.clause th
 
@@ -12018,7 +12057,7 @@
             in
               Metis_Print.sequence
                 (Metis_Print.blockProgram Metis_Print.Consistent (1 + size s)
-                   [Metis_Print.addString (s ^ " "),
+                   [Metis_Print.ppString (s ^ " "),
                     Metis_Thm.pp th,
                     Metis_Print.addBreak 2,
                     Metis_Print.ppBracket "[" "]" (ppInf (K Metis_Print.skip) ppThm) inf])
@@ -12026,10 +12065,10 @@
             end
       in
         Metis_Print.blockProgram Metis_Print.Consistent 0
-          [Metis_Print.addString "START OF PROOF",
+          [Metis_Print.ppString "START OF PROOF",
            Metis_Print.addNewline,
            Metis_Print.program (map ppStep prf),
-           Metis_Print.addString "END OF PROOF"]
+           Metis_Print.ppString "END OF PROOF"]
       end
 (*MetisDebug
       handle Error err => raise Bug ("Metis_Proof.pp: shouldn't fail:\n" ^ err);
@@ -12346,11 +12385,11 @@
 
 end
 
-(**** Original file: Rule.sig ****)
+(**** Original file: src/Rule.sig ****)
 
 (* ========================================================================= *)
 (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS                     *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Rule =
@@ -12623,11 +12662,11 @@
 
 end
 
-(**** Original file: Rule.sml ****)
+(**** Original file: src/Rule.sml ****)
 
 (* ========================================================================= *)
 (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS                     *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Rule :> Metis_Rule =
@@ -12783,17 +12822,19 @@
 
 val noConv : conv = fn _ => raise Error "noConv";
 
+(*MetisDebug
 fun traceConv s conv tm =
     let
       val res as (tm',th) = conv tm
-      val () = print (s ^ ": " ^ Metis_Term.toString tm ^ " --> " ^
+      val () = trace (s ^ ": " ^ Metis_Term.toString tm ^ " --> " ^
                       Metis_Term.toString tm' ^ " " ^ Metis_Thm.toString th ^ "\n")
     in
       res
     end
     handle Error err =>
-      (print (s ^ ": " ^ Metis_Term.toString tm ^ " --> Error: " ^ err ^ "\n");
+      (trace (s ^ ": " ^ Metis_Term.toString tm ^ " --> Error: " ^ err ^ "\n");
        raise Error (s ^ ": " ^ err));
+*)
 
 fun thenConvTrans tm (tm',th1) (tm'',th2) =
     let
@@ -13082,7 +13123,7 @@
       val reflTh = Metis_Thm.refl (Metis_Term.Fn (f,xs))
       val reflLit = Metis_Thm.destUnit reflTh
     in
-      fst (foldl cong (reflTh,reflLit) (enumerate ys))
+      fst (List.foldl cong (reflTh,reflLit) (enumerate ys))
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -13109,7 +13150,7 @@
       val assumeLit = (false,(R,xs))
       val assumeTh = Metis_Thm.assume assumeLit
     in
-      fst (foldl cong (assumeTh,assumeLit) (enumerate ys))
+      fst (List.foldl cong (assumeTh,assumeLit) (enumerate ys))
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -13409,11 +13450,11 @@
 
 end
 
-(**** Original file: Normalize.sig ****)
+(**** Original file: src/Normalize.sig ****)
 
 (* ========================================================================= *)
 (* NORMALIZING FORMULAS                                                      *)
-(* Copyright (c) 2001-2009 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Normalize =
@@ -13467,11 +13508,11 @@
 
 end
 
-(**** Original file: Normalize.sml ****)
+(**** Original file: src/Normalize.sml ****)
 
 (* ========================================================================= *)
 (* NORMALIZING FORMULAS                                                      *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Normalize :> Metis_Normalize =
@@ -14159,19 +14200,20 @@
 val newSkolemFunction =
     let
       val counter : int Metis_StringMap.map Unsynchronized.ref = Unsynchronized.ref (Metis_StringMap.new ())
-    in
-      (* MODIFIED by Jasmin Blanchette *)
-      fn n => CRITICAL (fn () =>
-         let
-           val Unsynchronized.ref m = counter
-           val s = Metis_Name.toString n
-           val i = Option.getOpt (Metis_StringMap.peek m s, 0)
-           val () = counter := Metis_StringMap.insert m (s, i + 1)
-           val i = if i = 0 then "" else "_" ^ Int.toString i
-           val s = skolemPrefix ^ "_" ^ s ^ i
-         in
-           Metis_Name.fromString s
-         end)
+
+      fun new n () =
+          let
+            val Unsynchronized.ref m = counter
+            val s = Metis_Name.toString n
+            val i = Option.getOpt (Metis_StringMap.peek m s, 0)
+            val () = counter := Metis_StringMap.insert m (s, i + 1)
+            val i = if i = 0 then "" else "_" ^ Int.toString i
+            val s = skolemPrefix ^ "_" ^ s ^ i
+          in
+            Metis_Name.fromString s
+          end
+    in
+      fn n => Metis_Portable.critical (new n) ()
     end;
 
 fun skolemize fv bv fm =
@@ -14286,8 +14328,8 @@
                     (c2', s2', (c1,s1,fm,c2,s2) :: l)
                   end
 
-              val (c1,_,fms) = foldl fwd (count0,empty,[]) fms
-              val (c2,_,fms) = foldl bwd (count0,empty,[]) fms
+              val (c1,_,fms) = List.foldl fwd (count0,empty,[]) fms
+              val (c2,_,fms) = List.foldl bwd (count0,empty,[]) fms
 
 (*MetisDebug
               val _ = countEquivish c1 c2 orelse
@@ -14332,10 +14374,10 @@
             let
               val fms = sortMap (measure o count) countCompare fms
             in
-              foldl breakSet1 best (cumulatives fms)
-            end
-
-        val best = foldl breakSing best (cumulatives fms)
+              List.foldl breakSet1 best (cumulatives fms)
+            end
+
+        val best = List.foldl breakSing best (cumulatives fms)
         val best = breakSet I best
         val best = breakSet countNegate best
         val best = breakSet countClauses best
@@ -14715,14 +14757,13 @@
     let
       val counter : int Unsynchronized.ref = Unsynchronized.ref 0
     in
-      (* MODIFIED by Jasmin Blanchette *)
-      fn () => CRITICAL (fn () =>
+      fn () => Metis_Portable.critical (fn () =>
          let
            val Unsynchronized.ref i = counter
            val () = counter := i + 1
          in
            definitionPrefix ^ "_" ^ Int.toString i
-         end)
+         end) ()
     end;
 
 fun newDefinition def =
@@ -14858,7 +14899,7 @@
 
 end
 
-(**** Original file: Model.sig ****)
+(**** Original file: src/Model.sig ****)
 
 (* ========================================================================= *)
 (* RANDOM FINITE MODELS                                                      *)
@@ -15138,7 +15179,7 @@
 
 end
 
-(**** Original file: Model.sml ****)
+(**** Original file: src/Model.sml ****)
 
 (* ========================================================================= *)
 (* RANDOM FINITE MODELS                                                      *)
@@ -15414,10 +15455,10 @@
 
   fun ppEntry (tag,source_arity,target) =
       Metis_Print.blockProgram Metis_Print.Inconsistent 2
-        [Metis_Print.addString tag,
+        [Metis_Print.ppString tag,
          Metis_Print.addBreak 1,
          Metis_NameArity.pp source_arity,
-         Metis_Print.addString " ->",
+         Metis_Print.ppString " ->",
          Metis_Print.addBreak 1,
          Metis_Name.pp target];
 in
@@ -16332,7 +16373,7 @@
                 let
                   fun add (y,acc) = FunctionPerturbation (func_xs,y) :: acc
                 in
-                  foldl add acc target
+                  List.foldl add acc target
                 end
         in
           pertTerms M onTarget tms xs acc
@@ -16413,17 +16454,17 @@
 
 fun pp M =
     Metis_Print.program
-      [Metis_Print.addString "Metis_Model{",
+      [Metis_Print.ppString "Metis_Model{",
        Metis_Print.ppInt (size M),
-       Metis_Print.addString "}"];
-
-end
-
-(**** Original file: Problem.sig ****)
+       Metis_Print.ppString "}"];
+
+end
+
+(**** Original file: src/Problem.sig ****)
 
 (* ========================================================================= *)
 (* CNF PROBLEMS                                                              *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Problem =
@@ -16485,11 +16526,11 @@
 
 end
 
-(**** Original file: Problem.sml ****)
+(**** Original file: src/Problem.sml ****)
 
 (* ========================================================================= *)
 (* CNF PROBLEMS                                                              *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Problem :> Metis_Problem =
@@ -16518,9 +16559,9 @@
       val cls = toClauses prob
     in
       {clauses = length cls,
-       literals = foldl lits 0 cls,
-       symbols = foldl syms 0 cls,
-       typedSymbols = foldl typedSyms 0 cls}
+       literals = List.foldl lits 0 cls,
+       symbols = List.foldl syms 0 cls,
+       typedSymbols = List.foldl typedSyms 0 cls}
     end;
 
 fun freeVars {axioms,conjecture} =
@@ -16648,11 +16689,11 @@
 
 end
 
-(**** Original file: TermNet.sig ****)
+(**** Original file: src/TermNet.sig ****)
 
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS              *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_TermNet =
@@ -16701,11 +16742,11 @@
 
 end
 
-(**** Original file: TermNet.sml ****)
+(**** Original file: src/TermNet.sml ****)
 
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS              *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_TermNet :> Metis_TermNet =
@@ -16863,7 +16904,7 @@
 
 fun null net = size net = 0;
 
-fun singles qtms a = foldr Single a qtms;
+fun singles qtms a = List.foldr Single a qtms;
 
 local
   fun pre NONE = (0,NONE)
@@ -16893,7 +16934,7 @@
       handle Error _ => raise Bug "Metis_TermNet.insert: should never fail";
 end;
 
-fun fromList parm l = foldl (fn (tm_a,n) => insert n tm_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (tm_a,n) => insert n tm_a) (new parm) l;
 
 fun filter pred =
     let
@@ -17146,7 +17187,7 @@
 
 local
   fun inc (qtm, Result l, acc) =
-      foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l
+      List.foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l
     | inc _ = raise Bug "Metis_TermNet.pp.inc";
 
   fun toList (Net (_,_,NONE)) = []
@@ -17159,11 +17200,11 @@
 
 end
 
-(**** Original file: AtomNet.sig ****)
+(**** Original file: src/AtomNet.sig ****)
 
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS              *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_AtomNet =
@@ -17210,11 +17251,11 @@
 
 end
 
-(**** Original file: AtomNet.sml ****)
+(**** Original file: src/AtomNet.sml ****)
 
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS              *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_AtomNet :> Metis_AtomNet =
@@ -17249,7 +17290,7 @@
 
 fun insert net (atm,a) = Metis_TermNet.insert net (atomToTerm atm, a);
 
-fun fromList parm l = foldl (fn (atm_a,n) => insert n atm_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (atm_a,n) => insert n atm_a) (new parm) l;
 
 val filter = Metis_TermNet.filter;
 
@@ -17272,11 +17313,11 @@
 
 end
 
-(**** Original file: LiteralNet.sig ****)
+(**** Original file: src/LiteralNet.sig ****)
 
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS           *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_LiteralNet =
@@ -17325,11 +17366,11 @@
 
 end
 
-(**** Original file: LiteralNet.sml ****)
+(**** Original file: src/LiteralNet.sml ****)
 
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS           *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_LiteralNet :> Metis_LiteralNet =
@@ -17368,7 +17409,7 @@
   | insert {positive,negative} ((false,atm),a) =
     {positive = positive, negative = Metis_AtomNet.insert negative (atm,a)};
 
-fun fromList parm l = foldl (fn (lit_a,n) => insert n lit_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (lit_a,n) => insert n lit_a) (new parm) l;
 
 fun filter pred {positive,negative} =
     {positive = Metis_AtomNet.filter pred positive,
@@ -17402,11 +17443,11 @@
 
 end
 
-(**** Original file: Subsume.sig ****)
+(**** Original file: src/Subsume.sig ****)
 
 (* ========================================================================= *)
 (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES                        *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Subsume =
@@ -17456,11 +17497,11 @@
 
 end
 
-(**** Original file: Subsume.sml ****)
+(**** Original file: src/Subsume.sml ****)
 
 (* ========================================================================= *)
 (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES                        *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Subsume :> Metis_Subsume =
@@ -17793,11 +17834,11 @@
 
 end
 
-(**** Original file: KnuthBendixOrder.sig ****)
+(**** Original file: src/KnuthBendixOrder.sig ****)
 
 (* ========================================================================= *)
 (* THE KNUTH-BENDIX TERM ORDERING                                            *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_KnuthBendixOrder =
@@ -17818,11 +17859,11 @@
 
 end
 
-(**** Original file: KnuthBendixOrder.sml ****)
+(**** Original file: src/KnuthBendixOrder.sml ****)
 
 (* ========================================================================= *)
 (* KNUTH-BENDIX TERM ORDERING CONSTRAINTS                                    *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_KnuthBendixOrder :> Metis_KnuthBendixOrder =
@@ -17921,7 +17962,7 @@
 val ppWeightList =
     let
       fun ppCoeff n =
-          if n < 0 then Metis_Print.sequence (Metis_Print.addString "~") (ppCoeff (~n))
+          if n < 0 then Metis_Print.sequence (Metis_Print.ppString "~") (ppCoeff (~n))
           else if n = 1 then Metis_Print.skip
           else Metis_Print.ppInt n
 
@@ -18020,11 +18061,11 @@
 
 end
 
-(**** Original file: Rewrite.sig ****)
+(**** Original file: src/Rewrite.sig ****)
 
 (* ========================================================================= *)
 (* ORDERED REWRITING FOR FIRST ORDER TERMS                                   *)
-(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Rewrite =
@@ -18122,11 +18163,11 @@
 
 end
 
-(**** Original file: Rewrite.sml ****)
+(**** Original file: src/Rewrite.sml ****)
 
 (* ========================================================================= *)
 (* ORDERED REWRITING FOR FIRST ORDER TERMS                                   *)
-(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Rewrite :> Metis_Rewrite =
@@ -18333,7 +18374,7 @@
         rw
       end;
 
-val addList = foldl (fn (eqn,rw) => add rw eqn);
+val addList = List.foldl (fn (eqn,rw) => add rw eqn);
 
 (* ------------------------------------------------------------------------- *)
 (* Rewriting (the order must be a refinement of the rewrite order).          *)
@@ -18394,7 +18435,7 @@
           val th = Metis_Rule.symmetryRule l r
         in
           fn tm =>
-             if Metis_Term.equal tm r then (l,th) else raise Error "mkNeqConv: Metis_RL"
+             if Metis_Term.equal tm r then (l,th) else raise Error "mkNeqConv: RL"
         end
       | SOME EQUAL => raise Error "irreflexive"
       | SOME GREATER =>
@@ -18560,8 +18601,9 @@
     let
       fun addSubterm b ((path,tm),net) = Metis_TermNet.insert net (tm,(id,b,path))
 
-      val subterms = foldl (addSubterm true) subterms (Metis_Term.subterms l)
-      val subterms = foldl (addSubterm false) subterms (Metis_Term.subterms r)
+      val subterms = List.foldl (addSubterm true) subterms (Metis_Term.subterms l)
+
+      val subterms = List.foldl (addSubterm false) subterms (Metis_Term.subterms r)
     in
       subterms
     end;
@@ -18786,7 +18828,7 @@
 in
   fun orderedRewrite order ths =
     let
-      val rw = foldl addEqn (new order) (enumerate ths)
+      val rw = List.foldl addEqn (new order) (enumerate ths)
     in
       rewriteRule rw order
     end;
@@ -18796,11 +18838,11 @@
 
 end
 
-(**** Original file: Units.sig ****)
+(**** Original file: src/Units.sig ****)
 
 (* ========================================================================= *)
 (* A STORE FOR UNIT THEOREMS                                                 *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Units =
@@ -18848,11 +18890,11 @@
 
 end
 
-(**** Original file: Units.sml ****)
+(**** Original file: src/Units.sml ****)
 
 (* ========================================================================= *)
 (* A STORE FOR UNIT THEOREMS                                                 *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Units :> Metis_Units =
@@ -18899,7 +18941,7 @@
         end
     end;
 
-val addList = foldl (fn (th,u) => add u th);
+val addList = List.foldl (fn (th,u) => add u th);
 
 (* ------------------------------------------------------------------------- *)
 (* Matching.                                                                 *)
@@ -18957,11 +18999,11 @@
 
 end
 
-(**** Original file: Clause.sig ****)
+(**** Original file: src/Clause.sig ****)
 
 (* ========================================================================= *)
 (* CLAUSE = ID + THEOREM                                                     *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Clause =
@@ -19067,11 +19109,11 @@
 
 end
 
-(**** Original file: Clause.sml ****)
+(**** Original file: src/Clause.sml ****)
 
 (* ========================================================================= *)
 (* CLAUSE = ID + THEOREM                                                     *)
-(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Clause :> Metis_Clause =
@@ -19086,10 +19128,17 @@
 val newId =
     let
       val r = Unsynchronized.ref 0
-    in
-      (* MODIFIED by Jasmin Blanchette *)
-      fn () => CRITICAL (fn () =>
-        case r of Unsynchronized.ref n => let val () = r := n + 1 in n end)
+
+      fun new () =
+          let
+            val Unsynchronized.ref n = r
+
+            val () = r := n + 1
+          in
+            n
+          end
+    in
+      fn () => Metis_Portable.critical new ()
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -19133,7 +19182,7 @@
 
 val default : parameters =
     {ordering = Metis_KnuthBendixOrder.default,
-     orderLiterals = UnsignedLiteralOrder (* PositiveLiteralOrder *) (* MODIFIED by Jasmin Blanchette *),
+     orderLiterals = UnsignedLiteralOrder,
      orderTerms = true};
 
 fun mk info = Metis_Clause info
@@ -19252,7 +19301,7 @@
       let
         fun addTm ((path,tm),acc) = (lit,path,tm) :: acc
       in
-        foldl addTm acc (Metis_Literal.nonVarTypedSubterms lit)
+        List.foldl addTm acc (Metis_Literal.nonVarTypedSubterms lit)
       end;
 in
   fun largestSubterms cl = Metis_LiteralSet.foldl addLit [] (largestLiterals cl);
@@ -19432,11 +19481,11 @@
 
 end
 
-(**** Original file: Active.sig ****)
+(**** Original file: src/Active.sig ****)
 
 (* ========================================================================= *)
 (* THE ACTIVE SET OF CLAUSES                                                 *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Active =
@@ -19490,11 +19539,11 @@
 
 end
 
-(**** Original file: Active.sml ****)
+(**** Original file: src/Active.sml ****)
 
 (* ========================================================================= *)
 (* THE ACTIVE SET OF CLAUSES                                                 *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Active :> Metis_Active =
@@ -19519,7 +19568,7 @@
               | NONE => rw
             end
       in
-        foldl add (Metis_Rewrite.new (Metis_KnuthBendixOrder.compare ordering))
+        List.foldl add (Metis_Rewrite.new (Metis_KnuthBendixOrder.compare ordering))
       end;
 
   fun allFactors red =
@@ -19787,8 +19836,8 @@
           end
 
       val cls = clauses active
-      val (cls,_) = foldl remove ([], Metis_Subsume.new ()) cls
-      val (cls,subs) = foldl remove ([], Metis_Subsume.new ()) cls
+      val (cls,_) = List.foldl remove ([], Metis_Subsume.new ()) cls
+      val (cls,subs) = List.foldl remove ([], Metis_Subsume.new ()) cls
 
 (*MetisDebug
       val Metis_Active {parameters,...} = active
@@ -19815,7 +19864,7 @@
 local
   fun ppField f ppA a =
       Metis_Print.blockProgram Metis_Print.Inconsistent 2
-        [Metis_Print.addString (f ^ " ="),
+        [Metis_Print.ppString (f ^ " ="),
          Metis_Print.addBreak 1,
          ppA a];
 
@@ -19836,21 +19885,21 @@
 in
   fun pp (Metis_Active {clauses,rewrite,subterms,...}) =
       Metis_Print.blockProgram Metis_Print.Inconsistent 2
-        [Metis_Print.addString "Metis_Active",
+        [Metis_Print.ppString "Metis_Active",
          Metis_Print.addBreak 1,
          Metis_Print.blockProgram Metis_Print.Inconsistent 1
-           [Metis_Print.addString "{",
+           [Metis_Print.ppString "{",
             ppClauses clauses,
-            Metis_Print.addString ",",
+            Metis_Print.ppString ",",
             Metis_Print.addBreak 1,
             ppRewrite rewrite,
 (*MetisTrace5
-            Metis_Print.addString ",",
+            Metis_Print.ppString ",",
             Metis_Print.addBreak 1,
             ppSubterms subterms,
 *)
             Metis_Print.skip],
-         Metis_Print.addString "}"];
+         Metis_Print.ppString "}"];
 end;
 *)
 
@@ -19968,7 +20017,7 @@
       fun add ((lit,ort,tm),equations) =
           Metis_TermNet.insert equations (tm,(cl,lit,ort,tm))
     in
-      foldl add equations (Metis_Clause.largestEquations cl)
+      List.foldl add equations (Metis_Clause.largestEquations cl)
     end;
 
 fun addSubterms subterms cl =
@@ -19976,7 +20025,7 @@
       fun add ((lit,path,tm),subterms) =
           Metis_TermNet.insert subterms (tm,(cl,lit,path,tm))
     in
-      foldl add subterms (Metis_Clause.largestSubterms cl)
+      List.foldl add subterms (Metis_Clause.largestSubterms cl)
     end;
 
 fun addAllSubterms allSubterms cl =
@@ -19984,7 +20033,7 @@
       fun add ((_,_,tm),allSubterms) =
           Metis_TermNet.insert allSubterms (tm,(cl,tm))
     in
-      foldl add allSubterms (Metis_Clause.allSubterms cl)
+      List.foldl add allSubterms (Metis_Clause.allSubterms cl)
     end;
 
 fun addClause active cl =
@@ -20035,7 +20084,8 @@
 *)
     in
       if Metis_Atom.isEq atm then acc
-      else foldl resolve acc (Metis_LiteralNet.unify literals (Metis_Literal.negate lit))
+      else
+        List.foldl resolve acc (Metis_LiteralNet.unify literals (Metis_Literal.negate lit))
     end;
 
 fun deduceParamodulationWith subterms cl ((lit,ort,tm),acc) =
@@ -20045,7 +20095,7 @@
             SOME cl' => cl' :: acc
           | NONE => acc
     in
-      foldl para acc (Metis_TermNet.unify subterms tm)
+      List.foldl para acc (Metis_TermNet.unify subterms tm)
     end;
 
 fun deduceParamodulationInto equations cl ((lit,path,tm),acc) =
@@ -20055,7 +20105,7 @@
             SOME cl' => cl' :: acc
           | NONE => acc
     in
-      foldl para acc (Metis_TermNet.unify equations tm)
+      List.foldl para acc (Metis_TermNet.unify equations tm)
     end;
 
 fun deduce active cl =
@@ -20081,8 +20131,8 @@
 
       val acc = []
       val acc = Metis_LiteralSet.foldl (deduceResolution literals cl) acc lits
-      val acc = foldl (deduceParamodulationWith subterms cl) acc eqns
-      val acc = foldl (deduceParamodulationInto equations cl) acc subtms
+      val acc = List.foldl (deduceParamodulationWith subterms cl) acc eqns
+      val acc = List.foldl (deduceParamodulationInto equations cl) acc subtms
       val acc = rev acc
 
 (*MetisTrace5
@@ -20328,7 +20378,7 @@
         let
           val cls = sort_utilitywise (cl :: Metis_Clause.factor cl)
         in
-          foldl factor_add active_subsume_acc cls
+          List.foldl factor_add active_subsume_acc cls
         end;
 
   fun factor' active acc [] = (active, rev acc)
@@ -20336,7 +20386,7 @@
       let
         val cls = sort_utilitywise cls
         val subsume = getSubsume active
-        val (active,_,acc) = foldl factor1 (active,subsume,acc) cls
+        val (active,_,acc) = List.foldl factor1 (active,subsume,acc) cls
         val (active,cls) = extract_rewritables active
       in
         factor' active acc cls
@@ -20404,11 +20454,11 @@
 
 end
 
-(**** Original file: Waiting.sig ****)
+(**** Original file: src/Waiting.sig ****)
 
 (* ========================================================================= *)
 (* THE WAITING SET OF CLAUSES                                                *)
-(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Waiting =
@@ -20484,11 +20534,11 @@
 
 end
 
-(**** Original file: Waiting.sml ****)
+(**** Original file: src/Waiting.sml ****)
 
 (* ========================================================================= *)
 (* THE WAITING SET OF CLAUSES                                                *)
-(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Waiting :> Metis_Waiting =
@@ -20528,17 +20578,16 @@
 (* ------------------------------------------------------------------------- *)
 
 val defaultModels : modelParameters list =
-    [(* MODIFIED by Jasmin Blanchette
-      {model = Metis_Model.default,
+    [{model = Metis_Model.default,
       initialPerturbations = 100,
       maxChecks = SOME 20,
       perturbations = 0,
-      weight = 1.0} *)];
+      weight = 1.0}];
 
 val default : parameters =
      {symbolsWeight = 1.0,
-      literalsWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *)
-      variablesWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *)
+      literalsWeight = 1.0,
+      variablesWeight = 1.0,
       models = defaultModels};
 
 fun size (Metis_Waiting {clauses,...}) = Metis_Heap.size clauses;
@@ -20651,7 +20700,7 @@
         val symbolsW = Math.pow (clauseSymbols lits, symbolsWeight)
         val variablesW = Math.pow (clauseVariables lits, variablesWeight)
         val literalsW = Math.pow (clauseLiterals lits, literalsWeight)
-        val modelsW = (* checkModels models mods mcl *) 1.0 (* MODIFIED by Jasmin Blanchette *)
+        val modelsW = checkModels models mods mcl
 (*MetisTrace4
         val () = trace ("Metis_Waiting.clauseWeight: dist = " ^
                         Real.toString dist ^ "\n")
@@ -20758,11 +20807,11 @@
 
 end
 
-(**** Original file: Resolution.sig ****)
+(**** Original file: src/Resolution.sig ****)
 
 (* ========================================================================= *)
 (* THE RESOLUTION PROOF PROCEDURE                                            *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Metis_Resolution =
@@ -20812,11 +20861,11 @@
 
 end
 
-(**** Original file: Resolution.sml ****)
+(**** Original file: src/Resolution.sml ****)
 
 (* ========================================================================= *)
 (* THE RESOLUTION PROOF PROCEDURE                                            *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Metis_Resolution :> Metis_Resolution =
--- a/src/Tools/Metis/scripts/mlpp	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/scripts/mlpp	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 #!/usr/bin/perl
 
-# Copyright (c) 2006 Joe Hurd, All Rights Reserved
+# Copyright (c) 2006 Joe Hurd, distributed under the BSD License
 
 use strict;
 use warnings;
@@ -78,6 +78,8 @@
         $text =~ s/Array\.modifyi/Array_modifyi/g;
         $text =~ s/OS\.Process\.isSuccess/OS_Process_isSuccess/g;
         $text =~ s/PP\.ppstream/ppstream/g;
+        $text =~ s/String\.concatWith/String_concatWith/g;
+        $text =~ s/String\.isSubstring/String_isSubstring/g;
         $text =~ s/String\.isSuffix/String_isSuffix/g;
         $text =~ s/Substring\.full/Substring_full/g;
         $text =~ s/TextIO\.inputLine/TextIO_inputLine/g;
--- a/src/Tools/Metis/src/Active.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Active.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* THE ACTIVE SET OF CLAUSES                                                 *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Active =
--- a/src/Tools/Metis/src/Active.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Active.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* THE ACTIVE SET OF CLAUSES                                                 *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Active :> Active =
@@ -25,7 +25,7 @@
               | NONE => rw
             end
       in
-        foldl add (Rewrite.new (KnuthBendixOrder.compare ordering))
+        List.foldl add (Rewrite.new (KnuthBendixOrder.compare ordering))
       end;
 
   fun allFactors red =
@@ -293,8 +293,8 @@
           end
 
       val cls = clauses active
-      val (cls,_) = foldl remove ([], Subsume.new ()) cls
-      val (cls,subs) = foldl remove ([], Subsume.new ()) cls
+      val (cls,_) = List.foldl remove ([], Subsume.new ()) cls
+      val (cls,subs) = List.foldl remove ([], Subsume.new ()) cls
 
 (*MetisDebug
       val Active {parameters,...} = active
@@ -321,7 +321,7 @@
 local
   fun ppField f ppA a =
       Print.blockProgram Print.Inconsistent 2
-        [Print.addString (f ^ " ="),
+        [Print.ppString (f ^ " ="),
          Print.addBreak 1,
          ppA a];
 
@@ -342,21 +342,21 @@
 in
   fun pp (Active {clauses,rewrite,subterms,...}) =
       Print.blockProgram Print.Inconsistent 2
-        [Print.addString "Active",
+        [Print.ppString "Active",
          Print.addBreak 1,
          Print.blockProgram Print.Inconsistent 1
-           [Print.addString "{",
+           [Print.ppString "{",
             ppClauses clauses,
-            Print.addString ",",
+            Print.ppString ",",
             Print.addBreak 1,
             ppRewrite rewrite,
 (*MetisTrace5
-            Print.addString ",",
+            Print.ppString ",",
             Print.addBreak 1,
             ppSubterms subterms,
 *)
             Print.skip],
-         Print.addString "}"];
+         Print.ppString "}"];
 end;
 *)
 
@@ -474,7 +474,7 @@
       fun add ((lit,ort,tm),equations) =
           TermNet.insert equations (tm,(cl,lit,ort,tm))
     in
-      foldl add equations (Clause.largestEquations cl)
+      List.foldl add equations (Clause.largestEquations cl)
     end;
 
 fun addSubterms subterms cl =
@@ -482,7 +482,7 @@
       fun add ((lit,path,tm),subterms) =
           TermNet.insert subterms (tm,(cl,lit,path,tm))
     in
-      foldl add subterms (Clause.largestSubterms cl)
+      List.foldl add subterms (Clause.largestSubterms cl)
     end;
 
 fun addAllSubterms allSubterms cl =
@@ -490,7 +490,7 @@
       fun add ((_,_,tm),allSubterms) =
           TermNet.insert allSubterms (tm,(cl,tm))
     in
-      foldl add allSubterms (Clause.allSubterms cl)
+      List.foldl add allSubterms (Clause.allSubterms cl)
     end;
 
 fun addClause active cl =
@@ -541,7 +541,8 @@
 *)
     in
       if Atom.isEq atm then acc
-      else foldl resolve acc (LiteralNet.unify literals (Literal.negate lit))
+      else
+        List.foldl resolve acc (LiteralNet.unify literals (Literal.negate lit))
     end;
 
 fun deduceParamodulationWith subterms cl ((lit,ort,tm),acc) =
@@ -551,7 +552,7 @@
             SOME cl' => cl' :: acc
           | NONE => acc
     in
-      foldl para acc (TermNet.unify subterms tm)
+      List.foldl para acc (TermNet.unify subterms tm)
     end;
 
 fun deduceParamodulationInto equations cl ((lit,path,tm),acc) =
@@ -561,7 +562,7 @@
             SOME cl' => cl' :: acc
           | NONE => acc
     in
-      foldl para acc (TermNet.unify equations tm)
+      List.foldl para acc (TermNet.unify equations tm)
     end;
 
 fun deduce active cl =
@@ -587,8 +588,8 @@
 
       val acc = []
       val acc = LiteralSet.foldl (deduceResolution literals cl) acc lits
-      val acc = foldl (deduceParamodulationWith subterms cl) acc eqns
-      val acc = foldl (deduceParamodulationInto equations cl) acc subtms
+      val acc = List.foldl (deduceParamodulationWith subterms cl) acc eqns
+      val acc = List.foldl (deduceParamodulationInto equations cl) acc subtms
       val acc = rev acc
 
 (*MetisTrace5
@@ -834,7 +835,7 @@
         let
           val cls = sort_utilitywise (cl :: Clause.factor cl)
         in
-          foldl factor_add active_subsume_acc cls
+          List.foldl factor_add active_subsume_acc cls
         end;
 
   fun factor' active acc [] = (active, rev acc)
@@ -842,7 +843,7 @@
       let
         val cls = sort_utilitywise cls
         val subsume = getSubsume active
-        val (active,_,acc) = foldl factor1 (active,subsume,acc) cls
+        val (active,_,acc) = List.foldl factor1 (active,subsume,acc) cls
         val (active,cls) = extract_rewritables active
       in
         factor' active acc cls
--- a/src/Tools/Metis/src/Atom.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Atom.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC ATOMS                                                   *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Atom =
--- a/src/Tools/Metis/src/Atom.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Atom.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC ATOMS                                                   *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Atom :> Atom =
@@ -34,14 +34,14 @@
     let
       fun f (tm,acc) = NameAritySet.union (Term.functions tm) acc
     in
-      fn atm => foldl f NameAritySet.empty (arguments atm)
+      fn atm => List.foldl f NameAritySet.empty (arguments atm)
     end;
 
 val functionNames =
     let
       fun f (tm,acc) = NameSet.union (Term.functionNames tm) acc
     in
-      fn atm => foldl f NameSet.empty (arguments atm)
+      fn atm => List.foldl f NameSet.empty (arguments atm)
     end;
 
 (* Binary relations *)
@@ -58,7 +58,8 @@
 (* The size of an atom in symbols.                                           *)
 (* ------------------------------------------------------------------------- *)
 
-fun symbols atm = foldl (fn (tm,z) => Term.symbols tm + z) 1 (arguments atm);
+fun symbols atm =
+    List.foldl (fn (tm,z) => Term.symbols tm + z) 1 (arguments atm);
 
 (* ------------------------------------------------------------------------- *)
 (* A total comparison function for atoms.                                    *)
@@ -85,7 +86,7 @@
     let
       fun f ((n,tm),l) = map (fn (p,s) => (n :: p, s)) (Term.subterms tm) @ l
     in
-      foldl f [] (enumerate tms)
+      List.foldl f [] (enumerate tms)
     end;
 
 fun replace _ ([],_) = raise Bug "Atom.replace: empty path"
@@ -120,7 +121,7 @@
     let
       fun f (tm,acc) = NameSet.union (Term.freeVars tm) acc
     in
-      fn atm => foldl f NameSet.empty (arguments atm)
+      fn atm => List.foldl f NameSet.empty (arguments atm)
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -146,7 +147,7 @@
         val _ = (Name.equal p1 p2 andalso length tms1 = length tms2) orelse
                 raise Error "Atom.match"
       in
-        foldl matchArg sub (zip tms1 tms2)
+        List.foldl matchArg sub (zip tms1 tms2)
       end;
 end;
 
@@ -162,7 +163,7 @@
         val _ = (Name.equal p1 p2 andalso length tms1 = length tms2) orelse
                 raise Error "Atom.unify"
       in
-        foldl unifyArg sub (zip tms1 tms2)
+        List.foldl unifyArg sub (zip tms1 tms2)
       end;
 end;
 
@@ -211,7 +212,7 @@
 (* ------------------------------------------------------------------------- *)
 
 fun typedSymbols ((_,tms) : atom) =
-    foldl (fn (tm,z) => Term.typedSymbols tm + z) 1 tms;
+    List.foldl (fn (tm,z) => Term.typedSymbols tm + z) 1 tms;
 
 fun nonVarTypedSubterms (_,tms) =
     let
@@ -219,10 +220,10 @@
           let
             fun addTm ((path,tm),acc) = (n :: path, tm) :: acc
           in
-            foldl addTm acc (Term.nonVarTypedSubterms arg)
+            List.foldl addTm acc (Term.nonVarTypedSubterms arg)
           end
     in
-      foldl addArg [] (enumerate tms)
+      List.foldl addArg [] (enumerate tms)
     end;
 
 (* ------------------------------------------------------------------------- *)
--- a/src/Tools/Metis/src/AtomNet.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/AtomNet.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS              *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature AtomNet =
--- a/src/Tools/Metis/src/AtomNet.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/AtomNet.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS              *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure AtomNet :> AtomNet =
@@ -35,7 +35,7 @@
 
 fun insert net (atm,a) = TermNet.insert net (atomToTerm atm, a);
 
-fun fromList parm l = foldl (fn (atm_a,n) => insert n atm_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (atm_a,n) => insert n atm_a) (new parm) l;
 
 val filter = TermNet.filter;
 
--- a/src/Tools/Metis/src/Clause.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Clause.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* CLAUSE = ID + THEOREM                                                     *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Clause =
--- a/src/Tools/Metis/src/Clause.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Clause.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* CLAUSE = ID + THEOREM                                                     *)
-(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Clause :> Clause =
@@ -15,10 +15,17 @@
 val newId =
     let
       val r = ref 0
+
+      fun new () =
+          let
+            val ref n = r
+
+            val () = r := n + 1
+          in
+            n
+          end
     in
-      (* MODIFIED by Jasmin Blanchette *)
-      fn () => CRITICAL (fn () =>
-        case r of ref n => let val () = r := n + 1 in n end)
+      fn () => Portable.critical new ()
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -62,7 +69,7 @@
 
 val default : parameters =
     {ordering = KnuthBendixOrder.default,
-     orderLiterals = UnsignedLiteralOrder (* PositiveLiteralOrder *) (* MODIFIED by Jasmin Blanchette *),
+     orderLiterals = UnsignedLiteralOrder,
      orderTerms = true};
 
 fun mk info = Clause info
@@ -181,7 +188,7 @@
       let
         fun addTm ((path,tm),acc) = (lit,path,tm) :: acc
       in
-        foldl addTm acc (Literal.nonVarTypedSubterms lit)
+        List.foldl addTm acc (Literal.nonVarTypedSubterms lit)
       end;
 in
   fun largestSubterms cl = LiteralSet.foldl addLit [] (largestLiterals cl);
--- a/src/Tools/Metis/src/Formula.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Formula.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC FORMULAS                                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Formula =
--- a/src/Tools/Metis/src/Formula.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Formula.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC FORMULAS                                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Formula :> Formula =
@@ -145,7 +145,7 @@
 (* Conjunctions *)
 
 fun listMkConj fms =
-    case rev fms of [] => True | fm :: fms => foldl And fm fms;
+    case rev fms of [] => True | fm :: fms => List.foldl And fm fms;
 
 local
   fun strip cs (And (p,q)) = strip (p :: cs) q
@@ -168,7 +168,7 @@
 (* Disjunctions *)
 
 fun listMkDisj fms =
-    case rev fms of [] => False | fm :: fms => foldl Or fm fms;
+    case rev fms of [] => False | fm :: fms => List.foldl Or fm fms;
 
 local
   fun strip cs (Or (p,q)) = strip (p :: cs) q
@@ -191,7 +191,7 @@
 (* Equivalences *)
 
 fun listMkEquiv fms =
-    case rev fms of [] => True | fm :: fms => foldl Iff fm fms;
+    case rev fms of [] => True | fm :: fms => List.foldl Iff fm fms;
 
 local
   fun strip cs (Iff (p,q)) = strip (p :: cs) q
--- a/src/Tools/Metis/src/Heap.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Heap.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* A HEAP DATATYPE FOR ML                                                    *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Heap =
--- a/src/Tools/Metis/src/Heap.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Heap.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* A HEAP DATATYPE FOR ML                                                    *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Heap :> Heap =
--- a/src/Tools/Metis/src/KeyMap.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/KeyMap.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -887,38 +887,40 @@
 (* ------------------------------------------------------------------------- *)
 
 datatype 'value iterator =
-    LR of (key * 'value) * 'value tree * 'value node list
-  | RL of (key * 'value) * 'value tree * 'value node list;
+    LeftToRightIterator of
+      (key * 'value) * 'value tree * 'value node list
+  | RightToLeftIterator of
+      (key * 'value) * 'value tree * 'value node list;
 
-fun fromSpineLR nodes =
+fun fromSpineLeftToRightIterator nodes =
     case nodes of
       [] => NONE
     | Node {key,value,right,...} :: nodes =>
-      SOME (LR ((key,value),right,nodes));
+      SOME (LeftToRightIterator ((key,value),right,nodes));
 
-fun fromSpineRL nodes =
+fun fromSpineRightToLeftIterator nodes =
     case nodes of
       [] => NONE
     | Node {key,value,left,...} :: nodes =>
-      SOME (RL ((key,value),left,nodes));
+      SOME (RightToLeftIterator ((key,value),left,nodes));
 
-fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
+fun addLeftToRightIterator nodes tree = fromSpineLeftToRightIterator (treeLeftSpine nodes tree);
 
-fun addRL nodes tree = fromSpineRL (treeRightSpine nodes tree);
+fun addRightToLeftIterator nodes tree = fromSpineRightToLeftIterator (treeRightSpine nodes tree);
 
-fun treeMkIterator tree = addLR [] tree;
+fun treeMkIterator tree = addLeftToRightIterator [] tree;
 
-fun treeMkRevIterator tree = addRL [] tree;
+fun treeMkRevIterator tree = addRightToLeftIterator [] tree;
 
 fun readIterator iter =
     case iter of
-      LR (key_value,_,_) => key_value
-    | RL (key_value,_,_) => key_value;
+      LeftToRightIterator (key_value,_,_) => key_value
+    | RightToLeftIterator (key_value,_,_) => key_value;
 
 fun advanceIterator iter =
     case iter of
-      LR (_,tree,nodes) => addLR nodes tree
-    | RL (_,tree,nodes) => addRL nodes tree;
+      LeftToRightIterator (_,tree,nodes) => addLeftToRightIterator nodes tree
+    | RightToLeftIterator (_,tree,nodes) => addRightToLeftIterator nodes tree;
 
 fun foldIterator f acc io =
     case io of
--- a/src/Tools/Metis/src/KnuthBendixOrder.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/KnuthBendixOrder.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* THE KNUTH-BENDIX TERM ORDERING                                            *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature KnuthBendixOrder =
--- a/src/Tools/Metis/src/KnuthBendixOrder.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/KnuthBendixOrder.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* KNUTH-BENDIX TERM ORDERING CONSTRAINTS                                    *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure KnuthBendixOrder :> KnuthBendixOrder =
@@ -99,7 +99,7 @@
 val ppWeightList =
     let
       fun ppCoeff n =
-          if n < 0 then Print.sequence (Print.addString "~") (ppCoeff (~n))
+          if n < 0 then Print.sequence (Print.ppString "~") (ppCoeff (~n))
           else if n = 1 then Print.skip
           else Print.ppInt n
 
--- a/src/Tools/Metis/src/Literal.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Literal.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC LITERALS                                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Literal =
--- a/src/Tools/Metis/src/Literal.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Literal.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC LITERALS                                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Literal :> Literal =
--- a/src/Tools/Metis/src/LiteralNet.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/LiteralNet.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS           *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature LiteralNet =
--- a/src/Tools/Metis/src/LiteralNet.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/LiteralNet.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS           *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure LiteralNet :> LiteralNet =
@@ -39,7 +39,7 @@
   | insert {positive,negative} ((false,atm),a) =
     {positive = positive, negative = AtomNet.insert negative (atm,a)};
 
-fun fromList parm l = foldl (fn (lit_a,n) => insert n lit_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (lit_a,n) => insert n lit_a) (new parm) l;
 
 fun filter pred {positive,negative} =
     {positive = AtomNet.filter pred positive,
--- a/src/Tools/Metis/src/Map.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Map.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -879,38 +879,40 @@
 (* ------------------------------------------------------------------------- *)
 
 datatype ('key,'value) iterator =
-    LR of ('key * 'value) * ('key,'value) tree * ('key,'value) node list
-  | RL of ('key * 'value) * ('key,'value) tree * ('key,'value) node list;
+    LeftToRightIterator of
+      ('key * 'value) * ('key,'value) tree * ('key,'value) node list
+  | RightToLeftIterator of
+      ('key * 'value) * ('key,'value) tree * ('key,'value) node list;
 
-fun fromSpineLR nodes =
+fun fromSpineLeftToRightIterator nodes =
     case nodes of
       [] => NONE
     | Node {key,value,right,...} :: nodes =>
-      SOME (LR ((key,value),right,nodes));
+      SOME (LeftToRightIterator ((key,value),right,nodes));
 
-fun fromSpineRL nodes =
+fun fromSpineRightToLeftIterator nodes =
     case nodes of
       [] => NONE
     | Node {key,value,left,...} :: nodes =>
-      SOME (RL ((key,value),left,nodes));
+      SOME (RightToLeftIterator ((key,value),left,nodes));
 
-fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
+fun addLeftToRightIterator nodes tree = fromSpineLeftToRightIterator (treeLeftSpine nodes tree);
 
-fun addRL nodes tree = fromSpineRL (treeRightSpine nodes tree);
+fun addRightToLeftIterator nodes tree = fromSpineRightToLeftIterator (treeRightSpine nodes tree);
 
-fun treeMkIterator tree = addLR [] tree;
+fun treeMkIterator tree = addLeftToRightIterator [] tree;
 
-fun treeMkRevIterator tree = addRL [] tree;
+fun treeMkRevIterator tree = addRightToLeftIterator [] tree;
 
 fun readIterator iter =
     case iter of
-      LR (key_value,_,_) => key_value
-    | RL (key_value,_,_) => key_value;
+      LeftToRightIterator (key_value,_,_) => key_value
+    | RightToLeftIterator (key_value,_,_) => key_value;
 
 fun advanceIterator iter =
     case iter of
-      LR (_,tree,nodes) => addLR nodes tree
-    | RL (_,tree,nodes) => addRL nodes tree;
+      LeftToRightIterator (_,tree,nodes) => addLeftToRightIterator nodes tree
+    | RightToLeftIterator (_,tree,nodes) => addRightToLeftIterator nodes tree;
 
 fun foldIterator f acc io =
     case io of
--- a/src/Tools/Metis/src/Model.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Model.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -272,10 +272,10 @@
 
   fun ppEntry (tag,source_arity,target) =
       Print.blockProgram Print.Inconsistent 2
-        [Print.addString tag,
+        [Print.ppString tag,
          Print.addBreak 1,
          NameArity.pp source_arity,
-         Print.addString " ->",
+         Print.ppString " ->",
          Print.addBreak 1,
          Name.pp target];
 in
@@ -1190,7 +1190,7 @@
                 let
                   fun add (y,acc) = FunctionPerturbation (func_xs,y) :: acc
                 in
-                  foldl add acc target
+                  List.foldl add acc target
                 end
         in
           pertTerms M onTarget tms xs acc
@@ -1271,8 +1271,8 @@
 
 fun pp M =
     Print.program
-      [Print.addString "Model{",
+      [Print.ppString "Model{",
        Print.ppInt (size M),
-       Print.addString "}"];
+       Print.ppString "}"];
 
 end
--- a/src/Tools/Metis/src/Name.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Name.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -84,15 +84,15 @@
 structure NameSet =
 struct
 
-  local
-    structure S = ElementSet (NameMap);
-  in
-    open S;
-  end;
+local
+  structure S = ElementSet (NameMap);
+in
+  open S;
+end;
 
-  val pp =
-      Print.ppMap
-        toList
-        (Print.ppBracket "{" "}" (Print.ppOpList "," Name.pp));
+val pp =
+    Print.ppMap
+      toList
+      (Print.ppBracket "{" "}" (Print.ppOpList "," Name.pp));
 
 end
--- a/src/Tools/Metis/src/NameArity.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/NameArity.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* NAME/ARITY PAIRS                                                          *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature NameArity =
--- a/src/Tools/Metis/src/NameArity.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/NameArity.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* NAME/ARITY PAIRS                                                          *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure NameArity :> NameArity =
@@ -46,7 +46,7 @@
 fun pp (n,i) =
     Print.blockProgram Print.Inconsistent 0
       [Name.pp n,
-       Print.addString "/",
+       Print.ppString "/",
        Print.ppInt i];
 
 end
@@ -57,35 +57,36 @@
 structure NameArityMap =
 struct
 
-  local
-    structure S = KeyMap (NameArityOrdered);
-  in
-    open S;
-  end;
+local
+  structure S = KeyMap (NameArityOrdered);
+in
+  open S;
+end;
 
-  fun compose m1 m2 =
-      let
-        fun pk ((_,a),n) = peek m2 (n,a)
-      in
-        mapPartial pk m1
-      end;
+fun compose m1 m2 =
+    let
+      fun pk ((_,a),n) = peek m2 (n,a)
+    in
+      mapPartial pk m1
+    end;
 
 end
 
 structure NameAritySet =
 struct
 
-  local
-    structure S = ElementSet (NameArityMap);
-  in
-    open S;
-  end;
+local
+  structure S = ElementSet (NameArityMap);
+in
+  open S;
+end;
 
-  val allNullary = all NameArity.nullary;
+val allNullary = all NameArity.nullary;
 
-  val pp =
-      Print.ppMap
-        toList
-        (Print.ppBracket "{" "}" (Print.ppOpList "," NameArity.pp));
+val pp =
+    Print.ppMap
+      toList
+      (Print.ppBracket "{" "}" (Print.ppOpList "," NameArity.pp));
+
 
 end
--- a/src/Tools/Metis/src/Normalize.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Normalize.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* NORMALIZING FORMULAS                                                      *)
-(* Copyright (c) 2001-2009 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Normalize =
--- a/src/Tools/Metis/src/Normalize.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Normalize.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* NORMALIZING FORMULAS                                                      *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Normalize :> Normalize =
@@ -688,19 +688,20 @@
 val newSkolemFunction =
     let
       val counter : int StringMap.map ref = ref (StringMap.new ())
+
+      fun new n () =
+          let
+            val ref m = counter
+            val s = Name.toString n
+            val i = Option.getOpt (StringMap.peek m s, 0)
+            val () = counter := StringMap.insert m (s, i + 1)
+            val i = if i = 0 then "" else "_" ^ Int.toString i
+            val s = skolemPrefix ^ "_" ^ s ^ i
+          in
+            Name.fromString s
+          end
     in
-      (* MODIFIED by Jasmin Blanchette *)
-      fn n => CRITICAL (fn () =>
-         let
-           val ref m = counter
-           val s = Name.toString n
-           val i = Option.getOpt (StringMap.peek m s, 0)
-           val () = counter := StringMap.insert m (s, i + 1)
-           val i = if i = 0 then "" else "_" ^ Int.toString i
-           val s = skolemPrefix ^ "_" ^ s ^ i
-         in
-           Name.fromString s
-         end)
+      fn n => Portable.critical (new n) ()
     end;
 
 fun skolemize fv bv fm =
@@ -815,8 +816,8 @@
                     (c2', s2', (c1,s1,fm,c2,s2) :: l)
                   end
 
-              val (c1,_,fms) = foldl fwd (count0,empty,[]) fms
-              val (c2,_,fms) = foldl bwd (count0,empty,[]) fms
+              val (c1,_,fms) = List.foldl fwd (count0,empty,[]) fms
+              val (c2,_,fms) = List.foldl bwd (count0,empty,[]) fms
 
 (*MetisDebug
               val _ = countEquivish c1 c2 orelse
@@ -861,10 +862,10 @@
             let
               val fms = sortMap (measure o count) countCompare fms
             in
-              foldl breakSet1 best (cumulatives fms)
+              List.foldl breakSet1 best (cumulatives fms)
             end
 
-        val best = foldl breakSing best (cumulatives fms)
+        val best = List.foldl breakSing best (cumulatives fms)
         val best = breakSet I best
         val best = breakSet countNegate best
         val best = breakSet countClauses best
@@ -1244,14 +1245,13 @@
     let
       val counter : int ref = ref 0
     in
-      (* MODIFIED by Jasmin Blanchette *)
-      fn () => CRITICAL (fn () =>
+      fn () => Portable.critical (fn () =>
          let
            val ref i = counter
            val () = counter := i + 1
          in
            definitionPrefix ^ "_" ^ Int.toString i
-         end)
+         end) ()
     end;
 
 fun newDefinition def =
--- a/src/Tools/Metis/src/Options.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Options.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* PROCESSING COMMAND LINE OPTIONS                                           *)
-(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Options =
--- a/src/Tools/Metis/src/Options.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Options.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* PROCESSING COMMAND LINE OPTIONS                                           *)
-(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Options :> Options =
@@ -146,7 +146,7 @@
           fun indent (s, "" :: l) = indent (s ^ "  ", l) | indent x = x
           val (res,n) = indent ("  ",n)
           val res = res ^ join ", " n
-          val res = foldl (fn (x,y) => y ^ " " ^ x) res r
+          val res = List.foldl (fn (x,y) => y ^ " " ^ x) res r
         in
           [res ^ " ...", " " ^ s]
         end
@@ -185,7 +185,7 @@
     exit allopts {message = SOME mesg, usage = true, success = false};
 
 fun version allopts =
-    (print (versionInformation allopts);
+    (TextIO.print (versionInformation allopts);
      exit allopts {message = NONE, usage = false, success = true});
 
 (* ------------------------------------------------------------------------- *)
@@ -220,7 +220,8 @@
       | process ("-v" :: _) = version allopts
       | process ("--version" :: _) = version allopts
       | process (x :: xs) =
-      if x = "" orelse x = "-" orelse hd (explode x) <> #"-" then ([], x :: xs)
+      if x = "" orelse x = "-" orelse hd (String.explode x) <> #"-" then
+        ([], x :: xs)
       else
         let
           val (r,f) = findOption x
@@ -233,7 +234,8 @@
     fn l =>
     let
       val (a,b) = process l
-      val a = foldl (fn ((x,xs),ys) => x :: xs @ ys) [] (rev a)
+
+      val a = List.foldl (fn ((x,xs),ys) => x :: xs @ ys) [] (rev a)
     in
       (a,b)
     end
--- a/src/Tools/Metis/src/Ordered.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Ordered.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* ORDERED TYPES                                                             *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Ordered =
--- a/src/Tools/Metis/src/Ordered.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Ordered.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* ORDERED TYPES                                                             *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure IntOrdered =
--- a/src/Tools/Metis/src/Parse.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Parse.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -99,8 +99,9 @@
 (* ------------------------------------------------------------------------- *)
 
 val parseInfixes :
-    Print.infixes -> (string * 'a * 'a -> 'a) -> (string,'a) parser ->
-    (string,'a) parser
+    Print.infixes ->
+    (Print.token * 'a * 'a -> 'a) -> ('b,Print.token) parser ->
+    ('b,'a) parser -> ('b,'a) parser
 
 (* ------------------------------------------------------------------------- *)
 (* Quotations.                                                               *)
--- a/src/Tools/Metis/src/Parse.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Parse.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -134,7 +134,7 @@
                   val ref (n,_,l2,l3) = lastLine
                   val () = lastLine := (n + 1, l2, l3, line)
                 in
-                  explode line
+                  String.explode line
                 end
           in
             Stream.memoize (Stream.map saveLast lines)
@@ -160,7 +160,7 @@
       [] => nothing
     | c :: cs => (exactChar c ++ exactCharList cs) >> snd;
 
-fun exactString s = exactCharList (explode s);
+fun exactString s = exactCharList (String.explode s);
 
 fun escapeString {escape} =
     let
@@ -183,7 +183,7 @@
           ((exactChar #"\\" ++ escapeParser) >> snd) ||
           some isNormal
     in
-      many charParser >> implode
+      many charParser >> String.implode
     end;
 
 local
@@ -196,46 +196,51 @@
   val atLeastOneSpace = atLeastOne space >> K ();
 end;
 
-fun fromString parser s = fromList parser (explode s);
+fun fromString parser s = fromList parser (String.explode s);
 
 (* ------------------------------------------------------------------------- *)
 (* Infix operators.                                                          *)
 (* ------------------------------------------------------------------------- *)
 
-fun parseGenInfix update sof toks parse inp =
+fun parseLayeredInfixes {tokens,assoc} mk tokParser subParser =
     let
-      val (e,rest) = parse inp
-
-      val continue =
-          case rest of
-            Stream.Nil => NONE
-          | Stream.Cons (h_t as (h,_)) =>
-            if StringSet.member h toks then SOME h_t else NONE
-    in
-      case continue of
-        NONE => (sof e, rest)
-      | SOME (h,t) => parseGenInfix update (update sof h e) toks parse (t ())
-    end;
+      fun layerTokParser inp =
+          let
+            val tok_rest as (tok,_) = tokParser inp
+          in
+            if StringSet.member tok tokens then tok_rest
+            else raise NoParse
+          end
 
-local
-  fun add ({leftSpaces = _, token = t, rightSpaces = _}, s) = StringSet.add s t;
+      fun layerMk (x,txs) =
+          case assoc of
+            Print.LeftAssoc =>
+            let
+              fun inc ((t,y),z) = mk (t,z,y)
+            in
+              List.foldl inc x txs
+            end
+          | Print.NonAssoc =>
+            (case txs of
+               [] => x
+             | [(t,y)] => mk (t,x,y)
+             | _ => raise NoParse)
+          | Print.RightAssoc =>
+            (case rev txs of
+               [] => x
+             | tx :: txs =>
+               let
+                 fun inc ((t,y),(u,z)) = (t, mk (u,y,z))
 
-  fun parse leftAssoc toks con =
-      let
-        val update =
-            if leftAssoc then (fn f => fn t => fn a => fn b => con (t, f a, b))
-            else (fn f => fn t => fn a => fn b => f (con (t, a, b)))
-      in
-        parseGenInfix update I toks
-      end;
-in
-  fun parseLayeredInfixes {tokens,leftAssoc} =
-      let
-        val toks = List.foldl add StringSet.empty tokens
-      in
-        parse leftAssoc toks
-      end;
-end;
+                 val (t,y) = List.foldl inc tx txs
+               in
+                 mk (t,x,y)
+               end)
+
+      val layerParser = subParser ++ many (layerTokParser ++ subParser)
+    in
+      layerParser >> layerMk
+    end;
 
 fun parseInfixes ops =
     let
@@ -243,7 +248,8 @@
 
       val iparsers = List.map parseLayeredInfixes layeredOps
     in
-      fn con => fn subparser => foldl (fn (p,sp) => p con sp) subparser iparsers
+      fn mk => fn tokParser => fn subParser =>
+         List.foldr (fn (p,sp) => p mk tokParser sp) subParser iparsers
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -257,7 +263,7 @@
     fun expand (QUOTE q, s) = s ^ q
       | expand (ANTIQUOTE a, s) = s ^ printer a
 
-    val string = foldl expand "" quote
+    val string = List.foldl expand "" quote
   in
     parser string
   end;
--- a/src/Tools/Metis/src/Portable.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Portable.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
-(* ML SPECIFIC FUNCTIONS                                                     *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
+(* ML COMPILER SPECIFIC FUNCTIONS                                            *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Portable =
@@ -19,17 +19,10 @@
 val pointerEqual : 'a * 'a -> bool
 
 (* ------------------------------------------------------------------------- *)
-(* Timing function applications.                                             *)
+(* Marking critical sections of code.                                        *)
 (* ------------------------------------------------------------------------- *)
 
-val time : ('a -> 'b) -> 'a -> 'b
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing)                                 *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-val CRITICAL: (unit -> 'a) -> 'a
+val critical : (unit -> 'a) -> unit -> 'a
 
 (* ------------------------------------------------------------------------- *)
 (* Generating random values.                                                 *)
@@ -43,4 +36,10 @@
 
 val randomWord : unit -> Word.word
 
+(* ------------------------------------------------------------------------- *)
+(* Timing function applications.                                             *)
+(* ------------------------------------------------------------------------- *)
+
+val time : ('a -> 'b) -> 'a -> 'b
+
 end
--- a/src/Tools/Metis/src/PortableMlton.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/PortableMlton.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* MLTON SPECIFIC FUNCTIONS                                                  *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Portable :> Portable =
@@ -19,6 +19,33 @@
 val pointerEqual = MLton.eq;
 
 (* ------------------------------------------------------------------------- *)
+(* Marking critical sections of code.                                        *)
+(* ------------------------------------------------------------------------- *)
+
+fun critical f () = f ();
+
+(* ------------------------------------------------------------------------- *)
+(* Generating random values.                                                 *)
+(* ------------------------------------------------------------------------- *)
+
+fun randomWord () = MLton.Random.rand ();
+
+fun randomBool () = Word.andb (randomWord (),0w1) = 0w0;
+
+fun randomInt 1 = 0
+  | randomInt 2 = Word.toInt (Word.andb (randomWord (), 0w1))
+  | randomInt 4 = Word.toInt (Word.andb (randomWord (), 0w3))
+  | randomInt n = Word.toInt (Word.mod (randomWord (), Word.fromInt n));
+
+local
+  fun wordToReal w = Real.fromInt (Word.toInt (Word.>> (w,0w1)))
+
+  val normalizer = 1.0 / wordToReal (Word.notb 0w0);
+in
+  fun randomReal () = normalizer * wordToReal (randomWord ());
+end;
+
+(* ------------------------------------------------------------------------- *)
 (* Timing function applications.                                             *)
 (* ------------------------------------------------------------------------- *)
 
@@ -56,34 +83,6 @@
       y
     end;
 
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing)                                 *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-fun CRITICAL e = e ();     (*dummy*)
-
-(* ------------------------------------------------------------------------- *)
-(* Generating random values.                                                 *)
-(* ------------------------------------------------------------------------- *)
-
-fun randomWord () = MLton.Random.rand ();
-
-fun randomBool () = Word.andb (randomWord (),0w1) = 0w0;
-
-fun randomInt 1 = 0
-  | randomInt 2 = Word.toInt (Word.andb (randomWord (), 0w1))
-  | randomInt 4 = Word.toInt (Word.andb (randomWord (), 0w3))
-  | randomInt n = Word.toInt (Word.mod (randomWord (), Word.fromInt n));
-
-local
-  fun wordToReal w = Real.fromInt (Word.toInt (Word.>> (w,0w1)))
-
-  val normalizer = 1.0 / wordToReal (Word.notb 0w0);
-in
-  fun randomReal () = normalizer * wordToReal (randomWord ());
-end;
-
 end
 
 (* ------------------------------------------------------------------------- *)
--- a/src/Tools/Metis/src/PortableMosml.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/PortableMosml.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -23,17 +23,10 @@
 end;
 
 (* ------------------------------------------------------------------------- *)
-(* Timing function applications.                                             *)
+(* Marking critical sections of code.                                        *)
 (* ------------------------------------------------------------------------- *)
 
-val time = Mosml.time;
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing)                                 *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-fun CRITICAL e = e ();     (*dummy*)
+fun critical f () = f ();
 
 (* ------------------------------------------------------------------------- *)
 (* Generating random values.                                                 *)
@@ -57,6 +50,12 @@
       Word.orb (Word.<< (h,0w16), l)
     end;
 
+(* ------------------------------------------------------------------------- *)
+(* Timing function applications.                                             *)
+(* ------------------------------------------------------------------------- *)
+
+val time = Mosml.time;
+
 end
 
 (* ------------------------------------------------------------------------- *)
@@ -68,6 +67,18 @@
 val _ = catch_interrupt true;
 
 (* ------------------------------------------------------------------------- *)
+(* Forcing fully qualified names of some key functions.                      *)
+(* ------------------------------------------------------------------------- *)
+
+(*BasicDebug
+val explode = ()
+and foldl = ()
+and foldr = ()
+and implode = ()
+and print = ();
+*)
+
+(* ------------------------------------------------------------------------- *)
 (* Ad-hoc upgrading of the Moscow ML basis library.                          *)
 (* ------------------------------------------------------------------------- *)
 
@@ -101,6 +112,28 @@
 
 fun OS_Process_isSuccess s = s = OS.Process.success;
 
+fun String_concatWith s l =
+    case l of
+      [] => ""
+    | h :: t => List.foldl (fn (x,y) => y ^ s ^ x) h t;
+
+fun String_isSubstring p s =
+    let
+      val sizeP = size p
+      and sizeS = size s
+    in
+      if sizeP > sizeS then false
+      else if sizeP = sizeS then p = s
+      else
+        let
+          fun check i = String.substring (s,i,sizeP) = p
+
+          fun checkn i = check i orelse (i > 0 andalso checkn (i - 1))
+        in
+          checkn (sizeS - sizeP)
+        end
+    end;
+
 fun String_isSuffix p s =
     let
       val sizeP = size p
--- a/src/Tools/Metis/src/PortablePolyml.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/PortablePolyml.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -19,6 +19,24 @@
 fun pointerEqual (x : 'a, y : 'a) = PolyML.pointerEq(x,y);
 
 (* ------------------------------------------------------------------------- *)
+(* Marking critical sections of code.                                        *)
+(* ------------------------------------------------------------------------- *)
+
+fun critical f () = f ();
+
+(* ------------------------------------------------------------------------- *)
+(* Generating random values.                                                 *)
+(* ------------------------------------------------------------------------- *)
+
+val randomWord = Random.nextWord;
+
+val randomBool = Random.nextBool;
+
+val randomInt = Random.nextInt;
+
+val randomReal = Random.nextReal;
+
+(* ------------------------------------------------------------------------- *)
 (* Timing function applications.                                             *)
 (* ------------------------------------------------------------------------- *)
 
@@ -44,7 +62,7 @@
             val {usr,sys} = Timer.checkCPUTimer c
             val real = Timer.checkRealTimer r
           in
-            TextIO.print (* MODIFIED by Jasmin Blanchette *)
+            print
               ("User: " ^ p usr ^ "  System: " ^ p sys ^
                "  Real: " ^ p real ^ "\n")
           end
@@ -56,26 +74,6 @@
       y
     end;
 
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing)                                 *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-fun CRITICAL e = NAMED_CRITICAL "metis" e;
-
-
-(* ------------------------------------------------------------------------- *)
-(* Generating random values.                                                 *)
-(* ------------------------------------------------------------------------- *)
-
-val randomWord = Random.nextWord;
-
-val randomBool = Random.nextBool;
-
-val randomInt = Random.nextInt;
-
-val randomReal = Random.nextReal;
-
 end
 
 (* ------------------------------------------------------------------------- *)
--- a/src/Tools/Metis/src/Print.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Print.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,27 +1,39 @@
 (* ========================================================================= *)
 (* PRETTY-PRINTING                                                           *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Print =
 sig
 
 (* ------------------------------------------------------------------------- *)
+(* Escaping strings.                                                         *)
+(* ------------------------------------------------------------------------- *)
+
+val escapeString : {escape : char list} -> string -> string
+
+(* ------------------------------------------------------------------------- *)
+(* A type of strings annotated with their size.                              *)
+(* ------------------------------------------------------------------------- *)
+
+type stringSize = string * int
+
+val mkStringSize : string -> stringSize
+
+(* ------------------------------------------------------------------------- *)
 (* A type of pretty-printers.                                                *)
 (* ------------------------------------------------------------------------- *)
 
 datatype breakStyle = Consistent | Inconsistent
 
-datatype ppStep =
+datatype step =
     BeginBlock of breakStyle * int
   | EndBlock
-  | AddString of string
+  | AddString of stringSize
   | AddBreak of int
   | AddNewline
 
-type ppstream = ppStep Stream.stream
-
-type 'a pp = 'a -> ppstream
+type ppstream = step Stream.stream
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty-printer primitives.                                                *)
@@ -31,7 +43,7 @@
 
 val endBlock : ppstream
 
-val addString : string -> ppstream
+val addString : stringSize -> ppstream
 
 val addBreak : int -> ppstream
 
@@ -51,18 +63,24 @@
 
 val blockProgram : breakStyle -> int -> ppstream list -> ppstream
 
-val bracket : string -> string -> ppstream -> ppstream
+(* ------------------------------------------------------------------------- *)
+(* Executing pretty-printers to generate lines.                              *)
+(* ------------------------------------------------------------------------- *)
 
-val field : string -> ppstream -> ppstream
-
-val record : (string * ppstream) list -> ppstream
+val execute :
+    {lineLength : int} -> ppstream ->
+    {indent : int, line : string} Stream.stream
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty-printer combinators.                                               *)
 (* ------------------------------------------------------------------------- *)
 
+type 'a pp = 'a -> ppstream
+
 val ppMap : ('a -> 'b) -> 'b pp -> 'a pp
 
+val ppString : string pp
+
 val ppBracket : string -> string -> 'a pp -> 'a pp
 
 val ppOp : string -> ppstream
@@ -81,8 +99,6 @@
 
 val ppChar : char pp
 
-val ppString : string pp
-
 val ppEscapeString : {escape : char list} -> string pp
 
 val ppUnit : unit pp
@@ -111,7 +127,7 @@
 
 val ppBreakStyle : breakStyle pp
 
-val ppPpStep : ppStep pp
+val ppStep : step pp
 
 val ppPpstream : ppstream pp
 
@@ -119,28 +135,27 @@
 (* Pretty-printing infix operators.                                          *)
 (* ------------------------------------------------------------------------- *)
 
+type token = string
+
+datatype assoc =
+    LeftAssoc
+  | NonAssoc
+  | RightAssoc
+
 datatype infixes =
     Infixes of
-      {token : string,
+      {token : token,
        precedence : int,
-       leftAssoc : bool} list
+       assoc : assoc} list
 
 val tokensInfixes : infixes -> StringSet.set
 
-val layerInfixes :
-    infixes ->
-    {tokens : {leftSpaces : int, token : string, rightSpaces : int} list,
-     leftAssoc : bool} list
+val layerInfixes : infixes -> {tokens : StringSet.set, assoc : assoc} list
 
 val ppInfixes :
-    infixes -> ('a -> (string * 'a * 'a) option) -> ('a * bool) pp ->
-    ('a * bool) pp
-
-(* ------------------------------------------------------------------------- *)
-(* Executing pretty-printers to generate lines.                              *)
-(* ------------------------------------------------------------------------- *)
-
-val execute : {lineLength : int} -> ppstream -> string Stream.stream
+    infixes ->
+    ('a -> (token * 'a * 'a) option) -> ('a * token) pp ->
+    ('a * bool) pp -> ('a * bool) pp
 
 (* ------------------------------------------------------------------------- *)
 (* Executing pretty-printers with a global line length.                      *)
--- a/src/Tools/Metis/src/Print.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Print.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* PRETTY-PRINTING                                                           *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Print :> Print =
@@ -29,23 +29,47 @@
     | Stream.Cons (h,t) => revAppend h (revConcat o t);
 
 local
-  fun join current prev = (prev ^ "\n", current);
+  fun calcSpaces n = nChars #" " n;
+
+  val cacheSize = 2 * initialLineLength;
+
+  val cachedSpaces = Vector.tabulate (cacheSize,calcSpaces);
 in
-  fun joinNewline strm =
-      case strm of
-        Stream.Nil => Stream.Nil
-      | Stream.Cons (h,t) => Stream.maps join Stream.singleton h (t ());
+  fun nSpaces n =
+      if n < cacheSize then Vector.sub (cachedSpaces,n)
+      else calcSpaces n;
 end;
 
-local
-  fun calcSpaces n = nChars #" " n;
+(* ------------------------------------------------------------------------- *)
+(* Escaping strings.                                                         *)
+(* ------------------------------------------------------------------------- *)
+
+fun escapeString {escape} =
+    let
+      val escapeMap = map (fn c => (c, "\\" ^ str c)) escape
 
-  val cachedSpaces = Vector.tabulate (initialLineLength,calcSpaces);
-in
-  fun nSpaces n =
-      if n < initialLineLength then Vector.sub (cachedSpaces,n)
-      else calcSpaces n;
-end;
+      fun escapeChar c =
+          case c of
+            #"\\" => "\\\\"
+          | #"\n" => "\\n"
+          | #"\t" => "\\t"
+          | _ =>
+            case List.find (equal c o fst) escapeMap of
+              SOME (_,s) => s
+            | NONE => str c
+    in
+      String.translate escapeChar
+    end;
+
+(* ------------------------------------------------------------------------- *)
+(* A type of strings annotated with their size.                              *)
+(* ------------------------------------------------------------------------- *)
+
+type stringSize = string * int;
+
+fun mkStringSize s = (s, size s);
+
+val emptyStringSize = mkStringSize "";
 
 (* ------------------------------------------------------------------------- *)
 (* A type of pretty-printers.                                                *)
@@ -53,23 +77,21 @@
 
 datatype breakStyle = Consistent | Inconsistent;
 
-datatype ppStep =
+datatype step =
     BeginBlock of breakStyle * int
   | EndBlock
-  | AddString of string
+  | AddString of stringSize
   | AddBreak of int
   | AddNewline;
 
-type ppstream = ppStep Stream.stream;
-
-type 'a pp = 'a -> ppstream;
+type ppstream = step Stream.stream;
 
 fun breakStyleToString style =
     case style of
       Consistent => "Consistent"
     | Inconsistent => "Inconsistent";
 
-fun ppStepToString step =
+fun stepToString step =
     case step of
       BeginBlock _ => "BeginBlock"
     | EndBlock => "EndBlock"
@@ -109,330 +131,6 @@
 
 fun blockProgram style indent pps = block style indent (program pps);
 
-fun bracket l r pp =
-    blockProgram Inconsistent (size l)
-      [addString l,
-       pp,
-       addString r];
-
-fun field f pp =
-    blockProgram Inconsistent 2
-      [addString (f ^ " ="),
-       addBreak 1,
-       pp];
-
-val record =
-    let
-      val sep = sequence (addString ",") (addBreak 1)
-
-      fun recordField (f,pp) = field f pp
-
-      fun sepField f = sequence sep (recordField f)
-
-      fun fields [] = []
-        | fields (f :: fs) = recordField f :: map sepField fs
-    in
-      bracket "{" "}" o blockProgram Consistent 0 o fields
-    end;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printer combinators.                                               *)
-(* ------------------------------------------------------------------------- *)
-
-fun ppMap f ppB a : ppstream = ppB (f a);
-
-fun ppBracket l r ppA a = bracket l r (ppA a);
-
-fun ppOp s = sequence (if s = "" then skip else addString s) (addBreak 1);
-
-fun ppOp2 ab ppA ppB (a,b) =
-    blockProgram Inconsistent 0
-      [ppA a,
-       ppOp ab,
-       ppB b];
-
-fun ppOp3 ab bc ppA ppB ppC (a,b,c) =
-    blockProgram Inconsistent 0
-      [ppA a,
-       ppOp ab,
-       ppB b,
-       ppOp bc,
-       ppC c];
-
-fun ppOpList s ppA =
-    let
-      fun ppOpA a = sequence (ppOp s) (ppA a)
-    in
-      fn [] => skip
-       | h :: t => blockProgram Inconsistent 0 (ppA h :: map ppOpA t)
-    end;
-
-fun ppOpStream s ppA =
-    let
-      fun ppOpA a = sequence (ppOp s) (ppA a)
-    in
-      fn Stream.Nil => skip
-       | Stream.Cons (h,t) =>
-         blockProgram Inconsistent 0
-           [ppA h,
-            Stream.concat (Stream.map ppOpA (t ()))]
-    end;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printers for common types.                                         *)
-(* ------------------------------------------------------------------------- *)
-
-fun ppChar c = addString (str c);
-
-val ppString = addString;
-
-fun ppEscapeString {escape} =
-    let
-      val escapeMap = map (fn c => (c, "\\" ^ str c)) escape
-
-      fun escapeChar c =
-          case c of
-            #"\\" => "\\\\"
-          | #"\n" => "\\n"
-          | #"\t" => "\\t"
-          | _ =>
-            case List.find (equal c o fst) escapeMap of
-              SOME (_,s) => s
-            | NONE => str c
-    in
-      fn s => addString (String.translate escapeChar s)
-    end;
-
-val ppUnit : unit pp = K (addString "()");
-
-fun ppBool b = addString (if b then "true" else "false");
-
-fun ppInt i = addString (Int.toString i);
-
-local
-  val ppNeg = addString "~"
-  and ppSep = addString ","
-  and ppZero = addString "0"
-  and ppZeroZero = addString "00";
-
-  fun ppIntBlock i =
-      if i < 10 then sequence ppZeroZero (ppInt i)
-      else if i < 100 then sequence ppZero (ppInt i)
-      else ppInt i;
-
-  fun ppIntBlocks i =
-      if i < 1000 then ppInt i
-      else sequence (ppIntBlocks (i div 1000))
-             (sequence ppSep (ppIntBlock (i mod 1000)));
-in
-  fun ppPrettyInt i =
-      if i < 0 then sequence ppNeg (ppIntBlocks (~i))
-      else ppIntBlocks i;
-end;
-
-fun ppReal r = addString (Real.toString r);
-
-fun ppPercent p = addString (percentToString p);
-
-fun ppOrder x =
-    addString
-      (case x of
-         LESS => "Less"
-       | EQUAL => "Equal"
-       | GREATER => "Greater");
-
-fun ppList ppA = ppBracket "[" "]" (ppOpList "," ppA);
-
-fun ppStream ppA = ppBracket "[" "]" (ppOpStream "," ppA);
-
-fun ppOption ppA ao =
-    case ao of
-      SOME a => ppA a
-    | NONE => addString "-";
-
-fun ppPair ppA ppB = ppBracket "(" ")" (ppOp2 "," ppA ppB);
-
-fun ppTriple ppA ppB ppC = ppBracket "(" ")" (ppOp3 "," "," ppA ppB ppC);
-
-fun ppBreakStyle style = addString (breakStyleToString style);
-
-fun ppPpStep step =
-    let
-      val cmd = ppStepToString step
-    in
-      blockProgram Inconsistent 2
-        (addString cmd ::
-         (case step of
-            BeginBlock style_indent =>
-              [addBreak 1,
-               ppPair ppBreakStyle ppInt style_indent]
-          | EndBlock => []
-          | AddString s =>
-              [addBreak 1,
-               addString ("\"" ^ s ^ "\"")]
-          | AddBreak n =>
-              [addBreak 1,
-               ppInt n]
-          | AddNewline => []))
-    end;
-
-val ppPpstream = ppStream ppPpStep;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printing infix operators.                                          *)
-(* ------------------------------------------------------------------------- *)
-
-datatype infixes =
-    Infixes of
-      {token : string,
-       precedence : int,
-       leftAssoc : bool} list;
-
-local
-  fun chop l =
-      case l of
-        #" " :: l => let val (n,l) = chop l in (n + 1, l) end
-      | _ => (0,l);
-in
-  fun opSpaces tok =
-      let
-        val tok = explode tok
-        val (r,tok) = chop (rev tok)
-        val (l,tok) = chop (rev tok)
-        val tok = implode tok
-      in
-        {leftSpaces = l, token = tok, rightSpaces = r}
-      end;
-end;
-
-fun ppOpSpaces {leftSpaces,token,rightSpaces} =
-    let
-      val leftSpacesToken =
-          if leftSpaces = 0 then token else nSpaces leftSpaces ^ token
-    in
-      sequence
-        (addString leftSpacesToken)
-        (addBreak rightSpaces)
-    end;
-
-local
-  fun new t l acc = {tokens = [opSpaces t], leftAssoc = l} :: acc;
-
-  fun add t l acc =
-      case acc of
-        [] => raise Bug "Print.layerInfixOps.layer"
-      | {tokens = ts, leftAssoc = l'} :: acc =>
-        if l = l' then {tokens = opSpaces t :: ts, leftAssoc = l} :: acc
-        else raise Bug "Print.layerInfixOps: mixed assocs";
-
-  fun layer ({token = t, precedence = p, leftAssoc = l}, (acc,p')) =
-      let
-        val acc = if p = p' then add t l acc else new t l acc
-      in
-        (acc,p)
-      end;
-in
-  fun layerInfixes (Infixes i) =
-      case sortMap #precedence Int.compare i of
-        [] => []
-      | {token = t, precedence = p, leftAssoc = l} :: i =>
-        let
-          val acc = new t l []
-
-          val (acc,_) = List.foldl layer (acc,p) i
-        in
-          acc
-        end;
-end;
-
-val tokensLayeredInfixes =
-    let
-      fun addToken ({leftSpaces = _, token = t, rightSpaces = _}, s) =
-          StringSet.add s t
-
-      fun addTokens ({tokens = t, leftAssoc = _}, s) =
-          List.foldl addToken s t
-    in
-      List.foldl addTokens StringSet.empty
-    end;
-
-val tokensInfixes = tokensLayeredInfixes o layerInfixes;
-
-local
-  val mkTokenMap =
-      let
-        fun add (token,m) =
-            let
-              val {leftSpaces = _, token = t, rightSpaces = _} = token
-            in
-              StringMap.insert m (t, ppOpSpaces token)
-            end
-      in
-        List.foldl add (StringMap.new ())
-      end;
-in
-  fun ppGenInfix {tokens,leftAssoc} =
-      let
-        val tokenMap = mkTokenMap tokens
-      in
-        fn dest => fn ppSub =>
-           let
-             fun dest' tm =
-                 case dest tm of
-                   NONE => NONE
-                 | SOME (t,a,b) =>
-                   case StringMap.peek tokenMap t of
-                     NONE => NONE
-                   | SOME p => SOME (p,a,b)
-
-             fun ppGo (tmr as (tm,r)) =
-                 case dest' tm of
-                   NONE => ppSub tmr
-                 | SOME (p,a,b) =>
-                   program
-                     [(if leftAssoc then ppGo else ppSub) (a,true),
-                      p,
-                      (if leftAssoc then ppSub else ppGo) (b,r)]
-           in
-             fn tmr as (tm,_) =>
-                if Option.isSome (dest' tm) then
-                  block Inconsistent 0 (ppGo tmr)
-                else
-                  ppSub tmr
-           end
-      end;
-end
-
-fun ppInfixes ops =
-    let
-      val layeredOps = layerInfixes ops
-
-      val toks = tokensLayeredInfixes layeredOps
-
-      val iprinters = List.map ppGenInfix layeredOps
-    in
-      fn dest => fn ppSub =>
-         let
-           fun printer sub = foldl (fn (ip,p) => ip dest p) sub iprinters
-
-           fun isOp t =
-               case dest t of
-                 SOME (x,_,_) => StringSet.member x toks
-               | NONE => false
-
-           fun subpr (tmr as (tm,_)) =
-               if isOp tm then
-                 blockProgram Inconsistent 1
-                   [addString "(",
-                    printer subpr (tm,false),
-                    addString ")"]
-               else
-                 ppSub tmr
-         in
-           fn tmr => block Inconsistent 0 (printer subpr tmr)
-         end
-    end;
-
 (* ------------------------------------------------------------------------- *)
 (* Executing pretty-printers to generate lines.                              *)
 (* ------------------------------------------------------------------------- *)
@@ -496,18 +194,22 @@
 val sizeChunks = List.foldl (fn (c,z) => sizeChunk c + z) 0;
 
 local
-  fun render acc [] = acc
-    | render acc (chunk :: chunks) =
-      case chunk of
-        StringChunk {string = s, ...} => render (acc ^ s) chunks
-      | BreakChunk n => render (acc ^ nSpaces n) chunks
-      | BlockChunk (Block {chunks = l, ...}) =>
-        render acc (List.revAppend (l,chunks));
+  fun flatten acc chunks =
+      case chunks of
+        [] => rev acc
+      | chunk :: chunks =>
+        case chunk of
+          StringChunk {string = s, ...} => flatten (s :: acc) chunks
+        | BreakChunk n => flatten (nSpaces n :: acc) chunks
+        | BlockChunk (Block {chunks = l, ...}) =>
+          flatten acc (List.revAppend (l,chunks));
 in
-  fun renderChunks indent chunks = render (nSpaces indent) (rev chunks);
+  fun renderChunks indent chunks =
+      {indent = indent,
+       line = String.concat (flatten [] (rev chunks))};
+end;
 
-  fun renderChunk indent chunk = renderChunks indent [chunk];
-end;
+fun renderChunk indent chunk = renderChunks indent [chunk];
 
 fun isEmptyBlock block =
     let
@@ -523,6 +225,7 @@
       empty
     end;
 
+(*BasicDebug
 fun checkBlock ind block =
     let
       val Block {indent, style = _, size, chunks} = block
@@ -558,6 +261,7 @@
     in
       check initialIndent o rev
     end;
+*)
 
 val initialBlock =
     let
@@ -969,12 +673,10 @@
         (lines,state)
       end;
 
-  fun executeAddString lineLength s lines state =
+  fun executeAddString lineLength (s,n) lines state =
       let
         val State {blocks,lineIndent,lineSize} = state
 
-        val n = size s
-
         val blocks =
             case blocks of
               [] => raise Bug "Print.executeAddString: no block"
@@ -1061,10 +763,13 @@
 
   fun executeAddNewline lineLength lines state =
       let
-        val (lines,state) = executeAddString lineLength "" lines state
-        val (lines,state) = executeBigBreak lineLength lines state
+        val (lines,state) =
+            executeAddString lineLength emptyStringSize lines state
+
+        val (lines,state) =
+            executeBigBreak lineLength lines state
       in
-        executeAddString lineLength "" lines state
+        executeAddString lineLength emptyStringSize lines state
       end;
 
   fun final lineLength lines state =
@@ -1108,7 +813,7 @@
               (lines,state)
             end
             handle Bug bug =>
-              raise Bug ("Print.advance: after " ^ ppStepToString step ^
+              raise Bug ("Print.advance: after " ^ stepToString step ^
                          " command:\n" ^ bug)
 *)
       in
@@ -1117,21 +822,372 @@
 end;
 
 (* ------------------------------------------------------------------------- *)
+(* Pretty-printer combinators.                                               *)
+(* ------------------------------------------------------------------------- *)
+
+type 'a pp = 'a -> ppstream;
+
+fun ppMap f ppB a : ppstream = ppB (f a);
+
+fun ppBracket' l r ppA a =
+    let
+      val (_,n) = l
+    in
+      blockProgram Inconsistent n
+        [addString l,
+         ppA a,
+         addString r]
+    end;
+
+fun ppOp' s = sequence (addString s) (addBreak 1);
+
+fun ppOp2' ab ppA ppB (a,b) =
+    blockProgram Inconsistent 0
+      [ppA a,
+       ppOp' ab,
+       ppB b];
+
+fun ppOp3' ab bc ppA ppB ppC (a,b,c) =
+    blockProgram Inconsistent 0
+      [ppA a,
+       ppOp' ab,
+       ppB b,
+       ppOp' bc,
+       ppC c];
+
+fun ppOpList' s ppA =
+    let
+      fun ppOpA a = sequence (ppOp' s) (ppA a)
+    in
+      fn [] => skip
+       | h :: t => blockProgram Inconsistent 0 (ppA h :: map ppOpA t)
+    end;
+
+fun ppOpStream' s ppA =
+    let
+      fun ppOpA a = sequence (ppOp' s) (ppA a)
+    in
+      fn Stream.Nil => skip
+       | Stream.Cons (h,t) =>
+         blockProgram Inconsistent 0
+           [ppA h,
+            Stream.concat (Stream.map ppOpA (t ()))]
+    end;
+
+fun ppBracket l r = ppBracket' (mkStringSize l) (mkStringSize r);
+
+fun ppOp s = ppOp' (mkStringSize s);
+
+fun ppOp2 ab = ppOp2' (mkStringSize ab);
+
+fun ppOp3 ab bc = ppOp3' (mkStringSize ab) (mkStringSize bc);
+
+fun ppOpList s = ppOpList' (mkStringSize s);
+
+fun ppOpStream s = ppOpStream' (mkStringSize s);
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printers for common types.                                         *)
+(* ------------------------------------------------------------------------- *)
+
+fun ppChar c = addString (str c, 1);
+
+fun ppString s = addString (mkStringSize s);
+
+fun ppEscapeString escape = ppMap (escapeString escape) ppString;
+
+local
+  val pp = ppString "()";
+in
+  fun ppUnit () = pp;
+end;
+
+local
+  val ppTrue = ppString "true"
+  and ppFalse = ppString "false";
+in
+  fun ppBool b = if b then ppTrue else ppFalse;
+end;
+
+val ppInt = ppMap Int.toString ppString;
+
+local
+  val ppNeg = ppString "~"
+  and ppSep = ppString ","
+  and ppZero = ppString "0"
+  and ppZeroZero = ppString "00";
+
+  fun ppIntBlock i =
+      if i < 10 then sequence ppZeroZero (ppInt i)
+      else if i < 100 then sequence ppZero (ppInt i)
+      else ppInt i;
+
+  fun ppIntBlocks i =
+      if i < 1000 then ppInt i
+      else sequence (ppIntBlocks (i div 1000))
+             (sequence ppSep (ppIntBlock (i mod 1000)));
+in
+  fun ppPrettyInt i =
+      if i < 0 then sequence ppNeg (ppIntBlocks (~i))
+      else ppIntBlocks i;
+end;
+
+val ppReal = ppMap Real.toString ppString;
+
+val ppPercent = ppMap percentToString ppString;
+
+local
+  val ppLess = ppString "Less"
+  and ppEqual = ppString "Equal"
+  and ppGreater = ppString "Greater";
+in
+  fun ppOrder ord =
+      case ord of
+        LESS => ppLess
+      | EQUAL => ppEqual
+      | GREATER => ppGreater;
+end;
+
+local
+  val left = mkStringSize "["
+  and right = mkStringSize "]"
+  and sep = mkStringSize ",";
+in
+  fun ppList ppX xs = ppBracket' left right (ppOpList' sep ppX) xs;
+
+  fun ppStream ppX xs = ppBracket' left right (ppOpStream' sep ppX) xs;
+end;
+
+local
+  val ppNone = ppString "-";
+in
+  fun ppOption ppX xo =
+      case xo of
+        SOME x => ppX x
+      | NONE => ppNone;
+end;
+
+local
+  val left = mkStringSize "("
+  and right = mkStringSize ")"
+  and sep = mkStringSize ",";
+in
+  fun ppPair ppA ppB =
+      ppBracket' left right (ppOp2' sep ppA ppB);
+
+  fun ppTriple ppA ppB ppC =
+      ppBracket' left right (ppOp3' sep sep ppA ppB ppC);
+end;
+
+val ppBreakStyle = ppMap breakStyleToString ppString;
+
+val ppStep = ppMap stepToString ppString;
+
+val ppStringSize =
+    let
+      val left = mkStringSize "\""
+      and right = mkStringSize "\""
+      and escape = {escape = [#"\""]}
+
+      val pp = ppBracket' left right (ppEscapeString escape)
+    in
+      fn (s,n) => if size s = n then pp s else ppPair pp ppInt (s,n)
+    end;
+
+fun ppStep step =
+    blockProgram Inconsistent 2
+      (ppStep step ::
+       (case step of
+          BeginBlock style_indent =>
+            [addBreak 1,
+             ppPair ppBreakStyle ppInt style_indent]
+        | EndBlock => []
+        | AddString s =>
+            [addBreak 1,
+             ppStringSize s]
+        | AddBreak n =>
+            [addBreak 1,
+             ppInt n]
+        | AddNewline => []));
+
+val ppPpstream = ppStream ppStep;
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing infix operators.                                          *)
+(* ------------------------------------------------------------------------- *)
+
+type token = string;
+
+datatype assoc =
+    LeftAssoc
+  | NonAssoc
+  | RightAssoc;
+
+datatype infixes =
+    Infixes of
+      {token : token,
+       precedence : int,
+       assoc : assoc} list;
+
+local
+  fun comparePrecedence (io1,io2) =
+      let
+        val {token = _, precedence = p1, assoc = _} = io1
+        and {token = _, precedence = p2, assoc = _} = io2
+      in
+        Int.compare (p2,p1)
+      end;
+
+  fun equalAssoc a a' =
+      case a of
+        LeftAssoc => (case a' of LeftAssoc => true | _ => false)
+      | NonAssoc => (case a' of NonAssoc => true | _ => false)
+      | RightAssoc => (case a' of RightAssoc => true | _ => false);
+
+  fun new t a acc = {tokens = StringSet.singleton t, assoc = a} :: acc;
+
+  fun add t a' acc =
+      case acc of
+        [] => raise Bug "Print.layerInfixes: null"
+      | {tokens = ts, assoc = a} :: acc =>
+        if equalAssoc a a' then {tokens = StringSet.add ts t, assoc = a} :: acc
+        else raise Bug "Print.layerInfixes: mixed assocs";
+
+  fun layer ({token = t, precedence = p, assoc = a}, (acc,p')) =
+      let
+        val acc = if p = p' then add t a acc else new t a acc
+      in
+        (acc,p)
+      end;
+in
+  fun layerInfixes (Infixes ios) =
+      case sort comparePrecedence ios of
+        [] => []
+      | {token = t, precedence = p, assoc = a} :: ios =>
+        let
+          val acc = new t a []
+
+          val (acc,_) = List.foldl layer (acc,p) ios
+        in
+          acc
+        end;
+end;
+
+local
+  fun add ({tokens = ts, assoc = _}, tokens) = StringSet.union ts tokens;
+in
+  fun tokensLayeredInfixes l = List.foldl add StringSet.empty l;
+end;
+
+fun tokensInfixes ios = tokensLayeredInfixes (layerInfixes ios);
+
+fun destInfixOp dest tokens tm =
+    case dest tm of
+      NONE => NONE
+    | s as SOME (t,a,b) => if StringSet.member t tokens then s else NONE;
+
+fun ppLayeredInfixes dest ppTok {tokens,assoc} ppLower ppSub =
+    let
+      fun isLayer t = StringSet.member t tokens
+
+      fun ppTerm aligned (tm,r) =
+          case dest tm of
+            NONE => ppSub (tm,r)
+          | SOME (t,a,b) =>
+            if aligned andalso isLayer t then ppLayer (tm,t,a,b,r)
+            else ppLower (tm,t,a,b,r)
+
+      and ppLeft tm_r =
+          let
+            val aligned = case assoc of LeftAssoc => true | _ => false
+          in
+            ppTerm aligned tm_r
+          end
+
+      and ppLayer (tm,t,a,b,r) =
+          program
+            [ppLeft (a,true),
+             ppTok (tm,t),
+             ppRight (b,r)]
+
+      and ppRight tm_r =
+          let
+            val aligned = case assoc of RightAssoc => true | _ => false
+          in
+            ppTerm aligned tm_r
+          end
+    in
+      fn tm_t_a_b_r as (_,t,_,_,_) =>
+         if isLayer t then block Inconsistent 0 (ppLayer tm_t_a_b_r)
+         else ppLower tm_t_a_b_r
+    end;
+
+local
+  val leftBrack = mkStringSize "("
+  and rightBrack = mkStringSize ")";
+in
+  fun ppInfixes ops =
+      let
+        val layers = layerInfixes ops
+
+        val toks = tokensLayeredInfixes layers
+      in
+        fn dest => fn ppTok => fn ppSub =>
+           let
+             fun destOp tm = destInfixOp dest toks tm
+
+             fun ppInfix tm_t_a_b_r = ppLayers layers tm_t_a_b_r
+
+             and ppLayers ls (tm,t,a,b,r) =
+                 case ls of
+                   [] =>
+                   ppBracket' leftBrack rightBrack ppInfix (tm,t,a,b,false)
+                 | l :: ls =>
+                   let
+                     val ppLower = ppLayers ls
+                   in
+                     ppLayeredInfixes destOp ppTok l ppLower ppSub (tm,t,a,b,r)
+                   end
+           in
+             fn (tm,r) =>
+                case destOp tm of
+                  SOME (t,a,b) => ppInfix (tm,t,a,b,r)
+                | NONE => ppSub (tm,r)
+           end
+      end;
+end;
+
+(* ------------------------------------------------------------------------- *)
 (* Executing pretty-printers with a global line length.                      *)
 (* ------------------------------------------------------------------------- *)
 
 val lineLength = ref initialLineLength;
 
 fun toStream ppA a =
-    Stream.map (fn s => s ^ "\n")
+    Stream.map (fn {indent,line} => nSpaces indent ^ line ^ "\n")
       (execute {lineLength = !lineLength} (ppA a));
 
-fun toString ppA a =
-    case execute {lineLength = !lineLength} (ppA a) of
-      Stream.Nil => ""
-    | Stream.Cons (h,t) => Stream.foldl (fn (s,z) => z ^ "\n" ^ s) h (t ());
+local
+  fun inc {indent,line} acc = line :: nSpaces indent :: acc;
 
-fun trace ppX nameX x =
-    Useful.trace (toString (ppOp2 " =" ppString ppX) (nameX,x) ^ "\n");
+  fun incn (indent_line,acc) = inc indent_line ("\n" :: acc);
+in
+  fun toString ppA a =
+      case execute {lineLength = !lineLength} (ppA a) of
+        Stream.Nil => ""
+      | Stream.Cons (h,t) =>
+        let
+          val lines = Stream.foldl incn (inc h []) (t ())
+        in
+          String.concat (rev lines)
+        end;
+end;
+
+local
+  val sep = mkStringSize " =";
+in
+  fun trace ppX nameX x =
+      Useful.trace (toString (ppOp2' sep ppString ppX) (nameX,x) ^ "\n");
+end;
 
 end
--- a/src/Tools/Metis/src/Problem.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Problem.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* CNF PROBLEMS                                                              *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Problem =
--- a/src/Tools/Metis/src/Problem.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Problem.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* CNF PROBLEMS                                                              *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Problem :> Problem =
@@ -29,9 +29,9 @@
       val cls = toClauses prob
     in
       {clauses = length cls,
-       literals = foldl lits 0 cls,
-       symbols = foldl syms 0 cls,
-       typedSymbols = foldl typedSyms 0 cls}
+       literals = List.foldl lits 0 cls,
+       symbols = List.foldl syms 0 cls,
+       typedSymbols = List.foldl typedSyms 0 cls}
     end;
 
 fun freeVars {axioms,conjecture} =
--- a/src/Tools/Metis/src/Proof.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Proof.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* PROOFS IN FIRST ORDER LOGIC                                               *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Proof =
--- a/src/Tools/Metis/src/Proof.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Proof.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* PROOFS IN FIRST ORDER LOGIC                                               *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Proof :> Proof =
@@ -39,40 +39,40 @@
   fun ppSubst ppThm (sub,thm) =
       Print.sequence (Print.addBreak 1)
         (Print.blockProgram Print.Inconsistent 1
-           [Print.addString "{",
+           [Print.ppString "{",
             Print.ppOp2 " =" Print.ppString Subst.pp ("sub",sub),
-            Print.addString ",",
+            Print.ppString ",",
             Print.addBreak 1,
             Print.ppOp2 " =" Print.ppString ppThm ("thm",thm),
-            Print.addString "}"]);
+            Print.ppString "}"]);
 
   fun ppResolve ppThm (res,pos,neg) =
       Print.sequence (Print.addBreak 1)
         (Print.blockProgram Print.Inconsistent 1
-           [Print.addString "{",
+           [Print.ppString "{",
             Print.ppOp2 " =" Print.ppString Atom.pp ("res",res),
-            Print.addString ",",
+            Print.ppString ",",
             Print.addBreak 1,
             Print.ppOp2 " =" Print.ppString ppThm ("pos",pos),
-            Print.addString ",",
+            Print.ppString ",",
             Print.addBreak 1,
             Print.ppOp2 " =" Print.ppString ppThm ("neg",neg),
-            Print.addString "}"]);
+            Print.ppString "}"]);
 
   fun ppRefl tm = Print.sequence (Print.addBreak 1) (Term.pp tm);
 
   fun ppEquality (lit,path,res) =
       Print.sequence (Print.addBreak 1)
         (Print.blockProgram Print.Inconsistent 1
-           [Print.addString "{",
+           [Print.ppString "{",
             Print.ppOp2 " =" Print.ppString Literal.pp ("lit",lit),
-            Print.addString ",",
+            Print.ppString ",",
             Print.addBreak 1,
             Print.ppOp2 " =" Print.ppString Term.ppPath ("path",path),
-            Print.addString ",",
+            Print.ppString ",",
             Print.addBreak 1,
             Print.ppOp2 " =" Print.ppString Term.pp ("res",res),
-            Print.addString "}"]);
+            Print.ppString "}"]);
 
   fun ppInf ppAxiom ppThm inf =
       let
@@ -80,7 +80,7 @@
       in
         Print.block Print.Inconsistent 2
           (Print.sequence
-             (Print.addString infString)
+             (Print.ppString infString)
              (case inf of
                 Axiom cl => ppAxiom cl
               | Assume x => ppAssume x
@@ -106,7 +106,7 @@
         val prf = enumerate prf
 
         fun ppThm th =
-            Print.addString
+            Print.ppString
             let
               val cl = Thm.clause th
 
@@ -123,7 +123,7 @@
             in
               Print.sequence
                 (Print.blockProgram Print.Consistent (1 + size s)
-                   [Print.addString (s ^ " "),
+                   [Print.ppString (s ^ " "),
                     Thm.pp th,
                     Print.addBreak 2,
                     Print.ppBracket "[" "]" (ppInf (K Print.skip) ppThm) inf])
@@ -131,10 +131,10 @@
             end
       in
         Print.blockProgram Print.Consistent 0
-          [Print.addString "START OF PROOF",
+          [Print.ppString "START OF PROOF",
            Print.addNewline,
            Print.program (map ppStep prf),
-           Print.addString "END OF PROOF"]
+           Print.ppString "END OF PROOF"]
       end
 (*MetisDebug
       handle Error err => raise Bug ("Proof.pp: shouldn't fail:\n" ^ err);
--- a/src/Tools/Metis/src/Random.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Random.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -8,6 +8,7 @@
 
 signature Random =
 sig
+
   val nextWord : unit -> word
 
   val nextBool : unit -> bool
--- a/src/Tools/Metis/src/Resolution.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Resolution.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* THE RESOLUTION PROOF PROCEDURE                                            *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Resolution =
--- a/src/Tools/Metis/src/Resolution.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Resolution.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* THE RESOLUTION PROOF PROCEDURE                                            *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Resolution :> Resolution =
--- a/src/Tools/Metis/src/Rewrite.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Rewrite.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* ORDERED REWRITING FOR FIRST ORDER TERMS                                   *)
-(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Rewrite =
--- a/src/Tools/Metis/src/Rewrite.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Rewrite.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* ORDERED REWRITING FOR FIRST ORDER TERMS                                   *)
-(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Rewrite :> Rewrite =
@@ -207,7 +207,7 @@
         rw
       end;
 
-val addList = foldl (fn (eqn,rw) => add rw eqn);
+val addList = List.foldl (fn (eqn,rw) => add rw eqn);
 
 (* ------------------------------------------------------------------------- *)
 (* Rewriting (the order must be a refinement of the rewrite order).          *)
@@ -434,8 +434,9 @@
     let
       fun addSubterm b ((path,tm),net) = TermNet.insert net (tm,(id,b,path))
 
-      val subterms = foldl (addSubterm true) subterms (Term.subterms l)
-      val subterms = foldl (addSubterm false) subterms (Term.subterms r)
+      val subterms = List.foldl (addSubterm true) subterms (Term.subterms l)
+
+      val subterms = List.foldl (addSubterm false) subterms (Term.subterms r)
     in
       subterms
     end;
@@ -660,7 +661,7 @@
 in
   fun orderedRewrite order ths =
     let
-      val rw = foldl addEqn (new order) (enumerate ths)
+      val rw = List.foldl addEqn (new order) (enumerate ths)
     in
       rewriteRule rw order
     end;
--- a/src/Tools/Metis/src/Rule.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Rule.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS                     *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Rule =
--- a/src/Tools/Metis/src/Rule.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Rule.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS                     *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Rule :> Rule =
@@ -156,17 +156,19 @@
 
 val noConv : conv = fn _ => raise Error "noConv";
 
+(*MetisDebug
 fun traceConv s conv tm =
     let
       val res as (tm',th) = conv tm
-      val () = print (s ^ ": " ^ Term.toString tm ^ " --> " ^
+      val () = trace (s ^ ": " ^ Term.toString tm ^ " --> " ^
                       Term.toString tm' ^ " " ^ Thm.toString th ^ "\n")
     in
       res
     end
     handle Error err =>
-      (print (s ^ ": " ^ Term.toString tm ^ " --> Error: " ^ err ^ "\n");
+      (trace (s ^ ": " ^ Term.toString tm ^ " --> Error: " ^ err ^ "\n");
        raise Error (s ^ ": " ^ err));
+*)
 
 fun thenConvTrans tm (tm',th1) (tm'',th2) =
     let
@@ -455,7 +457,7 @@
       val reflTh = Thm.refl (Term.Fn (f,xs))
       val reflLit = Thm.destUnit reflTh
     in
-      fst (foldl cong (reflTh,reflLit) (enumerate ys))
+      fst (List.foldl cong (reflTh,reflLit) (enumerate ys))
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -482,7 +484,7 @@
       val assumeLit = (false,(R,xs))
       val assumeTh = Thm.assume assumeLit
     in
-      fst (foldl cong (assumeTh,assumeLit) (enumerate ys))
+      fst (List.foldl cong (assumeTh,assumeLit) (enumerate ys))
     end;
 
 (* ------------------------------------------------------------------------- *)
--- a/src/Tools/Metis/src/Sharing.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Sharing.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* PRESERVING SHARING OF ML VALUES                                           *)
-(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2005 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Sharing =
--- a/src/Tools/Metis/src/Sharing.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Sharing.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* PRESERVING SHARING OF ML VALUES                                           *)
-(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2005 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Sharing :> Sharing =
--- a/src/Tools/Metis/src/Stream.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Stream.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Stream =
--- a/src/Tools/Metis/src/Stream.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Stream.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,13 +1,11 @@
 (* ========================================================================= *)
 (* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML                                *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Stream :> Stream =
 struct
 
-open Useful; (* MODIFIED by Jasmin Blanchette *)
-
 val K = Useful.K;
 
 val pair = Useful.pair;
@@ -201,9 +199,9 @@
 
 fun listConcat s = concat (map fromList s);
 
-fun toString s = implode (toList s);
+fun toString s = String.implode (toList s);
 
-fun fromString s = fromList (explode s);
+fun fromString s = fromList (String.explode s);
 
 fun toTextFile {filename = f} s =
     let
--- a/src/Tools/Metis/src/Subst.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Subst.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC SUBSTITUTIONS                                           *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Subst =
--- a/src/Tools/Metis/src/Subst.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Subst.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC SUBSTITUTIONS                                           *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Subst :> Subst =
--- a/src/Tools/Metis/src/Subsume.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Subsume.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES                        *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Subsume =
--- a/src/Tools/Metis/src/Subsume.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Subsume.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES                        *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Subsume :> Subsume =
--- a/src/Tools/Metis/src/Term.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Term.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -348,7 +348,7 @@
 
 val isApp = can destApp;
 
-fun listMkApp (f,l) = foldl mkApp f l;
+fun listMkApp (f,l) = List.foldl mkApp f l;
 
 local
   fun strip tms tm =
@@ -368,38 +368,38 @@
 val infixes =
     (ref o Print.Infixes)
       [(* ML symbols *)
-       {token = " / ", precedence = 7, leftAssoc = true},
-       {token = " div ", precedence = 7, leftAssoc = true},
-       {token = " mod ", precedence = 7, leftAssoc = true},
-       {token = " * ", precedence = 7, leftAssoc = true},
-       {token = " + ", precedence = 6, leftAssoc = true},
-       {token = " - ", precedence = 6, leftAssoc = true},
-       {token = " ^ ", precedence = 6, leftAssoc = true},
-       {token = " @ ", precedence = 5, leftAssoc = false},
-       {token = " :: ", precedence = 5, leftAssoc = false},
-       {token = " = ", precedence = 4, leftAssoc = true},
-       {token = " <> ", precedence = 4, leftAssoc = true},
-       {token = " <= ", precedence = 4, leftAssoc = true},
-       {token = " < ", precedence = 4, leftAssoc = true},
-       {token = " >= ", precedence = 4, leftAssoc = true},
-       {token = " > ", precedence = 4, leftAssoc = true},
-       {token = " o ", precedence = 3, leftAssoc = true},
-       {token = " -> ", precedence = 2, leftAssoc = false},  (* inferred prec *)
-       {token = " : ", precedence = 1, leftAssoc = false},  (* inferred prec *)
-       {token = ", ", precedence = 0, leftAssoc = false},  (* inferred prec *)
+       {token = "/", precedence = 7, assoc = Print.LeftAssoc},
+       {token = "div", precedence = 7, assoc = Print.LeftAssoc},
+       {token = "mod", precedence = 7, assoc = Print.LeftAssoc},
+       {token = "*", precedence = 7, assoc = Print.LeftAssoc},
+       {token = "+", precedence = 6, assoc = Print.LeftAssoc},
+       {token = "-", precedence = 6, assoc = Print.LeftAssoc},
+       {token = "^", precedence = 6, assoc = Print.LeftAssoc},
+       {token = "@", precedence = 5, assoc = Print.RightAssoc},
+       {token = "::", precedence = 5, assoc = Print.RightAssoc},
+       {token = "=", precedence = 4, assoc = Print.NonAssoc},
+       {token = "<>", precedence = 4, assoc = Print.NonAssoc},
+       {token = "<=", precedence = 4, assoc = Print.NonAssoc},
+       {token = "<", precedence = 4, assoc = Print.NonAssoc},
+       {token = ">=", precedence = 4, assoc = Print.NonAssoc},
+       {token = ">", precedence = 4, assoc = Print.NonAssoc},
+       {token = "o", precedence = 3, assoc = Print.LeftAssoc},
+       {token = "->", precedence = 2, assoc = Print.RightAssoc},
+       {token = ":", precedence = 1, assoc = Print.NonAssoc},
+       {token = ",", precedence = 0, assoc = Print.RightAssoc},
 
        (* Logical connectives *)
-       {token = " /\\ ", precedence = ~1, leftAssoc = false},
-       {token = " \\/ ", precedence = ~2, leftAssoc = false},
-       {token = " ==> ", precedence = ~3, leftAssoc = false},
-       {token = " <=> ", precedence = ~4, leftAssoc = false},
+       {token = "/\\", precedence = ~1, assoc = Print.RightAssoc},
+       {token = "\\/", precedence = ~2, assoc = Print.RightAssoc},
+       {token = "==>", precedence = ~3, assoc = Print.RightAssoc},
+       {token = "<=>", precedence = ~4, assoc = Print.RightAssoc},
 
        (* Other symbols *)
-       {token = " . ", precedence = 9, leftAssoc = true},  (* function app *)
-       {token = " ** ", precedence = 8, leftAssoc = true},
-       {token = " ++ ", precedence = 6, leftAssoc = true},
-       {token = " -- ", precedence = 6, leftAssoc = true},
-       {token = " == ", precedence = 4, leftAssoc = true}];
+       {token = ".", precedence = 9, assoc = Print.LeftAssoc},
+       {token = "**", precedence = 8, assoc = Print.LeftAssoc},
+       {token = "++", precedence = 6, assoc = Print.LeftAssoc},
+       {token = "--", precedence = 6, assoc = Print.LeftAssoc},
+       {token = "==", precedence = 4, assoc = Print.NonAssoc}];
 
 (* The negation symbol *)
 
@@ -422,9 +422,14 @@
       and neg = !negation
       and bracks = !brackets
 
-      val bracks = map (fn (b1,b2) => (b1 ^ b2, b1, b2)) bracks
+      val bMap =
+          let
+            fun f (b1,b2) = (b1 ^ b2, b1, b2)
+          in
+            List.map f bracks
+          end
 
-      val bTokens = map #2 bracks @ map #3 bracks
+      val bTokens = op@ (unzip bracks)
 
       val iTokens = Print.tokensInfixes iOps
 
@@ -438,7 +443,15 @@
             end
           | _ => NONE
 
-      val iPrinter = Print.ppInfixes iOps destI
+      fun isI tm = Option.isSome (destI tm)
+
+      fun iToken (_,tok) =
+          Print.program
+            [(if tok = "," then Print.skip else Print.ppString " "),
+             Print.ppString tok,
+             Print.addBreak 1];
+
+      val iPrinter = Print.ppInfixes iOps destI iToken
 
       val specialTokens =
           StringSet.addList iTokens (neg :: quants @ ["$","(",")"] @ bTokens)
@@ -466,8 +479,6 @@
 
       fun functionName bv = Print.ppMap (checkFunctionName bv) Print.ppString
 
-      fun isI tm = Option.isSome (destI tm)
-
       fun stripNeg tm =
           case tm of
             Fn (f,[a]) =>
@@ -494,7 +505,7 @@
           let
             val s = Name.toString b
           in
-            case List.find (fn (n,_,_) => n = s) bracks of
+            case List.find (fn (n,_,_) => n = s) bMap of
               NONE => NONE
             | SOME (_,b1,b2) => SOME (b1,tm,b2)
           end
@@ -527,11 +538,11 @@
               val bv = StringSet.addList bv (map Name.toString (v :: vs))
             in
               Print.program
-                [Print.addString q,
+                [Print.ppString q,
                  varName bv v,
                  Print.program
                    (map (Print.sequence (Print.addBreak 1) o varName bv) vs),
-                 Print.addString ".",
+                 Print.ppString ".",
                  Print.addBreak 1,
                  innerQuant bv tm]
             end
@@ -545,7 +556,7 @@
             val (n,tm) = stripNeg tm
           in
             Print.blockProgram Print.Inconsistent n
-              [Print.duplicate n (Print.addString neg),
+              [Print.duplicate n (Print.ppString neg),
                if isI tm orelse (r andalso isQuant tm) then bracket bv tm
                else quantifier bv tm]
           end
@@ -571,31 +582,32 @@
 
   val isAlphaNum =
       let
-        val alphaNumChars = explode "_'"
+        val alphaNumChars = String.explode "_'"
       in
         fn c => mem c alphaNumChars orelse Char.isAlphaNum c
       end;
 
   local
-    val alphaNumToken = atLeastOne (some isAlphaNum) >> implode;
+    val alphaNumToken = atLeastOne (some isAlphaNum) >> String.implode;
 
     val symbolToken =
         let
           fun isNeg c = str c = !negation
 
-          val symbolChars = explode "<>=-*+/\\?@|!$%&#^:;~"
+          val symbolChars = String.explode "<>=-*+/\\?@|!$%&#^:;~"
 
           fun isSymbol c = mem c symbolChars
 
           fun isNonNegSymbol c = not (isNeg c) andalso isSymbol c
         in
           some isNeg >> str ||
-          (some isNonNegSymbol ++ many (some isSymbol)) >> (implode o op::)
+          (some isNonNegSymbol ++ many (some isSymbol)) >>
+          (String.implode o op::)
         end;
 
     val punctToken =
         let
-          val punctChars = explode "()[]{}.,"
+          val punctChars = String.explode "()[]{}.,"
 
           fun isPunct c = mem c punctChars
         in
@@ -627,8 +639,9 @@
 
         val iTokens = Print.tokensInfixes iOps
 
-        val iParser =
-            parseInfixes iOps (fn (f,a,b) => Fn (Name.fromString f, [a,b]))
+        fun iMk (f,a,b) = Fn (Name.fromString f, [a,b])
+
+        val iParser = parseInfixes iOps iMk any
 
         val specialTokens =
             StringSet.addList iTokens (neg :: quants @ ["$"] @ bTokens)
@@ -667,7 +680,7 @@
                      some (Useful.equal ".")) >>++
                     (fn (_,(vs,_)) =>
                         term (StringSet.addList bv vs) >>
-                        (fn body => foldr bind body vs))
+                        (fn body => List.foldr bind body vs))
                   end
             in
               var ||
@@ -696,7 +709,7 @@
 in
   fun fromString input =
       let
-        val chars = Stream.fromList (explode input)
+        val chars = Stream.fromList (String.explode input)
 
         val tokens = everything (lexer >> singleton) chars
 
@@ -709,7 +722,8 @@
 end;
 
 local
-  val antiquotedTermToString = Print.toString (Print.ppBracket "(" ")" pp);
+  val antiquotedTermToString =
+      Print.toString (Print.ppBracket "(" ")" pp);
 in
   val parse = Parse.parseQuotation antiquotedTermToString fromString;
 end;
--- a/src/Tools/Metis/src/TermNet.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/TermNet.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS              *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature TermNet =
--- a/src/Tools/Metis/src/TermNet.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/TermNet.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS              *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure TermNet :> TermNet =
@@ -158,7 +158,7 @@
 
 fun null net = size net = 0;
 
-fun singles qtms a = foldr Single a qtms;
+fun singles qtms a = List.foldr Single a qtms;
 
 local
   fun pre NONE = (0,NONE)
@@ -188,7 +188,7 @@
       handle Error _ => raise Bug "TermNet.insert: should never fail";
 end;
 
-fun fromList parm l = foldl (fn (tm_a,n) => insert n tm_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (tm_a,n) => insert n tm_a) (new parm) l;
 
 fun filter pred =
     let
@@ -441,7 +441,7 @@
 
 local
   fun inc (qtm, Result l, acc) =
-      foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l
+      List.foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l
     | inc _ = raise Bug "TermNet.pp.inc";
 
   fun toList (Net (_,_,NONE)) = []
--- a/src/Tools/Metis/src/Thm.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Thm.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -111,7 +111,7 @@
 in
   fun pp th =
       Print.blockProgram Print.Inconsistent 3
-        [Print.addString "|- ",
+        [Print.ppString "|- ",
          Formula.pp (toFormula th)];
 end;
 
--- a/src/Tools/Metis/src/Tptp.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Tptp.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -158,7 +158,7 @@
   fun isTptpChar #"_" = true
     | isTptpChar c = Char.isAlphaNum c;
 
-  fun isTptpName l s = nonEmptyPred (existsPred l) (explode s);
+  fun isTptpName l s = nonEmptyPred (existsPred l) (String.explode s);
 
   fun isRegular (c,cs) =
       Char.isLower c andalso List.all isTptpChar cs;
@@ -175,14 +175,14 @@
   fun mkTptpVarName s =
       let
         val s =
-            case List.filter isTptpChar (explode s) of
+            case List.filter isTptpChar (String.explode s) of
               [] => [#"X"]
             | l as c :: cs =>
               if Char.isUpper c then l
               else if Char.isLower c then Char.toUpper c :: cs
               else #"X" :: l
       in
-        implode s
+        String.implode s
       end;
 
   val isTptpConstName = isTptpName [isRegular,isNumber,isDefined,isSystem]
@@ -615,10 +615,10 @@
     let
       val s = varToTptp mapping v
     in
-      Print.addString s
+      Print.ppString s
     end;
 
-fun ppFnName mapping fa = Print.addString (fnToTptp mapping fa);
+fun ppFnName mapping fa = Print.ppString (fnToTptp mapping fa);
 
 fun ppConst mapping c = ppFnName mapping (c,0);
 
@@ -633,14 +633,14 @@
             | a =>
               Print.blockProgram Print.Inconsistent 2
                 [ppFnName mapping (f,a),
-                 Print.addString "(",
+                 Print.ppString "(",
                  Print.ppOpList "," term tms,
-                 Print.addString ")"]
+                 Print.ppString ")"]
     in
       Print.block Print.Inconsistent 0 o term
     end;
 
-fun ppRelName mapping ra = Print.addString (relToTptp mapping ra);
+fun ppRelName mapping ra = Print.ppString (relToTptp mapping ra);
 
 fun ppProp mapping p = ppRelName mapping (p,0);
 
@@ -650,12 +650,12 @@
     | a =>
       Print.blockProgram Print.Inconsistent 2
         [ppRelName mapping (r,a),
-         Print.addString "(",
+         Print.ppString "(",
          Print.ppOpList "," (ppTerm mapping) tms,
-         Print.addString ")"];
+         Print.ppString ")"];
 
 local
-  val neg = Print.sequence (Print.addString "~") (Print.addBreak 1);
+  val neg = Print.sequence (Print.ppString "~") (Print.addBreak 1);
 
   fun fof mapping fm =
       case fm of
@@ -672,8 +672,8 @@
 
   and unitary mapping fm =
       case fm of
-        Formula.True => Print.addString "$true"
-      | Formula.False => Print.addString "$false"
+        Formula.True => Print.ppString "$true"
+      | Formula.False => Print.ppString "$false"
       | Formula.Forall _ => quantified mapping ("!", Formula.stripForall fm)
       | Formula.Exists _ => quantified mapping ("?", Formula.stripExists fm)
       | Formula.Not _ =>
@@ -693,21 +693,21 @@
          | NONE => ppAtom mapping atm)
       | _ =>
         Print.blockProgram Print.Inconsistent 1
-          [Print.addString "(",
+          [Print.ppString "(",
            fof mapping fm,
-           Print.addString ")"]
+           Print.ppString ")"]
 
   and quantified mapping (q,(vs,fm)) =
       let
         val mapping = addVarListMapping mapping vs
       in
         Print.blockProgram Print.Inconsistent 2
-          [Print.addString q,
-           Print.addString " ",
+          [Print.ppString q,
+           Print.ppString " ",
            Print.blockProgram Print.Inconsistent (String.size q)
-             [Print.addString "[",
+             [Print.ppString "[",
               Print.ppOpList "," (ppVar mapping) vs,
-              Print.addString "] :"],
+              Print.ppString "] :"],
            Print.addBreak 1,
            unitary mapping fm]
       end;
@@ -735,7 +735,8 @@
   infixr 7 >>
   infixr 6 ||
 
-  val alphaNumToken = atLeastOne (some isAlphaNum) >> (AlphaNum o implode);
+  val alphaNumToken =
+      atLeastOne (some isAlphaNum) >> (AlphaNum o String.implode);
 
   val punctToken =
       let
@@ -759,7 +760,7 @@
             some (not o stopOn) >> singleton
       in
         exactChar #"'" ++ many quotedParser ++ exactChar #"'" >>
-        (fn (_,(l,_)) => Quote (implode (List.concat l)))
+        (fn (_,(l,_)) => Quote (String.implode (List.concat l)))
       end;
 
   val lexToken = alphaNumToken || punctToken || quoteToken;
@@ -779,7 +780,7 @@
     let
       fun funcs (lit,acc) = NameAritySet.union (functionsLiteral lit) acc
     in
-      foldl funcs NameAritySet.empty
+      List.foldl funcs NameAritySet.empty
     end;
 
 val clauseRelations =
@@ -789,14 +790,14 @@
             NONE => acc
           | SOME r => NameAritySet.add acc r
     in
-      foldl rels NameAritySet.empty
+      List.foldl rels NameAritySet.empty
     end;
 
 val clauseFreeVars =
     let
       fun fvs (lit,acc) = NameSet.union (freeVarsLiteral lit) acc
     in
-      foldl fvs NameSet.empty
+      List.foldl fvs NameSet.empty
     end;
 
 fun clauseSubst sub lits = map (literalSubst sub) lits;
@@ -1013,7 +1014,7 @@
 
     fun ppLiteralInf mapping (pol,atm) =
         Print.sequence
-          (if pol then Print.skip else Print.addString "~ ")
+          (if pol then Print.skip else Print.ppString "~ ")
           (ppAtomInf mapping atm);
   in
     fun ppProofTerm mapping =
@@ -1032,7 +1033,7 @@
 
   fun ppProof mapping inf =
       Print.blockProgram Print.Inconsistent 1
-        [Print.addString "[",
+        [Print.ppString "[",
          (case inf of
             Proof.Axiom _ => Print.skip
           | Proof.Assume atm => ppProofAtom mapping atm
@@ -1042,13 +1043,13 @@
           | Proof.Equality (lit,path,tm) =>
             Print.program
               [ppProofLiteral mapping lit,
-               Print.addString ",",
+               Print.ppString ",",
                Print.addBreak 1,
                ppProofPath path,
-               Print.addString ",",
+               Print.ppString ",",
                Print.addBreak 1,
                ppProofTerm mapping tm]),
-         Print.addString "]"];
+         Print.ppString "]"];
 
   val ppParent = ppFormulaName;
 
@@ -1073,16 +1074,16 @@
           val name = nameStrip inference
         in
           Print.blockProgram Print.Inconsistent (size gen + 1)
-            [Print.addString gen,
-             Print.addString "(",
-             Print.addString name,
-             Print.addString ",",
+            [Print.ppString gen,
+             Print.ppString "(",
+             Print.ppString name,
+             Print.ppString ",",
              Print.addBreak 1,
              Print.ppBracket "[" "]" (ppStrip mapping) inference,
-             Print.addString ",",
+             Print.ppString ",",
              Print.addBreak 1,
              Print.ppList ppParent parents,
-             Print.addString ")"]
+             Print.ppString ")"]
         end
       | NormalizeFormulaSource {inference,parents} =>
         let
@@ -1091,16 +1092,16 @@
           val name = nameNormalize inference
         in
           Print.blockProgram Print.Inconsistent (size gen + 1)
-            [Print.addString gen,
-             Print.addString "(",
-             Print.addString name,
-             Print.addString ",",
+            [Print.ppString gen,
+             Print.ppString "(",
+             Print.ppString name,
+             Print.ppString ",",
              Print.addBreak 1,
              Print.ppBracket "[" "]" (ppNormalize mapping) inference,
-             Print.addString ",",
+             Print.ppString ",",
              Print.addBreak 1,
              Print.ppList ppParent parents,
-             Print.addString ")"]
+             Print.ppString ")"]
         end
       | ProofFormulaSource {inference,parents} =>
         let
@@ -1121,28 +1122,28 @@
               end
         in
           Print.blockProgram Print.Inconsistent (size gen + 1)
-            ([Print.addString gen,
-              Print.addString "("] @
+            ([Print.ppString gen,
+              Print.ppString "("] @
              (if isTaut then
-                [Print.addString "tautology",
-                 Print.addString ",",
+                [Print.ppString "tautology",
+                 Print.ppString ",",
                  Print.addBreak 1,
                  Print.blockProgram Print.Inconsistent 1
-                   [Print.addString "[",
-                    Print.addString name,
-                    Print.addString ",",
+                   [Print.ppString "[",
+                    Print.ppString name,
+                    Print.ppString ",",
                     Print.addBreak 1,
                     ppProof mapping inference,
-                    Print.addString "]"]]
+                    Print.ppString "]"]]
               else
-                [Print.addString name,
-                 Print.addString ",",
+                [Print.ppString name,
+                 Print.ppString ",",
                  Print.addBreak 1,
                  ppProof mapping inference,
-                 Print.addString ",",
+                 Print.ppString ",",
                  Print.addBreak 1,
                  Print.ppList (ppProofParent mapping) parents]) @
-             [Print.addString ")"])
+             [Print.ppString ")"])
         end
 end;
 
@@ -1208,14 +1209,14 @@
     let
       fun funcs (fm,acc) = NameAritySet.union (functionsFormula fm) acc
     in
-      foldl funcs NameAritySet.empty
+      List.foldl funcs NameAritySet.empty
     end;
 
 val formulasRelations =
     let
       fun rels (fm,acc) = NameAritySet.union (relationsFormula fm) acc
     in
-      foldl rels NameAritySet.empty
+      List.foldl rels NameAritySet.empty
     end;
 
 fun isCnfConjectureFormula fm =
@@ -1244,24 +1245,24 @@
           | FofFormulaBody _ => "fof"
     in
       Print.blockProgram Print.Inconsistent (size gen + 1)
-        ([Print.addString gen,
-          Print.addString "(",
+        ([Print.ppString gen,
+          Print.ppString "(",
           ppFormulaName name,
-          Print.addString ",",
+          Print.ppString ",",
           Print.addBreak 1,
           ppRole role,
-          Print.addString ",",
+          Print.ppString ",",
           Print.addBreak 1,
           Print.blockProgram Print.Consistent 1
-            [Print.addString "(",
+            [Print.ppString "(",
              ppFormulaBody mapping body,
-             Print.addString ")"]] @
+             Print.ppString ")"]] @
          (if isNoFormulaSource source then []
           else
-            [Print.addString ",",
+            [Print.ppString ",",
              Print.addBreak 1,
              ppFormulaSource mapping source]) @
-         [Print.addString ")."])
+         [Print.ppString ")."])
     end;
 
 fun formulaToString mapping = Print.toString (ppFormula mapping);
@@ -1285,7 +1286,7 @@
 
   val stringParser = lowerParser || upperParser;
 
-  val numberParser = someAlphaNum (List.all Char.isDigit o explode);
+  val numberParser = someAlphaNum (List.all Char.isDigit o String.explode);
 
   fun somePunct p =
       maybe (fn Punct c => if p c then SOME c else NONE | _ => NONE);
@@ -1299,7 +1300,7 @@
       | f [x] = x
       | f (h :: t) = (h ++ f t) >> K ();
   in
-    fun symbolParser s = f (map punctParser (explode s));
+    fun symbolParser s = f (map punctParser (String.explode s));
   end;
 
   val definedParser =
@@ -1641,9 +1642,9 @@
 
 fun ppInclude i =
     Print.blockProgram Print.Inconsistent 2
-      [Print.addString "include('",
-       Print.addString i,
-       Print.addString "')."];
+      [Print.ppString "include('",
+       Print.ppString i,
+       Print.ppString "')."];
 
 val includeToString = Print.toString ppInclude;
 
@@ -1745,8 +1746,8 @@
 fun destLineComment cs =
     case cs of
       [] => ""
-    | #"%" :: #" " :: rest => implode rest
-    | #"%" :: rest => implode rest
+    | #"%" :: #" " :: rest => String.implode rest
+    | #"%" :: rest => String.implode rest
     | _ => raise Error "Tptp.destLineComment";
 
 val isLineComment = can destLineComment;
--- a/src/Tools/Metis/src/Units.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Units.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* A STORE FOR UNIT THEOREMS                                                 *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Units =
--- a/src/Tools/Metis/src/Units.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Units.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* A STORE FOR UNIT THEOREMS                                                 *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Units :> Units =
@@ -47,7 +47,7 @@
         end
     end;
 
-val addList = foldl (fn (th,u) => add u th);
+val addList = List.foldl (fn (th,u) => add u th);
 
 (* ------------------------------------------------------------------------- *)
 (* Matching.                                                                 *)
--- a/src/Tools/Metis/src/Useful.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Useful.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -22,8 +22,6 @@
 (* Tracing.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-val print : string -> unit (* MODIFIED by Jasmin Blanchette *)
-
 val tracePrint : (string -> unit) ref
 
 val trace : string -> unit
@@ -109,10 +107,6 @@
 (* Lists: note we count elements from 0.                                     *)
 (* ------------------------------------------------------------------------- *)
 
-val foldl : ('a * 'b -> 'b) -> 'b -> 'a list -> 'b (* MODIFIED by Jasmin Blanchette *)
-
-val foldr : ('a * 'b -> 'b) -> 'b -> 'a list -> 'b (* MODIFIED by Jasmin Blanchette *)
-
 val cons : 'a -> 'a list -> 'a list
 
 val hdTl : 'a list -> 'a * 'a list
@@ -217,10 +211,6 @@
 (* Strings.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-val implode : char list -> string (* MODIFIED by Jasmin Blanchette *)
-
-val explode : string -> char list (* MODIFIED by Jasmin Blanchette *)
-
 val rot : int -> char -> char
 
 val charToInt : char -> int option
@@ -323,7 +313,9 @@
 
 val try : ('a -> 'b) -> 'a -> 'b
 
-val chat : string -> unit
+val chat : string -> unit  (* stdout *)
+
+val chide : string -> unit  (* stderr *)
 
 val warn : string -> unit
 
--- a/src/Tools/Metis/src/Useful.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Useful.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -50,9 +50,7 @@
 (* Tracing.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-val print = TextIO.print; (* MODIFIED by Jasmin Blanchette *)
-
-val tracePrint = ref print;
+val tracePrint = ref TextIO.print;
 
 fun trace mesg = !tracePrint mesg;
 
@@ -172,10 +170,6 @@
 (* Lists.                                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-val foldl = List.foldl; (* MODIFIED by Jasmin Blanchette *)
-
-val foldr = List.foldr; (* MODIFIED by Jasmin Blanchette *)
-
 fun cons x y = x :: y;
 
 fun hdTl l = (hd l, tl l);
@@ -211,19 +205,22 @@
 
 fun zip xs ys = zipWith pair xs ys;
 
-fun unzip ab =
-    foldl (fn ((x, y), (xs, ys)) => (x :: xs, y :: ys)) ([], []) (rev ab);
+local
+  fun inc ((x,y),(xs,ys)) = (x :: xs, y :: ys);
+in
+  fun unzip ab = List.foldl inc ([],[]) (rev ab);
+end;
 
 fun cartwith f =
-  let
-    fun aux _ res _ [] = res
-      | aux xsCopy res [] (y :: yt) = aux xsCopy res xsCopy yt
-      | aux xsCopy res (x :: xt) (ys as y :: _) =
-        aux xsCopy (f x y :: res) xt ys
-  in
-    fn xs => fn ys =>
-       let val xs' = rev xs in aux xs' [] xs' (rev ys) end
-  end;
+    let
+      fun aux _ res _ [] = res
+        | aux xsCopy res [] (y :: yt) = aux xsCopy res xsCopy yt
+        | aux xsCopy res (x :: xt) (ys as y :: _) =
+          aux xsCopy (f x y :: res) xt ys
+    in
+      fn xs => fn ys =>
+         let val xs' = rev xs in aux xs' [] xs' (rev ys) end
+    end;
 
 fun cart xs ys = cartwith pair xs ys;
 
@@ -342,15 +339,32 @@
 
 fun delete x s = List.filter (not o equal x) s;
 
-fun setify s = rev (foldl (fn (v,x) => if mem v x then x else v :: x) [] s);
+local
+  fun inc (v,x) = if mem v x then x else v :: x;
+in
+  fun setify s = rev (List.foldl inc [] s);
+end;
 
-fun union s t = foldl (fn (v,x) => if mem v t then x else v :: x) t (rev s);
+fun union s t =
+    let
+      fun inc (v,x) = if mem v t then x else v :: x
+    in
+      List.foldl inc t (rev s)
+    end;
 
 fun intersect s t =
-    foldl (fn (v,x) => if mem v t then v :: x else x) [] (rev s);
+    let
+      fun inc (v,x) = if mem v t then v :: x else x
+    in
+      List.foldl inc [] (rev s)
+    end;
 
 fun difference s t =
-    foldl (fn (v,x) => if mem v t then x else v :: x) [] (rev s);
+    let
+      fun inc (v,x) = if mem v t then x else v :: x
+    in
+      List.foldl inc [] (rev s)
+    end;
 
 fun subset s t = List.all (fn x => mem x t) s;
 
@@ -460,8 +474,7 @@
 
   val primesList = ref [2];
 in
-  (* MODIFIED by Jasmin Blanchette *)
-  fun primes n = CRITICAL (fn () =>
+  fun primes n = Portable.critical (fn () =>
       let
         val ref ps = primesList
 
@@ -476,11 +489,10 @@
           in
             ps
           end
-      end);
+      end) ();
 end;
 
-(* MODIFIED by Jasmin Blanchette *)
-fun primesUpTo n = CRITICAL (fn () =>
+fun primesUpTo n = Portable.critical (fn () =>
     let
       fun f k =
           let
@@ -492,22 +504,18 @@
           end
     in
       f 8
-    end);
+    end) ();
 
 (* ------------------------------------------------------------------------- *)
 (* Strings.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-val implode = String.implode (* MODIFIED by Jasmin Blanchette *)
-
-val explode = String.explode (* MODIFIED by Jasmin Blanchette *)
-
 local
   fun len l = (length l, l)
 
-  val upper = len (explode "ABCDEFGHIJKLMNOPQRSTUVWXYZ");
+  val upper = len (String.explode "ABCDEFGHIJKLMNOPQRSTUVWXYZ");
 
-  val lower = len (explode "abcdefghijklmnopqrstuvwxyz");
+  val lower = len (String.explode "abcdefghijklmnopqrstuvwxyz");
 
   fun rotate (n,l) c k =
       List.nth (l, (k + Option.valOf (index (equal c) l)) mod n);
@@ -546,7 +554,7 @@
     let
       fun dup 0 l = l | dup n l = dup (n - 1) (x :: l)
     in
-      fn n => implode (dup n [])
+      fn n => String.implode (dup n [])
     end;
 
 fun chomp s =
@@ -558,14 +566,15 @@
     end;
 
 local
-  fun chop [] = []
-    | chop (l as (h :: t)) = if Char.isSpace h then chop t else l;
+  fun chop l =
+      case l of
+        [] => []
+      | h :: t => if Char.isSpace h then chop t else l;
 in
-  val trim = implode o chop o rev o chop o rev o explode;
+  val trim = String.implode o chop o rev o chop o rev o String.explode;
 end;
 
-fun join _ [] = ""
-  | join s (h :: t) = foldl (fn (x,y) => y ^ s ^ x) h t;
+val join = String.concatWith;
 
 local
   fun match [] l = SOME l
@@ -573,18 +582,19 @@
     | match (x :: xs) (y :: ys) = if x = y then match xs ys else NONE;
 
   fun stringify acc [] = acc
-    | stringify acc (h :: t) = stringify (implode h :: acc) t;
+    | stringify acc (h :: t) = stringify (String.implode h :: acc) t;
 in
   fun split sep =
       let
         val pat = String.explode sep
+
         fun div1 prev recent [] = stringify [] (rev recent :: prev)
           | div1 prev recent (l as h :: t) =
             case match pat l of
               NONE => div1 prev (h :: recent) t
             | SOME rest => div1 (rev recent :: prev) [] rest
       in
-        fn s => div1 [] [] (explode s)
+        fn s => div1 [] [] (String.explode s)
       end;
 end;
 
@@ -714,24 +724,22 @@
 local
   val generator = ref 0
 in
-  (* MODIFIED by Jasmin Blanchette *)
-  fun newInt () = CRITICAL (fn () =>
+  fun newInt () = Portable.critical (fn () =>
       let
         val n = !generator
         val () = generator := n + 1
       in
         n
-      end);
+      end) ();
 
   fun newInts 0 = []
-      (* MODIFIED by Jasmin Blanchette *)
-    | newInts k = CRITICAL (fn () =>
+    | newInts k = Portable.critical (fn () =>
       let
         val n = !generator
         val () = generator := n + k
       in
         interval n k
-      end);
+      end) ();
 end;
 
 fun withRef (r,new) f x =
@@ -811,10 +819,12 @@
 (* Profiling and error reporting.                                            *)
 (* ------------------------------------------------------------------------- *)
 
-fun chat s = TextIO.output (TextIO.stdErr, s ^ "\n");
+fun chat s = TextIO.output (TextIO.stdOut, s ^ "\n");
+
+fun chide s = TextIO.output (TextIO.stdErr, s ^ "\n");
 
 local
-  fun err x s = chat (x ^ ": " ^ s);
+  fun err x s = chide (x ^ ": " ^ s);
 in
   fun try f x = f x
       handle e as Error _ => (err "try" (errorToString e); raise e)
--- a/src/Tools/Metis/src/Waiting.sig	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Waiting.sig	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* THE WAITING SET OF CLAUSES                                                *)
-(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 signature Waiting =
--- a/src/Tools/Metis/src/Waiting.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/Waiting.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* THE WAITING SET OF CLAUSES                                                *)
-(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 structure Waiting :> Waiting =
@@ -40,17 +40,16 @@
 (* ------------------------------------------------------------------------- *)
 
 val defaultModels : modelParameters list =
-    [(* MODIFIED by Jasmin Blanchette
-      {model = Model.default,
+    [{model = Model.default,
       initialPerturbations = 100,
       maxChecks = SOME 20,
       perturbations = 0,
-      weight = 1.0} *)];
+      weight = 1.0}];
 
 val default : parameters =
      {symbolsWeight = 1.0,
-      literalsWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *)
-      variablesWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *)
+      literalsWeight = 1.0,
+      variablesWeight = 1.0,
       models = defaultModels};
 
 fun size (Waiting {clauses,...}) = Heap.size clauses;
@@ -163,7 +162,7 @@
         val symbolsW = Math.pow (clauseSymbols lits, symbolsWeight)
         val variablesW = Math.pow (clauseVariables lits, variablesWeight)
         val literalsW = Math.pow (clauseLiterals lits, literalsWeight)
-        val modelsW = (* checkModels models mods mcl *) 1.0 (* MODIFIED by Jasmin Blanchette *)
+        val modelsW = checkModels models mods mcl
 (*MetisTrace4
         val () = trace ("Waiting.clauseWeight: dist = " ^
                         Real.toString dist ^ "\n")
--- a/src/Tools/Metis/src/metis.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/metis.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,21 +1,6 @@
 (* ========================================================================= *)
 (* METIS FIRST ORDER PROVER                                                  *)
-(*                                                                           *)
-(* Copyright (c) 2001 Joe Hurd                                               *)
-(*                                                                           *)
-(* Metis is free software; you can redistribute it and/or modify             *)
-(* it under the terms of the GNU General Public License as published by      *)
-(* the Free Software Foundation; either version 2 of the License, or         *)
-(* (at your option) any later version.                                       *)
-(*                                                                           *)
-(* Metis is distributed in the hope that it will be useful,                  *)
-(* but WITHOUT ANY WARRANTY; without even the implied warranty of            *)
-(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the             *)
-(* GNU General Public License for more details.                              *)
-(*                                                                           *)
-(* You should have received a copy of the GNU General Public License         *)
-(* along with Metis; if not, write to the Free Software                      *)
-(* Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA  02111-1307  USA *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 open Useful;
@@ -26,9 +11,9 @@
 
 val PROGRAM = "metis";
 
-val VERSION = "2.2";
+val VERSION = "2.3";
 
-val versionString = PROGRAM^" "^VERSION^" (release 20100825)"^"\n";
+val versionString = PROGRAM^" "^VERSION^" (release 20100915)"^"\n";
 
 (* ------------------------------------------------------------------------- *)
 (* Program options.                                                          *)
@@ -146,11 +131,11 @@
 local
   fun display_sep () =
       if notshowing_any () then ()
-      else print (nChars #"-" (!Print.lineLength) ^ "\n");
+      else TextIO.print (nChars #"-" (!Print.lineLength) ^ "\n");
 
   fun display_name filename =
       if notshowing "name" then ()
-      else print ("Problem: " ^ filename ^ "\n\n");
+      else TextIO.print ("Problem: " ^ filename ^ "\n\n");
 
   fun display_goal tptp =
       if notshowing "goal" then ()
@@ -158,12 +143,12 @@
         let
           val goal = Tptp.goal tptp
         in
-          print ("Goal:\n" ^ Formula.toString goal ^ "\n\n")
+          TextIO.print ("Goal:\n" ^ Formula.toString goal ^ "\n\n")
         end;
 
   fun display_clauses cls =
       if notshowing "clauses" then ()
-      else print ("Clauses:\n" ^ Problem.toString cls ^ "\n\n");
+      else TextIO.print ("Clauses:\n" ^ Problem.toString cls ^ "\n\n");
 
   fun display_size cls =
       if notshowing "size" then ()
@@ -174,7 +159,7 @@
 
           val {clauses,literals,symbols,typedSymbols} = Problem.size cls
         in
-          print
+          TextIO.print
             ("Size: " ^
              plural clauses "clause" ^ ", " ^
              plural literals "literal" ^ ", " ^
@@ -188,12 +173,12 @@
         let
           val cat = Problem.categorize cls
         in
-          print ("Category: " ^ Problem.categoryToString cat ^ ".\n\n")
+          TextIO.print ("Category: " ^ Problem.categoryToString cat ^ ".\n\n")
         end;
 
   local
     fun display_proof_start filename =
-        print ("\nSZS output start CNFRefutation for " ^ filename ^ "\n");
+        TextIO.print ("\nSZS output start CNFRefutation for " ^ filename ^ "\n");
 
     fun display_proof_body problem proofs =
         let
@@ -224,7 +209,7 @@
         end;
 
     fun display_proof_end filename =
-        print ("SZS output end CNFRefutation for " ^ filename ^ "\n\n");
+        TextIO.print ("SZS output end CNFRefutation for " ^ filename ^ "\n\n");
   in
     fun display_proof filename problem proofs =
         if notshowing "proof" then ()
@@ -264,17 +249,24 @@
                    filename = "saturation.tptp"}
               end
 *)
-          val () = print ("\nSZS output start Saturation for " ^ filename ^ "\n")
-          val () = app (fn th => print (Thm.toString th ^ "\n")) ths
-          val () = print ("SZS output end Saturation for " ^ filename ^ "\n\n")
+          val () =
+              TextIO.print
+                ("\nSZS output start Saturation for " ^ filename ^ "\n")
+
+          val () = app (fn th => TextIO.print (Thm.toString th ^ "\n")) ths
+
+          val () =
+              TextIO.print
+                ("SZS output end Saturation for " ^ filename ^ "\n\n")
         in
           ()
         end;
 
   fun display_status filename status =
       if notshowing "status" then ()
-      else print ("SZS status " ^ Tptp.toStringStatus status ^
-                  " for " ^ filename ^ "\n");
+      else
+        TextIO.print ("SZS status " ^ Tptp.toStringStatus status ^
+                      " for " ^ filename ^ "\n");
 
   fun display_problem filename cls =
       let
--- a/src/Tools/Metis/src/problems.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/problems.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -1,6 +1,6 @@
 (* ========================================================================= *)
 (* SOME SAMPLE PROBLEMS TO TEST PROOF PROCEDURES                             *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
 (* ========================================================================= *)
--- a/src/Tools/Metis/src/problems2tptp.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/problems2tptp.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -39,7 +39,7 @@
       dups names
     end;
 
-fun listProblem {name, comments = _, goal = _} = print (name ^ "\n");
+fun listProblem {name, comments = _, goal = _} = TextIO.print (name ^ "\n");
 
 fun outputProblem outputDir {name,comments,goal} =
     let
--- a/src/Tools/Metis/src/selftest.sml	Thu Sep 16 13:49:17 2010 +0200
+++ b/src/Tools/Metis/src/selftest.sml	Thu Sep 16 16:20:20 2010 +0200
@@ -52,11 +52,11 @@
   | partialOrderToString NONE = "NONE";
 
 fun SAY s =
-    print
+    TextIO.print
       ("-------------------------------------" ^
        "-------------------------------------\n" ^ s ^ "\n\n");
 
-fun printval p x = (print (Print.toString p x ^ "\n\n"); x);
+fun printval p x = (TextIO.print (Print.toString p x ^ "\n\n"); x);
 
 fun mkCl p th = Clause.mk {parameters = p, id = Clause.newId (), thm = th};
 
@@ -92,12 +92,13 @@
 
 fun test_fun eq p r a =
   if eq r a then p a ^ "\n" else
-    (print ("\n\n" ^
-            "test: should have\n-->" ^ p r ^ "<--\n\n" ^
-            "test: actually have\n-->" ^ p a ^ "<--\n\n");
+    (TextIO.print
+       ("\n\n" ^
+        "test: should have\n-->" ^ p r ^ "<--\n\n" ^
+        "test: actually have\n-->" ^ p a ^ "<--\n\n");
      raise Fail "test: failed a test");
 
-fun test eq p r a = print (test_fun eq p r a ^ "\n");
+fun test eq p r a = TextIO.print (test_fun eq p r a ^ "\n");
 
 val test_tm = test Term.equal Term.toString o Term.parse;
 
@@ -123,7 +124,7 @@
     (fn s => test_fun equal I s ((mini_print n o Formula.fromString) s))
       (prep q);
 
-fun test_pp q = print (testlen_pp 40 q ^ "\n");
+fun test_pp q = TextIO.print (testlen_pp 40 q ^ "\n");
 
 val () = test_pp `3 = f x`;
 
@@ -568,7 +569,7 @@
 
         val rows = alignTable format table
 
-        val () = print (join "\n" rows ^ "\n\n")
+        val () = TextIO.print (join "\n" rows ^ "\n\n")
       in
         ()
       end;
@@ -626,8 +627,8 @@
       val fm = LiteralSet.disjoin cl
     in
       Print.blockProgram Print.Consistent ind
-        [Print.addString p,
-         Print.addString (nChars #" " (ind - size p)),
+        [Print.ppString p,
+         Print.ppString (nChars #" " (ind - size p)),
          Formula.pp fm]
     end;
 
@@ -1108,13 +1109,19 @@
 
       val _ =
         test_fun equal I g (mini_print (!Print.lineLength) p)
-        handle e => (print ("Error in problem " ^ name ^ "\n\n"); raise e)
+        handle e =>
+          (TextIO.print ("Error in problem " ^ name ^ "\n\n");
+           raise e)
     in
       (name,p) :: acc
     end;
 in
   fun check_syntax (p : problem list) =
-      (foldl check [] p; print "ok\n\n");
+      let
+        val _ = List.foldl check [] p
+      in
+        TextIO.print "ok\n\n"
+      end;
 end;
 
 val () = check_syntax problems;
@@ -1125,11 +1132,11 @@
 
 fun tptp f =
     let
-      val () = print ("parsing " ^ f ^ "... ")
+      val () = TextIO.print ("parsing " ^ f ^ "... ")
       val filename = "tptp/" ^ f ^ ".tptp"
       val mapping = Tptp.defaultMapping
       val goal = Tptp.goal (Tptp.read {filename = filename, mapping = mapping})
-      val () = print "ok\n"
+      val () = TextIO.print "ok\n"
     in
       pvFm goal
     end;