regenerate "metis.ML", this time without manual hacks
authorblanchet
Wed, 15 Sep 2010 16:20:46 +0200
changeset 39417 0be01cad5df4
parent 39416 868f22264662
child 39418 27f5ee6b9101
regenerate "metis.ML", this time without manual hacks
src/Tools/Metis/metis.ML
--- a/src/Tools/Metis/metis.ML	Wed Sep 15 16:19:49 2010 +0200
+++ b/src/Tools/Metis/metis.ML	Wed Sep 15 16:20:46 2010 +0200
@@ -1,13 +1,7 @@
 (*
-   This file was generated by the "make-metis" script. A few changes were done
-   manually on the script's output; these are marked as follows:
-
-       MODIFIED by Jasmin Blanchette
-
-   Some of these changes are needed so that the ML files compiles at all. Others
-   are old tweaks by Lawrence C. Paulson that are needed, if nothing else, for
-   backward compatibility. The BSD License is used with Joe Hurd's kind
-   permission. Extract from a September 13, 2010 email written by Joe Hurd:
+   This file was generated by the "make_metis" script. The BSD License is used
+   with Joe Hurd's kind permission. Extract from a September 13, 2010 email
+   written by Joe Hurd:
 
        I hereby give permission to the Isabelle team to release Metis as part
        of Isabelle, with the Metis code covered under the Isabelle BSD
@@ -22,7 +16,6 @@
 
 print_depth 0;
 
-structure Metis = struct structure Word = Word structure Array = Array end;
 
 (**** Original file: Random.sig ****)
 
@@ -34,8 +27,9 @@
 randomness.
 *)
 
-signature Random =
+signature Metis_Random =
 sig
+  val CRITICAL: (unit -> 'a) -> 'a (* MODIFIED by Jasmin Blanchette *)
 
   val nextWord : unit -> word
 
@@ -49,15 +43,6 @@
 
 (**** Original file: Random.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (*  Title:      Tools/random_word.ML
     Author:     Makarius
 
@@ -66,9 +51,12 @@
 randomness.
 *)
 
-structure Random :> Random =
+structure Metis_Random :> Metis_Random =
 struct
 
+(* MODIFIED by Jasmin Blanchette *)
+fun CRITICAL e = NAMED_CRITICAL "metis" e;
+
 (* random words: 0w0 <= result <= max_word *)
 
 (*minimum length of unboxed words on all supported ML platforms*)
@@ -85,7 +73,8 @@
 
 fun change r f = r := f (!r);
 local val rand = (*Unsynchronized.*)Unsynchronized.ref 0w1
-in fun nextWord () = ((*Unsynchronized.*)change rand step; ! rand) end;
+(* MODIFIED by Jasmin Blanchette *)
+in fun nextWord () = CRITICAL (fn () => ((*Unsynchronized.*)change rand step; ! rand)) end;
 
 (*NB: higher bits are more random than lower ones*)
 fun nextBool () = Word.andb (nextWord (), top_bit) = 0w0;
@@ -106,7 +95,6 @@
 fun nextReal () = real (Word.toInt (nextWord ())) / scaling;
 
 end;
-end;
 
 (**** Original file: Portable.sig ****)
 
@@ -115,7 +103,7 @@
 (* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Portable =
+signature Metis_Portable =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -146,27 +134,18 @@
 
 val randomReal : unit -> real  (* () |-> [0,1] *)
 
-val randomWord : unit -> Metis.Word.word
+val randomWord : unit -> Word.word
 
 end
 
 (**** Original file: PortablePolyml.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* POLY/ML SPECIFIC FUNCTIONS                                                *)
 (* Copyright (c) 2008 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-structure Portable :> Portable =
+structure Metis_Portable :> Metis_Portable =
 struct
 
 (* ------------------------------------------------------------------------- *)
@@ -179,7 +158,7 @@
 (* Pointer equality using the run-time system.                               *)
 (* ------------------------------------------------------------------------- *)
 
-fun pointerEqual (x : 'a, y : 'a) = pointer_eq (x, y) (* MODIFIED by Jasmin Blanchette *)
+fun pointerEqual (x : 'a, y : 'a) = pointer_eq(x,y);
 
 (* ------------------------------------------------------------------------- *)
 (* Timing function applications.                                             *)
@@ -195,7 +174,7 @@
               3 => s
             | 2 => s ^ "0"
             | 1 => s ^ "00"
-            | _ => raise Fail "Portable.time"
+            | _ => raise Fail "Metis_Portable.time"
           end
 
       val c = Timer.startCPUTimer ()
@@ -207,7 +186,7 @@
             val {usr,sys} = Timer.checkCPUTimer c
             val real = Timer.checkRealTimer r
           in
-            print
+            TextIO.print (* MODIFIED by Jasmin Blanchette *)
               ("User: " ^ p usr ^ "  System: " ^ p sys ^
                "  Real: " ^ p real ^ "\n")
           end
@@ -219,17 +198,18 @@
       y
     end;
 
+
 (* ------------------------------------------------------------------------- *)
 (* Generating random values.                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-val randomWord = Random.nextWord;
-
-val randomBool = Random.nextBool;
-
-val randomInt = Random.nextInt;
-
-val randomReal = Random.nextReal;
+val randomWord = Metis_Random.nextWord;
+
+val randomBool = Metis_Random.nextBool;
+
+val randomInt = Metis_Random.nextInt;
+
+val randomReal = Metis_Random.nextReal;
 
 end
 
@@ -238,7 +218,6 @@
 (* ------------------------------------------------------------------------- *)
 
 datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a;
-end;
 
 (**** Original file: Useful.sig ****)
 
@@ -247,7 +226,7 @@
 (* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature Useful =
+signature Metis_Useful =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -266,6 +245,8 @@
 (* Tracing.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
+val print : string -> unit (* MODIFIED by Jasmin Blanchette *)
+
 val tracePrint : (string -> unit) Unsynchronized.ref
 
 val trace : string -> unit
@@ -351,6 +332,10 @@
 (* 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
@@ -455,6 +440,10 @@
 (* 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
@@ -524,7 +513,7 @@
 val isRight : ('a,'b) sum -> bool
 
 (* ------------------------------------------------------------------------- *)
-(* Useful impure features.                                                   *)
+(* Metis_Useful impure features.                                                   *)
 (* ------------------------------------------------------------------------- *)
 
 val newInt : unit -> int
@@ -533,7 +522,7 @@
 
 val withRef : 'r Unsynchronized.ref * 'r -> ('a -> 'b) -> 'a -> 'b
 
-val cloneArray : 'a Metis.Array.array -> 'a Metis.Array.array
+val cloneArray : 'a Array.array -> 'a Array.array
 
 (* ------------------------------------------------------------------------- *)
 (* The environment.                                                          *)
@@ -573,21 +562,12 @@
 
 (**** Original file: Useful.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* ML UTILITY FUNCTIONS                                                      *)
 (* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-structure Useful :> Useful =
+structure Metis_Useful :> Metis_Useful =
 struct
 
 (* ------------------------------------------------------------------------- *)
@@ -634,6 +614,8 @@
 (* Tracing.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
+val print = TextIO.print; (* MODIFIED by Jasmin Blanchette *)
+
 val tracePrint = Unsynchronized.ref print;
 
 fun trace mesg = !tracePrint mesg;
@@ -754,6 +736,10 @@
 (* 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);
@@ -1038,7 +1024,8 @@
 
   val primesList = Unsynchronized.ref [2];
 in
-  fun primes n =
+  (* MODIFIED by Jasmin Blanchette *)
+  fun primes n = CRITICAL (fn () =>
       let
         val Unsynchronized.ref ps = primesList
 
@@ -1053,10 +1040,11 @@
           in
             ps
           end
-      end;
-end;
-
-fun primesUpTo n =
+      end);
+end;
+
+(* MODIFIED by Jasmin Blanchette *)
+fun primesUpTo n = CRITICAL (fn () =>
     let
       fun f k =
           let
@@ -1068,12 +1056,16 @@
           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)
 
@@ -1280,28 +1272,30 @@
   | isRight (Right _) = true;
 
 (* ------------------------------------------------------------------------- *)
-(* Useful impure features.                                                   *)
+(* Metis_Useful impure features.                                                   *)
 (* ------------------------------------------------------------------------- *)
 
 local
   val generator = Unsynchronized.ref 0
 in
-  fun newInt () =
+  (* MODIFIED by Jasmin Blanchette *)
+  fun newInt () = CRITICAL (fn () =>
       let
         val n = !generator
         val () = generator := n + 1
       in
         n
-      end;
+      end);
 
   fun newInts 0 = []
-    | newInts k =
+      (* MODIFIED by Jasmin Blanchette *)
+    | newInts k = CRITICAL (fn () =>
       let
         val n = !generator
         val () = generator := n + k
       in
         interval n k
-      end;
+      end);
 end;
 
 fun withRef (r,new) f x =
@@ -1428,7 +1422,6 @@
     end;
 
 end
-end;
 
 (**** Original file: Lazy.sig ****)
 
@@ -1437,7 +1430,7 @@
 (* Copyright (c) 2007 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature Lazy =
+signature Metis_Lazy =
 sig
 
 type 'a lazy
@@ -1454,34 +1447,25 @@
 
 (**** Original file: Lazy.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* SUPPORT FOR LAZY EVALUATION                                               *)
 (* Copyright (c) 2007 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-structure Lazy :> Lazy =
+structure Metis_Lazy :> Metis_Lazy =
 struct
 
 datatype 'a thunk =
     Value of 'a
   | Thunk of unit -> 'a;
 
-datatype 'a lazy = Lazy of 'a thunk Unsynchronized.ref;
-
-fun quickly v = Lazy (Unsynchronized.ref (Value v));
-
-fun delay f = Lazy (Unsynchronized.ref (Thunk f));
-
-fun force (Lazy s) =
+datatype 'a lazy = Metis_Lazy of 'a thunk Unsynchronized.ref;
+
+fun quickly v = Metis_Lazy (Unsynchronized.ref (Value v));
+
+fun delay f = Metis_Lazy (Unsynchronized.ref (Thunk f));
+
+fun force (Metis_Lazy s) =
     case !s of
       Value v => v
     | Thunk f =>
@@ -1501,7 +1485,6 @@
     end;
 
 end
-end;
 
 (**** Original file: Stream.sig ****)
 
@@ -1510,7 +1493,7 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Stream =
+signature Metis_Stream =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -1520,10 +1503,10 @@
 datatype 'a stream = Nil | Cons of 'a * (unit -> 'a stream)
 
 (* If you're wondering how to create an infinite stream: *)
-(* val stream4 = let fun s4 () = Metis.Stream.Cons (4,s4) in s4 () end; *)
-
-(* ------------------------------------------------------------------------- *)
-(* Stream constructors.                                                      *)
+(* val stream4 = let fun s4 () = Metis_Stream.Cons (4,s4) in s4 () end; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Metis_Stream constructors.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
 val repeat : 'a -> 'a stream
@@ -1533,7 +1516,7 @@
 val funpows : ('a -> 'a) -> 'a -> 'a stream
 
 (* ------------------------------------------------------------------------- *)
-(* Stream versions of standard list operations: these should all terminate.  *)
+(* Metis_Stream versions of standard list operations: these should all terminate.  *)
 (* ------------------------------------------------------------------------- *)
 
 val cons : 'a -> (unit -> 'a stream) -> 'a stream
@@ -1564,7 +1547,7 @@
 val drop : int -> 'a stream -> 'a stream  (* raises Subscript *)
 
 (* ------------------------------------------------------------------------- *)
-(* Stream versions of standard list operations: these might not terminate.   *)
+(* Metis_Stream versions of standard list operations: these might not terminate.   *)
 (* ------------------------------------------------------------------------- *)
 
 val length : 'a stream -> int
@@ -1592,7 +1575,7 @@
     'a stream -> 'b stream
 
 (* ------------------------------------------------------------------------- *)
-(* Stream operations.                                                        *)
+(* Metis_Stream operations.                                                        *)
 (* ------------------------------------------------------------------------- *)
 
 val memoize : 'a stream -> 'a stream
@@ -1617,28 +1600,21 @@
 
 (**** Original file: Stream.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML                                *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Stream :> Stream =
+structure Metis_Stream :> Metis_Stream =
 struct
 
-val K = Useful.K;
-
-val pair = Useful.pair;
-
-val funpow = Useful.funpow;
+open Metis_Useful; (* MODIFIED by Jasmin Blanchette *)
+
+val K = Metis_Useful.K;
+
+val pair = Metis_Useful.pair;
+
+val funpow = Metis_Useful.funpow;
 
 (* ------------------------------------------------------------------------- *)
 (* The stream type.                                                          *)
@@ -1649,7 +1625,7 @@
   | Cons of 'a * (unit -> 'a stream);
 
 (* ------------------------------------------------------------------------- *)
-(* Stream constructors.                                                      *)
+(* Metis_Stream constructors.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
 fun repeat x = let fun rep () = Cons (x,rep) in rep () end;
@@ -1659,7 +1635,7 @@
 fun funpows f x = Cons (x, fn () => funpows f (f x));
 
 (* ------------------------------------------------------------------------- *)
-(* Stream versions of standard list operations: these should all terminate.  *)
+(* Metis_Stream versions of standard list operations: these should all terminate.  *)
 (* ------------------------------------------------------------------------- *)
 
 fun cons h t = Cons (h,t);
@@ -1721,7 +1697,7 @@
 fun drop n s = funpow n tl s handle Empty => raise Subscript;
 
 (* ------------------------------------------------------------------------- *)
-(* Stream versions of standard list operations: these might not terminate.   *)
+(* Metis_Stream versions of standard list operations: these might not terminate.   *)
 (* ------------------------------------------------------------------------- *)
 
 local
@@ -1806,11 +1782,11 @@
     end;
 
 (* ------------------------------------------------------------------------- *)
-(* Stream operations.                                                        *)
+(* Metis_Stream operations.                                                        *)
 (* ------------------------------------------------------------------------- *)
 
 fun memoize Nil = Nil
-  | memoize (Cons (h,t)) = Cons (h, Lazy.memoize (fn () => memoize (t ())));
+  | memoize (Cons (h,t)) = Cons (h, Metis_Lazy.memoize (fn () => memoize (t ())));
 
 fun concatList [] = Nil
   | concatList (h :: t) = append h (fn () => concatList t);
@@ -1860,7 +1836,6 @@
     end;
 
 end
-end;
 
 (**** Original file: Ordered.sig ****)
 
@@ -1869,7 +1844,7 @@
 (* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Ordered =
+signature Metis_Ordered =
 sig
 
 type t
@@ -1900,24 +1875,15 @@
 
 (**** Original file: Ordered.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* ORDERED TYPES                                                             *)
 (* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure IntOrdered =
+structure Metis_IntOrdered =
 struct type t = int val compare = Int.compare end;
 
-structure IntPairOrdered =
+structure Metis_IntPairOrdered =
 struct
 
 type t = int * int;
@@ -1930,9 +1896,8 @@
 
 end;
 
-structure StringOrdered =
+structure Metis_StringOrdered =
 struct type t = string val compare = String.compare end;
-end;
 
 (**** Original file: Map.sig ****)
 
@@ -1941,7 +1906,7 @@
 (* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature Map =
+signature Metis_Map =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -1959,7 +1924,7 @@
 val singleton : ('key * 'key -> order) -> 'key * 'a -> ('key,'a) map
 
 (* ------------------------------------------------------------------------- *)
-(* Map size.                                                                 *)
+(* Metis_Map size.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
 val null : ('key,'a) map -> bool
@@ -2023,7 +1988,7 @@
     ('key,'a) map -> ('key,'b) map -> ('key,'c) map
 
 (* ------------------------------------------------------------------------- *)
-(* Set operations on the domain.                                             *)
+(* Metis_Set operations on the domain.                                             *)
 (* ------------------------------------------------------------------------- *)
 
 val inDomain : 'key -> ('key,'a) map -> bool
@@ -2129,38 +2094,29 @@
 
 (**** Original file: Map.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES                      *)
 (* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-structure Map :> Map =
+structure Metis_Map :> Metis_Map =
 struct
 
 (* ------------------------------------------------------------------------- *)
 (* Importing useful functionality.                                           *)
 (* ------------------------------------------------------------------------- *)
 
-exception Bug = Useful.Bug;
-
-exception Error = Useful.Error;
-
-val pointerEqual = Portable.pointerEqual;
-
-val K = Useful.K;
-
-val randomInt = Portable.randomInt;
-
-val randomWord = Portable.randomWord;
+exception Bug = Metis_Useful.Bug;
+
+exception Error = Metis_Useful.Error;
+
+val pointerEqual = Metis_Portable.pointerEqual;
+
+val K = Metis_Useful.K;
+
+val randomInt = Metis_Portable.randomInt;
+
+val randomWord = Metis_Portable.randomWord;
 
 (* ------------------------------------------------------------------------- *)
 (* Converting a comparison function to an equality function.                 *)
@@ -2603,7 +2559,7 @@
 
 fun treeDelete compareKey dkey tree =
     case tree of
-      E => raise Bug "Map.delete: element not found"
+      E => raise Bug "Metis_Map.delete: element not found"
     | T node => nodeDelete compareKey dkey node
 
 and nodeDelete compareKey dkey node =
@@ -2921,7 +2877,7 @@
 
 fun treePick tree =
     case tree of
-      E => raise Bug "Map.treePick"
+      E => raise Bug "Metis_Map.treePick"
     | T node => nodePick node;
 
 (* ------------------------------------------------------------------------- *)
@@ -2937,7 +2893,7 @@
 
 fun treeDeletePick tree =
     case tree of
-      E => raise Bug "Map.treeDeletePick"
+      E => raise Bug "Metis_Map.treeDeletePick"
     | T node => nodeDeletePick node;
 
 (* ------------------------------------------------------------------------- *)
@@ -2946,7 +2902,7 @@
 
 fun treeNth n tree =
     case tree of
-      E => raise Bug "Map.treeNth"
+      E => raise Bug "Metis_Map.treeNth"
     | T node => nodeNth n node
 
 and nodeNth n node =
@@ -2966,7 +2922,7 @@
 
 fun treeDeleteNth n tree =
     case tree of
-      E => raise Bug "Map.treeDeleteNth"
+      E => raise Bug "Metis_Map.treeDeleteNth"
     | T node => nodeDeleteNth n node
 
 and nodeDeleteNth n node =
@@ -3020,7 +2976,7 @@
 
 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;
+  | Metis_RL of ('key * 'value) * ('key,'value) tree * ('key,'value) node list;
 
 fun fromSpineLR nodes =
     case nodes of
@@ -3032,7 +2988,7 @@
     case nodes of
       [] => NONE
     | Node {key,value,left,...} :: nodes =>
-      SOME (RL ((key,value),left,nodes));
+      SOME (Metis_RL ((key,value),left,nodes));
 
 fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
 
@@ -3045,12 +3001,12 @@
 fun readIterator iter =
     case iter of
       LR (key_value,_,_) => key_value
-    | RL (key_value,_,_) => key_value;
+    | Metis_RL (key_value,_,_) => key_value;
 
 fun advanceIterator iter =
     case iter of
       LR (_,tree,nodes) => addLR nodes tree
-    | RL (_,tree,nodes) => addRL nodes tree;
+    | Metis_RL (_,tree,nodes) => addRL nodes tree;
 
 fun foldIterator f acc io =
     case io of
@@ -3136,22 +3092,22 @@
 (* ------------------------------------------------------------------------- *)
 
 datatype ('key,'value) map =
-    Map of ('key * 'key -> order) * ('key,'value) tree;
-
-(* ------------------------------------------------------------------------- *)
-(* Map debugging functions.                                                  *)
+    Metis_Map of ('key * 'key -> order) * ('key,'value) tree;
+
+(* ------------------------------------------------------------------------- *)
+(* Metis_Map debugging functions.                                                  *)
 (* ------------------------------------------------------------------------- *)
 
 (*BasicDebug
 fun checkInvariants s m =
     let
-      val Map (compareKey,tree) = m
+      val Metis_Map (compareKey,tree) = m
 
       val _ = treeCheckInvariants compareKey tree
     in
       m
     end
-    handle Bug bug => raise Bug (s ^ "\n" ^ "Map.checkInvariants: " ^ bug);
+    handle Bug bug => raise Bug (s ^ "\n" ^ "Metis_Map.checkInvariants: " ^ bug);
 *)
 
 (* ------------------------------------------------------------------------- *)
@@ -3162,21 +3118,21 @@
     let
       val tree = treeNew ()
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 fun singleton compareKey key_value =
     let
       val tree = treeSingleton key_value
     in
-      Map (compareKey,tree)
-    end;
-
-(* ------------------------------------------------------------------------- *)
-(* Map size.                                                                 *)
-(* ------------------------------------------------------------------------- *)
-
-fun size (Map (_,tree)) = treeSize tree;
+      Metis_Map (compareKey,tree)
+    end;
+
+(* ------------------------------------------------------------------------- *)
+(* Metis_Map size.                                                                 *)
+(* ------------------------------------------------------------------------- *)
+
+fun size (Metis_Map (_,tree)) = treeSize tree;
 
 fun null m = size m = 0;
 
@@ -3184,26 +3140,26 @@
 (* Querying.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun peekKey (Map (compareKey,tree)) key = treePeekKey compareKey key tree;
-
-fun peek (Map (compareKey,tree)) key = treePeek compareKey key tree;
+fun peekKey (Metis_Map (compareKey,tree)) key = treePeekKey compareKey key tree;
+
+fun peek (Metis_Map (compareKey,tree)) key = treePeek compareKey key tree;
 
 fun inDomain key m = Option.isSome (peek m key);
 
 fun get m key =
     case peek m key of
-      NONE => raise Error "Map.get: element not found"
+      NONE => raise Error "Metis_Map.get: element not found"
     | SOME value => value;
 
-fun pick (Map (_,tree)) = treePick tree;
-
-fun nth (Map (_,tree)) n = treeNth n tree;
+fun pick (Metis_Map (_,tree)) = treePick tree;
+
+fun nth (Metis_Map (_,tree)) n = treeNth n tree;
 
 fun random m =
     let
       val n = size m
     in
-      if n = 0 then raise Bug "Map.random: empty"
+      if n = 0 then raise Bug "Metis_Map.random: empty"
       else nth m (randomInt n)
     end;
 
@@ -3211,17 +3167,17 @@
 (* Adding.                                                                   *)
 (* ------------------------------------------------------------------------- *)
 
-fun insert (Map (compareKey,tree)) key_value =
+fun insert (Metis_Map (compareKey,tree)) key_value =
     let
       val tree = treeInsert compareKey key_value tree
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 (*BasicDebug
 val insert = fn m => fn kv =>
-    checkInvariants "Map.insert: result"
-      (insert (checkInvariants "Map.insert: input" m) kv);
+    checkInvariants "Metis_Map.insert: result"
+      (insert (checkInvariants "Metis_Map.insert: input" m) kv);
 *)
 
 fun insertList m =
@@ -3235,50 +3191,50 @@
 (* Removing.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun delete (Map (compareKey,tree)) dkey =
+fun delete (Metis_Map (compareKey,tree)) dkey =
     let
       val tree = treeDelete compareKey dkey tree
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 (*BasicDebug
 val delete = fn m => fn k =>
-    checkInvariants "Map.delete: result"
-      (delete (checkInvariants "Map.delete: input" m) k);
+    checkInvariants "Metis_Map.delete: result"
+      (delete (checkInvariants "Metis_Map.delete: input" m) k);
 *)
 
 fun remove m key = if inDomain key m then delete m key else m;
 
-fun deletePick (Map (compareKey,tree)) =
+fun deletePick (Metis_Map (compareKey,tree)) =
     let
       val (key_value,tree) = treeDeletePick tree
     in
-      (key_value, Map (compareKey,tree))
+      (key_value, Metis_Map (compareKey,tree))
     end;
 
 (*BasicDebug
 val deletePick = fn m =>
     let
-      val (kv,m) = deletePick (checkInvariants "Map.deletePick: input" m)
-    in
-      (kv, checkInvariants "Map.deletePick: result" m)
-    end;
-*)
-
-fun deleteNth (Map (compareKey,tree)) n =
+      val (kv,m) = deletePick (checkInvariants "Metis_Map.deletePick: input" m)
+    in
+      (kv, checkInvariants "Metis_Map.deletePick: result" m)
+    end;
+*)
+
+fun deleteNth (Metis_Map (compareKey,tree)) n =
     let
       val (key_value,tree) = treeDeleteNth n tree
     in
-      (key_value, Map (compareKey,tree))
+      (key_value, Metis_Map (compareKey,tree))
     end;
 
 (*BasicDebug
 val deleteNth = fn m => fn n =>
     let
-      val (kv,m) = deleteNth (checkInvariants "Map.deleteNth: input" m) n
-    in
-      (kv, checkInvariants "Map.deleteNth: result" m)
+      val (kv,m) = deleteNth (checkInvariants "Metis_Map.deleteNth: input" m) n
+    in
+      (kv, checkInvariants "Metis_Map.deleteNth: result" m)
     end;
 *)
 
@@ -3286,7 +3242,7 @@
     let
       val n = size m
     in
-      if n = 0 then raise Bug "Map.deleteRandom: empty"
+      if n = 0 then raise Bug "Metis_Map.deleteRandom: empty"
       else deleteNth m (randomInt n)
     end;
 
@@ -3294,89 +3250,89 @@
 (* Joining (all join operations prefer keys in the second map).              *)
 (* ------------------------------------------------------------------------- *)
 
-fun merge {first,second,both} (Map (compareKey,tree1)) (Map (_,tree2)) =
+fun merge {first,second,both} (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) =
     let
       val tree = treeMerge compareKey first second both tree1 tree2
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 (*BasicDebug
 val merge = fn f => fn m1 => fn m2 =>
-    checkInvariants "Map.merge: result"
+    checkInvariants "Metis_Map.merge: result"
       (merge f
-         (checkInvariants "Map.merge: input 1" m1)
-         (checkInvariants "Map.merge: input 2" m2));
-*)
-
-fun union f (Map (compareKey,tree1)) (Map (_,tree2)) =
+         (checkInvariants "Metis_Map.merge: input 1" m1)
+         (checkInvariants "Metis_Map.merge: input 2" m2));
+*)
+
+fun union f (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) =
     let
       fun f2 kv = f (kv,kv)
 
       val tree = treeUnion compareKey f f2 tree1 tree2
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 (*BasicDebug
 val union = fn f => fn m1 => fn m2 =>
-    checkInvariants "Map.union: result"
+    checkInvariants "Metis_Map.union: result"
       (union f
-         (checkInvariants "Map.union: input 1" m1)
-         (checkInvariants "Map.union: input 2" m2));
-*)
-
-fun intersect f (Map (compareKey,tree1)) (Map (_,tree2)) =
+         (checkInvariants "Metis_Map.union: input 1" m1)
+         (checkInvariants "Metis_Map.union: input 2" m2));
+*)
+
+fun intersect f (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) =
     let
       val tree = treeIntersect compareKey f tree1 tree2
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 (*BasicDebug
 val intersect = fn f => fn m1 => fn m2 =>
-    checkInvariants "Map.intersect: result"
+    checkInvariants "Metis_Map.intersect: result"
       (intersect f
-         (checkInvariants "Map.intersect: input 1" m1)
-         (checkInvariants "Map.intersect: input 2" m2));
+         (checkInvariants "Metis_Map.intersect: input 1" m1)
+         (checkInvariants "Metis_Map.intersect: input 2" m2));
 *)
 
 (* ------------------------------------------------------------------------- *)
 (* Iterators over maps.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
-fun mkIterator (Map (_,tree)) = treeMkIterator tree;
-
-fun mkRevIterator (Map (_,tree)) = treeMkRevIterator tree;
+fun mkIterator (Metis_Map (_,tree)) = treeMkIterator tree;
+
+fun mkRevIterator (Metis_Map (_,tree)) = treeMkRevIterator tree;
 
 (* ------------------------------------------------------------------------- *)
 (* Mapping and folding.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
-fun mapPartial f (Map (compareKey,tree)) =
+fun mapPartial f (Metis_Map (compareKey,tree)) =
     let
       val tree = treeMapPartial f tree
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 (*BasicDebug
 val mapPartial = fn f => fn m =>
-    checkInvariants "Map.mapPartial: result"
-      (mapPartial f (checkInvariants "Map.mapPartial: input" m));
-*)
-
-fun map f (Map (compareKey,tree)) =
+    checkInvariants "Metis_Map.mapPartial: result"
+      (mapPartial f (checkInvariants "Metis_Map.mapPartial: input" m));
+*)
+
+fun map f (Metis_Map (compareKey,tree)) =
     let
       val tree = treeMap f tree
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 (*BasicDebug
 val map = fn f => fn m =>
-    checkInvariants "Map.map: result"
-      (map f (checkInvariants "Map.map: input" m));
+    checkInvariants "Metis_Map.map: result"
+      (map f (checkInvariants "Metis_Map.map: input" m));
 *)
 
 fun transform f = map (fn (_,value) => f value);
@@ -3441,7 +3397,7 @@
         LESS => LESS
       | EQUAL =>
         let
-          val Map (compareKey,_) = m1
+          val Metis_Map (compareKey,_) = m1
 
           val io1 = mkIterator m1
           and io2 = mkIterator m2
@@ -3454,7 +3410,7 @@
     pointerEqual (m1,m2) orelse
     (size m1 = size m2 andalso
      let
-       val Map (compareKey,_) = m1
+       val Metis_Map (compareKey,_) = m1
 
        val io1 = mkIterator m1
        and io2 = mkIterator m2
@@ -3463,22 +3419,22 @@
      end);
 
 (* ------------------------------------------------------------------------- *)
-(* Set operations on the domain.                                             *)
-(* ------------------------------------------------------------------------- *)
-
-fun unionDomain (Map (compareKey,tree1)) (Map (_,tree2)) =
+(* Metis_Set operations on the domain.                                             *)
+(* ------------------------------------------------------------------------- *)
+
+fun unionDomain (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) =
     let
       val tree = treeUnionDomain compareKey tree1 tree2
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 (*BasicDebug
 val unionDomain = fn m1 => fn m2 =>
-    checkInvariants "Map.unionDomain: result"
+    checkInvariants "Metis_Map.unionDomain: result"
       (unionDomain
-         (checkInvariants "Map.unionDomain: input 1" m1)
-         (checkInvariants "Map.unionDomain: input 2" m2));
+         (checkInvariants "Metis_Map.unionDomain: input 1" m1)
+         (checkInvariants "Metis_Map.unionDomain: input 2" m2));
 *)
 
 local
@@ -3486,23 +3442,23 @@
 in
   fun unionListDomain ms =
       case ms of
-        [] => raise Bug "Map.unionListDomain: no sets"
+        [] => raise Bug "Metis_Map.unionListDomain: no sets"
       | m :: ms => List.foldl uncurriedUnionDomain m ms;
 end;
 
-fun intersectDomain (Map (compareKey,tree1)) (Map (_,tree2)) =
+fun intersectDomain (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) =
     let
       val tree = treeIntersectDomain compareKey tree1 tree2
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 (*BasicDebug
 val intersectDomain = fn m1 => fn m2 =>
-    checkInvariants "Map.intersectDomain: result"
+    checkInvariants "Metis_Map.intersectDomain: result"
       (intersectDomain
-         (checkInvariants "Map.intersectDomain: input 1" m1)
-         (checkInvariants "Map.intersectDomain: input 2" m2));
+         (checkInvariants "Metis_Map.intersectDomain: input 1" m1)
+         (checkInvariants "Metis_Map.intersectDomain: input 2" m2));
 *)
 
 local
@@ -3510,23 +3466,23 @@
 in
   fun intersectListDomain ms =
       case ms of
-        [] => raise Bug "Map.intersectListDomain: no sets"
+        [] => raise Bug "Metis_Map.intersectListDomain: no sets"
       | m :: ms => List.foldl uncurriedIntersectDomain m ms;
 end;
 
-fun differenceDomain (Map (compareKey,tree1)) (Map (_,tree2)) =
+fun differenceDomain (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) =
     let
       val tree = treeDifferenceDomain compareKey tree1 tree2
     in
-      Map (compareKey,tree)
+      Metis_Map (compareKey,tree)
     end;
 
 (*BasicDebug
 val differenceDomain = fn m1 => fn m2 =>
-    checkInvariants "Map.differenceDomain: result"
+    checkInvariants "Metis_Map.differenceDomain: result"
       (differenceDomain
-         (checkInvariants "Map.differenceDomain: input 1" m1)
-         (checkInvariants "Map.differenceDomain: input 2" m2));
+         (checkInvariants "Metis_Map.differenceDomain: input 1" m1)
+         (checkInvariants "Metis_Map.differenceDomain: input 2" m2));
 *)
 
 fun symmetricDifferenceDomain m1 m2 =
@@ -3534,7 +3490,7 @@
 
 fun equalDomain m1 m2 = equal (K (K true)) m1 m2;
 
-fun subsetDomain (Map (compareKey,tree1)) (Map (_,tree2)) =
+fun subsetDomain (Metis_Map (compareKey,tree1)) (Metis_Map (_,tree2)) =
     treeSubsetDomain compareKey tree1 tree2;
 
 fun disjointDomain m1 m2 = null (intersectDomain m1 m2);
@@ -3563,7 +3519,6 @@
 fun toString m = "<" ^ (if null m then "" else Int.toString (size m)) ^ ">";
 
 end
-end;
 
 (**** Original file: KeyMap.sig ****)
 
@@ -3572,7 +3527,7 @@
 (* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature KeyMap =
+signature Metis_KeyMap =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -3596,7 +3551,7 @@
 val singleton : key * 'a -> 'a map
 
 (* ------------------------------------------------------------------------- *)
-(* Map size.                                                                 *)
+(* Metis_Map size.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
 val null : 'a map -> bool
@@ -3660,7 +3615,7 @@
     'a map -> 'b map -> 'c map
 
 (* ------------------------------------------------------------------------- *)
-(* Set operations on the domain.                                             *)
+(* Metis_Set operations on the domain.                                             *)
 (* ------------------------------------------------------------------------- *)
 
 val inDomain : key -> 'a map -> bool
@@ -3771,15 +3726,13 @@
 (* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-functor KeyMap (Key : Ordered) :> KeyMap where type key = Key.t =
+functor Metis_KeyMap (Key : Metis_Ordered) :> Metis_KeyMap where type key = Key.t =
 struct
 
 (* ------------------------------------------------------------------------- *)
 (* Importing from the input signature.                                       *)
 (* ------------------------------------------------------------------------- *)
 
-open Metis; (* MODIFIED by Jasmin Blanchette *)
-
 type key = Key.t;
 
 val compareKey = Key.compare;
@@ -3788,17 +3741,17 @@
 (* Importing useful functionality.                                           *)
 (* ------------------------------------------------------------------------- *)
 
-exception Bug = Useful.Bug;
-
-exception Error = Useful.Error;
-
-val pointerEqual = Portable.pointerEqual;
-
-val K = Useful.K;
-
-val randomInt = Portable.randomInt;
-
-val randomWord = Portable.randomWord;
+exception Bug = Metis_Useful.Bug;
+
+exception Error = Metis_Useful.Error;
+
+val pointerEqual = Metis_Portable.pointerEqual;
+
+val K = Metis_Useful.K;
+
+val randomInt = Metis_Portable.randomInt;
+
+val randomWord = Metis_Portable.randomWord;
 
 (* ------------------------------------------------------------------------- *)
 (* Converting a comparison function to an equality function.                 *)
@@ -4241,7 +4194,7 @@
 
 fun treeDelete dkey tree =
     case tree of
-      E => raise Bug "KeyMap.delete: element not found"
+      E => raise Bug "Metis_KeyMap.delete: element not found"
     | T node => nodeDelete dkey node
 
 and nodeDelete dkey node =
@@ -4369,7 +4322,7 @@
     end;
 
 (* ------------------------------------------------------------------------- *)
-(* A op union operation on trees.                                               *)
+(* A union operation on trees.                                               *)
 (* ------------------------------------------------------------------------- *)
 
 fun treeUnion f f2 tree1 tree2 =
@@ -4438,7 +4391,7 @@
     end;
 
 (* ------------------------------------------------------------------------- *)
-(* A op union operation on trees which simply chooses the second value.         *)
+(* A union operation on trees which simply chooses the second value.         *)
 (* ------------------------------------------------------------------------- *)
 
 fun treeUnionDomain tree1 tree2 =
@@ -4520,7 +4473,7 @@
       end;
 
 (* ------------------------------------------------------------------------- *)
-(* A op subset operation on trees.                                              *)
+(* A subset operation on trees.                                              *)
 (* ------------------------------------------------------------------------- *)
 
 fun treeSubsetDomain tree1 tree2 =
@@ -4559,7 +4512,7 @@
 
 fun treePick tree =
     case tree of
-      E => raise Bug "KeyMap.treePick"
+      E => raise Bug "Metis_KeyMap.treePick"
     | T node => nodePick node;
 
 (* ------------------------------------------------------------------------- *)
@@ -4575,7 +4528,7 @@
 
 fun treeDeletePick tree =
     case tree of
-      E => raise Bug "KeyMap.treeDeletePick"
+      E => raise Bug "Metis_KeyMap.treeDeletePick"
     | T node => nodeDeletePick node;
 
 (* ------------------------------------------------------------------------- *)
@@ -4584,7 +4537,7 @@
 
 fun treeNth n tree =
     case tree of
-      E => raise Bug "KeyMap.treeNth"
+      E => raise Bug "Metis_KeyMap.treeNth"
     | T node => nodeNth n node
 
 and nodeNth n node =
@@ -4604,7 +4557,7 @@
 
 fun treeDeleteNth n tree =
     case tree of
-      E => raise Bug "KeyMap.treeDeleteNth"
+      E => raise Bug "Metis_KeyMap.treeDeleteNth"
     | T node => nodeDeleteNth n node
 
 and nodeDeleteNth n node =
@@ -4658,7 +4611,7 @@
 
 datatype 'value iterator =
     LR of (key * 'value) * 'value tree * 'value node list
-  | RL of (key * 'value) * 'value tree * 'value node list;
+  | Metis_RL of (key * 'value) * 'value tree * 'value node list;
 
 fun fromSpineLR nodes =
     case nodes of
@@ -4670,7 +4623,7 @@
     case nodes of
       [] => NONE
     | Node {key,value,left,...} :: nodes =>
-      SOME (RL ((key,value),left,nodes));
+      SOME (Metis_RL ((key,value),left,nodes));
 
 fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
 
@@ -4683,12 +4636,12 @@
 fun readIterator iter =
     case iter of
       LR (key_value,_,_) => key_value
-    | RL (key_value,_,_) => key_value;
+    | Metis_RL (key_value,_,_) => key_value;
 
 fun advanceIterator iter =
     case iter of
       LR (_,tree,nodes) => addLR nodes tree
-    | RL (_,tree,nodes) => addRL nodes tree;
+    | Metis_RL (_,tree,nodes) => addRL nodes tree;
 
 fun foldIterator f acc io =
     case io of
@@ -4774,22 +4727,22 @@
 (* ------------------------------------------------------------------------- *)
 
 datatype 'value map =
-    Map of 'value tree;
-
-(* ------------------------------------------------------------------------- *)
-(* Map debugging functions.                                                  *)
+    Metis_Map of 'value tree;
+
+(* ------------------------------------------------------------------------- *)
+(* Metis_Map debugging functions.                                                  *)
 (* ------------------------------------------------------------------------- *)
 
 (*BasicDebug
 fun checkInvariants s m =
     let
-      val Map tree = m
+      val Metis_Map tree = m
 
       val _ = treeCheckInvariants tree
     in
       m
     end
-    handle Bug bug => raise Bug (s ^ "\n" ^ "KeyMap.checkInvariants: " ^ bug);
+    handle Bug bug => raise Bug (s ^ "\n" ^ "Metis_KeyMap.checkInvariants: " ^ bug);
 *)
 
 (* ------------------------------------------------------------------------- *)
@@ -4800,21 +4753,21 @@
     let
       val tree = treeNew ()
     in
-      Map tree
+      Metis_Map tree
     end;
 
 fun singleton key_value =
     let
       val tree = treeSingleton key_value
     in
-      Map tree
-    end;
-
-(* ------------------------------------------------------------------------- *)
-(* Map size.                                                                 *)
-(* ------------------------------------------------------------------------- *)
-
-fun size (Map tree) = treeSize tree;
+      Metis_Map tree
+    end;
+
+(* ------------------------------------------------------------------------- *)
+(* Metis_Map size.                                                                 *)
+(* ------------------------------------------------------------------------- *)
+
+fun size (Metis_Map tree) = treeSize tree;
 
 fun null m = size m = 0;
 
@@ -4822,26 +4775,26 @@
 (* Querying.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun peekKey (Map tree) key = treePeekKey key tree;
-
-fun peek (Map tree) key = treePeek key tree;
+fun peekKey (Metis_Map tree) key = treePeekKey key tree;
+
+fun peek (Metis_Map tree) key = treePeek key tree;
 
 fun inDomain key m = Option.isSome (peek m key);
 
 fun get m key =
     case peek m key of
-      NONE => raise Error "KeyMap.get: element not found"
+      NONE => raise Error "Metis_KeyMap.get: element not found"
     | SOME value => value;
 
-fun pick (Map tree) = treePick tree;
-
-fun nth (Map tree) n = treeNth n tree;
+fun pick (Metis_Map tree) = treePick tree;
+
+fun nth (Metis_Map tree) n = treeNth n tree;
 
 fun random m =
     let
       val n = size m
     in
-      if n = 0 then raise Bug "KeyMap.random: empty"
+      if n = 0 then raise Bug "Metis_KeyMap.random: empty"
       else nth m (randomInt n)
     end;
 
@@ -4849,17 +4802,17 @@
 (* Adding.                                                                   *)
 (* ------------------------------------------------------------------------- *)
 
-fun insert (Map tree) key_value =
+fun insert (Metis_Map tree) key_value =
     let
       val tree = treeInsert key_value tree
     in
-      Map tree
+      Metis_Map tree
     end;
 
 (*BasicDebug
 val insert = fn m => fn kv =>
-    checkInvariants "KeyMap.insert: result"
-      (insert (checkInvariants "KeyMap.insert: input" m) kv);
+    checkInvariants "Metis_KeyMap.insert: result"
+      (insert (checkInvariants "Metis_KeyMap.insert: input" m) kv);
 *)
 
 fun insertList m =
@@ -4873,50 +4826,50 @@
 (* Removing.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun delete (Map tree) dkey =
+fun delete (Metis_Map tree) dkey =
     let
       val tree = treeDelete dkey tree
     in
-      Map tree
+      Metis_Map tree
     end;
 
 (*BasicDebug
 val delete = fn m => fn k =>
-    checkInvariants "KeyMap.delete: result"
-      (delete (checkInvariants "KeyMap.delete: input" m) k);
+    checkInvariants "Metis_KeyMap.delete: result"
+      (delete (checkInvariants "Metis_KeyMap.delete: input" m) k);
 *)
 
 fun remove m key = if inDomain key m then delete m key else m;
 
-fun deletePick (Map tree) =
+fun deletePick (Metis_Map tree) =
     let
       val (key_value,tree) = treeDeletePick tree
     in
-      (key_value, Map tree)
+      (key_value, Metis_Map tree)
     end;
 
 (*BasicDebug
 val deletePick = fn m =>
     let
-      val (kv,m) = deletePick (checkInvariants "KeyMap.deletePick: input" m)
-    in
-      (kv, checkInvariants "KeyMap.deletePick: result" m)
-    end;
-*)
-
-fun deleteNth (Map tree) n =
+      val (kv,m) = deletePick (checkInvariants "Metis_KeyMap.deletePick: input" m)
+    in
+      (kv, checkInvariants "Metis_KeyMap.deletePick: result" m)
+    end;
+*)
+
+fun deleteNth (Metis_Map tree) n =
     let
       val (key_value,tree) = treeDeleteNth n tree
     in
-      (key_value, Map tree)
+      (key_value, Metis_Map tree)
     end;
 
 (*BasicDebug
 val deleteNth = fn m => fn n =>
     let
-      val (kv,m) = deleteNth (checkInvariants "KeyMap.deleteNth: input" m) n
-    in
-      (kv, checkInvariants "KeyMap.deleteNth: result" m)
+      val (kv,m) = deleteNth (checkInvariants "Metis_KeyMap.deleteNth: input" m) n
+    in
+      (kv, checkInvariants "Metis_KeyMap.deleteNth: result" m)
     end;
 *)
 
@@ -4924,7 +4877,7 @@
     let
       val n = size m
     in
-      if n = 0 then raise Bug "KeyMap.deleteRandom: empty"
+      if n = 0 then raise Bug "Metis_KeyMap.deleteRandom: empty"
       else deleteNth m (randomInt n)
     end;
 
@@ -4932,89 +4885,89 @@
 (* Joining (all join operations prefer keys in the second map).              *)
 (* ------------------------------------------------------------------------- *)
 
-fun merge {first,second,both} (Map tree1) (Map tree2) =
+fun merge {first,second,both} (Metis_Map tree1) (Metis_Map tree2) =
     let
       val tree = treeMerge first second both tree1 tree2
     in
-      Map tree
+      Metis_Map tree
     end;
 
 (*BasicDebug
 val merge = fn f => fn m1 => fn m2 =>
-    checkInvariants "KeyMap.merge: result"
+    checkInvariants "Metis_KeyMap.merge: result"
       (merge f
-         (checkInvariants "KeyMap.merge: input 1" m1)
-         (checkInvariants "KeyMap.merge: input 2" m2));
-*)
-
-fun op union f (Map tree1) (Map tree2) =
+         (checkInvariants "Metis_KeyMap.merge: input 1" m1)
+         (checkInvariants "Metis_KeyMap.merge: input 2" m2));
+*)
+
+fun union f (Metis_Map tree1) (Metis_Map tree2) =
     let
       fun f2 kv = f (kv,kv)
 
       val tree = treeUnion f f2 tree1 tree2
     in
-      Map tree
+      Metis_Map tree
     end;
 
 (*BasicDebug
-val op union = fn f => fn m1 => fn m2 =>
-    checkInvariants "KeyMap.union: result"
+val union = fn f => fn m1 => fn m2 =>
+    checkInvariants "Metis_KeyMap.union: result"
       (union f
-         (checkInvariants "KeyMap.union: input 1" m1)
-         (checkInvariants "KeyMap.union: input 2" m2));
-*)
-
-fun intersect f (Map tree1) (Map tree2) =
+         (checkInvariants "Metis_KeyMap.union: input 1" m1)
+         (checkInvariants "Metis_KeyMap.union: input 2" m2));
+*)
+
+fun intersect f (Metis_Map tree1) (Metis_Map tree2) =
     let
       val tree = treeIntersect f tree1 tree2
     in
-      Map tree
+      Metis_Map tree
     end;
 
 (*BasicDebug
 val intersect = fn f => fn m1 => fn m2 =>
-    checkInvariants "KeyMap.intersect: result"
+    checkInvariants "Metis_KeyMap.intersect: result"
       (intersect f
-         (checkInvariants "KeyMap.intersect: input 1" m1)
-         (checkInvariants "KeyMap.intersect: input 2" m2));
+         (checkInvariants "Metis_KeyMap.intersect: input 1" m1)
+         (checkInvariants "Metis_KeyMap.intersect: input 2" m2));
 *)
 
 (* ------------------------------------------------------------------------- *)
 (* Iterators over maps.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
-fun mkIterator (Map tree) = treeMkIterator tree;
-
-fun mkRevIterator (Map tree) = treeMkRevIterator tree;
+fun mkIterator (Metis_Map tree) = treeMkIterator tree;
+
+fun mkRevIterator (Metis_Map tree) = treeMkRevIterator tree;
 
 (* ------------------------------------------------------------------------- *)
 (* Mapping and folding.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
-fun mapPartial f (Map tree) =
+fun mapPartial f (Metis_Map tree) =
     let
       val tree = treeMapPartial f tree
     in
-      Map tree
+      Metis_Map tree
     end;
 
 (*BasicDebug
 val mapPartial = fn f => fn m =>
-    checkInvariants "KeyMap.mapPartial: result"
-      (mapPartial f (checkInvariants "KeyMap.mapPartial: input" m));
-*)
-
-fun map f (Map tree) =
+    checkInvariants "Metis_KeyMap.mapPartial: result"
+      (mapPartial f (checkInvariants "Metis_KeyMap.mapPartial: input" m));
+*)
+
+fun map f (Metis_Map tree) =
     let
       val tree = treeMap f tree
     in
-      Map tree
+      Metis_Map tree
     end;
 
 (*BasicDebug
 val map = fn f => fn m =>
-    checkInvariants "KeyMap.map: result"
-      (map f (checkInvariants "KeyMap.map: input" m));
+    checkInvariants "Metis_KeyMap.map: result"
+      (map f (checkInvariants "Metis_KeyMap.map: input" m));
 *)
 
 fun transform f = map (fn (_,value) => f value);
@@ -5079,7 +5032,7 @@
         LESS => LESS
       | EQUAL =>
         let
-          val Map _ = m1
+          val Metis_Map _ = m1
 
           val io1 = mkIterator m1
           and io2 = mkIterator m2
@@ -5092,7 +5045,7 @@
     pointerEqual (m1,m2) orelse
     (size m1 = size m2 andalso
      let
-       val Map _ = m1
+       val Metis_Map _ = m1
 
        val io1 = mkIterator m1
        and io2 = mkIterator m2
@@ -5101,22 +5054,22 @@
      end);
 
 (* ------------------------------------------------------------------------- *)
-(* Set operations on the domain.                                             *)
-(* ------------------------------------------------------------------------- *)
-
-fun unionDomain (Map tree1) (Map tree2) =
+(* Metis_Set operations on the domain.                                             *)
+(* ------------------------------------------------------------------------- *)
+
+fun unionDomain (Metis_Map tree1) (Metis_Map tree2) =
     let
       val tree = treeUnionDomain tree1 tree2
     in
-      Map tree
+      Metis_Map tree
     end;
 
 (*BasicDebug
 val unionDomain = fn m1 => fn m2 =>
-    checkInvariants "KeyMap.unionDomain: result"
+    checkInvariants "Metis_KeyMap.unionDomain: result"
       (unionDomain
-         (checkInvariants "KeyMap.unionDomain: input 1" m1)
-         (checkInvariants "KeyMap.unionDomain: input 2" m2));
+         (checkInvariants "Metis_KeyMap.unionDomain: input 1" m1)
+         (checkInvariants "Metis_KeyMap.unionDomain: input 2" m2));
 *)
 
 local
@@ -5124,23 +5077,23 @@
 in
   fun unionListDomain ms =
       case ms of
-        [] => raise Bug "KeyMap.unionListDomain: no sets"
+        [] => raise Bug "Metis_KeyMap.unionListDomain: no sets"
       | m :: ms => List.foldl uncurriedUnionDomain m ms;
 end;
 
-fun intersectDomain (Map tree1) (Map tree2) =
+fun intersectDomain (Metis_Map tree1) (Metis_Map tree2) =
     let
       val tree = treeIntersectDomain tree1 tree2
     in
-      Map tree
+      Metis_Map tree
     end;
 
 (*BasicDebug
 val intersectDomain = fn m1 => fn m2 =>
-    checkInvariants "KeyMap.intersectDomain: result"
+    checkInvariants "Metis_KeyMap.intersectDomain: result"
       (intersectDomain
-         (checkInvariants "KeyMap.intersectDomain: input 1" m1)
-         (checkInvariants "KeyMap.intersectDomain: input 2" m2));
+         (checkInvariants "Metis_KeyMap.intersectDomain: input 1" m1)
+         (checkInvariants "Metis_KeyMap.intersectDomain: input 2" m2));
 *)
 
 local
@@ -5148,23 +5101,23 @@
 in
   fun intersectListDomain ms =
       case ms of
-        [] => raise Bug "KeyMap.intersectListDomain: no sets"
+        [] => raise Bug "Metis_KeyMap.intersectListDomain: no sets"
       | m :: ms => List.foldl uncurriedIntersectDomain m ms;
 end;
 
-fun differenceDomain (Map tree1) (Map tree2) =
+fun differenceDomain (Metis_Map tree1) (Metis_Map tree2) =
     let
       val tree = treeDifferenceDomain tree1 tree2
     in
-      Map tree
+      Metis_Map tree
     end;
 
 (*BasicDebug
 val differenceDomain = fn m1 => fn m2 =>
-    checkInvariants "KeyMap.differenceDomain: result"
+    checkInvariants "Metis_KeyMap.differenceDomain: result"
       (differenceDomain
-         (checkInvariants "KeyMap.differenceDomain: input 1" m1)
-         (checkInvariants "KeyMap.differenceDomain: input 2" m2));
+         (checkInvariants "Metis_KeyMap.differenceDomain: input 1" m1)
+         (checkInvariants "Metis_KeyMap.differenceDomain: input 2" m2));
 *)
 
 fun symmetricDifferenceDomain m1 m2 =
@@ -5172,7 +5125,7 @@
 
 fun equalDomain m1 m2 = equal (K (K true)) m1 m2;
 
-fun subsetDomain (Map tree1) (Map tree2) =
+fun subsetDomain (Metis_Map tree1) (Metis_Map tree2) =
     treeSubsetDomain tree1 tree2;
 
 fun disjointDomain m1 m2 = null (intersectDomain m1 m2);
@@ -5202,14 +5155,14 @@
 
 end
 
-structure IntMap =
-KeyMap (Metis.IntOrdered); (* MODIFIED by Jasmin Blanchette *)
-
-structure IntPairMap =
-KeyMap (Metis.IntPairOrdered); (* MODIFIED by Jasmin Blanchette *)
-
-structure StringMap =
-KeyMap (Metis.StringOrdered); (* MODIFIED by Jasmin Blanchette *)
+structure Metis_IntMap =
+Metis_KeyMap (Metis_IntOrdered);
+
+structure Metis_IntPairMap =
+Metis_KeyMap (Metis_IntPairOrdered);
+
+structure Metis_StringMap =
+Metis_KeyMap (Metis_StringOrdered);
 
 (**** Original file: Set.sig ****)
 
@@ -5218,7 +5171,7 @@
 (* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature Set =
+signature Metis_Set =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -5236,7 +5189,7 @@
 val singleton : ('elt * 'elt -> order) -> 'elt -> 'elt set
 
 (* ------------------------------------------------------------------------- *)
-(* Set size.                                                                 *)
+(* Metis_Set size.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
 val null : 'elt set -> bool
@@ -5353,7 +5306,7 @@
 (* Converting to and from maps.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-type ('elt,'a) map = ('elt,'a) Metis.Map.map
+type ('elt,'a) map = ('elt,'a) Metis_Map.map
 
 val mapPartial : ('elt -> 'a option) -> 'elt set -> ('elt,'a) map
 
@@ -5385,97 +5338,88 @@
 
 (**** Original file: Set.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* FINITE SETS IMPLEMENTED WITH RANDOMLY BALANCED TREES                      *)
 (* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-structure Set :> Set =
+structure Metis_Set :> Metis_Set =
 struct
 
 (* ------------------------------------------------------------------------- *)
 (* A type of finite sets.                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-type ('elt,'a) map = ('elt,'a) Map.map;
-
-datatype 'elt set = Set of ('elt,unit) map;
+type ('elt,'a) map = ('elt,'a) Metis_Map.map;
+
+datatype 'elt set = Metis_Set of ('elt,unit) map;
 
 (* ------------------------------------------------------------------------- *)
 (* Converting to and from maps.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-fun dest (Set m) = m;
+fun dest (Metis_Set m) = m;
 
 fun mapPartial f =
     let
       fun mf (elt,()) = f elt
     in
-      fn Set m => Map.mapPartial mf m
+      fn Metis_Set m => Metis_Map.mapPartial mf m
     end;
 
 fun map f =
     let
       fun mf (elt,()) = f elt
     in
-      fn Set m => Map.map mf m
-    end;
-
-fun domain m = Set (Map.transform (fn _ => ()) m);
+      fn Metis_Set m => Metis_Map.map mf m
+    end;
+
+fun domain m = Metis_Set (Metis_Map.transform (fn _ => ()) m);
 
 (* ------------------------------------------------------------------------- *)
 (* Constructors.                                                             *)
 (* ------------------------------------------------------------------------- *)
 
-fun empty cmp = Set (Map.new cmp);
-
-fun singleton cmp elt = Set (Map.singleton cmp (elt,()));
-
-(* ------------------------------------------------------------------------- *)
-(* Set size.                                                                 *)
-(* ------------------------------------------------------------------------- *)
-
-fun null (Set m) = Map.null m;
-
-fun size (Set m) = Map.size m;
+fun empty cmp = Metis_Set (Metis_Map.new cmp);
+
+fun singleton cmp elt = Metis_Set (Metis_Map.singleton cmp (elt,()));
+
+(* ------------------------------------------------------------------------- *)
+(* Metis_Set size.                                                                 *)
+(* ------------------------------------------------------------------------- *)
+
+fun null (Metis_Set m) = Metis_Map.null m;
+
+fun size (Metis_Set m) = Metis_Map.size m;
 
 (* ------------------------------------------------------------------------- *)
 (* Querying.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun peek (Set m) elt =
-    case Map.peekKey m elt of
+fun peek (Metis_Set m) elt =
+    case Metis_Map.peekKey m elt of
       SOME (elt,()) => SOME elt
     | NONE => NONE;
 
-fun member elt (Set m) = Map.inDomain elt m;
-
-fun pick (Set m) =
-    let
-      val (elt,_) = Map.pick m
+fun member elt (Metis_Set m) = Metis_Map.inDomain elt m;
+
+fun pick (Metis_Set m) =
+    let
+      val (elt,_) = Metis_Map.pick m
     in
       elt
     end;
 
-fun nth (Set m) n =
-    let
-      val (elt,_) = Map.nth m n
+fun nth (Metis_Set m) n =
+    let
+      val (elt,_) = Metis_Map.nth m n
     in
       elt
     end;
 
-fun random (Set m) =
-    let
-      val (elt,_) = Map.random m
+fun random (Metis_Set m) =
+    let
+      val (elt,_) = Metis_Map.random m
     in
       elt
     end;
@@ -5484,11 +5428,11 @@
 (* Adding.                                                                   *)
 (* ------------------------------------------------------------------------- *)
 
-fun add (Set m) elt =
-    let
-      val m = Map.insert m (elt,())
-    in
-      Set m
+fun add (Metis_Set m) elt =
+    let
+      val m = Metis_Map.insert m (elt,())
+    in
+      Metis_Set m
     end;
 
 local
@@ -5501,68 +5445,68 @@
 (* Removing.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun delete (Set m) elt =
-    let
-      val m = Map.delete m elt
-    in
-      Set m
-    end;
-
-fun remove (Set m) elt =
-    let
-      val m = Map.remove m elt
-    in
-      Set m
-    end;
-
-fun deletePick (Set m) =
-    let
-      val ((elt,()),m) = Map.deletePick m
-    in
-      (elt, Set m)
-    end;
-
-fun deleteNth (Set m) n =
-    let
-      val ((elt,()),m) = Map.deleteNth m n
-    in
-      (elt, Set m)
-    end;
-
-fun deleteRandom (Set m) =
-    let
-      val ((elt,()),m) = Map.deleteRandom m
-    in
-      (elt, Set m)
+fun delete (Metis_Set m) elt =
+    let
+      val m = Metis_Map.delete m elt
+    in
+      Metis_Set m
+    end;
+
+fun remove (Metis_Set m) elt =
+    let
+      val m = Metis_Map.remove m elt
+    in
+      Metis_Set m
+    end;
+
+fun deletePick (Metis_Set m) =
+    let
+      val ((elt,()),m) = Metis_Map.deletePick m
+    in
+      (elt, Metis_Set m)
+    end;
+
+fun deleteNth (Metis_Set m) n =
+    let
+      val ((elt,()),m) = Metis_Map.deleteNth m n
+    in
+      (elt, Metis_Set m)
+    end;
+
+fun deleteRandom (Metis_Set m) =
+    let
+      val ((elt,()),m) = Metis_Map.deleteRandom m
+    in
+      (elt, Metis_Set m)
     end;
 
 (* ------------------------------------------------------------------------- *)
 (* Joining.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-fun union (Set m1) (Set m2) = Set (Map.unionDomain m1 m2);
+fun union (Metis_Set m1) (Metis_Set m2) = Metis_Set (Metis_Map.unionDomain m1 m2);
 
 fun unionList sets =
     let
       val ms = List.map dest sets
     in
-      Set (Map.unionListDomain ms)
-    end;
-
-fun intersect (Set m1) (Set m2) = Set (Map.intersectDomain m1 m2);
+      Metis_Set (Metis_Map.unionListDomain ms)
+    end;
+
+fun intersect (Metis_Set m1) (Metis_Set m2) = Metis_Set (Metis_Map.intersectDomain m1 m2);
 
 fun intersectList sets =
     let
       val ms = List.map dest sets
     in
-      Set (Map.intersectListDomain ms)
-    end;
-
-fun difference (Set m1) (Set m2) =
-    Set (Map.differenceDomain m1 m2);
-
-fun symmetricDifference (Set m1) (Set m2) =
-    Set (Map.symmetricDifferenceDomain m1 m2);
+      Metis_Set (Metis_Map.intersectListDomain ms)
+    end;
+
+fun difference (Metis_Set m1) (Metis_Set m2) =
+    Metis_Set (Metis_Map.differenceDomain m1 m2);
+
+fun symmetricDifference (Metis_Set m1) (Metis_Set m2) =
+    Metis_Set (Metis_Map.symmetricDifferenceDomain m1 m2);
 
 (* ------------------------------------------------------------------------- *)
 (* Mapping and folding.                                                      *)
@@ -5572,18 +5516,18 @@
     let
       fun mpred (elt,()) = pred elt
     in
-      fn Set m => Set (Map.filter mpred m)
+      fn Metis_Set m => Metis_Set (Metis_Map.filter mpred m)
     end;
 
 fun partition pred =
     let
       fun mpred (elt,()) = pred elt
     in
-      fn Set m =>
+      fn Metis_Set m =>
          let
-           val (m1,m2) = Map.partition mpred m
+           val (m1,m2) = Metis_Map.partition mpred m
          in
-           (Set m1, Set m2)
+           (Metis_Set m1, Metis_Set m2)
          end
     end;
 
@@ -5591,21 +5535,21 @@
     let
       fun mf (elt,()) = f elt
     in
-      fn Set m => Map.app mf m
+      fn Metis_Set m => Metis_Map.app mf m
     end;
 
 fun foldl f =
     let
       fun mf (elt,(),acc) = f (elt,acc)
     in
-      fn acc => fn Set m => Map.foldl mf acc m
+      fn acc => fn Metis_Set m => Metis_Map.foldl mf acc m
     end;
 
 fun foldr f =
     let
       fun mf (elt,(),acc) = f (elt,acc)
     in
-      fn acc => fn Set m => Map.foldr mf acc m
+      fn acc => fn Metis_Set m => Metis_Map.foldr mf acc m
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -5616,8 +5560,8 @@
     let
       fun mp (elt,()) = p elt
     in
-      fn Set m =>
-         case Map.findl mp m of
+      fn Metis_Set m =>
+         case Metis_Map.findl mp m of
            SOME (elt,()) => SOME elt
          | NONE => NONE
     end;
@@ -5626,8 +5570,8 @@
     let
       fun mp (elt,()) = p elt
     in
-      fn Set m =>
-         case Map.findr mp m of
+      fn Metis_Set m =>
+         case Metis_Map.findr mp m of
            SOME (elt,()) => SOME elt
          | NONE => NONE
     end;
@@ -5636,35 +5580,35 @@
     let
       fun mf (elt,()) = f elt
     in
-      fn Set m => Map.firstl mf m
+      fn Metis_Set m => Metis_Map.firstl mf m
     end;
 
 fun firstr f =
     let
       fun mf (elt,()) = f elt
     in
-      fn Set m => Map.firstr mf m
+      fn Metis_Set m => Metis_Map.firstr mf m
     end;
 
 fun exists p =
     let
       fun mp (elt,()) = p elt
     in
-      fn Set m => Map.exists mp m
+      fn Metis_Set m => Metis_Map.exists mp m
     end;
 
 fun all p =
     let
       fun mp (elt,()) = p elt
     in
-      fn Set m => Map.all mp m
+      fn Metis_Set m => Metis_Map.all mp m
     end;
 
 fun count p =
     let
       fun mp (elt,()) = p elt
     in
-      fn Set m => Map.count mp m
+      fn Metis_Set m => Metis_Map.count mp m
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -5675,13 +5619,13 @@
 
 fun equalValue () () = true;
 
-fun compare (Set m1, Set m2) = Map.compare compareValue (m1,m2);
-
-fun equal (Set m1) (Set m2) = Map.equal equalValue m1 m2;
-
-fun subset (Set m1) (Set m2) = Map.subsetDomain m1 m2;
-
-fun disjoint (Set m1) (Set m2) = Map.disjointDomain m1 m2;
+fun compare (Metis_Set m1, Metis_Set m2) = Metis_Map.compare compareValue (m1,m2);
+
+fun equal (Metis_Set m1) (Metis_Set m2) = Metis_Map.equal equalValue m1 m2;
+
+fun subset (Metis_Set m1) (Metis_Set m2) = Metis_Map.subsetDomain m1 m2;
+
+fun disjoint (Metis_Set m1) (Metis_Set m2) = Metis_Map.disjointDomain m1 m2;
 
 (* ------------------------------------------------------------------------- *)
 (* Converting to and from lists.                                             *)
@@ -5694,7 +5638,7 @@
       foldr inc []
     end;
 
-fun toList (Set m) = Map.keys m;
+fun toList (Metis_Set m) = Metis_Map.keys m;
 
 fun fromList cmp elts = addList (empty cmp) elts;
 
@@ -5709,23 +5653,22 @@
 (* Iterators over sets                                                       *)
 (* ------------------------------------------------------------------------- *)
 
-type 'elt iterator = ('elt,unit) Map.iterator;
-
-fun mkIterator (Set m) = Map.mkIterator m;
-
-fun mkRevIterator (Set m) = Map.mkRevIterator m;
+type 'elt iterator = ('elt,unit) Metis_Map.iterator;
+
+fun mkIterator (Metis_Set m) = Metis_Map.mkIterator m;
+
+fun mkRevIterator (Metis_Set m) = Metis_Map.mkRevIterator m;
 
 fun readIterator iter =
     let
-      val (elt,()) = Map.readIterator iter
+      val (elt,()) = Metis_Map.readIterator iter
     in
       elt
     end;
 
-fun advanceIterator iter = Map.advanceIterator iter;
-
-end
-end;
+fun advanceIterator iter = Metis_Map.advanceIterator iter;
+
+end
 
 (**** Original file: ElementSet.sig ****)
 
@@ -5734,7 +5677,7 @@
 (* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature ElementSet =
+signature Metis_ElementSet =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -5758,7 +5701,7 @@
 val singleton : element -> set
 
 (* ------------------------------------------------------------------------- *)
-(* Set size.                                                                 *)
+(* Metis_Set size.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
 val null : set -> bool
@@ -5912,7 +5855,7 @@
 (* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-functor ElementSet (KM : KeyMap) :> ElementSet
+functor Metis_ElementSet (KM : Metis_KeyMap) :> Metis_ElementSet
 where type element = KM.key and type 'a map = 'a KM.map =
 struct
 
@@ -5928,72 +5871,72 @@
 
 type 'a map = 'a KM.map;
 
-datatype set = Set of unit map;
+datatype set = Metis_Set of unit map;
 
 (* ------------------------------------------------------------------------- *)
 (* Converting to and from maps.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-fun dest (Set m) = m;
+fun dest (Metis_Set m) = m;
 
 fun mapPartial f =
     let
       fun mf (elt,()) = f elt
     in
-      fn Set m => KM.mapPartial mf m
+      fn Metis_Set m => KM.mapPartial mf m
     end;
 
 fun map f =
     let
       fun mf (elt,()) = f elt
     in
-      fn Set m => KM.map mf m
-    end;
-
-fun domain m = Set (KM.transform (fn _ => ()) m);
+      fn Metis_Set m => KM.map mf m
+    end;
+
+fun domain m = Metis_Set (KM.transform (fn _ => ()) m);
 
 (* ------------------------------------------------------------------------- *)
 (* Constructors.                                                             *)
 (* ------------------------------------------------------------------------- *)
 
-val empty = Set (KM.new ());
-
-fun singleton elt = Set (KM.singleton (elt,()));
-
-(* ------------------------------------------------------------------------- *)
-(* Set size.                                                                 *)
-(* ------------------------------------------------------------------------- *)
-
-fun null (Set m) = KM.null m;
-
-fun size (Set m) = KM.size m;
+val empty = Metis_Set (KM.new ());
+
+fun singleton elt = Metis_Set (KM.singleton (elt,()));
+
+(* ------------------------------------------------------------------------- *)
+(* Metis_Set size.                                                                 *)
+(* ------------------------------------------------------------------------- *)
+
+fun null (Metis_Set m) = KM.null m;
+
+fun size (Metis_Set m) = KM.size m;
 
 (* ------------------------------------------------------------------------- *)
 (* Querying.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun peek (Set m) elt =
+fun peek (Metis_Set m) elt =
     case KM.peekKey m elt of
       SOME (elt,()) => SOME elt
     | NONE => NONE;
 
-fun member elt (Set m) = KM.inDomain elt m;
-
-fun pick (Set m) =
+fun member elt (Metis_Set m) = KM.inDomain elt m;
+
+fun pick (Metis_Set m) =
     let
       val (elt,_) = KM.pick m
     in
       elt
     end;
 
-fun nth (Set m) n =
+fun nth (Metis_Set m) n =
     let
       val (elt,_) = KM.nth m n
     in
       elt
     end;
 
-fun random (Set m) =
+fun random (Metis_Set m) =
     let
       val (elt,_) = KM.random m
     in
@@ -6004,11 +5947,11 @@
 (* Adding.                                                                   *)
 (* ------------------------------------------------------------------------- *)
 
-fun add (Set m) elt =
+fun add (Metis_Set m) elt =
     let
       val m = KM.insert m (elt,())
     in
-      Set m
+      Metis_Set m
     end;
 
 local
@@ -6021,68 +5964,68 @@
 (* Removing.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun delete (Set m) elt =
+fun delete (Metis_Set m) elt =
     let
       val m = KM.delete m elt
     in
-      Set m
-    end;
-
-fun remove (Set m) elt =
+      Metis_Set m
+    end;
+
+fun remove (Metis_Set m) elt =
     let
       val m = KM.remove m elt
     in
-      Set m
-    end;
-
-fun deletePick (Set m) =
+      Metis_Set m
+    end;
+
+fun deletePick (Metis_Set m) =
     let
       val ((elt,()),m) = KM.deletePick m
     in
-      (elt, Set m)
-    end;
-
-fun deleteNth (Set m) n =
+      (elt, Metis_Set m)
+    end;
+
+fun deleteNth (Metis_Set m) n =
     let
       val ((elt,()),m) = KM.deleteNth m n
     in
-      (elt, Set m)
-    end;
-
-fun deleteRandom (Set m) =
+      (elt, Metis_Set m)
+    end;
+
+fun deleteRandom (Metis_Set m) =
     let
       val ((elt,()),m) = KM.deleteRandom m
     in
-      (elt, Set m)
+      (elt, Metis_Set m)
     end;
 
 (* ------------------------------------------------------------------------- *)
 (* Joining.                                                                  *)
 (* ------------------------------------------------------------------------- *)
 
-fun op union (Set m1) (Set m2) = Set (KM.unionDomain m1 m2);
+fun union (Metis_Set m1) (Metis_Set m2) = Metis_Set (KM.unionDomain m1 m2);
 
 fun unionList sets =
     let
       val ms = List.map dest sets
     in
-      Set (KM.unionListDomain ms)
-    end;
-
-fun intersect (Set m1) (Set m2) = Set (KM.intersectDomain m1 m2);
+      Metis_Set (KM.unionListDomain ms)
+    end;
+
+fun intersect (Metis_Set m1) (Metis_Set m2) = Metis_Set (KM.intersectDomain m1 m2);
 
 fun intersectList sets =
     let
       val ms = List.map dest sets
     in
-      Set (KM.intersectListDomain ms)
-    end;
-
-fun difference (Set m1) (Set m2) =
-    Set (KM.differenceDomain m1 m2);
-
-fun symmetricDifference (Set m1) (Set m2) =
-    Set (KM.symmetricDifferenceDomain m1 m2);
+      Metis_Set (KM.intersectListDomain ms)
+    end;
+
+fun difference (Metis_Set m1) (Metis_Set m2) =
+    Metis_Set (KM.differenceDomain m1 m2);
+
+fun symmetricDifference (Metis_Set m1) (Metis_Set m2) =
+    Metis_Set (KM.symmetricDifferenceDomain m1 m2);
 
 (* ------------------------------------------------------------------------- *)
 (* Mapping and folding.                                                      *)
@@ -6092,18 +6035,18 @@
     let
       fun mpred (elt,()) = pred elt
     in
-      fn Set m => Set (KM.filter mpred m)
+      fn Metis_Set m => Metis_Set (KM.filter mpred m)
     end;
 
 fun partition pred =
     let
       fun mpred (elt,()) = pred elt
     in
-      fn Set m =>
+      fn Metis_Set m =>
          let
            val (m1,m2) = KM.partition mpred m
          in
-           (Set m1, Set m2)
+           (Metis_Set m1, Metis_Set m2)
          end
     end;
 
@@ -6111,21 +6054,21 @@
     let
       fun mf (elt,()) = f elt
     in
-      fn Set m => KM.app mf m
+      fn Metis_Set m => KM.app mf m
     end;
 
 fun foldl f =
     let
       fun mf (elt,(),acc) = f (elt,acc)
     in
-      fn acc => fn Set m => KM.foldl mf acc m
+      fn acc => fn Metis_Set m => KM.foldl mf acc m
     end;
 
 fun foldr f =
     let
       fun mf (elt,(),acc) = f (elt,acc)
     in
-      fn acc => fn Set m => KM.foldr mf acc m
+      fn acc => fn Metis_Set m => KM.foldr mf acc m
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -6136,7 +6079,7 @@
     let
       fun mp (elt,()) = p elt
     in
-      fn Set m =>
+      fn Metis_Set m =>
          case KM.findl mp m of
            SOME (elt,()) => SOME elt
          | NONE => NONE
@@ -6146,7 +6089,7 @@
     let
       fun mp (elt,()) = p elt
     in
-      fn Set m =>
+      fn Metis_Set m =>
          case KM.findr mp m of
            SOME (elt,()) => SOME elt
          | NONE => NONE
@@ -6156,35 +6099,35 @@
     let
       fun mf (elt,()) = f elt
     in
-      fn Set m => KM.firstl mf m
+      fn Metis_Set m => KM.firstl mf m
     end;
 
 fun firstr f =
     let
       fun mf (elt,()) = f elt
     in
-      fn Set m => KM.firstr mf m
+      fn Metis_Set m => KM.firstr mf m
     end;
 
 fun exists p =
     let
       fun mp (elt,()) = p elt
     in
-      fn Set m => KM.exists mp m
+      fn Metis_Set m => KM.exists mp m
     end;
 
 fun all p =
     let
       fun mp (elt,()) = p elt
     in
-      fn Set m => KM.all mp m
+      fn Metis_Set m => KM.all mp m
     end;
 
 fun count p =
     let
       fun mp (elt,()) = p elt
     in
-      fn Set m => KM.count mp m
+      fn Metis_Set m => KM.count mp m
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -6195,13 +6138,13 @@
 
 fun equalValue () () = true;
 
-fun compare (Set m1, Set m2) = KM.compare compareValue (m1,m2);
-
-fun equal (Set m1) (Set m2) = KM.equal equalValue m1 m2;
-
-fun op subset (Set m1) (Set m2) = KM.subsetDomain m1 m2;
-
-fun disjoint (Set m1) (Set m2) = KM.disjointDomain m1 m2;
+fun compare (Metis_Set m1, Metis_Set m2) = KM.compare compareValue (m1,m2);
+
+fun equal (Metis_Set m1) (Metis_Set m2) = KM.equal equalValue m1 m2;
+
+fun subset (Metis_Set m1) (Metis_Set m2) = KM.subsetDomain m1 m2;
+
+fun disjoint (Metis_Set m1) (Metis_Set m2) = KM.disjointDomain m1 m2;
 
 (* ------------------------------------------------------------------------- *)
 (* Converting to and from lists.                                             *)
@@ -6214,7 +6157,7 @@
       foldr inc []
     end;
 
-fun toList (Set m) = KM.keys m;
+fun toList (Metis_Set m) = KM.keys m;
 
 fun fromList elts = addList empty elts;
 
@@ -6231,9 +6174,9 @@
 
 type iterator = unit KM.iterator;
 
-fun mkIterator (Set m) = KM.mkIterator m;
-
-fun mkRevIterator (Set m) = KM.mkRevIterator m;
+fun mkIterator (Metis_Set m) = KM.mkIterator m;
+
+fun mkRevIterator (Metis_Set m) = KM.mkRevIterator m;
 
 fun readIterator iter =
     let
@@ -6246,14 +6189,14 @@
 
 end
 
-structure IntSet =
-ElementSet (IntMap);
-
-structure IntPairSet =
-ElementSet (IntPairMap);
-
-structure StringSet =
-ElementSet (StringMap);
+structure Metis_IntSet =
+Metis_ElementSet (Metis_IntMap);
+
+structure Metis_IntPairSet =
+Metis_ElementSet (Metis_IntPairMap);
+
+structure Metis_StringSet =
+Metis_ElementSet (Metis_StringMap);
 
 (**** Original file: Sharing.sig ****)
 
@@ -6262,7 +6205,7 @@
 (* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Sharing =
+signature Metis_Sharing =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -6305,26 +6248,17 @@
 
 (**** Original file: Sharing.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* PRESERVING SHARING OF ML VALUES                                           *)
 (* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Sharing :> Sharing =
+structure Metis_Sharing :> Metis_Sharing =
 struct
 
 infix ==
 
-val op== = Portable.pointerEqual;
+val op== = Metis_Portable.pointerEqual;
 
 (* ------------------------------------------------------------------------- *)
 (* Option operations.                                                        *)
@@ -6431,7 +6365,7 @@
 
 fun updateNth (n,x) l =
     let
-      val (a,b) = Useful.revDivide l n
+      val (a,b) = Metis_Useful.revDivide l n
     in
       case b of
         [] => raise Subscript
@@ -6440,7 +6374,7 @@
 
 fun setify l =
     let
-      val l' = Useful.setify l
+      val l' = Metis_Useful.setify l
     in
       if length l' = length l then l else l'
     end;
@@ -6451,15 +6385,15 @@
 
 fun cache cmp f =
     let
-      val cache = Unsynchronized.ref (Map.new cmp)
+      val cache = Unsynchronized.ref (Metis_Map.new cmp)
     in
       fn a =>
-         case Map.peek (!cache) a of
+         case Metis_Map.peek (!cache) a of
            SOME b => b
          | NONE =>
            let
              val b = f a
-             val () = cache := Map.insert (!cache) (a,b)
+             val () = cache := Metis_Map.insert (!cache) (a,b)
            in
              b
            end
@@ -6469,10 +6403,9 @@
 (* Hash consing.                                                             *)
 (* ------------------------------------------------------------------------- *)
 
-fun hashCons cmp = cache cmp Useful.I;
-
-end
-end;
+fun hashCons cmp = cache cmp Metis_Useful.I;
+
+end
 
 (**** Original file: Heap.sig ****)
 
@@ -6481,7 +6414,7 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Heap =
+signature Metis_Heap =
 sig
 
 type 'a heap
@@ -6502,7 +6435,7 @@
 
 val toList : 'a heap -> 'a list
 
-val toStream : 'a heap -> 'a Metis.Stream.stream
+val toStream : 'a heap -> 'a Metis_Stream.stream
 
 val toString : 'a heap -> string
 
@@ -6510,28 +6443,19 @@
 
 (**** Original file: Heap.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* A HEAP DATATYPE FOR ML                                                    *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Heap :> Heap =
+structure Metis_Heap :> Metis_Heap =
 struct
 
 (* Leftist heaps as in Purely Functional Data Structures, by Chris Okasaki *)
 
 datatype 'a node = E | T of int * 'a * 'a node * 'a node;
 
-datatype 'a heap = Heap of ('a * 'a -> order) * int * 'a node;
+datatype 'a heap = Metis_Heap of ('a * 'a -> order) * int * 'a node;
 
 fun rank E = 0
   | rank (T (r,_,_,_)) = r;
@@ -6551,19 +6475,19 @@
       mrg
     end;
 
-fun new cmp = Heap (cmp,0,E);
-
-fun add (Heap (f,n,a)) x = Heap (f, n + 1, merge f (T (1,x,E,E), a));
-
-fun size (Heap (_, n, _)) = n;
+fun new cmp = Metis_Heap (cmp,0,E);
+
+fun add (Metis_Heap (f,n,a)) x = Metis_Heap (f, n + 1, merge f (T (1,x,E,E), a));
+
+fun size (Metis_Heap (_, n, _)) = n;
 
 fun null h = size h = 0;
 
-fun top (Heap (_,_,E)) = raise Empty
-  | top (Heap (_, _, T (_,x,_,_))) = x;
-
-fun remove (Heap (_,_,E)) = raise Empty
-  | remove (Heap (f, n, T (_,x,a,b))) = (x, Heap (f, n - 1, merge f (a,b)));
+fun top (Metis_Heap (_,_,E)) = raise Empty
+  | top (Metis_Heap (_, _, T (_,x,_,_))) = x;
+
+fun remove (Metis_Heap (_,_,E)) = raise Empty
+  | remove (Metis_Heap (f, n, T (_,x,a,b))) = (x, Metis_Heap (f, n - 1, merge f (a,b)));
 
 fun app f =
     let
@@ -6571,7 +6495,7 @@
         | ap (E :: rest) = ap rest
         | ap (T (_,d,a,b) :: rest) = (f d; ap (a :: b :: rest))
     in
-      fn Heap (_,_,a) => ap [a]
+      fn Metis_Heap (_,_,a) => ap [a]
     end;
 
 fun toList h =
@@ -6584,19 +6508,18 @@
       end;
 
 fun toStream h =
-    if null h then Stream.Nil
+    if null h then Metis_Stream.Nil
     else
       let
         val (x,h) = remove h
       in
-        Stream.Cons (x, fn () => toStream h)
+        Metis_Stream.Cons (x, fn () => toStream h)
       end;
 
 fun toString h =
-    "Heap[" ^ (if null h then "" else Int.toString (size h)) ^ "]";
-
-end
-end;
+    "Metis_Heap[" ^ (if null h then "" else Int.toString (size h)) ^ "]";
+
+end
 
 (**** Original file: Print.sig ****)
 
@@ -6605,7 +6528,7 @@
 (* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Print =
+signature Metis_Print =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -6621,7 +6544,7 @@
   | AddBreak of int
   | AddNewline
 
-type ppstream = ppStep Metis.Stream.stream
+type ppstream = ppStep Metis_Stream.stream
 
 type 'a pp = 'a -> ppstream
 
@@ -6647,7 +6570,7 @@
 
 val program : ppstream list -> ppstream
 
-val stream : ppstream Metis.Stream.stream -> ppstream
+val stream : ppstream Metis_Stream.stream -> ppstream
 
 val block : breakStyle -> int -> ppstream -> ppstream
 
@@ -6675,7 +6598,7 @@
 
 val ppOpList : string -> 'a pp -> 'a list pp
 
-val ppOpStream : string -> 'a pp -> 'a Metis.Stream.stream pp
+val ppOpStream : string -> 'a pp -> 'a Metis_Stream.stream pp
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty-printers for common types.                                         *)
@@ -6703,7 +6626,7 @@
 
 val ppList : 'a pp -> 'a list pp
 
-val ppStream : 'a pp -> 'a Metis.Stream.stream pp
+val ppStream : 'a pp -> 'a Metis_Stream.stream pp
 
 val ppOption : 'a pp -> 'a option pp
 
@@ -6727,7 +6650,7 @@
        precedence : int,
        leftAssoc : bool} list
 
-val tokensInfixes : infixes -> StringSet.set (* MODIFIED by Jasmin Blanchette *)
+val tokensInfixes : infixes -> Metis_StringSet.set
 
 val layerInfixes :
     infixes ->
@@ -6742,7 +6665,7 @@
 (* Executing pretty-printers to generate lines.                              *)
 (* ------------------------------------------------------------------------- *)
 
-val execute : {lineLength : int} -> ppstream -> string Metis.Stream.stream
+val execute : {lineLength : int} -> ppstream -> string Metis_Stream.stream
 
 (* ------------------------------------------------------------------------- *)
 (* Executing pretty-printers with a global line length.                      *)
@@ -6752,7 +6675,7 @@
 
 val toString : 'a pp -> 'a -> string
 
-val toStream : 'a pp -> 'a -> string Metis.Stream.stream
+val toStream : 'a pp -> 'a -> string Metis_Stream.stream
 
 val trace : 'a pp -> string -> 'a -> unit
 
@@ -6760,24 +6683,15 @@
 
 (**** Original file: Print.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* PRETTY-PRINTING                                                           *)
 (* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Print :> Print =
+structure Metis_Print :> Metis_Print =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Constants.                                                                *)
@@ -6792,20 +6706,20 @@
 fun revAppend xs s =
     case xs of
       [] => s ()
-    | x :: xs => revAppend xs (K (Stream.Cons (x,s)));
+    | x :: xs => revAppend xs (K (Metis_Stream.Cons (x,s)));
 
 fun revConcat strm =
     case strm of
-      Stream.Nil => Stream.Nil
-    | Stream.Cons (h,t) => revAppend h (revConcat o t);
+      Metis_Stream.Nil => Metis_Stream.Nil
+    | 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
-        Stream.Nil => Stream.Nil
-      | Stream.Cons (h,t) => Stream.maps join Stream.singleton h (t ());
+        Metis_Stream.Nil => Metis_Stream.Nil
+      | Metis_Stream.Cons (h,t) => Metis_Stream.maps join Metis_Stream.singleton h (t ());
 end;
 
 local
@@ -6831,7 +6745,7 @@
   | AddBreak of int
   | AddNewline;
 
-type ppstream = ppStep Stream.stream;
+type ppstream = ppStep Metis_Stream.stream;
 
 type 'a pp = 'a -> ppstream;
 
@@ -6852,29 +6766,29 @@
 (* Pretty-printer primitives.                                                *)
 (* ------------------------------------------------------------------------- *)
 
-fun beginBlock style indent = Stream.singleton (BeginBlock (style,indent));
-
-val endBlock = Stream.singleton EndBlock;
-
-fun addString s = Stream.singleton (AddString s);
-
-fun addBreak spaces = Stream.singleton (AddBreak spaces);
-
-val addNewline = Stream.singleton AddNewline;
-
-val skip : ppstream = Stream.Nil;
-
-fun sequence pp1 pp2 : ppstream = Stream.append pp1 (K pp2);
-
-local
-  fun dup pp n () = if n = 1 then pp else Stream.append pp (dup pp (n - 1));
+fun beginBlock style indent = Metis_Stream.singleton (BeginBlock (style,indent));
+
+val endBlock = Metis_Stream.singleton EndBlock;
+
+fun addString s = Metis_Stream.singleton (AddString s);
+
+fun addBreak spaces = Metis_Stream.singleton (AddBreak spaces);
+
+val addNewline = Metis_Stream.singleton AddNewline;
+
+val skip : ppstream = Metis_Stream.Nil;
+
+fun sequence pp1 pp2 : ppstream = Metis_Stream.append pp1 (K pp2);
+
+local
+  fun dup pp n () = if n = 1 then pp else Metis_Stream.append pp (dup pp (n - 1));
 in
   fun duplicate n pp = if n = 0 then skip else dup pp n ();
 end;
 
-val program : ppstream list -> ppstream = Stream.concatList;
-
-val stream : ppstream Stream.stream -> ppstream = Stream.concat;
+val program : ppstream list -> ppstream = Metis_Stream.concatList;
+
+val stream : ppstream Metis_Stream.stream -> ppstream = Metis_Stream.concat;
 
 fun block style indent pp = program [beginBlock style indent, pp, endBlock];
 
@@ -6942,11 +6856,11 @@
     let
       fun ppOpA a = sequence (ppOp s) (ppA a)
     in
-      fn Stream.Nil => skip
-       | Stream.Cons (h,t) =>
+      fn Metis_Stream.Nil => skip
+       | Metis_Stream.Cons (h,t) =>
          blockProgram Inconsistent 0
            [ppA h,
-            Stream.concat (Stream.map ppOpA (t ()))]
+            Metis_Stream.concat (Metis_Stream.map ppOpA (t ()))]
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -7091,10 +7005,10 @@
 
   fun add t l acc =
       case acc of
-        [] => raise Bug "Print.layerInfixOps.layer"
+        [] => 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 "Print.layerInfixOps: mixed assocs";
+        else raise Bug "Metis_Print.layerInfixOps: mixed assocs";
 
   fun layer ({token = t, precedence = p, leftAssoc = l}, (acc,p')) =
       let
@@ -7119,12 +7033,12 @@
 val tokensLayeredInfixes =
     let
       fun addToken ({leftSpaces = _, token = t, rightSpaces = _}, s) =
-          StringSet.add s t
+          Metis_StringSet.add s t
 
       fun addTokens ({tokens = t, leftAssoc = _}, s) =
           List.foldl addToken s t
     in
-      List.foldl addTokens StringSet.empty
+      List.foldl addTokens Metis_StringSet.empty
     end;
 
 val tokensInfixes = tokensLayeredInfixes o layerInfixes;
@@ -7136,10 +7050,10 @@
             let
               val {leftSpaces = _, token = t, rightSpaces = _} = token
             in
-              StringMap.insert m (t, ppOpSpaces token)
-            end
-      in
-        List.foldl add (StringMap.new ())
+              Metis_StringMap.insert m (t, ppOpSpaces token)
+            end
+      in
+        List.foldl add (Metis_StringMap.new ())
       end;
 in
   fun ppGenInfix {tokens,leftAssoc} =
@@ -7152,7 +7066,7 @@
                  case dest tm of
                    NONE => NONE
                  | SOME (t,a,b) =>
-                   case StringMap.peek tokenMap t of
+                   case Metis_StringMap.peek tokenMap t of
                      NONE => NONE
                    | SOME p => SOME (p,a,b)
 
@@ -7188,7 +7102,7 @@
 
            fun isOp t =
                case dest t of
-                 SOME (x,_,_) => StringSet.member x toks
+                 SOME (x,_,_) => Metis_StringSet.member x toks
                | NONE => false
 
            fun subpr (tmr as (tm,_)) =
@@ -7288,7 +7202,7 @@
 
 (*BasicDebug
       val _ = not empty orelse size = 0 orelse
-              raise Bug "Print.isEmptyBlock: bad size"
+              raise Bug "Metis_Print.isEmptyBlock: bad size"
 *)
     in
       empty
@@ -7297,9 +7211,9 @@
 fun checkBlock ind block =
     let
       val Block {indent, style = _, size, chunks} = block
-      val _ = indent >= ind orelse raise Bug "Print.checkBlock: bad indents"
+      val _ = indent >= ind orelse raise Bug "Metis_Print.checkBlock: bad indents"
       val size' = checkChunks indent chunks
-      val _ = size = size' orelse raise Bug "Print.checkBlock: wrong size"
+      val _ = size = size' orelse raise Bug "Metis_Print.checkBlock: wrong size"
     in
       size
     end
@@ -7367,7 +7281,7 @@
        ()
      end
      handle Error err => raise Bug err)
-    handle Bug bug => raise Bug ("Print.checkState: " ^ bug);
+    handle Bug bug => raise Bug ("Metis_Print.checkState: " ^ bug);
 *)
 
 fun isFinalState state =
@@ -7375,7 +7289,7 @@
       val State {blocks,lineIndent,lineSize} = state
     in
       case blocks of
-        [] => raise Bug "Print.isFinalState: no block"
+        [] => raise Bug "Metis_Print.isFinalState: no block"
       | [block] => isEmptyBlock block
       | _ :: _ :: _ => false
     end;
@@ -7392,7 +7306,7 @@
 
   fun renderBreaks lineIndent lineIndent' breaks lines =
       case rev breaks of
-        [] => raise Bug "Print.renderBreaks"
+        [] => raise Bug "Metis_Print.renderBreaks"
       | c :: cs =>
         let
           val lines = renderBreak lineIndent (c,lines)
@@ -7546,7 +7460,7 @@
   and forceBreakChunk cumulatingChunks chunk =
       case chunk of
         StringChunk _ => NONE
-      | BreakChunk _ => raise Bug "Print.forceBreakChunk: BreakChunk"
+      | BreakChunk _ => raise Bug "Metis_Print.forceBreakChunk: BreakChunk"
       | BlockChunk block =>
         case forceBreakBlock cumulatingChunks block of
           SOME (breaks,block,lineIndent,lineSize) =>
@@ -7564,12 +7478,12 @@
         let
           val cumulatingChunks =
               case cumulatingChunks of
-                [] => raise Bug "Print.forceBreak: null cumulatingChunks"
+                [] => raise Bug "Metis_Print.forceBreak: null cumulatingChunks"
               | _ :: cumulatingChunks => cumulatingChunks
 
           val blocks' =
               case blocks' of
-                [] => raise Bug "Print.forceBreak: null blocks'"
+                [] => raise Bug "Metis_Print.forceBreak: null blocks'"
               | _ :: blocks' => blocks'
         in
           case forceBreakBlock cumulatingChunks block of
@@ -7631,7 +7545,7 @@
         normalize lineLength lines state
       end
       handle Bug bug =>
-        raise Bug ("Print.normalize: before normalize:\n" ^ bug)
+        raise Bug ("Metis_Print.normalize: before normalize:\n" ^ bug)
 *)
 
   fun executeBeginBlock (style,ind) lines state =
@@ -7640,7 +7554,7 @@
 
         val Block {indent,...} =
             case blocks of
-              [] => raise Bug "Print.executeBeginBlock: no block"
+              [] => raise Bug "Metis_Print.executeBeginBlock: no block"
             | block :: _ => block
 
         val indent = indent + ind
@@ -7675,7 +7589,7 @@
 
         val (lineSize,blocks) =
             case blocks of
-              [] => raise Bug "Print.executeEndBlock: no block"
+              [] => raise Bug "Metis_Print.executeEndBlock: no block"
             | topBlock :: blocks =>
               let
                 val Block
@@ -7708,7 +7622,7 @@
                            chunks = topChunks}
                   in
                     case blocks of
-                      [] => raise Error "Print.executeEndBlock: no block"
+                      [] => raise Error "Metis_Print.executeEndBlock: no block"
                     | block :: blocks =>
                       let
                         val Block {indent,style,size,chunks} = block
@@ -7748,7 +7662,7 @@
 
         val blocks =
             case blocks of
-              [] => raise Bug "Print.executeAddString: no block"
+              [] => raise Bug "Metis_Print.executeAddString: no block"
             | Block {indent,style,size,chunks} :: blocks =>
               let
                 val size = size + n
@@ -7786,7 +7700,7 @@
 
         val (blocks,lineSize) =
             case blocks of
-              [] => raise Bug "Print.executeAddBreak: no block"
+              [] => raise Bug "Metis_Print.executeAddBreak: no block"
             | Block {indent,style,size,chunks} :: blocks' =>
               case chunks of
                 [] => (blocks,lineSize)
@@ -7847,13 +7761,13 @@
                 val (lines,state) = executeBigBreak lineLength lines state
 
 (*BasicDebug
-                val _ = isFinalState state orelse raise Bug "Print.final"
+                val _ = isFinalState state orelse raise Bug "Metis_Print.final"
 *)
               in
                 lines
               end
       in
-        if null lines then Stream.Nil else Stream.singleton lines
+        if null lines then Metis_Stream.Nil else Metis_Stream.singleton lines
       end;
 in
   fun execute {lineLength} =
@@ -7879,11 +7793,11 @@
               (lines,state)
             end
             handle Bug bug =>
-              raise Bug ("Print.advance: after " ^ ppStepToString step ^
+              raise Bug ("Metis_Print.advance: after " ^ ppStepToString step ^
                          " command:\n" ^ bug)
 *)
       in
-        revConcat o Stream.maps advance (final lineLength []) initialState
+        revConcat o Metis_Stream.maps advance (final lineLength []) initialState
       end;
 end;
 
@@ -7894,19 +7808,18 @@
 val lineLength = Unsynchronized.ref initialLineLength;
 
 fun toStream ppA a =
-    Stream.map (fn s => s ^ "\n")
+    Metis_Stream.map (fn s => s ^ "\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 ());
+      Metis_Stream.Nil => ""
+    | Metis_Stream.Cons (h,t) => Metis_Stream.foldl (fn (s,z) => z ^ "\n" ^ s) h (t ());
 
 fun trace ppX nameX x =
-    Useful.trace (toString (ppOp2 " =" ppString ppX) (nameX,x) ^ "\n");
-
-end
-end;
+    Metis_Useful.trace (toString (ppOp2 " =" ppString ppX) (nameX,x) ^ "\n");
+
+end
 
 (**** Original file: Parse.sig ****)
 
@@ -7915,7 +7828,7 @@
 (* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature Parse =
+signature Metis_Parse =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -7960,10 +7873,10 @@
 val optional : ('a -> 'b * 'a) -> 'a -> 'b option * 'a
 
 (* ------------------------------------------------------------------------- *)
-(* Stream-based parsers.                                                     *)
-(* ------------------------------------------------------------------------- *)
-
-type ('a,'b) parser = 'a Metis.Stream.stream -> 'b * 'a Metis.Stream.stream
+(* Metis_Stream-based parsers.                                                     *)
+(* ------------------------------------------------------------------------- *)
+
+type ('a,'b) parser = 'a Metis_Stream.stream -> 'b * 'a Metis_Stream.stream
 
 val maybe : ('a -> 'b option) -> ('a,'b) parser
 
@@ -7977,19 +7890,19 @@
 (* Parsing whole streams.                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-val fromStream : ('a,'b) parser -> 'a Metis.Stream.stream -> 'b
+val fromStream : ('a,'b) parser -> 'a Metis_Stream.stream -> 'b
 
 val fromList : ('a,'b) parser -> 'a list -> 'b
 
-val everything : ('a, 'b list) parser -> 'a Metis.Stream.stream -> 'b Metis.Stream.stream
+val everything : ('a, 'b list) parser -> 'a Metis_Stream.stream -> 'b Metis_Stream.stream
 
 (* ------------------------------------------------------------------------- *)
 (* Parsing lines of text.                                                    *)
 (* ------------------------------------------------------------------------- *)
 
 val initialize :
-    {lines : string Metis.Stream.stream} ->
-    {chars : char list Metis.Stream.stream,
+    {lines : string Metis_Stream.stream} ->
+    {chars : char list Metis_Stream.stream,
      parseErrorLocation : unit -> string}
 
 val exactChar : char -> (char,unit) parser
@@ -8011,14 +7924,14 @@
 (* ------------------------------------------------------------------------- *)
 
 val parseInfixes :
-    Metis.Print.infixes -> (string * 'a * 'a -> 'a) -> (string,'a) parser ->
+    Metis_Print.infixes -> (string * 'a * 'a -> 'a) -> (string,'a) parser ->
     (string,'a) parser
 
 (* ------------------------------------------------------------------------- *)
 (* Quotations.                                                               *)
 (* ------------------------------------------------------------------------- *)
 
-type 'a quotation = 'a Metis.frag list
+type 'a quotation = 'a frag list
 
 val parseQuotation : ('a -> string) -> (string -> 'b) -> 'a quotation -> 'b
 
@@ -8026,24 +7939,15 @@
 
 (**** Original file: Parse.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* PARSING                                                                   *)
 (* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-structure Parse :> Parse =
+structure Metis_Parse :> Metis_Parse =
 struct
 
-open Useful;
+open Metis_Useful;
 
 infixr 9 >>++
 infixr 8 ++
@@ -8112,17 +8016,17 @@
 fun optional p = (p >> SOME) || (nothing >> K NONE);
 
 (* ------------------------------------------------------------------------- *)
-(* Stream-based parsers.                                                     *)
-(* ------------------------------------------------------------------------- *)
-
-type ('a,'b) parser = 'a Stream.stream -> 'b * 'a Stream.stream
-
-fun maybe p Stream.Nil = raise NoParse
-  | maybe p (Stream.Cons (h,t)) =
+(* Metis_Stream-based parsers.                                                     *)
+(* ------------------------------------------------------------------------- *)
+
+type ('a,'b) parser = 'a Metis_Stream.stream -> 'b * 'a Metis_Stream.stream
+
+fun maybe p Metis_Stream.Nil = raise NoParse
+  | maybe p (Metis_Stream.Cons (h,t)) =
     case p h of SOME r => (r, t ()) | NONE => raise NoParse;
 
-fun finished Stream.Nil = ((), Stream.Nil)
-  | finished (Stream.Cons _) = raise NoParse;
+fun finished Metis_Stream.Nil = ((), Metis_Stream.Nil)
+  | finished (Metis_Stream.Cons _) = raise NoParse;
 
 fun some p = maybe (fn x => if p x then SOME x else NONE);
 
@@ -8139,19 +8043,19 @@
       res
     end;
 
-fun fromList parser l = fromStream parser (Stream.fromList l);
+fun fromList parser l = fromStream parser (Metis_Stream.fromList l);
 
 fun everything parser =
     let
       fun parserOption input =
           SOME (parser input)
-          handle e as NoParse => if Stream.null input then NONE else raise e
+          handle e as NoParse => if Metis_Stream.null input then NONE else raise e
 
       fun parserList input =
           case parserOption input of
-            NONE => Stream.Nil
+            NONE => Metis_Stream.Nil
           | SOME (result,input) =>
-            Stream.append (Stream.fromList result) (fn () => parserList input)
+            Metis_Stream.append (Metis_Stream.fromList result) (fn () => parserList input)
     in
       parserList
     end;
@@ -8174,7 +8078,7 @@
                   explode line
                 end
           in
-            Stream.memoize (Stream.map saveLast lines)
+            Metis_Stream.memoize (Metis_Stream.map saveLast lines)
           end
 
       fun parseErrorLocation () =
@@ -8245,9 +8149,9 @@
 
       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
+            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)
@@ -8255,7 +8159,7 @@
     end;
 
 local
-  fun add ({leftSpaces = _, token = t, rightSpaces = _}, s) = StringSet.add s t;
+  fun add ({leftSpaces = _, token = t, rightSpaces = _}, s) = Metis_StringSet.add s t;
 
   fun parse leftAssoc toks con =
       let
@@ -8268,7 +8172,7 @@
 in
   fun parseLayeredInfixes {tokens,leftAssoc} =
       let
-        val toks = List.foldl add StringSet.empty tokens
+        val toks = List.foldl add Metis_StringSet.empty tokens
       in
         parse leftAssoc toks
       end;
@@ -8276,7 +8180,7 @@
 
 fun parseInfixes ops =
     let
-      val layeredOps = Print.layerInfixes ops
+      val layeredOps = Metis_Print.layerInfixes ops
 
       val iparsers = List.map parseLayeredInfixes layeredOps
     in
@@ -8300,359 +8204,6 @@
   end;
 
 end
-end;
-
-(**** Original file: Options.sig ****)
-
-(* ========================================================================= *)
-(* PROCESSING COMMAND LINE OPTIONS                                           *)
-(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License       *)
-(* ========================================================================= *)
-
-signature Options =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* Option processors take an option with its associated arguments.           *)
-(* ------------------------------------------------------------------------- *)
-
-type proc = string * string list -> unit
-
-type ('a,'x) mkProc = ('x -> proc) -> ('a -> 'x) -> proc
-
-(* ------------------------------------------------------------------------- *)
-(* One command line option: names, arguments, description and a processor.   *)
-(* ------------------------------------------------------------------------- *)
-
-type opt =
-     {switches : string list, arguments : string list,
-      description : string, processor : proc}
-
-(* ------------------------------------------------------------------------- *)
-(* Option processors may raise an OptionExit exception.                      *)
-(* ------------------------------------------------------------------------- *)
-
-type optionExit = {message : string option, usage : bool, success : bool}
-
-exception OptionExit of optionExit
-
-(* ------------------------------------------------------------------------- *)
-(* Constructing option processors.                                           *)
-(* ------------------------------------------------------------------------- *)
-
-val beginOpt : (string,'x) mkProc
-
-val endOpt : unit -> proc
-
-val stringOpt : (string,'x) mkProc
-
-val intOpt : int option * int option -> (int,'x) mkProc
-
-val realOpt : real option * real option -> (real,'x) mkProc
-
-val enumOpt : string list -> (string,'x) mkProc
-
-val optionOpt : string * ('a,'x) mkProc -> ('a option,'x) mkProc
-
-(* ------------------------------------------------------------------------- *)
-(* Basic options useful for all programs.                                    *)
-(* ------------------------------------------------------------------------- *)
-
-val basicOptions : opt list
-
-(* ------------------------------------------------------------------------- *)
-(* All the command line options of a program.                                *)
-(* ------------------------------------------------------------------------- *)
-
-type allOptions =
-     {name : string, version : string, header : string,
-      footer : string, options : opt list}
-
-(* ------------------------------------------------------------------------- *)
-(* Usage information.                                                        *)
-(* ------------------------------------------------------------------------- *)
-
-val versionInformation : allOptions -> string
-
-val usageInformation : allOptions -> string
-
-(* ------------------------------------------------------------------------- *)
-(* Exit the program gracefully.                                              *)
-(* ------------------------------------------------------------------------- *)
-
-val exit : allOptions -> optionExit -> 'exit
-
-val succeed : allOptions -> 'exit
-
-val fail : allOptions -> string -> 'exit
-
-val usage : allOptions -> string -> 'exit
-
-val version : allOptions -> 'exit
-
-(* ------------------------------------------------------------------------- *)
-(* Process the command line options passed to the program.                   *)
-(* ------------------------------------------------------------------------- *)
-
-val processOptions : allOptions -> string list -> string list * string list
-
-end
-
-(**** Original file: Options.sml ****)
-
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
-(* ========================================================================= *)
-(* PROCESSING COMMAND LINE OPTIONS                                           *)
-(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License       *)
-(* ========================================================================= *)
-
-structure Options :> Options =
-struct
-
-infix ##
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* One command line option: names, arguments, description and a processor    *)
-(* ------------------------------------------------------------------------- *)
-
-type proc = string * string list -> unit;
-
-type ('a,'x) mkProc = ('x -> proc) -> ('a -> 'x) -> proc;
-
-type opt = {switches : string list, arguments : string list,
-            description : string, processor : proc};
-
-(* ------------------------------------------------------------------------- *)
-(* Option processors may raise an OptionExit exception                       *)
-(* ------------------------------------------------------------------------- *)
-
-type optionExit = {message : string option, usage : bool, success : bool};
-
-exception OptionExit of optionExit;
-
-(* ------------------------------------------------------------------------- *)
-(* Wrappers for option processors                                            *)
-(* ------------------------------------------------------------------------- *)
-
-fun beginOpt f p (s : string, l : string list) : unit = f (p s) (s,l);
-
-fun endOpt () (_ : string, [] : string list) = ()
-  | endOpt _ (_, _ :: _) = raise Bug "endOpt";
-
-fun stringOpt _ _ (_ : string, []) = raise Bug "stringOpt"
-  | stringOpt f p (s, (h : string) :: t) : unit = f (p h) (s,t);
-
-local
-  fun range NONE NONE = "Z"
-    | range (SOME i) NONE = "{n IN Z | " ^ Int.toString i ^ " <= n}"
-    | range NONE (SOME j) = "{n IN Z | n <= " ^ Int.toString j ^ "}"
-    | range (SOME i) (SOME j) =
-    "{n IN Z | " ^ Int.toString i ^ " <= n <= " ^ Int.toString j ^ "}";
-  fun oLeq (SOME x) (SOME y) = x <= y | oLeq _ _ = true;
-  fun argToInt arg omin omax x =
-    (case Int.fromString x of
-       SOME i =>
-       if oLeq omin (SOME i) andalso oLeq (SOME i) omax then i else
-         raise OptionExit
-           {success = false, usage = false, message =
-            SOME (arg ^ " option needs an integer argument in the range "
-                  ^ range omin omax ^ " (not " ^ x ^ ")")}
-     | NONE =>
-       raise OptionExit
-         {success = false, usage = false, message =
-          SOME (arg ^ " option needs an integer argument (not \"" ^ x ^ "\")")})
-    handle Overflow =>
-       raise OptionExit
-         {success = false, usage = false, message =
-          SOME (arg ^ " option suffered integer overflow on argument " ^ x)};
-in
-  fun intOpt _ _ _ (_,[]) = raise Bug "intOpt"
-    | intOpt (omin,omax) f p (s:string, h :: (t : string list)) : unit =
-      f (p (argToInt s omin omax h)) (s,t);
-end;
-
-local
-  fun range NONE NONE = "R"
-    | range (SOME i) NONE = "{n IN R | " ^ Real.toString i ^ " <= n}"
-    | range NONE (SOME j) = "{n IN R | n <= " ^ Real.toString j ^ "}"
-    | range (SOME i) (SOME j) =
-    "{n IN R | " ^ Real.toString i ^ " <= n <= " ^ Real.toString j ^ "}";
-  fun oLeq (SOME (x:real)) (SOME y) = x <= y | oLeq _ _ = true;
-  fun argToReal arg omin omax x =
-    (case Real.fromString x of
-       SOME i =>
-       if oLeq omin (SOME i) andalso oLeq (SOME i) omax then i else
-         raise OptionExit
-           {success = false, usage = false, message =
-            SOME (arg ^ " option needs an real argument in the range "
-                  ^ range omin omax ^ " (not " ^ x ^ ")")}
-     | NONE =>
-       raise OptionExit
-         {success = false, usage = false, message =
-          SOME (arg ^ " option needs an real argument (not \"" ^ x ^ "\")")})
-in
-  fun realOpt _ _ _ (_,[]) = raise Bug "realOpt"
-    | realOpt (omin,omax) f p (s:string, h :: (t : string list)) : unit =
-      f (p (argToReal s omin omax h)) (s,t);
-end;
-
-fun enumOpt _ _ _ (_,[]) = raise Bug "enumOpt"
-  | enumOpt (choices : string list) f p (s : string, h :: t) : unit =
-    if mem h choices then f (p h) (s,t) else
-      raise OptionExit
-        {success = false, usage = false,
-         message = SOME ("follow parameter " ^ s ^ " with one of {" ^
-                         join "," choices ^ "}, not \"" ^ h ^ "\"")};
-
-fun optionOpt _ _ _ (_,[]) = raise Bug "optionOpt"
-  | optionOpt (x : string, p) f q (s : string, l as h :: t) : unit =
-    if h = x then f (q NONE) (s,t) else p f (q o SOME) (s,l);
-
-(* ------------------------------------------------------------------------- *)
-(* Basic options useful for all programs                                     *)
-(* ------------------------------------------------------------------------- *)
-
-val basicOptions : opt list =
-  [{switches = ["--"], arguments = [],
-    description = "no more options",
-    processor = fn _ => raise Fail "basicOptions: --"},
-   {switches = ["-?","-h","--help"], arguments = [],
-    description = "display option information and exit",
-    processor = fn _ => raise OptionExit
-    {message = SOME "displaying option information",
-     usage = true, success = true}},
-   {switches = ["-v", "--version"], arguments = [],
-    description = "display version information",
-    processor = fn _ => raise Fail "basicOptions: -v, --version"}];
-
-(* ------------------------------------------------------------------------- *)
-(* All the command line options of a program                                 *)
-(* ------------------------------------------------------------------------- *)
-
-type allOptions =
-     {name : string, version : string, header : string,
-      footer : string, options : opt list};
-
-(* ------------------------------------------------------------------------- *)
-(* Usage information                                                         *)
-(* ------------------------------------------------------------------------- *)
-
-fun versionInformation ({version, ...} : allOptions) = version;
-
-fun usageInformation ({name,version,header,footer,options} : allOptions) =
-  let
-    fun listOpts {switches = n, arguments = r, description = s,
-                  processor = _} =
-        let
-          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
-        in
-          [res ^ " ...", " " ^ s]
-        end
-
-    val alignment =
-        [{leftAlign = true, padChar = #"."},
-         {leftAlign = true, padChar = #" "}]
-
-    val table = alignTable alignment (map listOpts options)
-  in
-    header ^ join "\n" table ^ "\n" ^ footer
-  end;
-
-(* ------------------------------------------------------------------------- *)
-(* Exit the program gracefully                                               *)
-(* ------------------------------------------------------------------------- *)
-
-fun exit (allopts : allOptions) (optexit : optionExit) =
-  let
-    val {name, options, ...} = allopts
-    val {message, usage, success} = optexit
-    fun err s = TextIO.output (TextIO.stdErr, s)
-  in
-    case message of NONE => () | SOME m => err (name ^ ": " ^ m ^ "\n");
-    if usage then err (usageInformation allopts) else ();
-    OS.Process.exit (if success then OS.Process.success else OS.Process.failure)
-  end;
-
-fun succeed allopts =
-    exit allopts {message = NONE, usage = false, success = true};
-
-fun fail allopts mesg =
-    exit allopts {message = SOME mesg, usage = false, success = false};
-
-fun usage allopts mesg =
-    exit allopts {message = SOME mesg, usage = true, success = false};
-
-fun version allopts =
-    (print (versionInformation allopts);
-     exit allopts {message = NONE, usage = false, success = true});
-
-(* ------------------------------------------------------------------------- *)
-(* Process the command line options passed to the program                    *)
-(* ------------------------------------------------------------------------- *)
-
-fun processOptions (allopts : allOptions) =
-  let
-    fun findOption x =
-      case List.find (fn {switches = n, ...} => mem x n) (#options allopts) of
-        NONE => raise OptionExit
-                        {message = SOME ("unknown switch \"" ^ x ^ "\""),
-                         usage = true, success = false}
-      | SOME {arguments = r, processor = f, ...} => (r,f)
-
-    fun getArgs x r xs =
-      let
-        fun f 1 = "a following argument"
-          | f m = Int.toString m ^ " following arguments"
-        val m = length r
-        val () =
-          if m <= length xs then () else
-            raise OptionExit
-              {usage = false, success = false, message = SOME
-               (x ^ " option needs " ^ f m ^ ": " ^ join " " r)}
-      in
-        divide xs m
-      end
-
-    fun process [] = ([], [])
-      | process ("--" :: xs) = ([("--",[])], xs)
-      | process ("-v" :: _) = version allopts
-      | process ("--version" :: _) = version allopts
-      | process (x :: xs) =
-      if x = "" orelse x = "-" orelse hd (explode x) <> #"-" then ([], x :: xs)
-      else
-        let
-          val (r,f) = findOption x
-          val (ys,xs) = getArgs x r xs
-          val () = f (x,ys)
-        in
-          (cons (x,ys) ## I) (process xs)
-        end
-  in
-    fn l =>
-    let
-      val (a,b) = process l
-      val a = foldl (fn ((x,xs),ys) => x :: xs @ ys) [] (rev a)
-    in
-      (a,b)
-    end
-    handle OptionExit x => exit allopts x
-  end;
-
-end
-end;
 
 (**** Original file: Name.sig ****)
 
@@ -8661,14 +8212,14 @@
 (* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature Name =
+signature Metis_Name =
 sig
 
 (* ------------------------------------------------------------------------- *)
 (* A type of names.                                                          *)
 (* ------------------------------------------------------------------------- *)
 
-type name = string (* MODIFIED by Jasmin Blanchette *)
+type name = string
 
 (* ------------------------------------------------------------------------- *)
 (* A total ordering.                                                         *)
@@ -8694,7 +8245,7 @@
 (* Parsing and pretty printing.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val pp : name Metis.Print.pp
+val pp : name Metis_Print.pp
 
 val toString : name -> string
 
@@ -8704,24 +8255,15 @@
 
 (**** Original file: Name.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* NAMES                                                                     *)
 (* Copyright (c) 2004 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-structure Name :> Name =
+structure Metis_Name :> Metis_Name =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of names.                                                          *)
@@ -8783,7 +8325,7 @@
 (* Parsing and pretty printing.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val pp = Print.ppString;
+val pp = Metis_Print.ppString;
 
 fun toString s : string = s;
 
@@ -8791,27 +8333,26 @@
 
 end
 
-structure NameOrdered =
-struct type t = Name.name val compare = Name.compare end
-
-structure NameMap = KeyMap (NameOrdered);
-
-structure NameSet =
+structure Metis_NameOrdered =
+struct type t = Metis_Name.name val compare = Metis_Name.compare end
+
+structure Metis_NameMap = Metis_KeyMap (Metis_NameOrdered);
+
+structure Metis_NameSet =
 struct
 
   local
-    structure S = ElementSet (NameMap);
+    structure S = Metis_ElementSet (Metis_NameMap);
   in
     open S;
   end;
 
   val pp =
-      Print.ppMap
+      Metis_Print.ppMap
         toList
-        (Print.ppBracket "{" "}" (Print.ppOpList "," Name.pp));
-
-end
-end;
+        (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_Name.pp));
+
+end
 
 (**** Original file: NameArity.sig ****)
 
@@ -8820,16 +8361,16 @@
 (* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature NameArity =
+signature Metis_NameArity =
 sig
 
 (* ------------------------------------------------------------------------- *)
 (* A type of name/arity pairs.                                               *)
 (* ------------------------------------------------------------------------- *)
 
-type nameArity = Metis.Name.name * int
-
-val name : nameArity -> Metis.Name.name
+type nameArity = Metis_Name.name * int
+
+val name : nameArity -> Metis_Name.name
 
 val arity : nameArity -> int
 
@@ -8859,34 +8400,25 @@
 (* Parsing and pretty printing.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val pp : nameArity Metis.Print.pp
+val pp : nameArity Metis_Print.pp
 
 end
 
 (**** Original file: NameArity.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* NAME/ARITY PAIRS                                                          *)
 (* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure NameArity :> NameArity =
+structure Metis_NameArity :> Metis_NameArity =
 struct
 
 (* ------------------------------------------------------------------------- *)
 (* A type of name/arity pairs.                                               *)
 (* ------------------------------------------------------------------------- *)
 
-type nameArity = Name.name * int;
+type nameArity = Metis_Name.name * int;
 
 fun name ((n,_) : nameArity) = n;
 
@@ -8908,33 +8440,33 @@
 (* ------------------------------------------------------------------------- *)
 
 fun compare ((n1,i1),(n2,i2)) =
-    case Name.compare (n1,n2) of
+    case Metis_Name.compare (n1,n2) of
       LESS => LESS
     | EQUAL => Int.compare (i1,i2)
     | GREATER => GREATER;
 
-fun equal (n1,i1) (n2,i2) = i1 = i2 andalso Name.equal n1 n2;
+fun equal (n1,i1) (n2,i2) = i1 = i2 andalso Metis_Name.equal n1 n2;
 
 (* ------------------------------------------------------------------------- *)
 (* Parsing and pretty printing.                                              *)
 (* ------------------------------------------------------------------------- *)
 
 fun pp (n,i) =
-    Print.blockProgram Print.Inconsistent 0
-      [Name.pp n,
-       Print.addString "/",
-       Print.ppInt i];
-
-end
-
-structure NameArityOrdered =
-struct type t = NameArity.nameArity val compare = NameArity.compare end
-
-structure NameArityMap =
+    Metis_Print.blockProgram Metis_Print.Inconsistent 0
+      [Metis_Name.pp n,
+       Metis_Print.addString "/",
+       Metis_Print.ppInt i];
+
+end
+
+structure Metis_NameArityOrdered =
+struct type t = Metis_NameArity.nameArity val compare = Metis_NameArity.compare end
+
+structure Metis_NameArityMap =
 struct
 
   local
-    structure S = KeyMap (NameArityOrdered);
+    structure S = Metis_KeyMap (Metis_NameArityOrdered);
   in
     open S;
   end;
@@ -8948,24 +8480,23 @@
 
 end
 
-structure NameAritySet =
+structure Metis_NameAritySet =
 struct
 
   local
-    structure S = ElementSet (NameArityMap);
+    structure S = Metis_ElementSet (Metis_NameArityMap);
   in
     open S;
   end;
 
-  val allNullary = all NameArity.nullary;
+  val allNullary = all Metis_NameArity.nullary;
 
   val pp =
-      Print.ppMap
+      Metis_Print.ppMap
         toList
-        (Print.ppBracket "{" "}" (Print.ppOpList "," NameArity.pp));
-
-end
-end;
+        (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_NameArity.pp));
+
+end
 
 (**** Original file: Term.sig ****)
 
@@ -8974,16 +8505,16 @@
 (* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature Term =
+signature Metis_Term =
 sig
 
 (* ------------------------------------------------------------------------- *)
 (* A type of first order logic terms.                                        *)
 (* ------------------------------------------------------------------------- *)
 
-type var = Metis.Name.name
-
-type functionName = Metis.Name.name
+type var = Metis_Name.name
+
+type functionName = Metis_Name.name
 
 type function = functionName * int
 
@@ -9019,9 +8550,9 @@
 
 val fnFunction : term -> function
 
-val functions : term -> Metis.NameAritySet.set
-
-val functionNames : term -> Metis.NameSet.set
+val functions : term -> Metis_NameAritySet.set
+
+val functionNames : term -> Metis_NameSet.set
 
 (* Constants *)
 
@@ -9067,7 +8598,7 @@
 
 val find : (term -> bool) -> term -> path option
 
-val ppPath : path Metis.Print.pp
+val ppPath : path Metis_Print.pp
 
 val pathToString : path -> string
 
@@ -9077,9 +8608,9 @@
 
 val freeIn : var -> term -> bool
 
-val freeVars : term -> Metis.NameSet.set
-
-val freeVarsList : term list -> Metis.NameSet.set
+val freeVars : term -> Metis_NameSet.set
+
+val freeVarsList : term list -> Metis_NameSet.set
 
 (* ------------------------------------------------------------------------- *)
 (* Fresh variables.                                                          *)
@@ -9089,9 +8620,9 @@
 
 val newVars : int -> term list
 
-val variantPrime : Metis.NameSet.set -> var -> var
-
-val variantNum : Metis.NameSet.set -> var -> var
+val variantPrime : Metis_NameSet.set -> var -> var
+
+val variantNum : Metis_NameSet.set -> var -> var
 
 (* ------------------------------------------------------------------------- *)
 (* Special support for terms with type annotations.                          *)
@@ -9111,7 +8642,7 @@
 (* Special support for terms with an explicit function application operator. *)
 (* ------------------------------------------------------------------------- *)
 
-val appName : Metis.Name.name
+val appName : Metis_Name.name
 
 val mkApp : term * term -> term
 
@@ -9129,7 +8660,7 @@
 
 (* Infix symbols *)
 
-val infixes : Metis.Print.infixes Unsynchronized.ref
+val infixes : Metis_Print.infixes Unsynchronized.ref
 
 (* The negation symbol *)
 
@@ -9145,7 +8676,7 @@
 
 (* Pretty printing *)
 
-val pp : term Metis.Print.pp
+val pp : term Metis_Print.pp
 
 val toString : term -> string
 
@@ -9153,46 +8684,37 @@
 
 val fromString : string -> term
 
-val parse : term Metis.Parse.quotation -> term
+val parse : term Metis_Parse.quotation -> term
 
 end
 
 (**** Original file: Term.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC TERMS                                                   *)
 (* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-structure Term :> Term =
+structure Metis_Term :> Metis_Term =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of first order logic terms.                                        *)
 (* ------------------------------------------------------------------------- *)
 
-type var = Name.name;
-
-type functionName = Name.name;
+type var = Metis_Name.name;
+
+type functionName = Metis_Name.name;
 
 type function = functionName * int;
 
 type const = functionName;
 
 datatype term =
-    Var of Name.name
-  | Fn of Name.name * term list;
+    Var of Metis_Name.name
+  | Fn of Metis_Name.name * term list;
 
 (* ------------------------------------------------------------------------- *)
 (* Constructors and destructors.                                             *)
@@ -9205,7 +8727,7 @@
 
 val isVar = can destVar;
 
-fun equalVar v (Var v') = Name.equal v v'
+fun equalVar v (Var v') = Metis_Name.equal v v'
   | equalVar _ _ = false;
 
 (* Functions *)
@@ -9227,17 +8749,17 @@
   fun func fs [] = fs
     | func fs (Var _ :: tms) = func fs tms
     | func fs (Fn (n,l) :: tms) =
-      func (NameAritySet.add fs (n, length l)) (l @ tms);
-in
-  fun functions tm = func NameAritySet.empty [tm];
+      func (Metis_NameAritySet.add fs (n, length l)) (l @ tms);
+in
+  fun functions tm = func Metis_NameAritySet.empty [tm];
 end;
 
 local
   fun func fs [] = fs
     | func fs (Var _ :: tms) = func fs tms
-    | func fs (Fn (n,l) :: tms) = func (NameSet.add fs n) (l @ tms);
-in
-  fun functionNames tm = func NameSet.empty [tm];
+    | func fs (Fn (n,l) :: tms) = func (Metis_NameSet.add fs n) (l @ tms);
+in
+  fun functionNames tm = func Metis_NameSet.empty [tm];
 end;
 
 (* Constants *)
@@ -9254,8 +8776,8 @@
 fun mkBinop f (a,b) = Fn (f,[a,b]);
 
 fun destBinop f (Fn (x,[a,b])) =
-    if Name.equal x f then (a,b) else raise Error "Term.destBinop: wrong binop"
-  | destBinop _ _ = raise Error "Term.destBinop: not a binop";
+    if Metis_Name.equal x f then (a,b) else raise Error "Metis_Term.destBinop: wrong binop"
+  | destBinop _ _ = raise Error "Metis_Term.destBinop: not a binop";
 
 fun isBinop f = can (destBinop f);
 
@@ -9285,18 +8807,18 @@
       let
         val tm1_tm2 = (tm1,tm2)
       in
-        if Portable.pointerEqual tm1_tm2 then cmp tms1 tms2
+        if Metis_Portable.pointerEqual tm1_tm2 then cmp tms1 tms2
         else
           case tm1_tm2 of
             (Var v1, Var v2) =>
-            (case Name.compare (v1,v2) of
+            (case Metis_Name.compare (v1,v2) of
                LESS => LESS
              | EQUAL => cmp tms1 tms2
              | GREATER => GREATER)
           | (Var _, Fn _) => LESS
           | (Fn _, Var _) => GREATER
           | (Fn (f1,a1), Fn (f2,a2)) =>
-            (case Name.compare (f1,f2) of
+            (case Metis_Name.compare (f1,f2) of
                LESS => LESS
              | EQUAL =>
                (case Int.compare (length a1, length a2) of
@@ -9305,7 +8827,7 @@
                 | GREATER => GREATER)
              | GREATER => GREATER)
       end
-    | cmp _ _ = raise Bug "Term.compare";
+    | cmp _ _ = raise Bug "Metis_Term.compare";
 in
   fun compare (tm1,tm2) = cmp [tm1] [tm2];
 end;
@@ -9319,9 +8841,9 @@
 type path = int list;
 
 fun subterm tm [] = tm
-  | subterm (Var _) (_ :: _) = raise Error "Term.subterm: Var"
+  | subterm (Var _) (_ :: _) = raise Error "Metis_Term.subterm: Var"
   | subterm (Fn (_,tms)) (h :: t) =
-    if h >= length tms then raise Error "Term.replace: Fn"
+    if h >= length tms then raise Error "Metis_Term.replace: Fn"
     else subterm (List.nth (tms,h)) t;
 
 local
@@ -9343,15 +8865,15 @@
 fun replace tm ([],res) = if equal res tm then tm else res
   | replace tm (h :: t, res) =
     case tm of
-      Var _ => raise Error "Term.replace: Var"
+      Var _ => raise Error "Metis_Term.replace: Var"
     | Fn (func,tms) =>
-      if h >= length tms then raise Error "Term.replace: Fn"
+      if h >= length tms then raise Error "Metis_Term.replace: Fn"
       else
         let
           val arg = List.nth (tms,h)
           val arg' = replace arg (t,res)
         in
-          if Portable.pointerEqual (arg',arg) then tm
+          if Metis_Portable.pointerEqual (arg',arg) then tm
           else Fn (func, updateNth (h,arg') tms)
         end;
 
@@ -9373,9 +8895,9 @@
       fn tm => search [([],tm)]
     end;
 
-val ppPath = Print.ppList Print.ppInt;
-
-val pathToString = Print.toString ppPath;
+val ppPath = Metis_Print.ppList Metis_Print.ppInt;
+
+val pathToString = Metis_Print.toString ppPath;
 
 (* ------------------------------------------------------------------------- *)
 (* Free variables.                                                           *)
@@ -9383,7 +8905,7 @@
 
 local
   fun free _ [] = false
-    | free v (Var w :: tms) = Name.equal v w orelse free v tms
+    | free v (Var w :: tms) = Metis_Name.equal v w orelse free v tms
     | free v (Fn (_,args) :: tms) = free v (args @ tms);
 in
   fun freeIn v tm = free v [tm];
@@ -9391,10 +8913,10 @@
 
 local
   fun free vs [] = vs
-    | free vs (Var v :: tms) = free (NameSet.add vs v) tms
+    | free vs (Var v :: tms) = free (Metis_NameSet.add vs v) tms
     | free vs (Fn (_,args) :: tms) = free vs (args @ tms);
 in
-  val freeVarsList = free NameSet.empty;
+  val freeVarsList = free Metis_NameSet.empty;
 
   fun freeVars tm = freeVarsList [tm];
 end;
@@ -9403,33 +8925,33 @@
 (* Fresh variables.                                                          *)
 (* ------------------------------------------------------------------------- *)
 
-fun newVar () = Var (Name.newName ());
-
-fun newVars n = map Var (Name.newNames n);
-
-local
-  fun avoidAcceptable avoid n = not (NameSet.member n avoid);
-in
-  fun variantPrime avoid = Name.variantPrime (avoidAcceptable avoid);
-
-  fun variantNum avoid = Name.variantNum (avoidAcceptable avoid);
+fun newVar () = Var (Metis_Name.newName ());
+
+fun newVars n = map Var (Metis_Name.newNames n);
+
+local
+  fun avoidAcceptable avoid n = not (Metis_NameSet.member n avoid);
+in
+  fun variantPrime avoid = Metis_Name.variantPrime (avoidAcceptable avoid);
+
+  fun variantNum avoid = Metis_Name.variantNum (avoidAcceptable avoid);
 end;
 
 (* ------------------------------------------------------------------------- *)
 (* Special support for terms with type annotations.                          *)
 (* ------------------------------------------------------------------------- *)
 
-val hasTypeFunctionName = Name.fromString ":";
+val hasTypeFunctionName = Metis_Name.fromString ":";
 
 val hasTypeFunction = (hasTypeFunctionName,2);
 
 fun destFnHasType ((f,a) : functionName * term list) =
-    if not (Name.equal f hasTypeFunctionName) then
-      raise Error "Term.destFnHasType"
+    if not (Metis_Name.equal f hasTypeFunctionName) then
+      raise Error "Metis_Term.destFnHasType"
     else
       case a of
         [tm,ty] => (tm,ty)
-      | _ => raise Error "Term.destFnHasType";
+      | _ => raise Error "Metis_Term.destFnHasType";
 
 val isFnHasType = can destFnHasType;
 
@@ -9496,24 +9018,24 @@
 (* Special support for terms with an explicit function application operator. *)
 (* ------------------------------------------------------------------------- *)
 
-val appName = Name.fromString ".";
+val appName = Metis_Name.fromString ".";
 
 fun mkFnApp (fTm,aTm) = (appName, [fTm,aTm]);
 
 fun mkApp f_a = Fn (mkFnApp f_a);
 
-fun destFnApp ((f,a) : Name.name * term list) =
-    if not (Name.equal f appName) then raise Error "Term.destFnApp"
+fun destFnApp ((f,a) : Metis_Name.name * term list) =
+    if not (Metis_Name.equal f appName) then raise Error "Metis_Term.destFnApp"
     else
       case a of
         [fTm,aTm] => (fTm,aTm)
-      | _ => raise Error "Term.destFnApp";
+      | _ => raise Error "Metis_Term.destFnApp";
 
 val isFnApp = can destFnApp;
 
 fun destApp tm =
     case tm of
-      Var _ => raise Error "Term.destApp"
+      Var _ => raise Error "Metis_Term.destApp"
     | Fn func => destFnApp func;
 
 val isApp = can destApp;
@@ -9536,7 +9058,7 @@
 (* Operators parsed and printed infix *)
 
 val infixes =
-    (Unsynchronized.ref o Print.Infixes)
+    (Unsynchronized.ref o Metis_Print.Infixes)
       [(* ML symbols *)
        {token = " / ", precedence = 7, leftAssoc = true},
        {token = " div ", precedence = 7, leftAssoc = true},
@@ -9596,73 +9118,73 @@
 
       val bTokens = map #2 bracks @ map #3 bracks
 
-      val iTokens = Print.tokensInfixes iOps
+      val iTokens = Metis_Print.tokensInfixes iOps
 
       fun destI tm =
           case tm of
             Fn (f,[a,b]) =>
             let
-              val f = Name.toString f
-            in
-              if StringSet.member f iTokens then SOME (f,a,b) else NONE
+              val f = Metis_Name.toString f
+            in
+              if Metis_StringSet.member f iTokens then SOME (f,a,b) else NONE
             end
           | _ => NONE
 
-      val iPrinter = Print.ppInfixes iOps destI
+      val iPrinter = Metis_Print.ppInfixes iOps destI
 
       val specialTokens =
-          StringSet.addList iTokens (neg :: quants @ ["$","(",")"] @ bTokens)
-
-      fun vName bv s = StringSet.member s bv
+          Metis_StringSet.addList iTokens (neg :: quants @ ["$","(",")"] @ bTokens)
+
+      fun vName bv s = Metis_StringSet.member s bv
 
       fun checkVarName bv n =
           let
-            val s = Name.toString n
+            val s = Metis_Name.toString n
           in
             if vName bv s then s else "$" ^ s
           end
 
-      fun varName bv = Print.ppMap (checkVarName bv) Print.ppString
+      fun varName bv = Metis_Print.ppMap (checkVarName bv) Metis_Print.ppString
 
       fun checkFunctionName bv n =
           let
-            val s = Name.toString n
-          in
-            if StringSet.member s specialTokens orelse vName bv s then
+            val s = Metis_Name.toString n
+          in
+            if Metis_StringSet.member s specialTokens orelse vName bv s then
               "(" ^ s ^ ")"
             else
               s
           end
 
-      fun functionName bv = Print.ppMap (checkFunctionName bv) Print.ppString
+      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]) =>
-            if Name.toString f <> neg then (0,tm)
+            if Metis_Name.toString f <> neg then (0,tm)
             else let val (n,tm) = stripNeg a in (n + 1, tm) end
           | _ => (0,tm)
 
       val destQuant =
           let
             fun dest q (Fn (q', [Var v, body])) =
-                if Name.toString q' <> q then NONE
+                if Metis_Name.toString q' <> q then NONE
                 else
                   (case dest q body of
                      NONE => SOME (q,v,[],body)
                    | SOME (_,v',vs,body) => SOME (q, v, v' :: vs, body))
               | dest _ _ = NONE
           in
-            fn tm => Useful.first (fn q => dest q tm) quants
+            fn tm => Metis_Useful.first (fn q => dest q tm) quants
           end
 
       fun isQuant tm = Option.isSome (destQuant tm)
 
       fun destBrack (Fn (b,[tm])) =
           let
-            val s = Name.toString b
+            val s = Metis_Name.toString b
           in
             case List.find (fn (n,_,_) => n = s) bracks of
               NONE => NONE
@@ -9673,20 +9195,20 @@
       fun isBrack tm = Option.isSome (destBrack tm)
 
       fun functionArgument bv tm =
-          Print.sequence
-            (Print.addBreak 1)
+          Metis_Print.sequence
+            (Metis_Print.addBreak 1)
             (if isBrack tm then customBracket bv tm
              else if isVar tm orelse isConst tm then basic bv tm
              else bracket bv tm)
 
       and basic bv (Var v) = varName bv v
         | basic bv (Fn (f,args)) =
-          Print.blockProgram Print.Inconsistent 2
+          Metis_Print.blockProgram Metis_Print.Inconsistent 2
             (functionName bv f :: map (functionArgument bv) args)
 
       and customBracket bv tm =
           case destBrack tm of
-            SOME (b1,tm,b2) => Print.ppBracket b1 b2 (term bv) tm
+            SOME (b1,tm,b2) => Metis_Print.ppBracket b1 b2 (term bv) tm
           | NONE => basic bv tm
 
       and innerQuant bv tm =
@@ -9694,45 +9216,45 @@
             NONE => term bv tm
           | SOME (q,v,vs,tm) =>
             let
-              val bv = StringSet.addList bv (map Name.toString (v :: vs))
-            in
-              Print.program
-                [Print.addString q,
+              val bv = Metis_StringSet.addList bv (map Metis_Name.toString (v :: vs))
+            in
+              Metis_Print.program
+                [Metis_Print.addString q,
                  varName bv v,
-                 Print.program
-                   (map (Print.sequence (Print.addBreak 1) o varName bv) vs),
-                 Print.addString ".",
-                 Print.addBreak 1,
+                 Metis_Print.program
+                   (map (Metis_Print.sequence (Metis_Print.addBreak 1) o varName bv) vs),
+                 Metis_Print.addString ".",
+                 Metis_Print.addBreak 1,
                  innerQuant bv tm]
             end
 
       and quantifier bv tm =
           if not (isQuant tm) then customBracket bv tm
-          else Print.block Print.Inconsistent 2 (innerQuant bv tm)
+          else Metis_Print.block Metis_Print.Inconsistent 2 (innerQuant bv tm)
 
       and molecule bv (tm,r) =
           let
             val (n,tm) = stripNeg tm
           in
-            Print.blockProgram Print.Inconsistent n
-              [Print.duplicate n (Print.addString neg),
+            Metis_Print.blockProgram Metis_Print.Inconsistent n
+              [Metis_Print.duplicate n (Metis_Print.addString neg),
                if isI tm orelse (r andalso isQuant tm) then bracket bv tm
                else quantifier bv tm]
           end
 
       and term bv tm = iPrinter (molecule bv) (tm,false)
 
-      and bracket bv tm = Print.ppBracket "(" ")" (term bv) tm
-    in
-      term StringSet.empty
+      and bracket bv tm = Metis_Print.ppBracket "(" ")" (term bv) tm
+    in
+      term Metis_StringSet.empty
     end inputTerm;
 
-val toString = Print.toString pp;
+val toString = Metis_Print.toString pp;
 
 (* Parsing *)
 
 local
-  open Parse;
+  open Metis_Parse;
 
   infixr 9 >>++
   infixr 8 ++
@@ -9793,50 +9315,50 @@
         fun possibleVarName "" = false
           | possibleVarName s = isAlphaNum (String.sub (s,0))
 
-        fun vName bv s = StringSet.member s bv
-
-        val iTokens = Print.tokensInfixes iOps
+        fun vName bv s = Metis_StringSet.member s bv
+
+        val iTokens = Metis_Print.tokensInfixes iOps
 
         val iParser =
-            parseInfixes iOps (fn (f,a,b) => Fn (Name.fromString f, [a,b]))
+            parseInfixes iOps (fn (f,a,b) => Fn (Metis_Name.fromString f, [a,b]))
 
         val specialTokens =
-            StringSet.addList iTokens (neg :: quants @ ["$"] @ bTokens)
+            Metis_StringSet.addList iTokens (neg :: quants @ ["$"] @ bTokens)
 
         fun varName bv =
             some (vName bv) ||
-            (some (Useful.equal "$") ++ some possibleVarName) >> snd
+            (some (Metis_Useful.equal "$") ++ some possibleVarName) >> snd
 
         fun fName bv s =
-            not (StringSet.member s specialTokens) andalso not (vName bv s)
+            not (Metis_StringSet.member s specialTokens) andalso not (vName bv s)
 
         fun functionName bv =
             some (fName bv) ||
-            (some (Useful.equal "(") ++ any ++ some (Useful.equal ")")) >>
+            (some (Metis_Useful.equal "(") ++ any ++ some (Metis_Useful.equal ")")) >>
             (fn (_,(s,_)) => s)
 
         fun basic bv tokens =
             let
-              val var = varName bv >> (Var o Name.fromString)
+              val var = varName bv >> (Var o Metis_Name.fromString)
 
               val const =
-                  functionName bv >> (fn f => Fn (Name.fromString f, []))
+                  functionName bv >> (fn f => Fn (Metis_Name.fromString f, []))
 
               fun bracket (ab,a,b) =
-                  (some (Useful.equal a) ++ term bv ++ some (Useful.equal b)) >>
+                  (some (Metis_Useful.equal a) ++ term bv ++ some (Metis_Useful.equal b)) >>
                   (fn (_,(tm,_)) =>
-                      if ab = "()" then tm else Fn (Name.fromString ab, [tm]))
+                      if ab = "()" then tm else Fn (Metis_Name.fromString ab, [tm]))
 
               fun quantifier q =
                   let
                     fun bind (v,t) =
-                        Fn (Name.fromString q, [Var (Name.fromString v), t])
+                        Fn (Metis_Name.fromString q, [Var (Metis_Name.fromString v), t])
                   in
-                    (some (Useful.equal q) ++
+                    (some (Metis_Useful.equal q) ++
                      atLeastOne (some possibleVarName) ++
-                     some (Useful.equal ".")) >>++
+                     some (Metis_Useful.equal ".")) >>++
                     (fn (_,(vs,_)) =>
-                        term (StringSet.addList bv vs) >>
+                        term (Metis_StringSet.addList bv vs) >>
                         (fn body => foldr bind body vs))
                   end
             in
@@ -9848,51 +9370,50 @@
 
         and molecule bv tokens =
             let
-              val negations = many (some (Useful.equal neg)) >> length
+              val negations = many (some (Metis_Useful.equal neg)) >> length
 
               val function =
                   (functionName bv ++ many (basic bv)) >>
-                  (fn (f,args) => Fn (Name.fromString f, args)) ||
+                  (fn (f,args) => Fn (Metis_Name.fromString f, args)) ||
                   basic bv
             in
               (negations ++ function) >>
-              (fn (n,tm) => funpow n (fn t => Fn (Name.fromString neg, [t])) tm)
+              (fn (n,tm) => funpow n (fn t => Fn (Metis_Name.fromString neg, [t])) tm)
             end tokens
 
         and term bv tokens = iParser (molecule bv) tokens
       in
-        term StringSet.empty
+        term Metis_StringSet.empty
       end inputStream;
 in
   fun fromString input =
       let
-        val chars = Stream.fromList (explode input)
+        val chars = Metis_Stream.fromList (explode input)
 
         val tokens = everything (lexer >> singleton) chars
 
         val terms = everything (termParser >> singleton) tokens
       in
-        case Stream.toList terms of
+        case Metis_Stream.toList terms of
           [tm] => tm
-        | _ => raise Error "Term.fromString"
-      end;
-end;
-
-local
-  val antiquotedTermToString = Print.toString (Print.ppBracket "(" ")" pp);
-in
-  val parse = Parse.parseQuotation antiquotedTermToString fromString;
-end;
-
-end
-
-structure TermOrdered =
-struct type t = Term.term val compare = Term.compare end
-
-structure TermMap = KeyMap (TermOrdered);
-
-structure TermSet = ElementSet (TermMap);
-end;
+        | _ => raise Error "Metis_Term.fromString"
+      end;
+end;
+
+local
+  val antiquotedTermToString = Metis_Print.toString (Metis_Print.ppBracket "(" ")" pp);
+in
+  val parse = Metis_Parse.parseQuotation antiquotedTermToString fromString;
+end;
+
+end
+
+structure Metis_TermOrdered =
+struct type t = Metis_Term.term val compare = Metis_Term.compare end
+
+structure Metis_TermMap = Metis_KeyMap (Metis_TermOrdered);
+
+structure Metis_TermSet = Metis_ElementSet (Metis_TermMap);
 
 (**** Original file: Subst.sig ****)
 
@@ -9901,7 +9422,7 @@
 (* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Subst =
+signature Metis_Subst =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -9920,21 +9441,21 @@
 
 val size : subst -> int
 
-val peek : subst -> Metis.Term.var -> Metis.Term.term option
-
-val insert : subst -> Metis.Term.var * Metis.Term.term -> subst
-
-val singleton : Metis.Term.var * Metis.Term.term -> subst
-
-val toList : subst -> (Metis.Term.var * Metis.Term.term) list
-
-val fromList : (Metis.Term.var * Metis.Term.term) list -> subst
-
-val foldl : (Metis.Term.var * Metis.Term.term * 'a -> 'a) -> 'a -> subst -> 'a
-
-val foldr : (Metis.Term.var * Metis.Term.term * 'a -> 'a) -> 'a -> subst -> 'a
-
-val pp : subst Metis.Print.pp
+val peek : subst -> Metis_Term.var -> Metis_Term.term option
+
+val insert : subst -> Metis_Term.var * Metis_Term.term -> subst
+
+val singleton : Metis_Term.var * Metis_Term.term -> subst
+
+val toList : subst -> (Metis_Term.var * Metis_Term.term) list
+
+val fromList : (Metis_Term.var * Metis_Term.term) list -> subst
+
+val foldl : (Metis_Term.var * Metis_Term.term * 'a -> 'a) -> 'a -> subst -> 'a
+
+val foldr : (Metis_Term.var * Metis_Term.term * 'a -> 'a) -> 'a -> subst -> 'a
+
+val pp : subst Metis_Print.pp
 
 val toString : subst -> string
 
@@ -9948,15 +9469,15 @@
 (* Applying a substitution to a first order logic term.                      *)
 (* ------------------------------------------------------------------------- *)
 
-val subst : subst -> Metis.Term.term -> Metis.Term.term
+val subst : subst -> Metis_Term.term -> Metis_Term.term
 
 (* ------------------------------------------------------------------------- *)
 (* Restricting a substitution to a smaller set of variables.                 *)
 (* ------------------------------------------------------------------------- *)
 
-val restrict : subst -> Metis.NameSet.set -> subst
-
-val remove : subst -> Metis.NameSet.set -> subst
+val restrict : subst -> Metis_NameSet.set -> subst
+
+val remove : subst -> Metis_NameSet.set -> subst
 
 (* ------------------------------------------------------------------------- *)
 (* Composing two substitutions so that the following identity holds:         *)
@@ -9984,108 +9505,99 @@
 (* Creating a substitution to freshen variables.                             *)
 (* ------------------------------------------------------------------------- *)
 
-val freshVars : Metis.NameSet.set -> subst
+val freshVars : Metis_NameSet.set -> subst
 
 (* ------------------------------------------------------------------------- *)
 (* Free variables.                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-val redexes : subst -> Metis.NameSet.set
-
-val residueFreeVars : subst -> Metis.NameSet.set
-
-val freeVars : subst -> Metis.NameSet.set
+val redexes : subst -> Metis_NameSet.set
+
+val residueFreeVars : subst -> Metis_NameSet.set
+
+val freeVars : subst -> Metis_NameSet.set
 
 (* ------------------------------------------------------------------------- *)
 (* Functions.                                                                *)
 (* ------------------------------------------------------------------------- *)
 
-val functions : subst -> Metis.NameAritySet.set
+val functions : subst -> Metis_NameAritySet.set
 
 (* ------------------------------------------------------------------------- *)
 (* Matching for first order logic terms.                                     *)
 (* ------------------------------------------------------------------------- *)
 
-val match : subst -> Metis.Term.term -> Metis.Term.term -> subst  (* raises Error *)
+val match : subst -> Metis_Term.term -> Metis_Term.term -> subst  (* raises Error *)
 
 (* ------------------------------------------------------------------------- *)
 (* Unification for first order logic terms.                                  *)
 (* ------------------------------------------------------------------------- *)
 
-val unify : subst -> Metis.Term.term -> Metis.Term.term -> subst  (* raises Error *)
+val unify : subst -> Metis_Term.term -> Metis_Term.term -> subst  (* raises Error *)
 
 end
 
 (**** Original file: Subst.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC SUBSTITUTIONS                                           *)
 (* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Subst :> Subst =
+structure Metis_Subst :> Metis_Subst =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of first order logic substitutions.                                *)
 (* ------------------------------------------------------------------------- *)
 
-datatype subst = Subst of Term.term NameMap.map;
+datatype subst = Metis_Subst of Metis_Term.term Metis_NameMap.map;
 
 (* ------------------------------------------------------------------------- *)
 (* Basic operations.                                                         *)
 (* ------------------------------------------------------------------------- *)
 
-val empty = Subst (NameMap.new ());
-
-fun null (Subst m) = NameMap.null m;
-
-fun size (Subst m) = NameMap.size m;
-
-fun peek (Subst m) v = NameMap.peek m v;
-
-fun insert (Subst m) v_tm = Subst (NameMap.insert m v_tm);
+val empty = Metis_Subst (Metis_NameMap.new ());
+
+fun null (Metis_Subst m) = Metis_NameMap.null m;
+
+fun size (Metis_Subst m) = Metis_NameMap.size m;
+
+fun peek (Metis_Subst m) v = Metis_NameMap.peek m v;
+
+fun insert (Metis_Subst m) v_tm = Metis_Subst (Metis_NameMap.insert m v_tm);
 
 fun singleton v_tm = insert empty v_tm;
 
-fun toList (Subst m) = NameMap.toList m;
-
-fun fromList l = Subst (NameMap.fromList l);
-
-fun foldl f b (Subst m) = NameMap.foldl f b m;
-
-fun foldr f b (Subst m) = NameMap.foldr f b m;
+fun toList (Metis_Subst m) = Metis_NameMap.toList m;
+
+fun fromList l = Metis_Subst (Metis_NameMap.fromList l);
+
+fun foldl f b (Metis_Subst m) = Metis_NameMap.foldl f b m;
+
+fun foldr f b (Metis_Subst m) = Metis_NameMap.foldr f b m;
 
 fun pp sub =
-    Print.ppBracket "<[" "]>"
-      (Print.ppOpList "," (Print.ppOp2 " |->" Name.pp Term.pp))
+    Metis_Print.ppBracket "<[" "]>"
+      (Metis_Print.ppOpList "," (Metis_Print.ppOp2 " |->" Metis_Name.pp Metis_Term.pp))
       (toList sub);
 
-val toString = Print.toString pp;
+val toString = Metis_Print.toString pp;
 
 (* ------------------------------------------------------------------------- *)
 (* Normalizing removes identity substitutions.                               *)
 (* ------------------------------------------------------------------------- *)
 
 local
-  fun isNotId (v,tm) = not (Term.equalVar v tm);
-in
-  fun normalize (sub as Subst m) =
-      let
-        val m' = NameMap.filter isNotId m
-      in
-        if NameMap.size m = NameMap.size m' then sub else Subst m'
+  fun isNotId (v,tm) = not (Metis_Term.equalVar v tm);
+in
+  fun normalize (sub as Metis_Subst m) =
+      let
+        val m' = Metis_NameMap.filter isNotId m
+      in
+        if Metis_NameMap.size m = Metis_NameMap.size m' then sub else Metis_Subst m'
       end;
 end;
 
@@ -10095,16 +9607,16 @@
 
 fun subst sub =
     let
-      fun tmSub (tm as Term.Var v) =
+      fun tmSub (tm as Metis_Term.Var v) =
           (case peek sub v of
-             SOME tm' => if Portable.pointerEqual (tm,tm') then tm else tm'
+             SOME tm' => if Metis_Portable.pointerEqual (tm,tm') then tm else tm'
            | NONE => tm)
-        | tmSub (tm as Term.Fn (f,args)) =
-          let
-            val args' = Sharing.map tmSub args
-          in
-            if Portable.pointerEqual (args,args') then tm
-            else Term.Fn (f,args')
+        | tmSub (tm as Metis_Term.Fn (f,args)) =
+          let
+            val args' = Metis_Sharing.map tmSub args
+          in
+            if Metis_Portable.pointerEqual (args,args') then tm
+            else Metis_Term.Fn (f,args')
           end
     in
       fn tm => if null sub then tm else tmSub tm
@@ -10114,22 +9626,22 @@
 (* Restricting a substitution to a given set of variables.                   *)
 (* ------------------------------------------------------------------------- *)
 
-fun restrict (sub as Subst m) varSet =
-    let
-      fun isRestrictedVar (v,_) = NameSet.member v varSet
-
-      val m' = NameMap.filter isRestrictedVar m
-    in
-      if NameMap.size m = NameMap.size m' then sub else Subst m'
-    end;
-
-fun remove (sub as Subst m) varSet =
-    let
-      fun isRestrictedVar (v,_) = not (NameSet.member v varSet)
-
-      val m' = NameMap.filter isRestrictedVar m
-    in
-      if NameMap.size m = NameMap.size m' then sub else Subst m'
+fun restrict (sub as Metis_Subst m) varSet =
+    let
+      fun isRestrictedVar (v,_) = Metis_NameSet.member v varSet
+
+      val m' = Metis_NameMap.filter isRestrictedVar m
+    in
+      if Metis_NameMap.size m = Metis_NameMap.size m' then sub else Metis_Subst m'
+    end;
+
+fun remove (sub as Metis_Subst m) varSet =
+    let
+      fun isRestrictedVar (v,_) = not (Metis_NameSet.member v varSet)
+
+      val m' = Metis_NameMap.filter isRestrictedVar m
+    in
+      if Metis_NameMap.size m = Metis_NameMap.size m' then sub else Metis_Subst m'
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -10138,11 +9650,11 @@
 (* subst (compose sub1 sub2) tm = subst sub2 (subst sub1 tm)                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun compose (sub1 as Subst m1) sub2 =
+fun compose (sub1 as Metis_Subst m1) sub2 =
     let
       fun f (v,tm,s) = insert s (v, subst sub2 tm)
     in
-      if null sub2 then sub1 else NameMap.foldl f sub2 m1
+      if null sub2 then sub1 else Metis_NameMap.foldl f sub2 m1
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -10151,13 +9663,13 @@
 
 local
   fun compatible ((_,tm1),(_,tm2)) =
-      if Term.equal tm1 tm2 then SOME tm1
-      else raise Error "Subst.union: incompatible";
-in
-  fun union (s1 as Subst m1) (s2 as Subst m2) =
-      if NameMap.null m1 then s2
-      else if NameMap.null m2 then s1
-      else Subst (NameMap.union compatible m1 m2);
+      if Metis_Term.equal tm1 tm2 then SOME tm1
+      else raise Error "Metis_Subst.union: incompatible";
+in
+  fun union (s1 as Metis_Subst m1) (s2 as Metis_Subst m2) =
+      if Metis_NameMap.null m1 then s2
+      else if Metis_NameMap.null m2 then s1
+      else Metis_Subst (Metis_NameMap.union compatible m1 m2);
 end;
 
 (* ------------------------------------------------------------------------- *)
@@ -10165,12 +9677,12 @@
 (* ------------------------------------------------------------------------- *)
 
 local
-  fun inv (v, Term.Var w, s) =
-      if NameMap.inDomain w s then raise Error "Subst.invert: non-injective"
-      else NameMap.insert s (w, Term.Var v)
-    | inv (_, Term.Fn _, _) = raise Error "Subst.invert: non-variable";
-in
-  fun invert (Subst m) = Subst (NameMap.foldl inv (NameMap.new ()) m);
+  fun inv (v, Metis_Term.Var w, s) =
+      if Metis_NameMap.inDomain w s then raise Error "Metis_Subst.invert: non-injective"
+      else Metis_NameMap.insert s (w, Metis_Term.Var v)
+    | inv (_, Metis_Term.Fn _, _) = raise Error "Metis_Subst.invert: non-variable";
+in
+  fun invert (Metis_Subst m) = Metis_Subst (Metis_NameMap.foldl inv (Metis_NameMap.new ()) m);
 end;
 
 val isRenaming = can invert;
@@ -10181,9 +9693,9 @@
 
 val freshVars =
     let
-      fun add (v,m) = insert m (v, Term.newVar ())
-    in
-      NameSet.foldl add empty
+      fun add (v,m) = insert m (v, Metis_Term.newVar ())
+    in
+      Metis_NameSet.foldl add empty
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -10192,23 +9704,23 @@
 
 val redexes =
     let
-      fun add (v,_,s) = NameSet.add s v
-    in
-      foldl add NameSet.empty
+      fun add (v,_,s) = Metis_NameSet.add s v
+    in
+      foldl add Metis_NameSet.empty
     end;
 
 val residueFreeVars =
     let
-      fun add (_,t,s) = NameSet.union s (Term.freeVars t)
-    in
-      foldl add NameSet.empty
+      fun add (_,t,s) = Metis_NameSet.union s (Metis_Term.freeVars t)
+    in
+      foldl add Metis_NameSet.empty
     end;
 
 val freeVars =
     let
-      fun add (v,t,s) = NameSet.union (NameSet.add s v) (Term.freeVars t)
-    in
-      foldl add NameSet.empty
+      fun add (v,t,s) = Metis_NameSet.union (Metis_NameSet.add s v) (Metis_Term.freeVars t)
+    in
+      foldl add Metis_NameSet.empty
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -10217,9 +9729,9 @@
 
 val functions =
     let
-      fun add (_,t,s) = NameAritySet.union s (Term.functions t)
-    in
-      foldl add NameAritySet.empty
+      fun add (_,t,s) = Metis_NameAritySet.union s (Metis_Term.functions t)
+    in
+      foldl add Metis_NameAritySet.empty
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -10228,22 +9740,22 @@
 
 local
   fun matchList sub [] = sub
-    | matchList sub ((Term.Var v, tm) :: rest) =
+    | matchList sub ((Metis_Term.Var v, tm) :: rest) =
       let
         val sub =
             case peek sub v of
               NONE => insert sub (v,tm)
             | SOME tm' =>
-              if Term.equal tm tm' then sub
-              else raise Error "Subst.match: incompatible matches"
+              if Metis_Term.equal tm tm' then sub
+              else raise Error "Metis_Subst.match: incompatible matches"
       in
         matchList sub rest
       end
-    | matchList sub ((Term.Fn (f1,args1), Term.Fn (f2,args2)) :: rest) =
-      if Name.equal f1 f2 andalso length args1 = length args2 then
+    | matchList sub ((Metis_Term.Fn (f1,args1), Metis_Term.Fn (f2,args2)) :: rest) =
+      if Metis_Name.equal f1 f2 andalso length args1 = length args2 then
         matchList sub (zip args1 args2 @ rest)
-      else raise Error "Subst.match: different structure"
-    | matchList _ _ = raise Error "Subst.match: functions can't match vars";
+      else raise Error "Metis_Subst.match: different structure"
+    | matchList _ _ = raise Error "Metis_Subst.match: functions can't match vars";
 in
   fun match sub tm1 tm2 = matchList sub [(tm1,tm2)];
 end;
@@ -10255,28 +9767,27 @@
 local
   fun solve sub [] = sub
     | solve sub ((tm1_tm2 as (tm1,tm2)) :: rest) =
-      if Portable.pointerEqual tm1_tm2 then solve sub rest
+      if Metis_Portable.pointerEqual tm1_tm2 then solve sub rest
       else solve' sub (subst sub tm1) (subst sub tm2) rest
 
-  and solve' sub (Term.Var v) tm rest =
-      if Term.equalVar v tm then solve sub rest
-      else if Term.freeIn v tm then raise Error "Subst.unify: occurs check"
+  and solve' sub (Metis_Term.Var v) tm rest =
+      if Metis_Term.equalVar v tm then solve sub rest
+      else if Metis_Term.freeIn v tm then raise Error "Metis_Subst.unify: occurs check"
       else
         (case peek sub v of
            NONE => solve (compose sub (singleton (v,tm))) rest
          | SOME tm' => solve' sub tm' tm rest)
-    | solve' sub tm1 (tm2 as Term.Var _) rest = solve' sub tm2 tm1 rest
-    | solve' sub (Term.Fn (f1,args1)) (Term.Fn (f2,args2)) rest =
-      if Name.equal f1 f2 andalso length args1 = length args2 then
+    | solve' sub tm1 (tm2 as Metis_Term.Var _) rest = solve' sub tm2 tm1 rest
+    | solve' sub (Metis_Term.Fn (f1,args1)) (Metis_Term.Fn (f2,args2)) rest =
+      if Metis_Name.equal f1 f2 andalso length args1 = length args2 then
         solve sub (zip args1 args2 @ rest)
       else
-        raise Error "Subst.unify: different structure";
+        raise Error "Metis_Subst.unify: different structure";
 in
   fun unify sub tm1 tm2 = solve sub [(tm1,tm2)];
 end;
 
 end
-end;
 
 (**** Original file: Atom.sig ****)
 
@@ -10285,18 +9796,18 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Atom =
+signature Metis_Atom =
 sig
 
 (* ------------------------------------------------------------------------- *)
 (* A type for storing first order logic atoms.                               *)
 (* ------------------------------------------------------------------------- *)
 
-type relationName = Metis.Name.name
+type relationName = Metis_Name.name
 
 type relation = relationName * int
 
-type atom = relationName * Metis.Term.term list
+type atom = relationName * Metis_Term.term list
 
 (* ------------------------------------------------------------------------- *)
 (* Constructors and destructors.                                             *)
@@ -10304,21 +9815,21 @@
 
 val name : atom -> relationName
 
-val arguments : atom -> Metis.Term.term list
+val arguments : atom -> Metis_Term.term list
 
 val arity : atom -> int
 
 val relation : atom -> relation
 
-val functions : atom -> Metis.NameAritySet.set
-
-val functionNames : atom -> Metis.NameSet.set
+val functions : atom -> Metis_NameAritySet.set
+
+val functionNames : atom -> Metis_NameSet.set
 
 (* Binary relations *)
 
-val mkBinop : relationName -> Metis.Term.term * Metis.Term.term -> atom
-
-val destBinop : relationName -> atom -> Metis.Term.term * Metis.Term.term
+val mkBinop : relationName -> Metis_Term.term * Metis_Term.term -> atom
+
+val destBinop : relationName -> atom -> Metis_Term.term * Metis_Term.term
 
 val isBinop : relationName -> atom -> bool
 
@@ -10340,39 +9851,39 @@
 (* Subterms.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-val subterm : atom -> Metis.Term.path -> Metis.Term.term
-
-val subterms : atom -> (Metis.Term.path * Metis.Term.term) list
-
-val replace : atom -> Metis.Term.path * Metis.Term.term -> atom
-
-val find : (Metis.Term.term -> bool) -> atom -> Metis.Term.path option
+val subterm : atom -> Metis_Term.path -> Metis_Term.term
+
+val subterms : atom -> (Metis_Term.path * Metis_Term.term) list
+
+val replace : atom -> Metis_Term.path * Metis_Term.term -> atom
+
+val find : (Metis_Term.term -> bool) -> atom -> Metis_Term.path option
 
 (* ------------------------------------------------------------------------- *)
 (* Free variables.                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-val freeIn : Metis.Term.var -> atom -> bool
-
-val freeVars : atom -> Metis.NameSet.set
+val freeIn : Metis_Term.var -> atom -> bool
+
+val freeVars : atom -> Metis_NameSet.set
 
 (* ------------------------------------------------------------------------- *)
 (* Substitutions.                                                            *)
 (* ------------------------------------------------------------------------- *)
 
-val subst : Metis.Subst.subst -> atom -> atom
+val subst : Metis_Subst.subst -> atom -> atom
 
 (* ------------------------------------------------------------------------- *)
 (* Matching.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-val match : Metis.Subst.subst -> atom -> atom -> Metis.Subst.subst  (* raises Error *)
+val match : Metis_Subst.subst -> atom -> atom -> Metis_Subst.subst  (* raises Error *)
 
 (* ------------------------------------------------------------------------- *)
 (* Unification.                                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val unify : Metis.Subst.subst -> atom -> atom -> Metis.Subst.subst  (* raises Error *)
+val unify : Metis_Subst.subst -> atom -> atom -> Metis_Subst.subst  (* raises Error *)
 
 (* ------------------------------------------------------------------------- *)
 (* The equality relation.                                                    *)
@@ -10382,23 +9893,23 @@
 
 val eqRelation : relation
 
-val mkEq : Metis.Term.term * Metis.Term.term -> atom
-
-val destEq : atom -> Metis.Term.term * Metis.Term.term
+val mkEq : Metis_Term.term * Metis_Term.term -> atom
+
+val destEq : atom -> Metis_Term.term * Metis_Term.term
 
 val isEq : atom -> bool
 
-val mkRefl : Metis.Term.term -> atom
-
-val destRefl : atom -> Metis.Term.term
+val mkRefl : Metis_Term.term -> atom
+
+val destRefl : atom -> Metis_Term.term
 
 val isRefl : atom -> bool
 
 val sym : atom -> atom  (* raises Error if given a refl *)
 
-val lhs : atom -> Metis.Term.term
-
-val rhs : atom -> Metis.Term.term
+val lhs : atom -> Metis_Term.term
+
+val rhs : atom -> Metis_Term.term
 
 (* ------------------------------------------------------------------------- *)
 (* Special support for terms with type annotations.                          *)
@@ -10406,52 +9917,43 @@
 
 val typedSymbols : atom -> int
 
-val nonVarTypedSubterms : atom -> (Metis.Term.path * Metis.Term.term) list
+val nonVarTypedSubterms : atom -> (Metis_Term.path * Metis_Term.term) list
 
 (* ------------------------------------------------------------------------- *)
 (* Parsing and pretty printing.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val pp : atom Metis.Print.pp
+val pp : atom Metis_Print.pp
 
 val toString : atom -> string
 
 val fromString : string -> atom
 
-val parse : Metis.Term.term Metis.Parse.quotation -> atom
+val parse : Metis_Term.term Metis_Parse.quotation -> atom
 
 end
 
 (**** Original file: Atom.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC ATOMS                                                   *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Atom :> Atom =
+structure Metis_Atom :> Metis_Atom =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type for storing first order logic atoms.                               *)
 (* ------------------------------------------------------------------------- *)
 
-type relationName = Name.name;
+type relationName = Metis_Name.name;
 
 type relation = relationName * int;
 
-type atom = relationName * Term.term list;
+type atom = relationName * Metis_Term.term list;
 
 (* ------------------------------------------------------------------------- *)
 (* Constructors and destructors.                                             *)
@@ -10467,16 +9969,16 @@
 
 val functions =
     let
-      fun f (tm,acc) = NameAritySet.union (Term.functions tm) acc
-    in
-      fn atm => foldl f NameAritySet.empty (arguments atm)
+      fun f (tm,acc) = Metis_NameAritySet.union (Metis_Term.functions tm) acc
+    in
+      fn atm => foldl f Metis_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)
+      fun f (tm,acc) = Metis_NameSet.union (Metis_Term.functionNames tm) acc
+    in
+      fn atm => foldl f Metis_NameSet.empty (arguments atm)
     end;
 
 (* Binary relations *)
@@ -10484,8 +9986,8 @@
 fun mkBinop p (a,b) : atom = (p,[a,b]);
 
 fun destBinop p (x,[a,b]) =
-    if Name.equal x p then (a,b) else raise Error "Atom.destBinop: wrong binop"
-  | destBinop _ _ = raise Error "Atom.destBinop: not a binop";
+    if Metis_Name.equal x p then (a,b) else raise Error "Metis_Atom.destBinop: wrong binop"
+  | destBinop _ _ = raise Error "Metis_Atom.destBinop: not a binop";
 
 fun isBinop p = can (destBinop p);
 
@@ -10493,16 +9995,16 @@
 (* The size of an atom in symbols.                                           *)
 (* ------------------------------------------------------------------------- *)
 
-fun symbols atm = foldl (fn (tm,z) => Term.symbols tm + z) 1 (arguments atm);
+fun symbols atm = foldl (fn (tm,z) => Metis_Term.symbols tm + z) 1 (arguments atm);
 
 (* ------------------------------------------------------------------------- *)
 (* A total comparison function for atoms.                                    *)
 (* ------------------------------------------------------------------------- *)
 
 fun compare ((p1,tms1),(p2,tms2)) =
-    case Name.compare (p1,p2) of
+    case Metis_Name.compare (p1,p2) of
       LESS => LESS
-    | EQUAL => lexCompare Term.compare (tms1,tms2)
+    | EQUAL => lexCompare Metis_Term.compare (tms1,tms2)
     | GREATER => GREATER;
 
 fun equal atm1 atm2 = compare (atm1,atm2) = EQUAL;
@@ -10511,34 +10013,34 @@
 (* Subterms.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun subterm _ [] = raise Bug "Atom.subterm: empty path"
+fun subterm _ [] = raise Bug "Metis_Atom.subterm: empty path"
   | subterm ((_,tms) : atom) (h :: t) =
-    if h >= length tms then raise Error "Atom.subterm: bad path"
-    else Term.subterm (List.nth (tms,h)) t;
+    if h >= length tms then raise Error "Metis_Atom.subterm: bad path"
+    else Metis_Term.subterm (List.nth (tms,h)) t;
 
 fun subterms ((_,tms) : atom) =
     let
-      fun f ((n,tm),l) = map (fn (p,s) => (n :: p, s)) (Term.subterms tm) @ l
+      fun f ((n,tm),l) = map (fn (p,s) => (n :: p, s)) (Metis_Term.subterms tm) @ l
     in
       foldl f [] (enumerate tms)
     end;
 
-fun replace _ ([],_) = raise Bug "Atom.replace: empty path"
+fun replace _ ([],_) = raise Bug "Metis_Atom.replace: empty path"
   | replace (atm as (rel,tms)) (h :: t, res) : atom =
-    if h >= length tms then raise Error "Atom.replace: bad path"
+    if h >= length tms then raise Error "Metis_Atom.replace: bad path"
     else
       let
         val tm = List.nth (tms,h)
-        val tm' = Term.replace tm (t,res)
-      in
-        if Portable.pointerEqual (tm,tm') then atm
+        val tm' = Metis_Term.replace tm (t,res)
+      in
+        if Metis_Portable.pointerEqual (tm,tm') then atm
         else (rel, updateNth (h,tm') tms)
       end;
 
 fun find pred =
     let
       fun f (i,tm) =
-          case Term.find pred tm of
+          case Metis_Term.find pred tm of
             SOME path => SOME (i :: path)
           | NONE => NONE
     in
@@ -10549,13 +10051,13 @@
 (* Free variables.                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-fun freeIn v atm = List.exists (Term.freeIn v) (arguments atm);
+fun freeIn v atm = List.exists (Metis_Term.freeIn v) (arguments atm);
 
 val freeVars =
     let
-      fun f (tm,acc) = NameSet.union (Term.freeVars tm) acc
-    in
-      fn atm => foldl f NameSet.empty (arguments atm)
+      fun f (tm,acc) = Metis_NameSet.union (Metis_Term.freeVars tm) acc
+    in
+      fn atm => foldl f Metis_NameSet.empty (arguments atm)
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -10564,9 +10066,9 @@
 
 fun subst sub (atm as (p,tms)) : atom =
     let
-      val tms' = Sharing.map (Subst.subst sub) tms
-    in
-      if Portable.pointerEqual (tms',tms) then atm else (p,tms')
+      val tms' = Metis_Sharing.map (Metis_Subst.subst sub) tms
+    in
+      if Metis_Portable.pointerEqual (tms',tms) then atm else (p,tms')
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -10574,12 +10076,12 @@
 (* ------------------------------------------------------------------------- *)
 
 local
-  fun matchArg ((tm1,tm2),sub) = Subst.match sub tm1 tm2;
+  fun matchArg ((tm1,tm2),sub) = Metis_Subst.match sub tm1 tm2;
 in
   fun match sub (p1,tms1) (p2,tms2) =
       let
-        val _ = (Name.equal p1 p2 andalso length tms1 = length tms2) orelse
-                raise Error "Atom.match"
+        val _ = (Metis_Name.equal p1 p2 andalso length tms1 = length tms2) orelse
+                raise Error "Metis_Atom.match"
       in
         foldl matchArg sub (zip tms1 tms2)
       end;
@@ -10590,12 +10092,12 @@
 (* ------------------------------------------------------------------------- *)
 
 local
-  fun unifyArg ((tm1,tm2),sub) = Subst.unify sub tm1 tm2;
+  fun unifyArg ((tm1,tm2),sub) = Metis_Subst.unify sub tm1 tm2;
 in
   fun unify sub (p1,tms1) (p2,tms2) =
       let
-        val _ = (Name.equal p1 p2 andalso length tms1 = length tms2) orelse
-                raise Error "Atom.unify"
+        val _ = (Metis_Name.equal p1 p2 andalso length tms1 = length tms2) orelse
+                raise Error "Metis_Atom.unify"
       in
         foldl unifyArg sub (zip tms1 tms2)
       end;
@@ -10605,7 +10107,7 @@
 (* The equality relation.                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-val eqRelationName = Name.fromString "=";
+val eqRelationName = Metis_Name.fromString "=";
 
 val eqRelationArity = 2;
 
@@ -10622,7 +10124,7 @@
 fun destRefl atm =
     let
       val (l,r) = destEq atm
-      val _ = Term.equal l r orelse raise Error "Atom.destRefl"
+      val _ = Metis_Term.equal l r orelse raise Error "Metis_Atom.destRefl"
     in
       l
     end;
@@ -10632,7 +10134,7 @@
 fun sym atm =
     let
       val (l,r) = destEq atm
-      val _ = not (Term.equal l r) orelse raise Error "Atom.sym: refl"
+      val _ = not (Metis_Term.equal l r) orelse raise Error "Metis_Atom.sym: refl"
     in
       mkEq (r,l)
     end;
@@ -10646,7 +10148,7 @@
 (* ------------------------------------------------------------------------- *)
 
 fun typedSymbols ((_,tms) : atom) =
-    foldl (fn (tm,z) => Term.typedSymbols tm + z) 1 tms;
+    foldl (fn (tm,z) => Metis_Term.typedSymbols tm + z) 1 tms;
 
 fun nonVarTypedSubterms (_,tms) =
     let
@@ -10654,7 +10156,7 @@
           let
             fun addTm ((path,tm),acc) = (n :: path, tm) :: acc
           in
-            foldl addTm acc (Term.nonVarTypedSubterms arg)
+            foldl addTm acc (Metis_Term.nonVarTypedSubterms arg)
           end
     in
       foldl addArg [] (enumerate tms)
@@ -10664,23 +10166,22 @@
 (* Parsing and pretty printing.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val pp = Print.ppMap Term.Fn Term.pp;
-
-val toString = Print.toString pp;
-
-fun fromString s = Term.destFn (Term.fromString s);
-
-val parse = Parse.parseQuotation Term.toString fromString;
-
-end
-
-structure AtomOrdered =
-struct type t = Atom.atom val compare = Atom.compare end
-
-structure AtomMap = KeyMap (AtomOrdered);
-
-structure AtomSet = ElementSet (AtomMap);
-end;
+val pp = Metis_Print.ppMap Metis_Term.Fn Metis_Term.pp;
+
+val toString = Metis_Print.toString pp;
+
+fun fromString s = Metis_Term.destFn (Metis_Term.fromString s);
+
+val parse = Metis_Parse.parseQuotation Metis_Term.toString fromString;
+
+end
+
+structure Metis_AtomOrdered =
+struct type t = Metis_Atom.atom val compare = Metis_Atom.compare end
+
+structure Metis_AtomMap = Metis_KeyMap (Metis_AtomOrdered);
+
+structure Metis_AtomSet = Metis_ElementSet (Metis_AtomMap);
 
 (**** Original file: Formula.sig ****)
 
@@ -10689,7 +10190,7 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Formula =
+signature Metis_Formula =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -10699,14 +10200,14 @@
 datatype formula =
     True
   | False
-  | Atom of Metis.Atom.atom
+  | Metis_Atom of Metis_Atom.atom
   | Not of formula
   | And of formula * formula
   | Or of formula * formula
   | Imp of formula * formula
   | Iff of formula * formula
-  | Forall of Metis.Term.var * formula
-  | Exists of Metis.Term.var * formula
+  | Forall of Metis_Term.var * formula
+  | Exists of Metis_Term.var * formula
 
 (* ------------------------------------------------------------------------- *)
 (* Constructors and destructors.                                             *)
@@ -10726,19 +10227,19 @@
 
 (* Functions *)
 
-val functions : formula -> Metis.NameAritySet.set
-
-val functionNames : formula -> Metis.NameSet.set
+val functions : formula -> Metis_NameAritySet.set
+
+val functionNames : formula -> Metis_NameSet.set
 
 (* Relations *)
 
-val relations : formula -> Metis.NameAritySet.set
-
-val relationNames : formula -> Metis.NameSet.set
+val relations : formula -> Metis_NameAritySet.set
+
+val relationNames : formula -> Metis_NameSet.set
 
 (* Atoms *)
 
-val destAtom : formula -> Metis.Atom.atom
+val destAtom : formula -> Metis_Atom.atom
 
 val isAtom : formula -> bool
 
@@ -10776,27 +10277,27 @@
 
 (* Universal quantification *)
 
-val destForall : formula -> Metis.Term.var * formula
+val destForall : formula -> Metis_Term.var * formula
 
 val isForall : formula -> bool
 
-val listMkForall : Metis.Term.var list * formula -> formula
-
-val setMkForall : Metis.NameSet.set * formula -> formula
-
-val stripForall : formula -> Metis.Term.var list * formula
+val listMkForall : Metis_Term.var list * formula -> formula
+
+val setMkForall : Metis_NameSet.set * formula -> formula
+
+val stripForall : formula -> Metis_Term.var list * formula
 
 (* Existential quantification *)
 
-val destExists : formula -> Metis.Term.var * formula
+val destExists : formula -> Metis_Term.var * formula
 
 val isExists : formula -> bool
 
-val listMkExists : Metis.Term.var list * formula -> formula
-
-val setMkExists : Metis.NameSet.set * formula -> formula
-
-val stripExists : formula -> Metis.Term.var list * formula
+val listMkExists : Metis_Term.var list * formula -> formula
+
+val setMkExists : Metis_NameSet.set * formula -> formula
+
+val stripExists : formula -> Metis_Term.var list * formula
 
 (* ------------------------------------------------------------------------- *)
 (* The size of a formula in symbols.                                         *)
@@ -10816,11 +10317,11 @@
 (* Free variables.                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-val freeIn : Metis.Term.var -> formula -> bool
-
-val freeVars : formula -> Metis.NameSet.set
-
-val freeVarsList : formula list -> Metis.NameSet.set
+val freeIn : Metis_Term.var -> formula -> bool
+
+val freeVars : formula -> Metis_NameSet.set
+
+val freeVarsList : formula list -> Metis_NameSet.set
 
 val specialize : formula -> formula
 
@@ -10830,35 +10331,35 @@
 (* Substitutions.                                                            *)
 (* ------------------------------------------------------------------------- *)
 
-val subst : Metis.Subst.subst -> formula -> formula
+val subst : Metis_Subst.subst -> formula -> formula
 
 (* ------------------------------------------------------------------------- *)
 (* The equality relation.                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-val mkEq : Metis.Term.term * Metis.Term.term -> formula
-
-val destEq : formula -> Metis.Term.term * Metis.Term.term
+val mkEq : Metis_Term.term * Metis_Term.term -> formula
+
+val destEq : formula -> Metis_Term.term * Metis_Term.term
 
 val isEq : formula -> bool
 
-val mkNeq : Metis.Term.term * Metis.Term.term -> formula
-
-val destNeq : formula -> Metis.Term.term * Metis.Term.term
+val mkNeq : Metis_Term.term * Metis_Term.term -> formula
+
+val destNeq : formula -> Metis_Term.term * Metis_Term.term
 
 val isNeq : formula -> bool
 
-val mkRefl : Metis.Term.term -> formula
-
-val destRefl : formula -> Metis.Term.term
+val mkRefl : Metis_Term.term -> formula
+
+val destRefl : formula -> Metis_Term.term
 
 val isRefl : formula -> bool
 
 val sym : formula -> formula  (* raises Error if given a refl *)
 
-val lhs : formula -> Metis.Term.term
-
-val rhs : formula -> Metis.Term.term
+val lhs : formula -> Metis_Term.term
+
+val rhs : formula -> Metis_Term.term
 
 (* ------------------------------------------------------------------------- *)
 (* Splitting goals.                                                          *)
@@ -10870,9 +10371,9 @@
 (* Parsing and pretty-printing.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-type quotation = formula Metis.Parse.quotation
-
-val pp : formula Metis.Print.pp
+type quotation = formula Metis_Parse.quotation
+
+val pp : formula Metis_Print.pp
 
 val toString : formula -> string
 
@@ -10884,24 +10385,15 @@
 
 (**** Original file: Formula.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC FORMULAS                                                *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Formula :> Formula =
+structure Metis_Formula :> Metis_Formula =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of first order logic formulas.                                     *)
@@ -10910,14 +10402,14 @@
 datatype formula =
     True
   | False
-  | Atom of Atom.atom
+  | Metis_Atom of Metis_Atom.atom
   | Not of formula
   | And of formula * formula
   | Or of formula * formula
   | Imp of formula * formula
   | Iff of formula * formula
-  | Forall of Term.var * formula
-  | Exists of Term.var * formula;
+  | Forall of Metis_Term.var * formula
+  | Exists of Metis_Term.var * formula;
 
 (* ------------------------------------------------------------------------- *)
 (* Constructors and destructors.                                             *)
@@ -10950,8 +10442,8 @@
   fun funcs fs [] = fs
     | funcs fs (True :: fms) = funcs fs fms
     | funcs fs (False :: fms) = funcs fs fms
-    | funcs fs (Atom atm :: fms) =
-      funcs (NameAritySet.union (Atom.functions atm) fs) fms
+    | funcs fs (Metis_Atom atm :: fms) =
+      funcs (Metis_NameAritySet.union (Metis_Atom.functions atm) fs) fms
     | funcs fs (Not p :: fms) = funcs fs (p :: fms)
     | funcs fs (And (p,q) :: fms) = funcs fs (p :: q :: fms)
     | funcs fs (Or (p,q) :: fms) = funcs fs (p :: q :: fms)
@@ -10960,15 +10452,15 @@
     | funcs fs (Forall (_,p) :: fms) = funcs fs (p :: fms)
     | funcs fs (Exists (_,p) :: fms) = funcs fs (p :: fms);
 in
-  fun functions fm = funcs NameAritySet.empty [fm];
+  fun functions fm = funcs Metis_NameAritySet.empty [fm];
 end;
 
 local
   fun funcs fs [] = fs
     | funcs fs (True :: fms) = funcs fs fms
     | funcs fs (False :: fms) = funcs fs fms
-    | funcs fs (Atom atm :: fms) =
-      funcs (NameSet.union (Atom.functionNames atm) fs) fms
+    | funcs fs (Metis_Atom atm :: fms) =
+      funcs (Metis_NameSet.union (Metis_Atom.functionNames atm) fs) fms
     | funcs fs (Not p :: fms) = funcs fs (p :: fms)
     | funcs fs (And (p,q) :: fms) = funcs fs (p :: q :: fms)
     | funcs fs (Or (p,q) :: fms) = funcs fs (p :: q :: fms)
@@ -10977,7 +10469,7 @@
     | funcs fs (Forall (_,p) :: fms) = funcs fs (p :: fms)
     | funcs fs (Exists (_,p) :: fms) = funcs fs (p :: fms);
 in
-  fun functionNames fm = funcs NameSet.empty [fm];
+  fun functionNames fm = funcs Metis_NameSet.empty [fm];
 end;
 
 (* Relations *)
@@ -10986,8 +10478,8 @@
   fun rels fs [] = fs
     | rels fs (True :: fms) = rels fs fms
     | rels fs (False :: fms) = rels fs fms
-    | rels fs (Atom atm :: fms) =
-      rels (NameAritySet.add fs (Atom.relation atm)) fms
+    | rels fs (Metis_Atom atm :: fms) =
+      rels (Metis_NameAritySet.add fs (Metis_Atom.relation atm)) fms
     | rels fs (Not p :: fms) = rels fs (p :: fms)
     | rels fs (And (p,q) :: fms) = rels fs (p :: q :: fms)
     | rels fs (Or (p,q) :: fms) = rels fs (p :: q :: fms)
@@ -10996,14 +10488,14 @@
     | rels fs (Forall (_,p) :: fms) = rels fs (p :: fms)
     | rels fs (Exists (_,p) :: fms) = rels fs (p :: fms);
 in
-  fun relations fm = rels NameAritySet.empty [fm];
+  fun relations fm = rels Metis_NameAritySet.empty [fm];
 end;
 
 local
   fun rels fs [] = fs
     | rels fs (True :: fms) = rels fs fms
     | rels fs (False :: fms) = rels fs fms
-    | rels fs (Atom atm :: fms) = rels (NameSet.add fs (Atom.name atm)) fms
+    | rels fs (Metis_Atom atm :: fms) = rels (Metis_NameSet.add fs (Metis_Atom.name atm)) fms
     | rels fs (Not p :: fms) = rels fs (p :: fms)
     | rels fs (And (p,q) :: fms) = rels fs (p :: q :: fms)
     | rels fs (Or (p,q) :: fms) = rels fs (p :: q :: fms)
@@ -11012,20 +10504,20 @@
     | rels fs (Forall (_,p) :: fms) = rels fs (p :: fms)
     | rels fs (Exists (_,p) :: fms) = rels fs (p :: fms);
 in
-  fun relationNames fm = rels NameSet.empty [fm];
+  fun relationNames fm = rels Metis_NameSet.empty [fm];
 end;
 
 (* Atoms *)
 
-fun destAtom (Atom atm) = atm
-  | destAtom _ = raise Error "Formula.destAtom";
+fun destAtom (Metis_Atom atm) = atm
+  | destAtom _ = raise Error "Metis_Formula.destAtom";
 
 val isAtom = can destAtom;
 
 (* Negations *)
 
 fun destNeg (Not p) = p
-  | destNeg _ = raise Error "Formula.destNeg";
+  | destNeg _ = raise Error "Metis_Formula.destNeg";
 
 val isNeg = can destNeg;
 
@@ -11116,7 +10608,7 @@
 fun listMkForall ([],body) = body
   | listMkForall (v :: vs, body) = Forall (v, listMkForall (vs,body));
 
-fun setMkForall (vs,body) = NameSet.foldr Forall body vs;
+fun setMkForall (vs,body) = Metis_NameSet.foldr Forall body vs;
 
 local
   fun strip vs (Forall (v,b)) = strip (v :: vs) b
@@ -11135,7 +10627,7 @@
 fun listMkExists ([],body) = body
   | listMkExists (v :: vs, body) = Exists (v, listMkExists (vs,body));
 
-fun setMkExists (vs,body) = NameSet.foldr Exists body vs;
+fun setMkExists (vs,body) = Metis_NameSet.foldr Exists body vs;
 
 local
   fun strip vs (Exists (v,b)) = strip (v :: vs) b
@@ -11152,7 +10644,7 @@
   fun sz n [] = n
     | sz n (True :: fms) = sz (n + 1) fms
     | sz n (False :: fms) = sz (n + 1) fms
-    | sz n (Atom atm :: fms) = sz (n + Atom.symbols atm) fms
+    | sz n (Metis_Atom atm :: fms) = sz (n + Metis_Atom.symbols atm) fms
     | sz n (Not p :: fms) = sz (n + 1) (p :: fms)
     | sz n (And (p,q) :: fms) = sz (n + 1) (p :: q :: fms)
     | sz n (Or (p,q) :: fms) = sz (n + 1) (p :: q :: fms)
@@ -11171,7 +10663,7 @@
 local
   fun cmp [] = EQUAL
     | cmp (f1_f2 :: fs) =
-      if Portable.pointerEqual f1_f2 then cmp fs
+      if Metis_Portable.pointerEqual f1_f2 then cmp fs
       else
         case f1_f2 of
           (True,True) => cmp fs
@@ -11180,13 +10672,13 @@
         | (False,False) => cmp fs
         | (False,_) => LESS
         | (_,False) => GREATER
-        | (Atom atm1, Atom atm2) =>
-          (case Atom.compare (atm1,atm2) of
+        | (Metis_Atom atm1, Metis_Atom atm2) =>
+          (case Metis_Atom.compare (atm1,atm2) of
              LESS => LESS
            | EQUAL => cmp fs
            | GREATER => GREATER)
-        | (Atom _, _) => LESS
-        | (_, Atom _) => GREATER
+        | (Metis_Atom _, _) => LESS
+        | (_, Metis_Atom _) => GREATER
         | (Not p1, Not p2) => cmp ((p1,p2) :: fs)
         | (Not _, _) => LESS
         | (_, Not _) => GREATER
@@ -11203,14 +10695,14 @@
         | (Iff _, _) => LESS
         | (_, Iff _) => GREATER
         | (Forall (v1,p1), Forall (v2,p2)) =>
-          (case Name.compare (v1,v2) of
+          (case Metis_Name.compare (v1,v2) of
              LESS => LESS
            | EQUAL => cmp ((p1,p2) :: fs)
            | GREATER => GREATER)
         | (Forall _, Exists _) => LESS
         | (Exists _, Forall _) => GREATER
         | (Exists (v1,p1), Exists (v2,p2)) =>
-          (case Name.compare (v1,v2) of
+          (case Metis_Name.compare (v1,v2) of
              LESS => LESS
            | EQUAL => cmp ((p1,p2) :: fs)
            | GREATER => GREATER);
@@ -11229,16 +10721,16 @@
       fun f [] = false
         | f (True :: fms) = f fms
         | f (False :: fms) = f fms
-        | f (Atom atm :: fms) = Atom.freeIn v atm orelse f fms
+        | f (Metis_Atom atm :: fms) = Metis_Atom.freeIn v atm orelse f fms
         | f (Not p :: fms) = f (p :: fms)
         | f (And (p,q) :: fms) = f (p :: q :: fms)
         | f (Or (p,q) :: fms) = f (p :: q :: fms)
         | f (Imp (p,q) :: fms) = f (p :: q :: fms)
         | f (Iff (p,q) :: fms) = f (p :: q :: fms)
         | f (Forall (w,p) :: fms) =
-          if Name.equal v w then f fms else f (p :: fms)
+          if Metis_Name.equal v w then f fms else f (p :: fms)
         | f (Exists (w,p) :: fms) =
-          if Name.equal v w then f fms else f (p :: fms)
+          if Metis_Name.equal v w then f fms else f (p :: fms)
     in
       fn fm => f [fm]
     end;
@@ -11247,49 +10739,49 @@
   fun fv vs [] = vs
     | fv vs ((_,True) :: fms) = fv vs fms
     | fv vs ((_,False) :: fms) = fv vs fms
-    | fv vs ((bv, Atom atm) :: fms) =
-      fv (NameSet.union vs (NameSet.difference (Atom.freeVars atm) bv)) fms
+    | fv vs ((bv, Metis_Atom atm) :: fms) =
+      fv (Metis_NameSet.union vs (Metis_NameSet.difference (Metis_Atom.freeVars atm) bv)) fms
     | fv vs ((bv, Not p) :: fms) = fv vs ((bv,p) :: fms)
     | fv vs ((bv, And (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms)
     | fv vs ((bv, Or (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms)
     | fv vs ((bv, Imp (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms)
     | fv vs ((bv, Iff (p,q)) :: fms) = fv vs ((bv,p) :: (bv,q) :: fms)
-    | fv vs ((bv, Forall (v,p)) :: fms) = fv vs ((NameSet.add bv v, p) :: fms)
-    | fv vs ((bv, Exists (v,p)) :: fms) = fv vs ((NameSet.add bv v, p) :: fms);
-
-  fun add (fm,vs) = fv vs [(NameSet.empty,fm)];
-in
-  fun freeVars fm = add (fm,NameSet.empty);
-
-  fun freeVarsList fms = List.foldl add NameSet.empty fms;
+    | fv vs ((bv, Forall (v,p)) :: fms) = fv vs ((Metis_NameSet.add bv v, p) :: fms)
+    | fv vs ((bv, Exists (v,p)) :: fms) = fv vs ((Metis_NameSet.add bv v, p) :: fms);
+
+  fun add (fm,vs) = fv vs [(Metis_NameSet.empty,fm)];
+in
+  fun freeVars fm = add (fm,Metis_NameSet.empty);
+
+  fun freeVarsList fms = List.foldl add Metis_NameSet.empty fms;
 end;
 
 fun specialize fm = snd (stripForall fm);
 
-fun generalize fm = listMkForall (NameSet.toList (freeVars fm), fm);
+fun generalize fm = listMkForall (Metis_NameSet.toList (freeVars fm), fm);
 
 (* ------------------------------------------------------------------------- *)
 (* Substitutions.                                                            *)
 (* ------------------------------------------------------------------------- *)
 
 local
-  fun substCheck sub fm = if Subst.null sub then fm else substFm sub fm
+  fun substCheck sub fm = if Metis_Subst.null sub then fm else substFm sub fm
 
   and substFm sub fm =
       case fm of
         True => fm
       | False => fm
-      | Atom (p,tms) =>
-        let
-          val tms' = Sharing.map (Subst.subst sub) tms
-        in
-          if Portable.pointerEqual (tms,tms') then fm else Atom (p,tms')
+      | Metis_Atom (p,tms) =>
+        let
+          val tms' = Metis_Sharing.map (Metis_Subst.subst sub) tms
+        in
+          if Metis_Portable.pointerEqual (tms,tms') then fm else Metis_Atom (p,tms')
         end
       | Not p =>
         let
           val p' = substFm sub p
         in
-          if Portable.pointerEqual (p,p') then fm else Not p'
+          if Metis_Portable.pointerEqual (p,p') then fm else Not p'
         end
       | And (p,q) => substConn sub fm And p q
       | Or (p,q) => substConn sub fm Or p q
@@ -11303,8 +10795,8 @@
         val p' = substFm sub p
         and q' = substFm sub q
       in
-        if Portable.pointerEqual (p,p') andalso
-           Portable.pointerEqual (q,q')
+        if Metis_Portable.pointerEqual (p,p') andalso
+           Metis_Portable.pointerEqual (q,q')
         then fm
         else conn (p',q')
       end
@@ -11314,25 +10806,25 @@
         val v' =
             let
               fun f (w,s) =
-                  if Name.equal w v then s
+                  if Metis_Name.equal w v then s
                   else
-                    case Subst.peek sub w of
-                      NONE => NameSet.add s w
-                    | SOME tm => NameSet.union s (Term.freeVars tm)
+                    case Metis_Subst.peek sub w of
+                      NONE => Metis_NameSet.add s w
+                    | SOME tm => Metis_NameSet.union s (Metis_Term.freeVars tm)
 
               val vars = freeVars p
-              val vars = NameSet.foldl f NameSet.empty vars
-            in
-              Term.variantPrime vars v
+              val vars = Metis_NameSet.foldl f Metis_NameSet.empty vars
+            in
+              Metis_Term.variantPrime vars v
             end
 
         val sub =
-            if Name.equal v v' then Subst.remove sub (NameSet.singleton v)
-            else Subst.insert sub (v, Term.Var v')
+            if Metis_Name.equal v v' then Metis_Subst.remove sub (Metis_NameSet.singleton v)
+            else Metis_Subst.insert sub (v, Metis_Term.Var v')
 
         val p' = substCheck sub p
       in
-        if Name.equal v v' andalso Portable.pointerEqual (p,p') then fm
+        if Metis_Name.equal v v' andalso Metis_Portable.pointerEqual (p,p') then fm
         else quant (v',p')
       end;
 in
@@ -11343,26 +10835,26 @@
 (* The equality relation.                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-fun mkEq a_b = Atom (Atom.mkEq a_b);
-
-fun destEq fm = Atom.destEq (destAtom fm);
+fun mkEq a_b = Metis_Atom (Metis_Atom.mkEq a_b);
+
+fun destEq fm = Metis_Atom.destEq (destAtom fm);
 
 val isEq = can destEq;
 
 fun mkNeq a_b = Not (mkEq a_b);
 
 fun destNeq (Not fm) = destEq fm
-  | destNeq _ = raise Error "Formula.destNeq";
+  | destNeq _ = raise Error "Metis_Formula.destNeq";
 
 val isNeq = can destNeq;
 
-fun mkRefl tm = Atom (Atom.mkRefl tm);
-
-fun destRefl fm = Atom.destRefl (destAtom fm);
+fun mkRefl tm = Metis_Atom (Metis_Atom.mkRefl tm);
+
+fun destRefl fm = Metis_Atom.destRefl (destAtom fm);
 
 val isRefl = can destRefl;
 
-fun sym fm = Atom (Atom.sym (destAtom fm));
+fun sym fm = Metis_Atom (Metis_Atom.sym (destAtom fm));
 
 fun lhs fm = fst (destEq fm);
 
@@ -11372,71 +10864,71 @@
 (* Parsing and pretty-printing.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-type quotation = formula Parse.quotation;
-
-val truthName = Name.fromString "T"
-and falsityName = Name.fromString "F"
-and conjunctionName = Name.fromString "/\\"
-and disjunctionName = Name.fromString "\\/"
-and implicationName = Name.fromString "==>"
-and equivalenceName = Name.fromString "<=>"
-and universalName = Name.fromString "!"
-and existentialName = Name.fromString "?";
-
-local
-  fun demote True = Term.Fn (truthName,[])
-    | demote False = Term.Fn (falsityName,[])
-    | demote (Atom (p,tms)) = Term.Fn (p,tms)
+type quotation = formula Metis_Parse.quotation;
+
+val truthName = Metis_Name.fromString "T"
+and falsityName = Metis_Name.fromString "F"
+and conjunctionName = Metis_Name.fromString "/\\"
+and disjunctionName = Metis_Name.fromString "\\/"
+and implicationName = Metis_Name.fromString "==>"
+and equivalenceName = Metis_Name.fromString "<=>"
+and universalName = Metis_Name.fromString "!"
+and existentialName = Metis_Name.fromString "?";
+
+local
+  fun demote True = Metis_Term.Fn (truthName,[])
+    | demote False = Metis_Term.Fn (falsityName,[])
+    | demote (Metis_Atom (p,tms)) = Metis_Term.Fn (p,tms)
     | demote (Not p) =
       let
-        val Unsynchronized.ref s = Term.negation
-      in
-        Term.Fn (Name.fromString s, [demote p])
+        val Unsynchronized.ref s = Metis_Term.negation
+      in
+        Metis_Term.Fn (Metis_Name.fromString s, [demote p])
       end
-    | demote (And (p,q)) = Term.Fn (conjunctionName, [demote p, demote q])
-    | demote (Or (p,q)) = Term.Fn (disjunctionName, [demote p, demote q])
-    | demote (Imp (p,q)) = Term.Fn (implicationName, [demote p, demote q])
-    | demote (Iff (p,q)) = Term.Fn (equivalenceName, [demote p, demote q])
-    | demote (Forall (v,b)) = Term.Fn (universalName, [Term.Var v, demote b])
+    | demote (And (p,q)) = Metis_Term.Fn (conjunctionName, [demote p, demote q])
+    | demote (Or (p,q)) = Metis_Term.Fn (disjunctionName, [demote p, demote q])
+    | demote (Imp (p,q)) = Metis_Term.Fn (implicationName, [demote p, demote q])
+    | demote (Iff (p,q)) = Metis_Term.Fn (equivalenceName, [demote p, demote q])
+    | demote (Forall (v,b)) = Metis_Term.Fn (universalName, [Metis_Term.Var v, demote b])
     | demote (Exists (v,b)) =
-      Term.Fn (existentialName, [Term.Var v, demote b]);
-in
-  fun pp fm = Term.pp (demote fm);
-end;
-
-val toString = Print.toString pp;
-
-local
-  fun isQuant [Term.Var _, _] = true
+      Metis_Term.Fn (existentialName, [Metis_Term.Var v, demote b]);
+in
+  fun pp fm = Metis_Term.pp (demote fm);
+end;
+
+val toString = Metis_Print.toString pp;
+
+local
+  fun isQuant [Metis_Term.Var _, _] = true
     | isQuant _ = false;
 
-  fun promote (Term.Var v) = Atom (v,[])
-    | promote (Term.Fn (f,tms)) =
-      if Name.equal f truthName andalso null tms then
+  fun promote (Metis_Term.Var v) = Metis_Atom (v,[])
+    | promote (Metis_Term.Fn (f,tms)) =
+      if Metis_Name.equal f truthName andalso null tms then
         True
-      else if Name.equal f falsityName andalso null tms then
+      else if Metis_Name.equal f falsityName andalso null tms then
         False
-      else if Name.toString f = !Term.negation andalso length tms = 1 then
+      else if Metis_Name.toString f = !Metis_Term.negation andalso length tms = 1 then
         Not (promote (hd tms))
-      else if Name.equal f conjunctionName andalso length tms = 2 then
+      else if Metis_Name.equal f conjunctionName andalso length tms = 2 then
         And (promote (hd tms), promote (List.nth (tms,1)))
-      else if Name.equal f disjunctionName andalso length tms = 2 then
+      else if Metis_Name.equal f disjunctionName andalso length tms = 2 then
         Or (promote (hd tms), promote (List.nth (tms,1)))
-      else if Name.equal f implicationName andalso length tms = 2 then
+      else if Metis_Name.equal f implicationName andalso length tms = 2 then
         Imp (promote (hd tms), promote (List.nth (tms,1)))
-      else if Name.equal f equivalenceName andalso length tms = 2 then
+      else if Metis_Name.equal f equivalenceName andalso length tms = 2 then
         Iff (promote (hd tms), promote (List.nth (tms,1)))
-      else if Name.equal f universalName andalso isQuant tms then
-        Forall (Term.destVar (hd tms), promote (List.nth (tms,1)))
-      else if Name.equal f existentialName andalso isQuant tms then
-        Exists (Term.destVar (hd tms), promote (List.nth (tms,1)))
+      else if Metis_Name.equal f universalName andalso isQuant tms then
+        Forall (Metis_Term.destVar (hd tms), promote (List.nth (tms,1)))
+      else if Metis_Name.equal f existentialName andalso isQuant tms then
+        Exists (Metis_Term.destVar (hd tms), promote (List.nth (tms,1)))
       else
-        Atom (f,tms);
-in
-  fun fromString s = promote (Term.fromString s);
-end;
-
-val parse = Parse.parseQuotation toString fromString;
+        Metis_Atom (f,tms);
+in
+  fun fromString s = promote (Metis_Term.fromString s);
+end;
+
+val parse = Metis_Parse.parseQuotation toString fromString;
 
 (* ------------------------------------------------------------------------- *)
 (* Splitting goals.                                                          *)
@@ -11479,8 +10971,8 @@
 val splitGoal = fn fm =>
     let
       val result = splitGoal fm
-      val () = Print.trace pp "Formula.splitGoal: fm" fm
-      val () = Print.trace (Print.ppList pp) "Formula.splitGoal: result" result
+      val () = Metis_Print.trace pp "Metis_Formula.splitGoal: fm" fm
+      val () = Metis_Print.trace (Metis_Print.ppList pp) "Metis_Formula.splitGoal: result" result
     in
       result
     end;
@@ -11488,13 +10980,12 @@
 
 end
 
-structure FormulaOrdered =
-struct type t = Formula.formula val compare = Formula.compare end
-
-structure FormulaMap = KeyMap (FormulaOrdered);
-
-structure FormulaSet = ElementSet (FormulaMap);
-end;
+structure Metis_FormulaOrdered =
+struct type t = Metis_Formula.formula val compare = Metis_Formula.compare end
+
+structure Metis_FormulaMap = Metis_KeyMap (Metis_FormulaOrdered);
+
+structure Metis_FormulaSet = Metis_ElementSet (Metis_FormulaMap);
 
 (**** Original file: Literal.sig ****)
 
@@ -11503,7 +10994,7 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Literal =
+signature Metis_Literal =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -11512,7 +11003,7 @@
 
 type polarity = bool
 
-type literal = polarity * Metis.Atom.atom
+type literal = polarity * Metis_Atom.atom
 
 (* ------------------------------------------------------------------------- *)
 (* Constructors and destructors.                                             *)
@@ -11520,11 +11011,11 @@
 
 val polarity : literal -> polarity
 
-val atom : literal -> Metis.Atom.atom
-
-val name : literal -> Metis.Atom.relationName
-
-val arguments : literal -> Metis.Term.term list
+val atom : literal -> Metis_Atom.atom
+
+val name : literal -> Metis_Atom.relationName
+
+val arguments : literal -> Metis_Term.term list
 
 val arity : literal -> int
 
@@ -11534,25 +11025,25 @@
 
 val negate : literal -> literal
 
-val relation : literal -> Metis.Atom.relation
-
-val functions : literal -> Metis.NameAritySet.set
-
-val functionNames : literal -> Metis.NameSet.set
+val relation : literal -> Metis_Atom.relation
+
+val functions : literal -> Metis_NameAritySet.set
+
+val functionNames : literal -> Metis_NameSet.set
 
 (* Binary relations *)
 
-val mkBinop : Metis.Atom.relationName -> polarity * Metis.Term.term * Metis.Term.term -> literal
-
-val destBinop : Metis.Atom.relationName -> literal -> polarity * Metis.Term.term * Metis.Term.term
-
-val isBinop : Metis.Atom.relationName -> literal -> bool
+val mkBinop : Metis_Atom.relationName -> polarity * Metis_Term.term * Metis_Term.term -> literal
+
+val destBinop : Metis_Atom.relationName -> literal -> polarity * Metis_Term.term * Metis_Term.term
+
+val isBinop : Metis_Atom.relationName -> literal -> bool
 
 (* Formulas *)
 
-val toFormula : literal -> Metis.Formula.formula
-
-val fromFormula : Metis.Formula.formula -> literal
+val toFormula : literal -> Metis_Formula.formula
+
+val fromFormula : Metis_Formula.formula -> literal
 
 (* ------------------------------------------------------------------------- *)
 (* The size of a literal in symbols.                                         *)
@@ -11572,65 +11063,65 @@
 (* Subterms.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-val subterm : literal -> Metis.Term.path -> Metis.Term.term
-
-val subterms : literal -> (Metis.Term.path * Metis.Term.term) list
-
-val replace : literal -> Metis.Term.path * Metis.Term.term -> literal
+val subterm : literal -> Metis_Term.path -> Metis_Term.term
+
+val subterms : literal -> (Metis_Term.path * Metis_Term.term) list
+
+val replace : literal -> Metis_Term.path * Metis_Term.term -> literal
 
 (* ------------------------------------------------------------------------- *)
 (* Free variables.                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-val freeIn : Metis.Term.var -> literal -> bool
-
-val freeVars : literal -> Metis.NameSet.set
+val freeIn : Metis_Term.var -> literal -> bool
+
+val freeVars : literal -> Metis_NameSet.set
 
 (* ------------------------------------------------------------------------- *)
 (* Substitutions.                                                            *)
 (* ------------------------------------------------------------------------- *)
 
-val subst : Metis.Subst.subst -> literal -> literal
+val subst : Metis_Subst.subst -> literal -> literal
 
 (* ------------------------------------------------------------------------- *)
 (* Matching.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
 val match :  (* raises Error *)
-    Metis.Subst.subst -> literal -> literal -> Metis.Subst.subst
+    Metis_Subst.subst -> literal -> literal -> Metis_Subst.subst
 
 (* ------------------------------------------------------------------------- *)
 (* Unification.                                                              *)
 (* ------------------------------------------------------------------------- *)
 
 val unify :  (* raises Error *)
-    Metis.Subst.subst -> literal -> literal -> Metis.Subst.subst
+    Metis_Subst.subst -> literal -> literal -> Metis_Subst.subst
 
 (* ------------------------------------------------------------------------- *)
 (* The equality relation.                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-val mkEq : Metis.Term.term * Metis.Term.term -> literal
-
-val destEq : literal -> Metis.Term.term * Metis.Term.term
+val mkEq : Metis_Term.term * Metis_Term.term -> literal
+
+val destEq : literal -> Metis_Term.term * Metis_Term.term
 
 val isEq : literal -> bool
 
-val mkNeq : Metis.Term.term * Metis.Term.term -> literal
-
-val destNeq : literal -> Metis.Term.term * Metis.Term.term
+val mkNeq : Metis_Term.term * Metis_Term.term -> literal
+
+val destNeq : literal -> Metis_Term.term * Metis_Term.term
 
 val isNeq : literal -> bool
 
-val mkRefl : Metis.Term.term -> literal
-
-val destRefl : literal -> Metis.Term.term
+val mkRefl : Metis_Term.term -> literal
+
+val destRefl : literal -> Metis_Term.term
 
 val isRefl : literal -> bool
 
-val mkIrrefl : Metis.Term.term -> literal
-
-val destIrrefl : literal -> Metis.Term.term
+val mkIrrefl : Metis_Term.term -> literal
+
+val destIrrefl : literal -> Metis_Term.term
 
 val isIrrefl : literal -> bool
 
@@ -11638,9 +11129,9 @@
 
 val sym : literal -> literal  (* raises Error if given a refl or irrefl *)
 
-val lhs : literal -> Metis.Term.term
-
-val rhs : literal -> Metis.Term.term
+val lhs : literal -> Metis_Term.term
+
+val rhs : literal -> Metis_Term.term
 
 (* ------------------------------------------------------------------------- *)
 (* Special support for terms with type annotations.                          *)
@@ -11648,42 +11139,33 @@
 
 val typedSymbols : literal -> int
 
-val nonVarTypedSubterms : literal -> (Metis.Term.path * Metis.Term.term) list
+val nonVarTypedSubterms : literal -> (Metis_Term.path * Metis_Term.term) list
 
 (* ------------------------------------------------------------------------- *)
 (* Parsing and pretty-printing.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val pp : literal Metis.Print.pp
+val pp : literal Metis_Print.pp
 
 val toString : literal -> string
 
 val fromString : string -> literal
 
-val parse : Metis.Term.term Metis.Parse.quotation -> literal
+val parse : Metis_Term.term Metis_Parse.quotation -> literal
 
 end
 
 (**** Original file: Literal.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* FIRST ORDER LOGIC LITERALS                                                *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Literal :> Literal =
+structure Metis_Literal :> Metis_Literal =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type for storing first order logic literals.                            *)
@@ -11691,7 +11173,7 @@
 
 type polarity = bool;
 
-type literal = polarity * Atom.atom;
+type literal = polarity * Metis_Atom.atom;
 
 (* ------------------------------------------------------------------------- *)
 (* Constructors and destructors.                                             *)
@@ -11701,11 +11183,11 @@
 
 fun atom ((_,atm) : literal) = atm;
 
-fun name lit = Atom.name (atom lit);
-
-fun arguments lit = Atom.arguments (atom lit);
-
-fun arity lit = Atom.arity (atom lit);
+fun name lit = Metis_Atom.name (atom lit);
+
+fun arguments lit = Metis_Atom.arguments (atom lit);
+
+fun arity lit = Metis_Atom.arity (atom lit);
 
 fun positive lit = polarity lit;
 
@@ -11713,66 +11195,66 @@
 
 fun negate (pol,atm) : literal = (not pol, atm)
 
-fun relation lit = Atom.relation (atom lit);
-
-fun functions lit = Atom.functions (atom lit);
-
-fun functionNames lit = Atom.functionNames (atom lit);
+fun relation lit = Metis_Atom.relation (atom lit);
+
+fun functions lit = Metis_Atom.functions (atom lit);
+
+fun functionNames lit = Metis_Atom.functionNames (atom lit);
 
 (* Binary relations *)
 
-fun mkBinop rel (pol,a,b) : literal = (pol, Atom.mkBinop rel (a,b));
+fun mkBinop rel (pol,a,b) : literal = (pol, Metis_Atom.mkBinop rel (a,b));
 
 fun destBinop rel ((pol,atm) : literal) =
-    case Atom.destBinop rel atm of (a,b) => (pol,a,b);
+    case Metis_Atom.destBinop rel atm of (a,b) => (pol,a,b);
 
 fun isBinop rel = can (destBinop rel);
 
 (* Formulas *)
 
-fun toFormula (true,atm) = Formula.Atom atm
-  | toFormula (false,atm) = Formula.Not (Formula.Atom atm);
-
-fun fromFormula (Formula.Atom atm) = (true,atm)
-  | fromFormula (Formula.Not (Formula.Atom atm)) = (false,atm)
-  | fromFormula _ = raise Error "Literal.fromFormula";
+fun toFormula (true,atm) = Metis_Formula.Metis_Atom atm
+  | toFormula (false,atm) = Metis_Formula.Not (Metis_Formula.Metis_Atom atm);
+
+fun fromFormula (Metis_Formula.Metis_Atom atm) = (true,atm)
+  | fromFormula (Metis_Formula.Not (Metis_Formula.Metis_Atom atm)) = (false,atm)
+  | fromFormula _ = raise Error "Metis_Literal.fromFormula";
 
 (* ------------------------------------------------------------------------- *)
 (* The size of a literal in symbols.                                         *)
 (* ------------------------------------------------------------------------- *)
 
-fun symbols ((_,atm) : literal) = Atom.symbols atm;
+fun symbols ((_,atm) : literal) = Metis_Atom.symbols atm;
 
 (* ------------------------------------------------------------------------- *)
 (* A total comparison function for literals.                                 *)
 (* ------------------------------------------------------------------------- *)
 
-val compare = prodCompare boolCompare Atom.compare;
-
-fun equal (p1,atm1) (p2,atm2) = p1 = p2 andalso Atom.equal atm1 atm2;
+val compare = prodCompare boolCompare Metis_Atom.compare;
+
+fun equal (p1,atm1) (p2,atm2) = p1 = p2 andalso Metis_Atom.equal atm1 atm2;
 
 (* ------------------------------------------------------------------------- *)
 (* Subterms.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun subterm lit path = Atom.subterm (atom lit) path;
-
-fun subterms lit = Atom.subterms (atom lit);
+fun subterm lit path = Metis_Atom.subterm (atom lit) path;
+
+fun subterms lit = Metis_Atom.subterms (atom lit);
 
 fun replace (lit as (pol,atm)) path_tm =
     let
-      val atm' = Atom.replace atm path_tm
-    in
-      if Portable.pointerEqual (atm,atm') then lit else (pol,atm')
+      val atm' = Metis_Atom.replace atm path_tm
+    in
+      if Metis_Portable.pointerEqual (atm,atm') then lit else (pol,atm')
     end;
 
 (* ------------------------------------------------------------------------- *)
 (* Free variables.                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-fun freeIn v lit = Atom.freeIn v (atom lit);
-
-fun freeVars lit = Atom.freeVars (atom lit);
+fun freeIn v lit = Metis_Atom.freeIn v (atom lit);
+
+fun freeVars lit = Metis_Atom.freeVars (atom lit);
 
 (* ------------------------------------------------------------------------- *)
 (* Substitutions.                                                            *)
@@ -11780,9 +11262,9 @@
 
 fun subst sub (lit as (pol,atm)) : literal =
     let
-      val atm' = Atom.subst sub atm
-    in
-      if Portable.pointerEqual (atm',atm) then lit else (pol,atm')
+      val atm' = Metis_Atom.subst sub atm
+    in
+      if Metis_Portable.pointerEqual (atm',atm) then lit else (pol,atm')
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -11791,9 +11273,9 @@
 
 fun match sub ((pol1,atm1) : literal) (pol2,atm2) =
     let
-      val _ = pol1 = pol2 orelse raise Error "Literal.match"
-    in
-      Atom.match sub atm1 atm2
+      val _ = pol1 = pol2 orelse raise Error "Metis_Literal.match"
+    in
+      Metis_Atom.match sub atm1 atm2
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -11802,134 +11284,134 @@
 
 fun unify sub ((pol1,atm1) : literal) (pol2,atm2) =
     let
-      val _ = pol1 = pol2 orelse raise Error "Literal.unify"
-    in
-      Atom.unify sub atm1 atm2
+      val _ = pol1 = pol2 orelse raise Error "Metis_Literal.unify"
+    in
+      Metis_Atom.unify sub atm1 atm2
     end;
 
 (* ------------------------------------------------------------------------- *)
 (* The equality relation.                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-fun mkEq l_r : literal = (true, Atom.mkEq l_r);
-
-fun destEq ((true,atm) : literal) = Atom.destEq atm
-  | destEq (false,_) = raise Error "Literal.destEq";
+fun mkEq l_r : literal = (true, Metis_Atom.mkEq l_r);
+
+fun destEq ((true,atm) : literal) = Metis_Atom.destEq atm
+  | destEq (false,_) = raise Error "Metis_Literal.destEq";
 
 val isEq = can destEq;
 
-fun mkNeq l_r : literal = (false, Atom.mkEq l_r);
-
-fun destNeq ((false,atm) : literal) = Atom.destEq atm
-  | destNeq (true,_) = raise Error "Literal.destNeq";
+fun mkNeq l_r : literal = (false, Metis_Atom.mkEq l_r);
+
+fun destNeq ((false,atm) : literal) = Metis_Atom.destEq atm
+  | destNeq (true,_) = raise Error "Metis_Literal.destNeq";
 
 val isNeq = can destNeq;
 
-fun mkRefl tm = (true, Atom.mkRefl tm);
-
-fun destRefl (true,atm) = Atom.destRefl atm
-  | destRefl (false,_) = raise Error "Literal.destRefl";
+fun mkRefl tm = (true, Metis_Atom.mkRefl tm);
+
+fun destRefl (true,atm) = Metis_Atom.destRefl atm
+  | destRefl (false,_) = raise Error "Metis_Literal.destRefl";
 
 val isRefl = can destRefl;
 
-fun mkIrrefl tm = (false, Atom.mkRefl tm);
-
-fun destIrrefl (true,_) = raise Error "Literal.destIrrefl"
-  | destIrrefl (false,atm) = Atom.destRefl atm;
+fun mkIrrefl tm = (false, Metis_Atom.mkRefl tm);
+
+fun destIrrefl (true,_) = raise Error "Metis_Literal.destIrrefl"
+  | destIrrefl (false,atm) = Metis_Atom.destRefl atm;
 
 val isIrrefl = can destIrrefl;
 
-fun sym (pol,atm) : literal = (pol, Atom.sym atm);
-
-fun lhs ((_,atm) : literal) = Atom.lhs atm;
-
-fun rhs ((_,atm) : literal) = Atom.rhs atm;
+fun sym (pol,atm) : literal = (pol, Metis_Atom.sym atm);
+
+fun lhs ((_,atm) : literal) = Metis_Atom.lhs atm;
+
+fun rhs ((_,atm) : literal) = Metis_Atom.rhs atm;
 
 (* ------------------------------------------------------------------------- *)
 (* Special support for terms with type annotations.                          *)
 (* ------------------------------------------------------------------------- *)
 
-fun typedSymbols ((_,atm) : literal) = Atom.typedSymbols atm;
-
-fun nonVarTypedSubterms ((_,atm) : literal) = Atom.nonVarTypedSubterms atm;
+fun typedSymbols ((_,atm) : literal) = Metis_Atom.typedSymbols atm;
+
+fun nonVarTypedSubterms ((_,atm) : literal) = Metis_Atom.nonVarTypedSubterms atm;
 
 (* ------------------------------------------------------------------------- *)
 (* Parsing and pretty-printing.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val pp = Print.ppMap toFormula Formula.pp;
-
-val toString = Print.toString pp;
-
-fun fromString s = fromFormula (Formula.fromString s);
-
-val parse = Parse.parseQuotation Term.toString fromString;
-
-end
-
-structure LiteralOrdered =
-struct type t = Literal.literal val compare = Literal.compare end
-
-structure LiteralMap = KeyMap (LiteralOrdered);
-
-structure LiteralSet =
+val pp = Metis_Print.ppMap toFormula Metis_Formula.pp;
+
+val toString = Metis_Print.toString pp;
+
+fun fromString s = fromFormula (Metis_Formula.fromString s);
+
+val parse = Metis_Parse.parseQuotation Metis_Term.toString fromString;
+
+end
+
+structure Metis_LiteralOrdered =
+struct type t = Metis_Literal.literal val compare = Metis_Literal.compare end
+
+structure Metis_LiteralMap = Metis_KeyMap (Metis_LiteralOrdered);
+
+structure Metis_LiteralSet =
 struct
 
   local
-    structure S = ElementSet (LiteralMap);
+    structure S = Metis_ElementSet (Metis_LiteralMap);
   in
     open S;
   end;
 
-  fun negateMember lit set = member (Literal.negate lit) set;
+  fun negateMember lit set = member (Metis_Literal.negate lit) set;
 
   val negate =
       let
-        fun f (lit,set) = add set (Literal.negate lit)
+        fun f (lit,set) = add set (Metis_Literal.negate lit)
       in
         foldl f empty
       end;
 
   val relations =
       let
-        fun f (lit,set) = NameAritySet.add set (Literal.relation lit)
-      in
-        foldl f NameAritySet.empty
+        fun f (lit,set) = Metis_NameAritySet.add set (Metis_Literal.relation lit)
+      in
+        foldl f Metis_NameAritySet.empty
       end;
 
   val functions =
       let
-        fun f (lit,set) = NameAritySet.union set (Literal.functions lit)
-      in
-        foldl f NameAritySet.empty
-      end;
-
-  fun freeIn v = exists (Literal.freeIn v);
+        fun f (lit,set) = Metis_NameAritySet.union set (Metis_Literal.functions lit)
+      in
+        foldl f Metis_NameAritySet.empty
+      end;
+
+  fun freeIn v = exists (Metis_Literal.freeIn v);
 
   val freeVars =
       let
-        fun f (lit,set) = NameSet.union set (Literal.freeVars lit)
-      in
-        foldl f NameSet.empty
+        fun f (lit,set) = Metis_NameSet.union set (Metis_Literal.freeVars lit)
+      in
+        foldl f Metis_NameSet.empty
       end;
 
   val freeVarsList =
       let
-        fun f (lits,set) = NameSet.union set (freeVars lits)
-      in
-        List.foldl f NameSet.empty
+        fun f (lits,set) = Metis_NameSet.union set (freeVars lits)
+      in
+        List.foldl f Metis_NameSet.empty
       end;
 
   val symbols =
       let
-        fun f (lit,z) = Literal.symbols lit + z
+        fun f (lit,z) = Metis_Literal.symbols lit + z
       in
         foldl f 0
       end;
 
   val typedSymbols =
       let
-        fun f (lit,z) = Literal.typedSymbols lit + z
+        fun f (lit,z) = Metis_Literal.typedSymbols lit + z
       in
         foldl f 0
       end;
@@ -11938,8 +11420,8 @@
       let
         fun substLit (lit,(eq,lits')) =
             let
-              val lit' = Literal.subst sub lit
-              val eq = eq andalso Portable.pointerEqual (lit,lit')
+              val lit' = Metis_Literal.subst sub lit
+              val eq = eq andalso Metis_Portable.pointerEqual (lit,lit')
             in
               (eq, add lits' lit')
             end
@@ -11950,25 +11432,24 @@
       end;
 
   fun conjoin set =
-      Formula.listMkConj (List.map Literal.toFormula (toList set));
+      Metis_Formula.listMkConj (List.map Metis_Literal.toFormula (toList set));
 
   fun disjoin set =
-      Formula.listMkDisj (List.map Literal.toFormula (toList set));
+      Metis_Formula.listMkDisj (List.map Metis_Literal.toFormula (toList set));
 
   val pp =
-      Print.ppMap
+      Metis_Print.ppMap
         toList
-        (Print.ppBracket "{" "}" (Print.ppOpList "," Literal.pp));
-
-end
-
-structure LiteralSetOrdered =
-struct type t = LiteralSet.set val compare = LiteralSet.compare end
-
-structure LiteralSetMap = KeyMap (LiteralSetOrdered);
-
-structure LiteralSetSet = ElementSet (LiteralSetMap);
-end;
+        (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_Literal.pp));
+
+end
+
+structure Metis_LiteralSetOrdered =
+struct type t = Metis_LiteralSet.set val compare = Metis_LiteralSet.compare end
+
+structure Metis_LiteralSetMap = Metis_KeyMap (Metis_LiteralSetOrdered);
+
+structure Metis_LiteralSetSet = Metis_ElementSet (Metis_LiteralSetMap);
 
 (**** Original file: Thm.sig ****)
 
@@ -11977,7 +11458,7 @@
 (* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature Thm =
+signature Metis_Thm =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -11990,12 +11471,12 @@
 (* Theorem destructors.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
-type clause = Metis.LiteralSet.set
+type clause = Metis_LiteralSet.set
 
 datatype inferenceType =
     Axiom
   | Assume
-  | Subst
+  | Metis_Subst
   | Factor
   | Resolve
   | Refl
@@ -12017,21 +11498,21 @@
 
 (* Unit theorems *)
 
-val destUnit : thm -> Metis.Literal.literal
+val destUnit : thm -> Metis_Literal.literal
 
 val isUnit : thm -> bool
 
 (* Unit equality theorems *)
 
-val destUnitEq : thm -> Metis.Term.term * Metis.Term.term
+val destUnitEq : thm -> Metis_Term.term * Metis_Term.term
 
 val isUnitEq : thm -> bool
 
 (* Literals *)
 
-val member : Metis.Literal.literal -> thm -> bool
-
-val negateMember : Metis.Literal.literal -> thm -> bool
+val member : Metis_Literal.literal -> thm -> bool
+
+val negateMember : Metis_Literal.literal -> thm -> bool
 
 (* ------------------------------------------------------------------------- *)
 (* A total order.                                                            *)
@@ -12045,19 +11526,19 @@
 (* Free variables.                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-val freeIn : Metis.Term.var -> thm -> bool
-
-val freeVars : thm -> Metis.NameSet.set
+val freeIn : Metis_Term.var -> thm -> bool
+
+val freeVars : thm -> Metis_NameSet.set
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty-printing.                                                          *)
 (* ------------------------------------------------------------------------- *)
 
-val ppInferenceType : inferenceType Metis.Print.pp
+val ppInferenceType : inferenceType Metis_Print.pp
 
 val inferenceTypeToString : inferenceType -> string
 
-val pp : thm Metis.Print.pp
+val pp : thm Metis_Print.pp
 
 val toString : thm -> string
 
@@ -12079,7 +11560,7 @@
 (*   L \/ ~L                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-val assume : Metis.Literal.literal -> thm
+val assume : Metis_Literal.literal -> thm
 
 (* ------------------------------------------------------------------------- *)
 (*    C                                                                      *)
@@ -12087,7 +11568,7 @@
 (*   C[s]                                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-val subst : Metis.Subst.subst -> thm -> thm
+val subst : Metis_Subst.subst -> thm -> thm
 
 (* ------------------------------------------------------------------------- *)
 (*   L \/ C    ~L \/ D                                                       *)
@@ -12098,7 +11579,7 @@
 (* occur in the second theorem.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val resolve : Metis.Literal.literal -> thm -> thm -> thm
+val resolve : Metis_Literal.literal -> thm -> thm -> thm
 
 (* ------------------------------------------------------------------------- *)
 (*                                                                           *)
@@ -12106,7 +11587,7 @@
 (*   t = t                                                                   *)
 (* ------------------------------------------------------------------------- *)
 
-val refl : Metis.Term.term -> thm
+val refl : Metis_Term.term -> thm
 
 (* ------------------------------------------------------------------------- *)
 (*                                                                           *)
@@ -12117,47 +11598,38 @@
 (* path p being replaced by t.                                               *)
 (* ------------------------------------------------------------------------- *)
 
-val equality : Metis.Literal.literal -> Metis.Term.path -> Metis.Term.term -> thm
+val equality : Metis_Literal.literal -> Metis_Term.path -> Metis_Term.term -> thm
 
 end
 
 (**** Original file: Thm.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS                         *)
 (* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-structure Thm :> Thm =
+structure Metis_Thm :> Metis_Thm =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* An abstract type of first order logic theorems.                           *)
 (* ------------------------------------------------------------------------- *)
 
-type clause = LiteralSet.set;
+type clause = Metis_LiteralSet.set;
 
 datatype inferenceType =
     Axiom
   | Assume
-  | Subst
+  | Metis_Subst
   | Factor
   | Resolve
   | Refl
   | Equality;
 
-datatype thm = Thm of clause * (inferenceType * thm list);
+datatype thm = Metis_Thm of clause * (inferenceType * thm list);
 
 type inference = inferenceType * thm list;
 
@@ -12165,63 +11637,63 @@
 (* Theorem destructors.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
-fun clause (Thm (cl,_)) = cl;
-
-fun inference (Thm (_,inf)) = inf;
+fun clause (Metis_Thm (cl,_)) = cl;
+
+fun inference (Metis_Thm (_,inf)) = inf;
 
 (* Tautologies *)
 
 local
   fun chk (_,NONE) = NONE
     | chk ((pol,atm), SOME set) =
-      if (pol andalso Atom.isRefl atm) orelse AtomSet.member atm set then NONE
-      else SOME (AtomSet.add set atm);
+      if (pol andalso Metis_Atom.isRefl atm) orelse Metis_AtomSet.member atm set then NONE
+      else SOME (Metis_AtomSet.add set atm);
 in
   fun isTautology th =
-      case LiteralSet.foldl chk (SOME AtomSet.empty) (clause th) of
+      case Metis_LiteralSet.foldl chk (SOME Metis_AtomSet.empty) (clause th) of
         SOME _ => false
       | NONE => true;
 end;
 
 (* Contradictions *)
 
-fun isContradiction th = LiteralSet.null (clause th);
+fun isContradiction th = Metis_LiteralSet.null (clause th);
 
 (* Unit theorems *)
 
-fun destUnit (Thm (cl,_)) =
-    if LiteralSet.size cl = 1 then LiteralSet.pick cl
-    else raise Error "Thm.destUnit";
+fun destUnit (Metis_Thm (cl,_)) =
+    if Metis_LiteralSet.size cl = 1 then Metis_LiteralSet.pick cl
+    else raise Error "Metis_Thm.destUnit";
 
 val isUnit = can destUnit;
 
 (* Unit equality theorems *)
 
-fun destUnitEq th = Literal.destEq (destUnit th);
+fun destUnitEq th = Metis_Literal.destEq (destUnit th);
 
 val isUnitEq = can destUnitEq;
 
 (* Literals *)
 
-fun member lit (Thm (cl,_)) = LiteralSet.member lit cl;
-
-fun negateMember lit (Thm (cl,_)) = LiteralSet.negateMember lit cl;
+fun member lit (Metis_Thm (cl,_)) = Metis_LiteralSet.member lit cl;
+
+fun negateMember lit (Metis_Thm (cl,_)) = Metis_LiteralSet.negateMember lit cl;
 
 (* ------------------------------------------------------------------------- *)
 (* A total order.                                                            *)
 (* ------------------------------------------------------------------------- *)
 
-fun compare (th1,th2) = LiteralSet.compare (clause th1, clause th2);
-
-fun equal th1 th2 = LiteralSet.equal (clause th1) (clause th2);
+fun compare (th1,th2) = Metis_LiteralSet.compare (clause th1, clause th2);
+
+fun equal th1 th2 = Metis_LiteralSet.equal (clause th1) (clause th2);
 
 (* ------------------------------------------------------------------------- *)
 (* Free variables.                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-fun freeIn v (Thm (cl,_)) = LiteralSet.freeIn v cl;
-
-fun freeVars (Thm (cl,_)) = LiteralSet.freeVars cl;
+fun freeIn v (Metis_Thm (cl,_)) = Metis_LiteralSet.freeIn v cl;
+
+fun freeVars (Metis_Thm (cl,_)) = Metis_LiteralSet.freeVars cl;
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty-printing.                                                          *)
@@ -12229,27 +11701,27 @@
 
 fun inferenceTypeToString Axiom = "Axiom"
   | inferenceTypeToString Assume = "Assume"
-  | inferenceTypeToString Subst = "Subst"
+  | inferenceTypeToString Metis_Subst = "Metis_Subst"
   | inferenceTypeToString Factor = "Factor"
   | inferenceTypeToString Resolve = "Resolve"
   | inferenceTypeToString Refl = "Refl"
   | inferenceTypeToString Equality = "Equality";
 
 fun ppInferenceType inf =
-    Print.ppString (inferenceTypeToString inf);
+    Metis_Print.ppString (inferenceTypeToString inf);
 
 local
   fun toFormula th =
-      Formula.listMkDisj
-        (map Literal.toFormula (LiteralSet.toList (clause th)));
+      Metis_Formula.listMkDisj
+        (map Metis_Literal.toFormula (Metis_LiteralSet.toList (clause th)));
 in
   fun pp th =
-      Print.blockProgram Print.Inconsistent 3
-        [Print.addString "|- ",
-         Formula.pp (toFormula th)];
-end;
-
-val toString = Print.toString pp;
+      Metis_Print.blockProgram Metis_Print.Inconsistent 3
+        [Metis_Print.addString "|- ",
+         Metis_Formula.pp (toFormula th)];
+end;
+
+val toString = Metis_Print.toString pp;
 
 (* ------------------------------------------------------------------------- *)
 (* Primitive rules of inference.                                             *)
@@ -12261,7 +11733,7 @@
 (*   C                                                                       *)
 (* ------------------------------------------------------------------------- *)
 
-fun axiom cl = Thm (cl,(Axiom,[]));
+fun axiom cl = Metis_Thm (cl,(Axiom,[]));
 
 (* ------------------------------------------------------------------------- *)
 (*                                                                           *)
@@ -12270,7 +11742,7 @@
 (* ------------------------------------------------------------------------- *)
 
 fun assume lit =
-    Thm (LiteralSet.fromList [lit, Literal.negate lit], (Assume,[]));
+    Metis_Thm (Metis_LiteralSet.fromList [lit, Metis_Literal.negate lit], (Assume,[]));
 
 (* ------------------------------------------------------------------------- *)
 (*    C                                                                      *)
@@ -12278,15 +11750,15 @@
 (*   C[s]                                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-fun subst sub (th as Thm (cl,inf)) =
-    let
-      val cl' = LiteralSet.subst sub cl
-    in
-      if Portable.pointerEqual (cl,cl') then th
+fun subst sub (th as Metis_Thm (cl,inf)) =
+    let
+      val cl' = Metis_LiteralSet.subst sub cl
+    in
+      if Metis_Portable.pointerEqual (cl,cl') then th
       else
         case inf of
-          (Subst,_) => Thm (cl',inf)
-        | _ => Thm (cl',(Subst,[th]))
+          (Metis_Subst,_) => Metis_Thm (cl',inf)
+        | _ => Metis_Thm (cl',(Metis_Subst,[th]))
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -12298,19 +11770,19 @@
 (* occur in the second theorem.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-fun resolve lit (th1 as Thm (cl1,_)) (th2 as Thm (cl2,_)) =
-    let
-      val cl1' = LiteralSet.delete cl1 lit
-      and cl2' = LiteralSet.delete cl2 (Literal.negate lit)
-    in
-      Thm (LiteralSet.union cl1' cl2', (Resolve,[th1,th2]))
+fun resolve lit (th1 as Metis_Thm (cl1,_)) (th2 as Metis_Thm (cl2,_)) =
+    let
+      val cl1' = Metis_LiteralSet.delete cl1 lit
+      and cl2' = Metis_LiteralSet.delete cl2 (Metis_Literal.negate lit)
+    in
+      Metis_Thm (Metis_LiteralSet.union cl1' cl2', (Resolve,[th1,th2]))
     end;
 
 (*MetisDebug
 val resolve = fn lit => fn pos => fn neg =>
     resolve lit pos neg
     handle Error err =>
-      raise Error ("Thm.resolve:\nlit = " ^ Literal.toString lit ^
+      raise Error ("Metis_Thm.resolve:\nlit = " ^ Metis_Literal.toString lit ^
                    "\npos = " ^ toString pos ^
                    "\nneg = " ^ toString neg ^ "\n" ^ err);
 *)
@@ -12321,7 +11793,7 @@
 (*   t = t                                                                   *)
 (* ------------------------------------------------------------------------- *)
 
-fun refl tm = Thm (LiteralSet.singleton (true, Atom.mkRefl tm), (Refl,[]));
+fun refl tm = Metis_Thm (Metis_LiteralSet.singleton (true, Metis_Atom.mkRefl tm), (Refl,[]));
 
 (* ------------------------------------------------------------------------- *)
 (*                                                                           *)
@@ -12334,19 +11806,18 @@
 
 fun equality lit path t =
     let
-      val s = Literal.subterm lit path
-
-      val lit' = Literal.replace lit (path,t)
-
-      val eqLit = Literal.mkNeq (s,t)
-
-      val cl = LiteralSet.fromList [eqLit, Literal.negate lit, lit']
-    in
-      Thm (cl,(Equality,[]))
-    end;
-
-end
-end;
+      val s = Metis_Literal.subterm lit path
+
+      val lit' = Metis_Literal.replace lit (path,t)
+
+      val eqLit = Metis_Literal.mkNeq (s,t)
+
+      val cl = Metis_LiteralSet.fromList [eqLit, Metis_Literal.negate lit, lit']
+    in
+      Metis_Thm (cl,(Equality,[]))
+    end;
+
+end
 
 (**** Original file: Proof.sig ****)
 
@@ -12355,7 +11826,7 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Proof =
+signature Metis_Proof =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -12363,50 +11834,50 @@
 (* ------------------------------------------------------------------------- *)
 
 datatype inference =
-    Axiom of Metis.LiteralSet.set
-  | Assume of Metis.Atom.atom
-  | Subst of Metis.Subst.subst * Metis.Thm.thm
-  | Resolve of Metis.Atom.atom * Metis.Thm.thm * Metis.Thm.thm
-  | Refl of Metis.Term.term
-  | Equality of Metis.Literal.literal * Metis.Term.path * Metis.Term.term
-
-type proof = (Metis.Thm.thm * inference) list
+    Axiom of Metis_LiteralSet.set
+  | Assume of Metis_Atom.atom
+  | Metis_Subst of Metis_Subst.subst * Metis_Thm.thm
+  | Resolve of Metis_Atom.atom * Metis_Thm.thm * Metis_Thm.thm
+  | Refl of Metis_Term.term
+  | Equality of Metis_Literal.literal * Metis_Term.path * Metis_Term.term
+
+type proof = (Metis_Thm.thm * inference) list
 
 (* ------------------------------------------------------------------------- *)
 (* Reconstructing single inferences.                                         *)
 (* ------------------------------------------------------------------------- *)
 
-val inferenceType : inference -> Metis.Thm.inferenceType
-
-val parents : inference -> Metis.Thm.thm list
-
-val inferenceToThm : inference -> Metis.Thm.thm
-
-val thmToInference : Metis.Thm.thm -> inference
+val inferenceType : inference -> Metis_Thm.inferenceType
+
+val parents : inference -> Metis_Thm.thm list
+
+val inferenceToThm : inference -> Metis_Thm.thm
+
+val thmToInference : Metis_Thm.thm -> inference
 
 (* ------------------------------------------------------------------------- *)
 (* Reconstructing whole proofs.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val proof : Metis.Thm.thm -> proof
+val proof : Metis_Thm.thm -> proof
 
 (* ------------------------------------------------------------------------- *)
 (* Free variables.                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-val freeIn : Metis.Term.var -> proof -> bool
-
-val freeVars : proof -> Metis.NameSet.set
+val freeIn : Metis_Term.var -> proof -> bool
+
+val freeVars : proof -> Metis_NameSet.set
 
 (* ------------------------------------------------------------------------- *)
 (* Printing.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-val ppInference : inference Metis.Print.pp
+val ppInference : inference Metis_Print.pp
 
 val inferenceToString : inference -> string
 
-val pp : proof Metis.Print.pp
+val pp : proof Metis_Print.pp
 
 val toString : proof -> string
 
@@ -12414,115 +11885,106 @@
 
 (**** Original file: Proof.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* PROOFS IN FIRST ORDER LOGIC                                               *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Proof :> Proof =
+structure Metis_Proof :> Metis_Proof =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of first order logic proofs.                                       *)
 (* ------------------------------------------------------------------------- *)
 
 datatype inference =
-    Axiom of LiteralSet.set
-  | Assume of Atom.atom
-  | Subst of Subst.subst * Thm.thm
-  | Resolve of Atom.atom * Thm.thm * Thm.thm
-  | Refl of Term.term
-  | Equality of Literal.literal * Term.path * Term.term;
-
-type proof = (Thm.thm * inference) list;
+    Axiom of Metis_LiteralSet.set
+  | Assume of Metis_Atom.atom
+  | Metis_Subst of Metis_Subst.subst * Metis_Thm.thm
+  | Resolve of Metis_Atom.atom * Metis_Thm.thm * Metis_Thm.thm
+  | Refl of Metis_Term.term
+  | Equality of Metis_Literal.literal * Metis_Term.path * Metis_Term.term;
+
+type proof = (Metis_Thm.thm * inference) list;
 
 (* ------------------------------------------------------------------------- *)
 (* Printing.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun inferenceType (Axiom _) = Thm.Axiom
-  | inferenceType (Assume _) = Thm.Assume
-  | inferenceType (Subst _) = Thm.Subst
-  | inferenceType (Resolve _) = Thm.Resolve
-  | inferenceType (Refl _) = Thm.Refl
-  | inferenceType (Equality _) = Thm.Equality;
-
-local
-  fun ppAssume atm = Print.sequence (Print.addBreak 1) (Atom.pp atm);
+fun inferenceType (Axiom _) = Metis_Thm.Axiom
+  | inferenceType (Assume _) = Metis_Thm.Assume
+  | inferenceType (Metis_Subst _) = Metis_Thm.Metis_Subst
+  | inferenceType (Resolve _) = Metis_Thm.Resolve
+  | inferenceType (Refl _) = Metis_Thm.Refl
+  | inferenceType (Equality _) = Metis_Thm.Equality;
+
+local
+  fun ppAssume atm = Metis_Print.sequence (Metis_Print.addBreak 1) (Metis_Atom.pp atm);
 
   fun ppSubst ppThm (sub,thm) =
-      Print.sequence (Print.addBreak 1)
-        (Print.blockProgram Print.Inconsistent 1
-           [Print.addString "{",
-            Print.ppOp2 " =" Print.ppString Subst.pp ("sub",sub),
-            Print.addString ",",
-            Print.addBreak 1,
-            Print.ppOp2 " =" Print.ppString ppThm ("thm",thm),
-            Print.addString "}"]);
+      Metis_Print.sequence (Metis_Print.addBreak 1)
+        (Metis_Print.blockProgram Metis_Print.Inconsistent 1
+           [Metis_Print.addString "{",
+            Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Subst.pp ("sub",sub),
+            Metis_Print.addString ",",
+            Metis_Print.addBreak 1,
+            Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("thm",thm),
+            Metis_Print.addString "}"]);
 
   fun ppResolve ppThm (res,pos,neg) =
-      Print.sequence (Print.addBreak 1)
-        (Print.blockProgram Print.Inconsistent 1
-           [Print.addString "{",
-            Print.ppOp2 " =" Print.ppString Atom.pp ("res",res),
-            Print.addString ",",
-            Print.addBreak 1,
-            Print.ppOp2 " =" Print.ppString ppThm ("pos",pos),
-            Print.addString ",",
-            Print.addBreak 1,
-            Print.ppOp2 " =" Print.ppString ppThm ("neg",neg),
-            Print.addString "}"]);
-
-  fun ppRefl tm = Print.sequence (Print.addBreak 1) (Term.pp tm);
+      Metis_Print.sequence (Metis_Print.addBreak 1)
+        (Metis_Print.blockProgram Metis_Print.Inconsistent 1
+           [Metis_Print.addString "{",
+            Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Atom.pp ("res",res),
+            Metis_Print.addString ",",
+            Metis_Print.addBreak 1,
+            Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("pos",pos),
+            Metis_Print.addString ",",
+            Metis_Print.addBreak 1,
+            Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("neg",neg),
+            Metis_Print.addString "}"]);
+
+  fun ppRefl tm = Metis_Print.sequence (Metis_Print.addBreak 1) (Metis_Term.pp tm);
 
   fun ppEquality (lit,path,res) =
-      Print.sequence (Print.addBreak 1)
-        (Print.blockProgram Print.Inconsistent 1
-           [Print.addString "{",
-            Print.ppOp2 " =" Print.ppString Literal.pp ("lit",lit),
-            Print.addString ",",
-            Print.addBreak 1,
-            Print.ppOp2 " =" Print.ppString Term.ppPath ("path",path),
-            Print.addString ",",
-            Print.addBreak 1,
-            Print.ppOp2 " =" Print.ppString Term.pp ("res",res),
-            Print.addString "}"]);
+      Metis_Print.sequence (Metis_Print.addBreak 1)
+        (Metis_Print.blockProgram Metis_Print.Inconsistent 1
+           [Metis_Print.addString "{",
+            Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Literal.pp ("lit",lit),
+            Metis_Print.addString ",",
+            Metis_Print.addBreak 1,
+            Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Term.ppPath ("path",path),
+            Metis_Print.addString ",",
+            Metis_Print.addBreak 1,
+            Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Term.pp ("res",res),
+            Metis_Print.addString "}"]);
 
   fun ppInf ppAxiom ppThm inf =
       let
-        val infString = Thm.inferenceTypeToString (inferenceType inf)
-      in
-        Print.block Print.Inconsistent 2
-          (Print.sequence
-             (Print.addString infString)
+        val infString = Metis_Thm.inferenceTypeToString (inferenceType inf)
+      in
+        Metis_Print.block Metis_Print.Inconsistent 2
+          (Metis_Print.sequence
+             (Metis_Print.addString infString)
              (case inf of
                 Axiom cl => ppAxiom cl
               | Assume x => ppAssume x
-              | Subst x => ppSubst ppThm x
+              | Metis_Subst x => ppSubst ppThm x
               | Resolve x => ppResolve ppThm x
               | Refl x => ppRefl x
               | Equality x => ppEquality x))
       end;
 
   fun ppAxiom cl =
-      Print.sequence
-        (Print.addBreak 1)
-        (Print.ppMap
-           LiteralSet.toList
-           (Print.ppBracket "{" "}" (Print.ppOpList "," Literal.pp)) cl);
-in
-  val ppInference = ppInf ppAxiom Thm.pp;
+      Metis_Print.sequence
+        (Metis_Print.addBreak 1)
+        (Metis_Print.ppMap
+           Metis_LiteralSet.toList
+           (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_Literal.pp)) cl);
+in
+  val ppInference = ppInf ppAxiom Metis_Thm.pp;
 
   fun pp prf =
       let
@@ -12531,11 +11993,11 @@
         val prf = enumerate prf
 
         fun ppThm th =
-            Print.addString
-            let
-              val cl = Thm.clause th
-
-              fun pred (_,(th',_)) = LiteralSet.equal (Thm.clause th') cl
+            Metis_Print.addString
+            let
+              val cl = Metis_Thm.clause th
+
+              fun pred (_,(th',_)) = Metis_LiteralSet.equal (Metis_Thm.clause th') cl
             in
               case List.find pred prf of
                 NONE => "(?)"
@@ -12546,30 +12008,30 @@
             let
               val s = thmString n
             in
-              Print.sequence
-                (Print.blockProgram Print.Consistent (1 + size s)
-                   [Print.addString (s ^ " "),
-                    Thm.pp th,
-                    Print.addBreak 2,
-                    Print.ppBracket "[" "]" (ppInf (K Print.skip) ppThm) inf])
-                Print.addNewline
-            end
-      in
-        Print.blockProgram Print.Consistent 0
-          [Print.addString "START OF PROOF",
-           Print.addNewline,
-           Print.program (map ppStep prf),
-           Print.addString "END OF PROOF"]
+              Metis_Print.sequence
+                (Metis_Print.blockProgram Metis_Print.Consistent (1 + size s)
+                   [Metis_Print.addString (s ^ " "),
+                    Metis_Thm.pp th,
+                    Metis_Print.addBreak 2,
+                    Metis_Print.ppBracket "[" "]" (ppInf (K Metis_Print.skip) ppThm) inf])
+                Metis_Print.addNewline
+            end
+      in
+        Metis_Print.blockProgram Metis_Print.Consistent 0
+          [Metis_Print.addString "START OF PROOF",
+           Metis_Print.addNewline,
+           Metis_Print.program (map ppStep prf),
+           Metis_Print.addString "END OF PROOF"]
       end
 (*MetisDebug
-      handle Error err => raise Bug ("Proof.pp: shouldn't fail:\n" ^ err);
-*)
-
-end;
-
-val inferenceToString = Print.toString ppInference;
-
-val toString = Print.toString pp;
+      handle Error err => raise Bug ("Metis_Proof.pp: shouldn't fail:\n" ^ err);
+*)
+
+end;
+
+val inferenceToString = Metis_Print.toString ppInference;
+
+val toString = Metis_Print.toString pp;
 
 (* ------------------------------------------------------------------------- *)
 (* Reconstructing single inferences.                                         *)
@@ -12577,17 +12039,17 @@
 
 fun parents (Axiom _) = []
   | parents (Assume _) = []
-  | parents (Subst (_,th)) = [th]
+  | parents (Metis_Subst (_,th)) = [th]
   | parents (Resolve (_,th,th')) = [th,th']
   | parents (Refl _) = []
   | parents (Equality _) = [];
 
-fun inferenceToThm (Axiom cl) = Thm.axiom cl
-  | inferenceToThm (Assume atm) = Thm.assume (true,atm)
-  | inferenceToThm (Subst (sub,th)) = Thm.subst sub th
-  | inferenceToThm (Resolve (atm,th,th')) = Thm.resolve (true,atm) th th'
-  | inferenceToThm (Refl tm) = Thm.refl tm
-  | inferenceToThm (Equality (lit,path,r)) = Thm.equality lit path r;
+fun inferenceToThm (Axiom cl) = Metis_Thm.axiom cl
+  | inferenceToThm (Assume atm) = Metis_Thm.assume (true,atm)
+  | inferenceToThm (Metis_Subst (sub,th)) = Metis_Thm.subst sub th
+  | inferenceToThm (Resolve (atm,th,th')) = Metis_Thm.resolve (true,atm) th th'
+  | inferenceToThm (Refl tm) = Metis_Thm.refl tm
+  | inferenceToThm (Equality (lit,path,r)) = Metis_Thm.equality lit path r;
 
 local
   fun reconstructSubst cl cl' =
@@ -12595,74 +12057,74 @@
         fun recon [] =
             let
 (*MetisTrace3
-              val () = Print.trace LiteralSet.pp "reconstructSubst: cl" cl
-              val () = Print.trace LiteralSet.pp "reconstructSubst: cl'" cl'
-*)
-            in
-              raise Bug "can't reconstruct Subst rule"
+              val () = Metis_Print.trace Metis_LiteralSet.pp "reconstructSubst: cl" cl
+              val () = Metis_Print.trace Metis_LiteralSet.pp "reconstructSubst: cl'" cl'
+*)
+            in
+              raise Bug "can't reconstruct Metis_Subst rule"
             end
           | recon (([],sub) :: others) =
-            if LiteralSet.equal (LiteralSet.subst sub cl) cl' then sub
+            if Metis_LiteralSet.equal (Metis_LiteralSet.subst sub cl) cl' then sub
             else recon others
           | recon ((lit :: lits, sub) :: others) =
             let
               fun checkLit (lit',acc) =
-                  case total (Literal.match sub lit) lit' of
+                  case total (Metis_Literal.match sub lit) lit' of
                     NONE => acc
                   | SOME sub => (lits,sub) :: acc
             in
-              recon (LiteralSet.foldl checkLit others cl')
-            end
-      in
-        Subst.normalize (recon [(LiteralSet.toList cl, Subst.empty)])
+              recon (Metis_LiteralSet.foldl checkLit others cl')
+            end
+      in
+        Metis_Subst.normalize (recon [(Metis_LiteralSet.toList cl, Metis_Subst.empty)])
       end
 (*MetisDebug
       handle Error err =>
-        raise Bug ("Proof.recontructSubst: shouldn't fail:\n" ^ err);
+        raise Bug ("Metis_Proof.recontructSubst: shouldn't fail:\n" ^ err);
 *)
 
   fun reconstructResolvant cl1 cl2 cl =
-      (if not (LiteralSet.subset cl1 cl) then
-         LiteralSet.pick (LiteralSet.difference cl1 cl)
-       else if not (LiteralSet.subset cl2 cl) then
-         Literal.negate (LiteralSet.pick (LiteralSet.difference cl2 cl))
+      (if not (Metis_LiteralSet.subset cl1 cl) then
+         Metis_LiteralSet.pick (Metis_LiteralSet.difference cl1 cl)
+       else if not (Metis_LiteralSet.subset cl2 cl) then
+         Metis_Literal.negate (Metis_LiteralSet.pick (Metis_LiteralSet.difference cl2 cl))
        else
          (* A useless resolution, but we must reconstruct it anyway *)
          let
-           val cl1' = LiteralSet.negate cl1
-           and cl2' = LiteralSet.negate cl2
-           val lits = LiteralSet.intersectList [cl1,cl1',cl2,cl2']
+           val cl1' = Metis_LiteralSet.negate cl1
+           and cl2' = Metis_LiteralSet.negate cl2
+           val lits = Metis_LiteralSet.intersectList [cl1,cl1',cl2,cl2']
          in
-           if not (LiteralSet.null lits) then LiteralSet.pick lits
+           if not (Metis_LiteralSet.null lits) then Metis_LiteralSet.pick lits
            else raise Bug "can't reconstruct Resolve rule"
          end)
 (*MetisDebug
       handle Error err =>
-        raise Bug ("Proof.recontructResolvant: shouldn't fail:\n" ^ err);
+        raise Bug ("Metis_Proof.recontructResolvant: shouldn't fail:\n" ^ err);
 *)
 
   fun reconstructEquality cl =
       let
 (*MetisTrace3
-        val () = Print.trace LiteralSet.pp "Proof.reconstructEquality: cl" cl
+        val () = Metis_Print.trace Metis_LiteralSet.pp "Metis_Proof.reconstructEquality: cl" cl
 *)
 
         fun sync s t path (f,a) (f',a') =
-            if not (Name.equal f f' andalso length a = length a') then NONE
+            if not (Metis_Name.equal f f' andalso length a = length a') then NONE
             else
               let
                 val itms = enumerate (zip a a')
               in
-                case List.filter (not o uncurry Term.equal o snd) itms of
+                case List.filter (not o uncurry Metis_Term.equal o snd) itms of
                   [(i,(tm,tm'))] =>
                   let
                     val path = i :: path
                   in
-                    if Term.equal tm s andalso Term.equal tm' t then
+                    if Metis_Term.equal tm s andalso Metis_Term.equal tm' t then
                       SOME (rev path)
                     else
                       case (tm,tm') of
-                        (Term.Fn f_a, Term.Fn f_a') => sync s t path f_a f_a'
+                        (Metis_Term.Fn f_a, Metis_Term.Fn f_a') => sync s t path f_a f_a'
                       | _ => NONE
                   end
                 | _ => NONE
@@ -12672,12 +12134,12 @@
             if pol = pol' then NONE
             else
               let
-                val (s,t) = Literal.destNeq neq
+                val (s,t) = Metis_Literal.destNeq neq
 
                 val path =
-                    if not (Term.equal s t) then sync s t [] atm atm'
-                    else if not (Atom.equal atm atm') then NONE
-                    else Atom.find (Term.equal s) atm
+                    if not (Metis_Term.equal s t) then sync s t [] atm atm'
+                    else if not (Metis_Atom.equal atm atm') then NONE
+                    else Metis_Atom.find (Metis_Term.equal s) atm
               in
                 case path of
                   SOME path => SOME ((pol',atm),path,t)
@@ -12685,7 +12147,7 @@
               end
 
         val candidates =
-            case List.partition Literal.isNeq (LiteralSet.toList cl) of
+            case List.partition Metis_Literal.isNeq (Metis_LiteralSet.toList cl) of
               ([l1],[l2,l3]) => [(l1,l2,l3),(l1,l3,l2)]
             | ([l1,l2],[l3]) => [(l1,l2,l3),(l1,l3,l2),(l2,l1,l3),(l2,l3,l1)]
             | ([l1],[l2]) => [(l1,l1,l2),(l1,l2,l1)]
@@ -12693,9 +12155,9 @@
 
 (*MetisTrace3
         val ppCands =
-            Print.ppList (Print.ppTriple Literal.pp Literal.pp Literal.pp)
-        val () = Print.trace ppCands
-                   "Proof.reconstructEquality: candidates" candidates
+            Metis_Print.ppList (Metis_Print.ppTriple Metis_Literal.pp Metis_Literal.pp Metis_Literal.pp)
+        val () = Metis_Print.trace ppCands
+                   "Metis_Proof.reconstructEquality: candidates" candidates
 *)
       in
         case first recon candidates of
@@ -12704,64 +12166,64 @@
       end
 (*MetisDebug
       handle Error err =>
-        raise Bug ("Proof.recontructEquality: shouldn't fail:\n" ^ err);
-*)
-
-  fun reconstruct cl (Thm.Axiom,[]) = Axiom cl
-    | reconstruct cl (Thm.Assume,[]) =
-      (case LiteralSet.findl Literal.positive cl of
+        raise Bug ("Metis_Proof.recontructEquality: shouldn't fail:\n" ^ err);
+*)
+
+  fun reconstruct cl (Metis_Thm.Axiom,[]) = Axiom cl
+    | reconstruct cl (Metis_Thm.Assume,[]) =
+      (case Metis_LiteralSet.findl Metis_Literal.positive cl of
          SOME (_,atm) => Assume atm
        | NONE => raise Bug "malformed Assume inference")
-    | reconstruct cl (Thm.Subst,[th]) =
-      Subst (reconstructSubst (Thm.clause th) cl, th)
-    | reconstruct cl (Thm.Resolve,[th1,th2]) =
-      let
-        val cl1 = Thm.clause th1
-        and cl2 = Thm.clause th2
+    | reconstruct cl (Metis_Thm.Metis_Subst,[th]) =
+      Metis_Subst (reconstructSubst (Metis_Thm.clause th) cl, th)
+    | reconstruct cl (Metis_Thm.Resolve,[th1,th2]) =
+      let
+        val cl1 = Metis_Thm.clause th1
+        and cl2 = Metis_Thm.clause th2
         val (pol,atm) = reconstructResolvant cl1 cl2 cl
       in
         if pol then Resolve (atm,th1,th2) else Resolve (atm,th2,th1)
       end
-    | reconstruct cl (Thm.Refl,[]) =
-      (case LiteralSet.findl (K true) cl of
-         SOME lit => Refl (Literal.destRefl lit)
+    | reconstruct cl (Metis_Thm.Refl,[]) =
+      (case Metis_LiteralSet.findl (K true) cl of
+         SOME lit => Refl (Metis_Literal.destRefl lit)
        | NONE => raise Bug "malformed Refl inference")
-    | reconstruct cl (Thm.Equality,[]) = Equality (reconstructEquality cl)
+    | reconstruct cl (Metis_Thm.Equality,[]) = Equality (reconstructEquality cl)
     | reconstruct _ _ = raise Bug "malformed inference";
 in
   fun thmToInference th =
       let
 (*MetisTrace3
-        val () = Print.trace Thm.pp "Proof.thmToInference: th" th
-*)
-
-        val cl = Thm.clause th
-
-        val thmInf = Thm.inference th
+        val () = Metis_Print.trace Metis_Thm.pp "Metis_Proof.thmToInference: th" th
+*)
+
+        val cl = Metis_Thm.clause th
+
+        val thmInf = Metis_Thm.inference th
 
 (*MetisTrace3
-        val ppThmInf = Print.ppPair Thm.ppInferenceType (Print.ppList Thm.pp)
-        val () = Print.trace ppThmInf "Proof.thmToInference: thmInf" thmInf
+        val ppThmInf = Metis_Print.ppPair Metis_Thm.ppInferenceType (Metis_Print.ppList Metis_Thm.pp)
+        val () = Metis_Print.trace ppThmInf "Metis_Proof.thmToInference: thmInf" thmInf
 *)
 
         val inf = reconstruct cl thmInf
 
 (*MetisTrace3
-        val () = Print.trace ppInference "Proof.thmToInference: inf" inf
+        val () = Metis_Print.trace ppInference "Metis_Proof.thmToInference: inf" inf
 *)
 (*MetisDebug
         val () =
             let
               val th' = inferenceToThm inf
             in
-              if LiteralSet.equal (Thm.clause th') cl then ()
+              if Metis_LiteralSet.equal (Metis_Thm.clause th') cl then ()
               else
                 raise
                   Bug
-                    ("Proof.thmToInference: bad inference reconstruction:" ^
-                     "\n  th = " ^ Thm.toString th ^
+                    ("Metis_Proof.thmToInference: bad inference reconstruction:" ^
+                     "\n  th = " ^ Metis_Thm.toString th ^
                      "\n  inf = " ^ inferenceToString inf ^
-                     "\n  inf th = " ^ Thm.toString th')
+                     "\n  inf th = " ^ Metis_Thm.toString th')
             end
 *)
       in
@@ -12769,7 +12231,7 @@
       end
 (*MetisDebug
       handle Error err =>
-        raise Bug ("Proof.thmToInference: shouldn't fail:\n" ^ err);
+        raise Bug ("Metis_Proof.thmToInference: shouldn't fail:\n" ^ err);
 *)
 end;
 
@@ -12778,20 +12240,20 @@
 (* ------------------------------------------------------------------------- *)
 
 local
-  val emptyThms : Thm.thm LiteralSetMap.map = LiteralSetMap.new ();
+  val emptyThms : Metis_Thm.thm Metis_LiteralSetMap.map = Metis_LiteralSetMap.new ();
 
   fun addThms (th,ths) =
       let
-        val cl = Thm.clause th
-      in
-        if LiteralSetMap.inDomain cl ths then ths
+        val cl = Metis_Thm.clause th
+      in
+        if Metis_LiteralSetMap.inDomain cl ths then ths
         else
           let
-            val (_,pars) = Thm.inference th
+            val (_,pars) = Metis_Thm.inference th
             val ths = List.foldl addThms ths pars
           in
-            if LiteralSetMap.inDomain cl ths then ths
-            else LiteralSetMap.insert ths (cl,th)
+            if Metis_LiteralSetMap.inDomain cl ths then ths
+            else Metis_LiteralSetMap.insert ths (cl,th)
           end
       end;
 
@@ -12799,15 +12261,15 @@
 
   fun addProof (th,(ths,acc)) =
       let
-        val cl = Thm.clause th
-      in
-        case LiteralSetMap.peek ths cl of
+        val cl = Metis_Thm.clause th
+      in
+        case Metis_LiteralSetMap.peek ths cl of
           NONE => (ths,acc)
         | SOME th =>
           let
-            val (_,pars) = Thm.inference th
+            val (_,pars) = Metis_Thm.inference th
             val (ths,acc) = List.foldl addProof (ths,acc) pars
-            val ths = LiteralSetMap.delete ths cl
+            val ths = Metis_LiteralSetMap.delete ths cl
             val acc = (th, thmToInference th) :: acc
           in
             (ths,acc)
@@ -12818,7 +12280,7 @@
       let
         val (ths,acc) = addProof (th,(ths,[]))
 (*MetisTrace4
-        val () = Print.trace Print.ppInt "Proof.proof: unnecessary clauses" (LiteralSetMap.size ths)
+        val () = Metis_Print.trace Metis_Print.ppInt "Metis_Proof.proof: unnecessary clauses" (Metis_LiteralSetMap.size ths)
 *)
       in
         rev acc
@@ -12827,12 +12289,12 @@
   fun proof th =
       let
 (*MetisTrace3
-        val () = Print.trace Thm.pp "Proof.proof: th" th
+        val () = Metis_Print.trace Metis_Thm.pp "Metis_Proof.proof: th" th
 *)
         val ths = mkThms th
         val infs = mkProof ths th
 (*MetisTrace3
-        val () = Print.trace Print.ppInt "Proof.proof: size" (length infs)
+        val () = Metis_Print.trace Metis_Print.ppInt "Metis_Proof.proof: size" (length infs)
 *)
       in
         infs
@@ -12847,13 +12309,13 @@
     let
       fun free th_inf =
           case th_inf of
-            (_, Axiom lits) => LiteralSet.freeIn v lits
-          | (_, Assume atm) => Atom.freeIn v atm
-          | (th, Subst _) => Thm.freeIn v th
+            (_, Axiom lits) => Metis_LiteralSet.freeIn v lits
+          | (_, Assume atm) => Metis_Atom.freeIn v atm
+          | (th, Metis_Subst _) => Metis_Thm.freeIn v th
           | (_, Resolve _) => false
-          | (_, Refl tm) => Term.freeIn v tm
+          | (_, Refl tm) => Metis_Term.freeIn v tm
           | (_, Equality (lit,_,tm)) =>
-            Literal.freeIn v lit orelse Term.freeIn v tm
+            Metis_Literal.freeIn v lit orelse Metis_Term.freeIn v tm
     in
       List.exists free
     end;
@@ -12861,21 +12323,20 @@
 val freeVars =
     let
       fun inc (th_inf,set) =
-          NameSet.union set
+          Metis_NameSet.union set
           (case th_inf of
-             (_, Axiom lits) => LiteralSet.freeVars lits
-           | (_, Assume atm) => Atom.freeVars atm
-           | (th, Subst _) => Thm.freeVars th
-           | (_, Resolve _) => NameSet.empty
-           | (_, Refl tm) => Term.freeVars tm
+             (_, Axiom lits) => Metis_LiteralSet.freeVars lits
+           | (_, Assume atm) => Metis_Atom.freeVars atm
+           | (th, Metis_Subst _) => Metis_Thm.freeVars th
+           | (_, Resolve _) => Metis_NameSet.empty
+           | (_, Refl tm) => Metis_Term.freeVars tm
            | (_, Equality (lit,_,tm)) =>
-             NameSet.union (Literal.freeVars lit) (Term.freeVars tm))
-    in
-      List.foldl inc NameSet.empty
-    end;
-
-end
-end;
+             Metis_NameSet.union (Metis_Literal.freeVars lit) (Metis_Term.freeVars tm))
+    in
+      List.foldl inc Metis_NameSet.empty
+    end;
+
+end
 
 (**** Original file: Rule.sig ****)
 
@@ -12884,7 +12345,7 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Rule =
+signature Metis_Rule =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -12892,16 +12353,16 @@
 (* t = u \/ C.                                                               *)
 (* ------------------------------------------------------------------------- *)
 
-type equation = (Metis.Term.term * Metis.Term.term) * Metis.Thm.thm
-
-val ppEquation : equation Metis.Print.pp
+type equation = (Metis_Term.term * Metis_Term.term) * Metis_Thm.thm
+
+val ppEquation : equation Metis_Print.pp
 
 val equationToString : equation -> string
 
 (* Returns t = u if the equation theorem contains this literal *)
-val equationLiteral : equation -> Metis.Literal.literal option
-
-val reflEqn : Metis.Term.term -> equation
+val equationLiteral : equation -> Metis_Literal.literal option
+
+val reflEqn : Metis_Term.term -> equation
 
 val symEqn : equation -> equation
 
@@ -12913,7 +12374,7 @@
 (* 2. Raises an Error exception.                                             *)
 (* ------------------------------------------------------------------------- *)
 
-type conv = Metis.Term.term -> Metis.Term.term * Metis.Thm.thm
+type conv = Metis_Term.term -> Metis_Term.term * Metis_Thm.thm
 
 val allConv : conv
 
@@ -12931,9 +12392,9 @@
 
 val everyConv : conv list -> conv
 
-val rewrConv : equation -> Metis.Term.path -> conv
-
-val pathConv : conv -> Metis.Term.path -> conv
+val rewrConv : equation -> Metis_Term.path -> conv
+
+val pathConv : conv -> Metis_Term.path -> conv
 
 val subtermConv : conv -> int -> conv
 
@@ -12955,7 +12416,7 @@
 (* 2. Raises an Error exception.                                             *)
 (* ------------------------------------------------------------------------- *)
 
-type literule = Metis.Literal.literal -> Metis.Literal.literal * Metis.Thm.thm
+type literule = Metis_Literal.literal -> Metis_Literal.literal * Metis_Thm.thm
 
 val allLiterule : literule
 
@@ -12973,9 +12434,9 @@
 
 val everyLiterule : literule list -> literule
 
-val rewrLiterule : equation -> Metis.Term.path -> literule
-
-val pathLiterule : conv -> Metis.Term.path -> literule
+val rewrLiterule : equation -> Metis_Term.path -> literule
+
+val pathLiterule : conv -> Metis_Term.path -> literule
 
 val argumentLiterule : conv -> int -> literule
 
@@ -12986,7 +12447,7 @@
 (* exception.                                                                *)
 (* ------------------------------------------------------------------------- *)
 
-type rule = Metis.Thm.thm -> Metis.Thm.thm
+type rule = Metis_Thm.thm -> Metis_Thm.thm
 
 val allRule : rule
 
@@ -13006,13 +12467,13 @@
 
 val everyRule : rule list -> rule
 
-val literalRule : literule -> Metis.Literal.literal -> rule
-
-val rewrRule : equation -> Metis.Literal.literal -> Metis.Term.path -> rule
-
-val pathRule : conv -> Metis.Literal.literal -> Metis.Term.path -> rule
-
-val literalsRule : literule -> Metis.LiteralSet.set -> rule
+val literalRule : literule -> Metis_Literal.literal -> rule
+
+val rewrRule : equation -> Metis_Literal.literal -> Metis_Term.path -> rule
+
+val pathRule : conv -> Metis_Literal.literal -> Metis_Term.path -> rule
+
+val literalsRule : literule -> Metis_LiteralSet.set -> rule
 
 val allLiteralsRule : literule -> rule
 
@@ -13024,9 +12485,9 @@
 (*   x = x                                                                   *)
 (* ------------------------------------------------------------------------- *)
 
-val reflexivityRule : Metis.Term.term -> Metis.Thm.thm
-
-val reflexivity : Metis.Thm.thm
+val reflexivityRule : Metis_Term.term -> Metis_Thm.thm
+
+val reflexivity : Metis_Thm.thm
 
 (* ------------------------------------------------------------------------- *)
 (*                                                                           *)
@@ -13034,9 +12495,9 @@
 (*   ~(x = y) \/ y = x                                                       *)
 (* ------------------------------------------------------------------------- *)
 
-val symmetryRule : Metis.Term.term -> Metis.Term.term -> Metis.Thm.thm
-
-val symmetry : Metis.Thm.thm
+val symmetryRule : Metis_Term.term -> Metis_Term.term -> Metis_Thm.thm
+
+val symmetry : Metis_Thm.thm
 
 (* ------------------------------------------------------------------------- *)
 (*                                                                           *)
@@ -13044,7 +12505,7 @@
 (*   ~(x = y) \/ ~(y = z) \/ x = z                                           *)
 (* ------------------------------------------------------------------------- *)
 
-val transitivity : Metis.Thm.thm
+val transitivity : Metis_Thm.thm
 
 (* ------------------------------------------------------------------------- *)
 (*                                                                           *)
@@ -13053,7 +12514,7 @@
 (*   f x0 ... x{n-1} = f y0 ... y{n-1}                                       *)
 (* ------------------------------------------------------------------------- *)
 
-val functionCongruence : Metis.Term.function -> Metis.Thm.thm
+val functionCongruence : Metis_Term.function -> Metis_Thm.thm
 
 (* ------------------------------------------------------------------------- *)
 (*                                                                           *)
@@ -13062,7 +12523,7 @@
 (*   ~R x0 ... x{n-1} \/ R y0 ... y{n-1}                                     *)
 (* ------------------------------------------------------------------------- *)
 
-val relationCongruence : Metis.Atom.relation -> Metis.Thm.thm
+val relationCongruence : Metis_Atom.relation -> Metis_Thm.thm
 
 (* ------------------------------------------------------------------------- *)
 (*   x = y \/ C                                                              *)
@@ -13070,7 +12531,7 @@
 (*   y = x \/ C                                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val symEq : Metis.Literal.literal -> rule
+val symEq : Metis_Literal.literal -> rule
 
 (* ------------------------------------------------------------------------- *)
 (*   ~(x = y) \/ C                                                           *)
@@ -13078,13 +12539,13 @@
 (*   ~(y = x) \/ C                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-val symNeq : Metis.Literal.literal -> rule
+val symNeq : Metis_Literal.literal -> rule
 
 (* ------------------------------------------------------------------------- *)
 (* sym (x = y) = symEq (x = y)  /\  sym ~(x = y) = symNeq ~(x = y)           *)
 (* ------------------------------------------------------------------------- *)
 
-val sym : Metis.Literal.literal -> rule
+val sym : Metis_Literal.literal -> rule
 
 (* ------------------------------------------------------------------------- *)
 (*   ~(x = x) \/ C                                                           *)
@@ -13120,7 +12581,7 @@
 (* simplify = isTautology + expandAbbrevs + removeSym                        *)
 (* ------------------------------------------------------------------------- *)
 
-val simplify : Metis.Thm.thm -> Metis.Thm.thm option
+val simplify : Metis_Thm.thm -> Metis_Thm.thm option
 
 (* ------------------------------------------------------------------------- *)
 (*    C                                                                      *)
@@ -13140,7 +12601,7 @@
 (*                                                                           *)
 (* where each s_i is a substitution that factors C, meaning that the theorem *)
 (*                                                                           *)
-(*   C_s_i = (removeIrrefl o removeSym o Metis.Thm.subst s_i) C                    *)
+(*   C_s_i = (removeIrrefl o removeSym o Metis_Thm.subst s_i) C                    *)
 (*                                                                           *)
 (* has fewer literals than C.                                                *)
 (*                                                                           *)
@@ -13148,51 +12609,42 @@
 (* result in a theorem C_s_i that strictly subsumes the theorem C_s.         *)
 (* ------------------------------------------------------------------------- *)
 
-val factor' : Metis.Thm.clause -> Metis.Subst.subst list
-
-val factor : Metis.Thm.thm -> Metis.Thm.thm list
+val factor' : Metis_Thm.clause -> Metis_Subst.subst list
+
+val factor : Metis_Thm.thm -> Metis_Thm.thm list
 
 end
 
 (**** Original file: Rule.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS                     *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Rule :> Rule =
+structure Metis_Rule :> Metis_Rule =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Variable names.                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-val xVarName = Name.fromString "x";
-val xVar = Term.Var xVarName;
-
-val yVarName = Name.fromString "y";
-val yVar = Term.Var yVarName;
-
-val zVarName = Name.fromString "z";
-val zVar = Term.Var zVarName;
-
-fun xIVarName i = Name.fromString ("x" ^ Int.toString i);
-fun xIVar i = Term.Var (xIVarName i);
-
-fun yIVarName i = Name.fromString ("y" ^ Int.toString i);
-fun yIVar i = Term.Var (yIVarName i);
+val xVarName = Metis_Name.fromString "x";
+val xVar = Metis_Term.Var xVarName;
+
+val yVarName = Metis_Name.fromString "y";
+val yVar = Metis_Term.Var yVarName;
+
+val zVarName = Metis_Name.fromString "z";
+val zVar = Metis_Term.Var zVarName;
+
+fun xIVarName i = Metis_Name.fromString ("x" ^ Int.toString i);
+fun xIVar i = Metis_Term.Var (xIVarName i);
+
+fun yIVarName i = Metis_Name.fromString ("y" ^ Int.toString i);
+fun yIVar i = Metis_Term.Var (yIVarName i);
 
 (* ------------------------------------------------------------------------- *)
 (*                                                                           *)
@@ -13200,7 +12652,7 @@
 (*   x = x                                                                   *)
 (* ------------------------------------------------------------------------- *)
 
-fun reflexivityRule x = Thm.refl x;
+fun reflexivityRule x = Metis_Thm.refl x;
 
 val reflexivity = reflexivityRule xVar;
 
@@ -13213,10 +12665,10 @@
 fun symmetryRule x y =
     let
       val reflTh = reflexivityRule x
-      val reflLit = Thm.destUnit reflTh
-      val eqTh = Thm.equality reflLit [0] y
-    in
-      Thm.resolve reflLit reflTh eqTh
+      val reflLit = Metis_Thm.destUnit reflTh
+      val eqTh = Metis_Thm.equality reflLit [0] y
+    in
+      Metis_Thm.resolve reflLit reflTh eqTh
     end;
 
 val symmetry = symmetryRule xVar yVar;
@@ -13229,9 +12681,9 @@
 
 val transitivity =
     let
-      val eqTh = Thm.equality (Literal.mkEq (yVar,zVar)) [0] xVar
-    in
-      Thm.resolve (Literal.mkEq (yVar,xVar)) symmetry eqTh
+      val eqTh = Metis_Thm.equality (Metis_Literal.mkEq (yVar,zVar)) [0] xVar
+    in
+      Metis_Thm.resolve (Metis_Literal.mkEq (yVar,xVar)) symmetry eqTh
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -13242,15 +12694,15 @@
 
 fun symEq lit th =
     let
-      val (x,y) = Literal.destEq lit
-    in
-      if Term.equal x y then th
+      val (x,y) = Metis_Literal.destEq lit
+    in
+      if Metis_Term.equal x y then th
       else
         let
-          val sub = Subst.fromList [(xVarName,x),(yVarName,y)]
-          val symTh = Thm.subst sub symmetry
-        in
-          Thm.resolve lit th symTh
+          val sub = Metis_Subst.fromList [(xVarName,x),(yVarName,y)]
+          val symTh = Metis_Thm.subst sub symmetry
+        in
+          Metis_Thm.resolve lit th symTh
         end
     end;
 
@@ -13259,23 +12711,23 @@
 (* t = u \/ C.                                                               *)
 (* ------------------------------------------------------------------------- *)
 
-type equation = (Term.term * Term.term) * Thm.thm;
-
-fun ppEquation (_,th) = Thm.pp th;
-
-val equationToString = Print.toString ppEquation;
+type equation = (Metis_Term.term * Metis_Term.term) * Metis_Thm.thm;
+
+fun ppEquation (_,th) = Metis_Thm.pp th;
+
+val equationToString = Metis_Print.toString ppEquation;
 
 fun equationLiteral (t_u,th) =
     let
-      val lit = Literal.mkEq t_u
-    in
-      if LiteralSet.member lit (Thm.clause th) then SOME lit else NONE
-    end;
-
-fun reflEqn t = ((t,t), Thm.refl t);
+      val lit = Metis_Literal.mkEq t_u
+    in
+      if Metis_LiteralSet.member lit (Metis_Thm.clause th) then SOME lit else NONE
+    end;
+
+fun reflEqn t = ((t,t), Metis_Thm.refl t);
 
 fun symEqn (eqn as ((t,u), th)) =
-    if Term.equal t u then eqn
+    if Metis_Term.equal t u then eqn
     else
       ((u,t),
        case equationLiteral eqn of
@@ -13283,9 +12735,9 @@
        | NONE => th);
 
 fun transEqn (eqn1 as ((x,y), th1)) (eqn2 as ((_,z), th2)) =
-    if Term.equal x y then eqn2
-    else if Term.equal y z then eqn1
-    else if Term.equal x z then reflEqn x
+    if Metis_Term.equal x y then eqn2
+    else if Metis_Term.equal y z then eqn1
+    else if Metis_Term.equal x z then reflEqn x
     else
       ((x,z),
        case equationLiteral eqn1 of
@@ -13295,10 +12747,10 @@
            NONE => th2
          | SOME y_z =>
            let
-             val sub = Subst.fromList [(xVarName,x),(yVarName,y),(zVarName,z)]
-             val th = Thm.subst sub transitivity
-             val th = Thm.resolve x_y th1 th
-             val th = Thm.resolve y_z th2 th
+             val sub = Metis_Subst.fromList [(xVarName,x),(yVarName,y),(zVarName,z)]
+             val th = Metis_Thm.subst sub transitivity
+             val th = Metis_Thm.resolve x_y th1 th
+             val th = Metis_Thm.resolve y_z th2 th
            in
              th
            end);
@@ -13307,7 +12759,7 @@
 val transEqn = fn eqn1 => fn eqn2 =>
     transEqn eqn1 eqn2
     handle Error err =>
-      raise Error ("Rule.transEqn:\neqn1 = " ^ equationToString eqn1 ^
+      raise Error ("Metis_Rule.transEqn:\neqn1 = " ^ equationToString eqn1 ^
                    "\neqn2 = " ^ equationToString eqn2 ^ "\n" ^ err);
 *)
 
@@ -13317,22 +12769,22 @@
 (* 2. Raises an Error exception.                                             *)
 (* ------------------------------------------------------------------------- *)
 
-type conv = Term.term -> Term.term * Thm.thm;
-
-fun allConv tm = (tm, Thm.refl tm);
+type conv = Metis_Term.term -> Metis_Term.term * Metis_Thm.thm;
+
+fun allConv tm = (tm, Metis_Thm.refl tm);
 
 val noConv : conv = fn _ => raise Error "noConv";
 
 fun traceConv s conv tm =
     let
       val res as (tm',th) = conv tm
-      val () = print (s ^ ": " ^ Term.toString tm ^ " --> " ^
-                      Term.toString tm' ^ " " ^ Thm.toString th ^ "\n")
+      val () = print (s ^ ": " ^ Metis_Term.toString tm ^ " --> " ^
+                      Metis_Term.toString tm' ^ " " ^ Metis_Thm.toString th ^ "\n")
     in
       res
     end
     handle Error err =>
-      (print (s ^ ": " ^ Term.toString tm ^ " --> Error: " ^ err ^ "\n");
+      (print (s ^ ": " ^ Metis_Term.toString tm ^ " --> Error: " ^ err ^ "\n");
        raise Error (s ^ ": " ^ err));
 
 fun thenConvTrans tm (tm',th1) (tm'',th2) =
@@ -13374,19 +12826,19 @@
   | everyConv (conv :: convs) tm = thenConv conv (everyConv convs) tm;
 
 fun rewrConv (eqn as ((x,y), eqTh)) path tm =
-    if Term.equal x y then allConv tm
+    if Metis_Term.equal x y then allConv tm
     else if null path then (y,eqTh)
     else
       let
-        val reflTh = Thm.refl tm
-        val reflLit = Thm.destUnit reflTh
-        val th = Thm.equality reflLit (1 :: path) y
-        val th = Thm.resolve reflLit reflTh th
+        val reflTh = Metis_Thm.refl tm
+        val reflLit = Metis_Thm.destUnit reflTh
+        val th = Metis_Thm.equality reflLit (1 :: path) y
+        val th = Metis_Thm.resolve reflLit reflTh th
         val th =
             case equationLiteral eqn of
               NONE => th
-            | SOME x_y => Thm.resolve x_y eqTh th
-        val tm' = Term.replace tm (path,y)
+            | SOME x_y => Metis_Thm.resolve x_y eqTh th
+        val tm' = Metis_Term.replace tm (path,y)
       in
         (tm',th)
       end;
@@ -13395,16 +12847,16 @@
 val rewrConv = fn eqn as ((x,y),eqTh) => fn path => fn tm =>
     rewrConv eqn path tm
     handle Error err =>
-      raise Error ("Rule.rewrConv:\nx = " ^ Term.toString x ^
-                   "\ny = " ^ Term.toString y ^
-                   "\neqTh = " ^ Thm.toString eqTh ^
-                   "\npath = " ^ Term.pathToString path ^
-                   "\ntm = " ^ Term.toString tm ^ "\n" ^ err);
+      raise Error ("Metis_Rule.rewrConv:\nx = " ^ Metis_Term.toString x ^
+                   "\ny = " ^ Metis_Term.toString y ^
+                   "\neqTh = " ^ Metis_Thm.toString eqTh ^
+                   "\npath = " ^ Metis_Term.pathToString path ^
+                   "\ntm = " ^ Metis_Term.toString tm ^ "\n" ^ err);
 *)
 
 fun pathConv conv path tm =
     let
-      val x = Term.subterm tm path
+      val x = Metis_Term.subterm tm path
       val (y,th) = conv x
     in
       rewrConv ((x,y),th) path tm
@@ -13412,8 +12864,8 @@
 
 fun subtermConv conv i = pathConv conv [i];
 
-fun subtermsConv _ (tm as Term.Var _) = allConv tm
-  | subtermsConv conv (tm as Term.Fn (_,a)) =
+fun subtermsConv _ (tm as Metis_Term.Var _) = allConv tm
+  | subtermsConv conv (tm as Metis_Term.Fn (_,a)) =
     everyConv (map (subtermConv conv) (interval 0 (length a))) tm;
 
 (* ------------------------------------------------------------------------- *)
@@ -13447,9 +12899,9 @@
 (* 2. Raises an Error exception.                                             *)
 (* ------------------------------------------------------------------------- *)
 
-type literule = Literal.literal -> Literal.literal * Thm.thm;
-
-fun allLiterule lit = (lit, Thm.assume lit);
+type literule = Metis_Literal.literal -> Metis_Literal.literal * Metis_Thm.thm;
+
+fun allLiterule lit = (lit, Metis_Thm.assume lit);
 
 val noLiterule : literule = fn _ => raise Error "noLiterule";
 
@@ -13458,14 +12910,14 @@
       val res1 as (lit',th1) = literule1 lit
       val res2 as (lit'',th2) = literule2 lit'
     in
-      if Literal.equal lit lit' then res2
-      else if Literal.equal lit' lit'' then res1
-      else if Literal.equal lit lit'' then allLiterule lit
+      if Metis_Literal.equal lit lit' then res2
+      else if Metis_Literal.equal lit' lit'' then res1
+      else if Metis_Literal.equal lit lit'' then allLiterule lit
       else
         (lit'',
-         if not (Thm.member lit' th1) then th1
-         else if not (Thm.negateMember lit' th2) then th2
-         else Thm.resolve lit' th1 th2)
+         if not (Metis_Thm.member lit' th1) then th1
+         else if not (Metis_Thm.negateMember lit' th2) then th2
+         else Metis_Thm.resolve lit' th1 th2)
     end;
 
 fun orelseLiterule (literule1 : literule) literule2 lit =
@@ -13494,15 +12946,15 @@
     thenLiterule literule (everyLiterule literules) lit;
 
 fun rewrLiterule (eqn as ((x,y),eqTh)) path lit =
-    if Term.equal x y then allLiterule lit
+    if Metis_Term.equal x y then allLiterule lit
     else
       let
-        val th = Thm.equality lit path y
+        val th = Metis_Thm.equality lit path y
         val th =
             case equationLiteral eqn of
               NONE => th
-            | SOME x_y => Thm.resolve x_y eqTh th
-        val lit' = Literal.replace lit (path,y)
+            | SOME x_y => Metis_Thm.resolve x_y eqTh th
+        val lit' = Metis_Literal.replace lit (path,y)
       in
         (lit',th)
       end;
@@ -13511,14 +12963,14 @@
 val rewrLiterule = fn eqn => fn path => fn lit =>
     rewrLiterule eqn path lit
     handle Error err =>
-      raise Error ("Rule.rewrLiterule:\neqn = " ^ equationToString eqn ^
-                   "\npath = " ^ Term.pathToString path ^
-                   "\nlit = " ^ Literal.toString lit ^ "\n" ^ err);
+      raise Error ("Metis_Rule.rewrLiterule:\neqn = " ^ equationToString eqn ^
+                   "\npath = " ^ Metis_Term.pathToString path ^
+                   "\nlit = " ^ Metis_Literal.toString lit ^ "\n" ^ err);
 *)
 
 fun pathLiterule conv path lit =
     let
-      val tm = Literal.subterm lit path
+      val tm = Metis_Literal.subterm lit path
       val (tm',th) = conv tm
     in
       rewrLiterule ((tm,tm'),th) path lit
@@ -13528,14 +12980,14 @@
 
 fun allArgumentsLiterule conv lit =
     everyLiterule
-      (map (argumentLiterule conv) (interval 0 (Literal.arity lit))) lit;
+      (map (argumentLiterule conv) (interval 0 (Metis_Literal.arity lit))) lit;
 
 (* ------------------------------------------------------------------------- *)
 (* A rule takes one theorem and either deduces another or raises an Error    *)
 (* exception.                                                                *)
 (* ------------------------------------------------------------------------- *)
 
-type rule = Thm.thm -> Thm.thm;
+type rule = Metis_Thm.thm -> Metis_Thm.thm;
 
 val allRule : rule = fn th => th;
 
@@ -13551,7 +13003,7 @@
     let
       val th' = rule th
     in
-      if not (LiteralSet.equal (Thm.clause th) (Thm.clause th')) then th'
+      if not (Metis_LiteralSet.equal (Metis_Thm.clause th) (Metis_Thm.clause th')) then th'
       else raise Error "changedRule"
     end;
 
@@ -13569,17 +13021,17 @@
     let
       val (lit',litTh) = literule lit
     in
-      if Literal.equal lit lit' then th
-      else if not (Thm.negateMember lit litTh) then litTh
-      else Thm.resolve lit th litTh
+      if Metis_Literal.equal lit lit' then th
+      else if not (Metis_Thm.negateMember lit litTh) then litTh
+      else Metis_Thm.resolve lit th litTh
     end;
 
 (*MetisDebug
 val literalRule = fn literule => fn lit => fn th =>
     literalRule literule lit th
     handle Error err =>
-      raise Error ("Rule.literalRule:\nlit = " ^ Literal.toString lit ^
-                   "\nth = " ^ Thm.toString th ^ "\n" ^ err);
+      raise Error ("Metis_Rule.literalRule:\nlit = " ^ Metis_Literal.toString lit ^
+                   "\nth = " ^ Metis_Thm.toString th ^ "\n" ^ err);
 *)
 
 fun rewrRule eqTh lit path = literalRule (rewrLiterule eqTh path) lit;
@@ -13589,12 +13041,12 @@
 fun literalsRule literule =
     let
       fun f (lit,th) =
-          if Thm.member lit th then literalRule literule lit th else th
-    in
-      fn lits => fn th => LiteralSet.foldl f th lits
-    end;
-
-fun allLiteralsRule literule th = literalsRule literule (Thm.clause th) th;
+          if Metis_Thm.member lit th then literalRule literule lit th else th
+    in
+      fn lits => fn th => Metis_LiteralSet.foldl f th lits
+    end;
+
+fun allLiteralsRule literule th = literalsRule literule (Metis_Thm.clause th) th;
 
 fun convRule conv = allLiteralsRule (allArgumentsLiterule conv);
 
@@ -13613,14 +13065,14 @@
       fun cong ((i,yi),(th,lit)) =
           let
             val path = [1,i]
-            val th = Thm.resolve lit th (Thm.equality lit path yi)
-            val lit = Literal.replace lit (path,yi)
+            val th = Metis_Thm.resolve lit th (Metis_Thm.equality lit path yi)
+            val lit = Metis_Literal.replace lit (path,yi)
           in
             (th,lit)
           end
 
-      val reflTh = Thm.refl (Term.Fn (f,xs))
-      val reflLit = Thm.destUnit reflTh
+      val reflTh = Metis_Thm.refl (Metis_Term.Fn (f,xs))
+      val reflLit = Metis_Thm.destUnit reflTh
     in
       fst (foldl cong (reflTh,reflLit) (enumerate ys))
     end;
@@ -13640,14 +13092,14 @@
       fun cong ((i,yi),(th,lit)) =
           let
             val path = [i]
-            val th = Thm.resolve lit th (Thm.equality lit path yi)
-            val lit = Literal.replace lit (path,yi)
+            val th = Metis_Thm.resolve lit th (Metis_Thm.equality lit path yi)
+            val lit = Metis_Literal.replace lit (path,yi)
           in
             (th,lit)
           end
 
       val assumeLit = (false,(R,xs))
-      val assumeTh = Thm.assume assumeLit
+      val assumeTh = Metis_Thm.assume assumeLit
     in
       fst (foldl cong (assumeTh,assumeLit) (enumerate ys))
     end;
@@ -13660,15 +13112,15 @@
 
 fun symNeq lit th =
     let
-      val (x,y) = Literal.destNeq lit
-    in
-      if Term.equal x y then th
+      val (x,y) = Metis_Literal.destNeq lit
+    in
+      if Metis_Term.equal x y then th
       else
         let
-          val sub = Subst.fromList [(xVarName,y),(yVarName,x)]
-          val symTh = Thm.subst sub symmetry
-        in
-          Thm.resolve lit th symTh
+          val sub = Metis_Subst.fromList [(xVarName,y),(yVarName,x)]
+          val symTh = Metis_Thm.subst sub symmetry
+        in
+          Metis_Thm.resolve lit th symTh
         end
     end;
 
@@ -13689,11 +13141,11 @@
 local
   fun irrefl ((true,_),th) = th
     | irrefl (lit as (false,atm), th) =
-      case total Atom.destRefl atm of
-        SOME x => Thm.resolve lit th (Thm.refl x)
+      case total Metis_Atom.destRefl atm of
+        SOME x => Metis_Thm.resolve lit th (Metis_Thm.refl x)
       | NONE => th;
 in
-  fun removeIrrefl th = LiteralSet.foldl irrefl th (Thm.clause th);
+  fun removeIrrefl th = Metis_LiteralSet.foldl irrefl th (Metis_Thm.clause th);
 end;
 
 (* ------------------------------------------------------------------------- *)
@@ -13706,16 +13158,16 @@
 
 local
   fun rem (lit as (pol,atm), eqs_th as (eqs,th)) =
-      case total Atom.sym atm of
+      case total Metis_Atom.sym atm of
         NONE => eqs_th
       | SOME atm' =>
-        if LiteralSet.member lit eqs then
+        if Metis_LiteralSet.member lit eqs then
           (eqs, if pol then symEq lit th else symNeq lit th)
         else
-          (LiteralSet.add eqs (pol,atm'), th);
+          (Metis_LiteralSet.add eqs (pol,atm'), th);
 in
   fun removeSym th =
-      snd (LiteralSet.foldl rem (LiteralSet.empty,th) (Thm.clause th));
+      snd (Metis_LiteralSet.foldl rem (Metis_LiteralSet.empty,th) (Metis_Thm.clause th));
 end;
 
 (* ------------------------------------------------------------------------- *)
@@ -13729,19 +13181,19 @@
 local
   fun expand lit =
       let
-        val (x,y) = Literal.destNeq lit
-        val _ = Term.isTypedVar x orelse Term.isTypedVar y orelse
-                raise Error "Rule.expandAbbrevs: no vars"
-        val _ = not (Term.equal x y) orelse
-                raise Error "Rule.expandAbbrevs: equal vars"
-      in
-        Subst.unify Subst.empty x y
+        val (x,y) = Metis_Literal.destNeq lit
+        val _ = Metis_Term.isTypedVar x orelse Metis_Term.isTypedVar y orelse
+                raise Error "Metis_Rule.expandAbbrevs: no vars"
+        val _ = not (Metis_Term.equal x y) orelse
+                raise Error "Metis_Rule.expandAbbrevs: equal vars"
+      in
+        Metis_Subst.unify Metis_Subst.empty x y
       end;
 in
   fun expandAbbrevs th =
-      case LiteralSet.firstl (total expand) (Thm.clause th) of
+      case Metis_LiteralSet.firstl (total expand) (Metis_Thm.clause th) of
         NONE => removeIrrefl th
-      | SOME sub => expandAbbrevs (Thm.subst sub th);
+      | SOME sub => expandAbbrevs (Metis_Thm.subst sub th);
 end;
 
 (* ------------------------------------------------------------------------- *)
@@ -13749,14 +13201,14 @@
 (* ------------------------------------------------------------------------- *)
 
 fun simplify th =
-    if Thm.isTautology th then NONE
+    if Metis_Thm.isTautology th then NONE
     else
       let
         val th' = th
         val th' = expandAbbrevs th'
         val th' = removeSym th'
       in
-        if Thm.equal th th' then SOME th else simplify th'
+        if Metis_Thm.equal th th' then SOME th else simplify th'
       end;
 
 (* ------------------------------------------------------------------------- *)
@@ -13768,7 +13220,7 @@
 (* C are replaced by fresh variables.                                        *)
 (* ------------------------------------------------------------------------- *)
 
-fun freshVars th = Thm.subst (Subst.freshVars (Thm.freeVars th)) th;
+fun freshVars th = Metis_Thm.subst (Metis_Subst.freshVars (Metis_Thm.freeVars th)) th;
 
 (* ------------------------------------------------------------------------- *)
 (*               C                                                           *)
@@ -13777,7 +13229,7 @@
 (*                                                                           *)
 (* where each s_i is a substitution that factors C, meaning that the theorem *)
 (*                                                                           *)
-(*   C_s_i = (removeIrrefl o removeSym o Thm.subst s_i) C                    *)
+(*   C_s_i = (removeIrrefl o removeSym o Metis_Thm.subst s_i) C                    *)
 (*                                                                           *)
 (* has fewer literals than C.                                                *)
 (*                                                                           *)
@@ -13787,28 +13239,28 @@
 
 local
   datatype edge =
-      FactorEdge of Atom.atom * Atom.atom
-    | ReflEdge of Term.term * Term.term;
-
-  fun ppEdge (FactorEdge atm_atm') = Print.ppPair Atom.pp Atom.pp atm_atm'
-    | ppEdge (ReflEdge tm_tm') = Print.ppPair Term.pp Term.pp tm_tm';
+      FactorEdge of Metis_Atom.atom * Metis_Atom.atom
+    | ReflEdge of Metis_Term.term * Metis_Term.term;
+
+  fun ppEdge (FactorEdge atm_atm') = Metis_Print.ppPair Metis_Atom.pp Metis_Atom.pp atm_atm'
+    | ppEdge (ReflEdge tm_tm') = Metis_Print.ppPair Metis_Term.pp Metis_Term.pp tm_tm';
 
   datatype joinStatus =
       Joined
-    | Joinable of Subst.subst
+    | Joinable of Metis_Subst.subst
     | Apart;
 
   fun joinEdge sub edge =
       let
         val result =
             case edge of
-              FactorEdge (atm,atm') => total (Atom.unify sub atm) atm'
-            | ReflEdge (tm,tm') => total (Subst.unify sub tm) tm'
+              FactorEdge (atm,atm') => total (Metis_Atom.unify sub atm) atm'
+            | ReflEdge (tm,tm') => total (Metis_Subst.unify sub tm) tm'
       in
         case result of
           NONE => Apart
         | SOME sub' =>
-          if Portable.pointerEqual (sub,sub') then Joined else Joinable sub'
+          if Metis_Portable.pointerEqual (sub,sub') then Joined else Joinable sub'
       end;
 
   fun updateApart sub =
@@ -13829,7 +13281,7 @@
         let
           val edge = FactorEdge (atm,atm')
         in
-          case joinEdge Subst.empty edge of
+          case joinEdge Metis_Subst.empty edge of
             Joined => raise Bug "addFactorEdge: joined"
           | Joinable sub => (sub,edge) :: acc
           | Apart => acc
@@ -13838,9 +13290,9 @@
   fun addReflEdge (false,_) acc = acc
     | addReflEdge (true,atm) acc =
       let
-        val edge = ReflEdge (Atom.destEq atm)
-      in
-        case joinEdge Subst.empty edge of
+        val edge = ReflEdge (Metis_Atom.destEq atm)
+      in
+        case joinEdge Metis_Subst.empty edge of
           Joined => raise Bug "addRefl: joined"
         | Joinable _ => edge :: acc
         | Apart => acc
@@ -13849,9 +13301,9 @@
   fun addIrreflEdge (true,_) acc = acc
     | addIrreflEdge (false,atm) acc =
       let
-        val edge = ReflEdge (Atom.destEq atm)
-      in
-        case joinEdge Subst.empty edge of
+        val edge = ReflEdge (Metis_Atom.destEq atm)
+      in
+        case joinEdge Metis_Subst.empty edge of
           Joined => raise Bug "addRefl: joined"
         | Joinable sub => (sub,edge) :: acc
         | Apart => acc
@@ -13867,8 +13319,8 @@
     | init_edges acc apart ((sub,edge) :: sub_edges) =
       let
 (*MetisDebug
-        val () = if not (Subst.null sub) then ()
-                 else raise Bug "Rule.factor.init_edges: empty subst"
+        val () = if not (Metis_Subst.null sub) then ()
+                 else raise Bug "Metis_Rule.factor.init_edges: empty subst"
 *)
         val (acc,apart) =
             case updateApart sub apart of
@@ -13884,7 +13336,7 @@
         val sub_edges = List.foldl (addFactorEdge lit) sub_edges lits
 
         val (apart,sub_edges) =
-            case total Literal.sym lit of
+            case total Metis_Literal.sym lit of
               NONE => (apart,sub_edges)
             | SOME lit' =>
               let
@@ -13920,20 +13372,20 @@
   fun factor' cl =
       let
 (*MetisTrace6
-        val () = Print.trace LiteralSet.pp "Rule.factor': cl" cl
-*)
-        val edges = mk_edges [] [] (LiteralSet.toList cl)
+        val () = Metis_Print.trace Metis_LiteralSet.pp "Metis_Rule.factor': cl" cl
+*)
+        val edges = mk_edges [] [] (Metis_LiteralSet.toList cl)
 (*MetisTrace6
-        val ppEdgesSize = Print.ppMap length Print.ppInt
-        val ppEdgel = Print.ppList ppEdge
-        val ppEdges = Print.ppList (Print.ppTriple ppEdgel Subst.pp ppEdgel)
-        val () = Print.trace ppEdgesSize "Rule.factor': |edges|" edges
-        val () = Print.trace ppEdges "Rule.factor': edges" edges
+        val ppEdgesSize = Metis_Print.ppMap length Metis_Print.ppInt
+        val ppEdgel = Metis_Print.ppList ppEdge
+        val ppEdges = Metis_Print.ppList (Metis_Print.ppTriple ppEdgel Metis_Subst.pp ppEdgel)
+        val () = Metis_Print.trace ppEdgesSize "Metis_Rule.factor': |edges|" edges
+        val () = Metis_Print.trace ppEdges "Metis_Rule.factor': edges" edges
 *)
         val result = fact [] edges
 (*MetisTrace6
-        val ppResult = Print.ppList Subst.pp
-        val () = Print.trace ppResult "Rule.factor': result" result
+        val ppResult = Metis_Print.ppList Metis_Subst.pp
+        val () = Metis_Print.trace ppResult "Metis_Rule.factor': result" result
 *)
       in
         result
@@ -13942,13 +13394,12 @@
 
 fun factor th =
     let
-      fun fact sub = removeSym (Thm.subst sub th)
-    in
-      map fact (factor' (Thm.clause th))
-    end;
-
-end
-end;
+      fun fact sub = removeSym (Metis_Thm.subst sub th)
+    in
+      map fact (factor' (Metis_Thm.clause th))
+    end;
+
+end
 
 (**** Original file: Normalize.sig ****)
 
@@ -13957,14 +13408,14 @@
 (* Copyright (c) 2001-2009 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Normalize =
+signature Metis_Normalize =
 sig
 
 (* ------------------------------------------------------------------------- *)
 (* Negation normal form.                                                     *)
 (* ------------------------------------------------------------------------- *)
 
-val nnf : Metis.Formula.formula -> Metis.Formula.formula
+val nnf : Metis_Formula.formula -> Metis_Formula.formula
 
 (* ------------------------------------------------------------------------- *)
 (* Conjunctive normal form derivations.                                      *)
@@ -13973,24 +13424,24 @@
 type thm
 
 datatype inference =
-    Axiom of Metis.Formula.formula
-  | Definition of string * Metis.Formula.formula
+    Axiom of Metis_Formula.formula
+  | Definition of string * Metis_Formula.formula
   | Simplify of thm * thm list
   | Conjunct of thm
   | Specialize of thm
   | Skolemize of thm
   | Clausify of thm
 
-val mkAxiom : Metis.Formula.formula -> thm
-
-val destThm : thm -> Metis.Formula.formula * inference
+val mkAxiom : Metis_Formula.formula -> thm
+
+val destThm : thm -> Metis_Formula.formula * inference
 
 val proveThms :
-    thm list -> (Metis.Formula.formula * inference * Metis.Formula.formula list) list
+    thm list -> (Metis_Formula.formula * inference * Metis_Formula.formula list) list
 
 val toStringInference : inference -> string
 
-val ppInference : inference Metis.Print.pp
+val ppInference : inference Metis_Print.pp
 
 (* ------------------------------------------------------------------------- *)
 (* Conjunctive normal form.                                                  *)
@@ -14000,34 +13451,25 @@
 
 val initialCnf : cnf
 
-val addCnf : thm -> cnf -> (Metis.Thm.clause * thm) list * cnf
-
-val proveCnf : thm list -> (Metis.Thm.clause * thm) list
-
-val cnf : Metis.Formula.formula -> Metis.Thm.clause list
+val addCnf : thm -> cnf -> (Metis_Thm.clause * thm) list * cnf
+
+val proveCnf : thm list -> (Metis_Thm.clause * thm) list
+
+val cnf : Metis_Formula.formula -> Metis_Thm.clause list
 
 end
 
 (**** Original file: Normalize.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* NORMALIZING FORMULAS                                                      *)
 (* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Normalize :> Normalize =
+structure Metis_Normalize :> Metis_Normalize =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Constants.                                                                *)
@@ -14162,7 +13604,7 @@
          "(+" ^ rToS p ^ ",-" ^ rToS n ^ ")"
     end;
 
-val ppCount = Print.ppMap countToString Print.ppString;
+val ppCount = Metis_Print.ppMap countToString Metis_Print.ppString;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of normalized formula.                                             *)
@@ -14171,15 +13613,15 @@
 datatype formula =
     True
   | False
-  | Literal of NameSet.set * Literal.literal
-  | And of NameSet.set * count * formula Set.set
-  | Or of NameSet.set * count * formula Set.set
-  | Xor of NameSet.set * count * bool * formula Set.set
-  | Exists of NameSet.set * count * NameSet.set * formula
-  | Forall of NameSet.set * count * NameSet.set * formula;
+  | Metis_Literal of Metis_NameSet.set * Metis_Literal.literal
+  | And of Metis_NameSet.set * count * formula Metis_Set.set
+  | Or of Metis_NameSet.set * count * formula Metis_Set.set
+  | Xor of Metis_NameSet.set * count * bool * formula Metis_Set.set
+  | Exists of Metis_NameSet.set * count * Metis_NameSet.set * formula
+  | Forall of Metis_NameSet.set * count * Metis_NameSet.set * formula;
 
 fun compare f1_f2 =
-    if Portable.pointerEqual f1_f2 then EQUAL
+    if Metis_Portable.pointerEqual f1_f2 then EQUAL
     else
       case f1_f2 of
         (True,True) => EQUAL
@@ -14188,71 +13630,71 @@
       | (False,False) => EQUAL
       | (False,_) => LESS
       | (_,False) => GREATER
-      | (Literal (_,l1), Literal (_,l2)) => Literal.compare (l1,l2)
-      | (Literal _, _) => LESS
-      | (_, Literal _) => GREATER
-      | (And (_,_,s1), And (_,_,s2)) => Set.compare (s1,s2)
+      | (Metis_Literal (_,l1), Metis_Literal (_,l2)) => Metis_Literal.compare (l1,l2)
+      | (Metis_Literal _, _) => LESS
+      | (_, Metis_Literal _) => GREATER
+      | (And (_,_,s1), And (_,_,s2)) => Metis_Set.compare (s1,s2)
       | (And _, _) => LESS
       | (_, And _) => GREATER
-      | (Or (_,_,s1), Or (_,_,s2)) => Set.compare (s1,s2)
+      | (Or (_,_,s1), Or (_,_,s2)) => Metis_Set.compare (s1,s2)
       | (Or _, _) => LESS
       | (_, Or _) => GREATER
       | (Xor (_,_,p1,s1), Xor (_,_,p2,s2)) =>
         (case boolCompare (p1,p2) of
            LESS => LESS
-         | EQUAL => Set.compare (s1,s2)
+         | EQUAL => Metis_Set.compare (s1,s2)
          | GREATER => GREATER)
       | (Xor _, _) => LESS
       | (_, Xor _) => GREATER
       | (Exists (_,_,n1,f1), Exists (_,_,n2,f2)) =>
-        (case NameSet.compare (n1,n2) of
+        (case Metis_NameSet.compare (n1,n2) of
            LESS => LESS
          | EQUAL => compare (f1,f2)
          | GREATER => GREATER)
       | (Exists _, _) => LESS
       | (_, Exists _) => GREATER
       | (Forall (_,_,n1,f1), Forall (_,_,n2,f2)) =>
-        (case NameSet.compare (n1,n2) of
+        (case Metis_NameSet.compare (n1,n2) of
            LESS => LESS
          | EQUAL => compare (f1,f2)
          | GREATER => GREATER);
 
-val empty = Set.empty compare;
-
-val singleton = Set.singleton compare;
+val empty = Metis_Set.empty compare;
+
+val singleton = Metis_Set.singleton compare;
 
 local
   fun neg True = False
     | neg False = True
-    | neg (Literal (fv,lit)) = Literal (fv, Literal.negate lit)
+    | neg (Metis_Literal (fv,lit)) = Metis_Literal (fv, Metis_Literal.negate lit)
     | neg (And (fv,c,s)) = Or (fv, countNegate c, neg_set s)
     | neg (Or (fv,c,s)) = And (fv, countNegate c, neg_set s)
     | neg (Xor (fv,c,p,s)) = Xor (fv, c, not p, s)
     | neg (Exists (fv,c,n,f)) = Forall (fv, countNegate c, n, neg f)
     | neg (Forall (fv,c,n,f)) = Exists (fv, countNegate c, n, neg f)
 
-  and neg_set s = Set.foldl neg_elt empty s
-
-  and neg_elt (f,s) = Set.add s (neg f);
+  and neg_set s = Metis_Set.foldl neg_elt empty s
+
+  and neg_elt (f,s) = Metis_Set.add s (neg f);
 in
   val negate = neg;
 
   val negateSet = neg_set;
 end;
 
-fun negateMember x s = Set.member (negate x) s;
+fun negateMember x s = Metis_Set.member (negate x) s;
 
 local
   fun member s x = negateMember x s;
 in
   fun negateDisjoint s1 s2 =
-      if Set.size s1 < Set.size s2 then not (Set.exists (member s2) s1)
-      else not (Set.exists (member s1) s2);
+      if Metis_Set.size s1 < Metis_Set.size s2 then not (Metis_Set.exists (member s2) s1)
+      else not (Metis_Set.exists (member s1) s2);
 end;
 
 fun polarity True = true
   | polarity False = false
-  | polarity (Literal (_,(pol,_))) = not pol
+  | polarity (Metis_Literal (_,(pol,_))) = not pol
   | polarity (And _) = true
   | polarity (Or _) = false
   | polarity (Xor (_,_,pol,_)) = pol
@@ -14273,27 +13715,27 @@
 fun applyPolarity true fm = fm
   | applyPolarity false fm = negate fm;
 
-fun freeVars True = NameSet.empty
-  | freeVars False = NameSet.empty
-  | freeVars (Literal (fv,_)) = fv
+fun freeVars True = Metis_NameSet.empty
+  | freeVars False = Metis_NameSet.empty
+  | freeVars (Metis_Literal (fv,_)) = fv
   | freeVars (And (fv,_,_)) = fv
   | freeVars (Or (fv,_,_)) = fv
   | freeVars (Xor (fv,_,_,_)) = fv
   | freeVars (Exists (fv,_,_,_)) = fv
   | freeVars (Forall (fv,_,_,_)) = fv;
 
-fun freeIn v fm = NameSet.member v (freeVars fm);
+fun freeIn v fm = Metis_NameSet.member v (freeVars fm);
 
 val freeVarsSet =
     let
-      fun free (fm,acc) = NameSet.union (freeVars fm) acc
-    in
-      Set.foldl free NameSet.empty
+      fun free (fm,acc) = Metis_NameSet.union (freeVars fm) acc
+    in
+      Metis_Set.foldl free Metis_NameSet.empty
     end;
 
 fun count True = countTrue
   | count False = countFalse
-  | count (Literal _) = countLiteral
+  | count (Metis_Literal _) = countLiteral
   | count (And (_,c,_)) = c
   | count (Or (_,c,_)) = c
   | count (Xor (_,c,p,_)) = if p then c else countNegate c
@@ -14304,21 +13746,21 @@
     let
       fun countAnd (fm,c) = countAnd2 (count fm, c)
     in
-      Set.foldl countAnd countTrue
+      Metis_Set.foldl countAnd countTrue
     end;
 
 val countOrSet =
     let
       fun countOr (fm,c) = countOr2 (count fm, c)
     in
-      Set.foldl countOr countFalse
+      Metis_Set.foldl countOr countFalse
     end;
 
 val countXorSet =
     let
       fun countXor (fm,c) = countXor2 (count fm, c)
     in
-      Set.foldl countXor countFalse
+      Metis_Set.foldl countXor countFalse
     end;
 
 fun And2 (False,_) = False
@@ -14340,14 +13782,14 @@
       if not (negateDisjoint s1 s2) then False
       else
         let
-          val s = Set.union s1 s2
-        in
-          case Set.size s of
+          val s = Metis_Set.union s1 s2
+        in
+          case Metis_Set.size s of
             0 => True
-          | 1 => Set.pick s
+          | 1 => Metis_Set.pick s
           | n =>
-            if n = Set.size s1 + Set.size s2 then
-              And (NameSet.union fv1 fv2, countAnd2 (c1,c2), s)
+            if n = Metis_Set.size s1 + Metis_Set.size s2 then
+              And (Metis_NameSet.union fv1 fv2, countAnd2 (c1,c2), s)
             else
               And (freeVarsSet s, countAndSet s, s)
         end
@@ -14355,7 +13797,7 @@
 
 val AndList = List.foldl And2 True;
 
-val AndSet = Set.foldl And2 True;
+val AndSet = Metis_Set.foldl And2 True;
 
 fun Or2 (True,_) = True
   | Or2 (_,True) = True
@@ -14376,14 +13818,14 @@
       if not (negateDisjoint s1 s2) then True
       else
         let
-          val s = Set.union s1 s2
-        in
-          case Set.size s of
+          val s = Metis_Set.union s1 s2
+        in
+          case Metis_Set.size s of
             0 => False
-          | 1 => Set.pick s
+          | 1 => Metis_Set.pick s
           | n =>
-            if n = Set.size s1 + Set.size s2 then
-              Or (NameSet.union fv1 fv2, countOr2 (c1,c2), s)
+            if n = Metis_Set.size s1 + Metis_Set.size s2 then
+              Or (Metis_NameSet.union fv1 fv2, countOr2 (c1,c2), s)
             else
               Or (freeVarsSet s, countOrSet s, s)
         end
@@ -14391,7 +13833,7 @@
 
 val OrList = List.foldl Or2 False;
 
-val OrSet = Set.foldl Or2 False;
+val OrSet = Metis_Set.foldl Or2 False;
 
 fun pushOr2 (f1,f2) =
     let
@@ -14400,9 +13842,9 @@
 
       fun g x1 (x2,acc) = And2 (Or2 (x1,x2), acc)
 
-      fun f (x1,acc) = Set.foldl (g x1) acc s2
-    in
-      Set.foldl f True s1
+      fun f (x1,acc) = Metis_Set.foldl (g x1) acc s2
+    in
+      Metis_Set.foldl f True s1
     end;
 
 val pushOrList = List.foldl pushOr2 False;
@@ -14425,15 +13867,15 @@
         val (fv1,c1,p1,s1) = case f1 of Xor x => x | _ => normalize f1
         and (fv2,c2,p2,s2) = case f2 of Xor x => x | _ => normalize f2
 
-        val s = Set.symmetricDifference s1 s2
+        val s = Metis_Set.symmetricDifference s1 s2
 
         val fm =
-            case Set.size s of
+            case Metis_Set.size s of
               0 => False
-            | 1 => Set.pick s
+            | 1 => Metis_Set.pick s
             | n =>
-              if n = Set.size s1 + Set.size s2 then
-                Xor (NameSet.union fv1 fv2, countXor2 (c1,c2), true, s)
+              if n = Metis_Set.size s1 + Metis_Set.size s2 then
+                Xor (Metis_NameSet.union fv1 fv2, countXor2 (c1,c2), true, s)
               else
                 Xor (freeVarsSet s, countXorSet s, true, s)
 
@@ -14445,7 +13887,7 @@
 
 val XorList = List.foldl Xor2 False;
 
-val XorSet = Set.foldl Xor2 False;
+val XorSet = Metis_Set.foldl Xor2 False;
 
 fun XorPolarityList (p,l) = applyPolarity p (XorList l);
 
@@ -14453,9 +13895,9 @@
 
 fun destXor (Xor (_,_,p,s)) =
     let
-      val (fm1,s) = Set.deletePick s
+      val (fm1,s) = Metis_Set.deletePick s
       val fm2 =
-          if Set.size s = 1 then applyPolarity p (Set.pick s)
+          if Metis_Set.size s = 1 then applyPolarity p (Metis_Set.pick s)
           else Xor (freeVarsSet s, countXorSet s, p, s)
     in
       (fm1,fm2)
@@ -14475,31 +13917,31 @@
     let
       fun exists_gen fm =
           let
-            val fv = NameSet.delete (freeVars fm) v
+            val fv = Metis_NameSet.delete (freeVars fm) v
             val c = count fm
-            val n = NameSet.singleton v
+            val n = Metis_NameSet.singleton v
           in
             Exists (fv,c,n,fm)
           end
 
       fun exists fm = if freeIn v fm then exists_free fm else fm
 
-      and exists_free (Or (_,_,s)) = OrList (Set.transform exists s)
+      and exists_free (Or (_,_,s)) = OrList (Metis_Set.transform exists s)
         | exists_free (fm as And (_,_,s)) =
           let
-            val sv = Set.filter (freeIn v) s
-          in
-            if Set.size sv <> 1 then exists_gen fm
+            val sv = Metis_Set.filter (freeIn v) s
+          in
+            if Metis_Set.size sv <> 1 then exists_gen fm
             else
               let
-                val fm = Set.pick sv
-                val s = Set.delete s fm
+                val fm = Metis_Set.pick sv
+                val s = Metis_Set.delete s fm
               in
                 And2 (exists_free fm, AndSet s)
               end
           end
         | exists_free (Exists (fv,c,n,f)) =
-          Exists (NameSet.delete fv v, c, NameSet.add n v, f)
+          Exists (Metis_NameSet.delete fv v, c, Metis_NameSet.add n v, f)
         | exists_free fm = exists_gen fm
     in
       exists init_fm
@@ -14507,37 +13949,37 @@
 
 fun ExistsList (vs,f) = List.foldl Exists1 f vs;
 
-fun ExistsSet (n,f) = NameSet.foldl Exists1 f n;
+fun ExistsSet (n,f) = Metis_NameSet.foldl Exists1 f n;
 
 fun Forall1 (v,init_fm) =
     let
       fun forall_gen fm =
           let
-            val fv = NameSet.delete (freeVars fm) v
+            val fv = Metis_NameSet.delete (freeVars fm) v
             val c = count fm
-            val n = NameSet.singleton v
+            val n = Metis_NameSet.singleton v
           in
             Forall (fv,c,n,fm)
           end
 
       fun forall fm = if freeIn v fm then forall_free fm else fm
 
-      and forall_free (And (_,_,s)) = AndList (Set.transform forall s)
+      and forall_free (And (_,_,s)) = AndList (Metis_Set.transform forall s)
         | forall_free (fm as Or (_,_,s)) =
           let
-            val sv = Set.filter (freeIn v) s
-          in
-            if Set.size sv <> 1 then forall_gen fm
+            val sv = Metis_Set.filter (freeIn v) s
+          in
+            if Metis_Set.size sv <> 1 then forall_gen fm
             else
               let
-                val fm = Set.pick sv
-                val s = Set.delete s fm
+                val fm = Metis_Set.pick sv
+                val s = Metis_Set.delete s fm
               in
                 Or2 (forall_free fm, OrSet s)
               end
           end
         | forall_free (Forall (fv,c,n,f)) =
-          Forall (NameSet.delete fv v, c, NameSet.add n v, f)
+          Forall (Metis_NameSet.delete fv v, c, Metis_NameSet.add n v, f)
         | forall_free fm = forall_gen fm
     in
       forall init_fm
@@ -14545,54 +13987,54 @@
 
 fun ForallList (vs,f) = List.foldl Forall1 f vs;
 
-fun ForallSet (n,f) = NameSet.foldl Forall1 f n;
+fun ForallSet (n,f) = Metis_NameSet.foldl Forall1 f n;
 
 fun generalize f = ForallSet (freeVars f, f);
 
 local
   fun subst_fv fvSub =
       let
-        fun add_fv (v,s) = NameSet.union (NameMap.get fvSub v) s
-      in
-        NameSet.foldl add_fv NameSet.empty
+        fun add_fv (v,s) = Metis_NameSet.union (Metis_NameMap.get fvSub v) s
+      in
+        Metis_NameSet.foldl add_fv Metis_NameSet.empty
       end;
 
   fun subst_rename (v,(avoid,bv,sub,domain,fvSub)) =
       let
-        val v' = Term.variantPrime avoid v
-        val avoid = NameSet.add avoid v'
-        val bv = NameSet.add bv v'
-        val sub = Subst.insert sub (v, Term.Var v')
-        val domain = NameSet.add domain v
-        val fvSub = NameMap.insert fvSub (v, NameSet.singleton v')
+        val v' = Metis_Term.variantPrime avoid v
+        val avoid = Metis_NameSet.add avoid v'
+        val bv = Metis_NameSet.add bv v'
+        val sub = Metis_Subst.insert sub (v, Metis_Term.Var v')
+        val domain = Metis_NameSet.add domain v
+        val fvSub = Metis_NameMap.insert fvSub (v, Metis_NameSet.singleton v')
       in
         (avoid,bv,sub,domain,fvSub)
       end;
 
   fun subst_check sub domain fvSub fm =
       let
-        val domain = NameSet.intersect domain (freeVars fm)
-      in
-        if NameSet.null domain then fm
+        val domain = Metis_NameSet.intersect domain (freeVars fm)
+      in
+        if Metis_NameSet.null domain then fm
         else subst_domain sub domain fvSub fm
       end
 
   and subst_domain sub domain fvSub fm =
       case fm of
-        Literal (fv,lit) =>
-        let
-          val fv = NameSet.difference fv domain
-          val fv = NameSet.union fv (subst_fv fvSub domain)
-          val lit = Literal.subst sub lit
-        in
-          Literal (fv,lit)
+        Metis_Literal (fv,lit) =>
+        let
+          val fv = Metis_NameSet.difference fv domain
+          val fv = Metis_NameSet.union fv (subst_fv fvSub domain)
+          val lit = Metis_Literal.subst sub lit
+        in
+          Metis_Literal (fv,lit)
         end
       | And (_,_,s) =>
-        AndList (Set.transform (subst_check sub domain fvSub) s)
+        AndList (Metis_Set.transform (subst_check sub domain fvSub) s)
       | Or (_,_,s) =>
-        OrList (Set.transform (subst_check sub domain fvSub) s)
+        OrList (Metis_Set.transform (subst_check sub domain fvSub) s)
       | Xor (_,_,p,s) =>
-        XorPolarityList (p, Set.transform (subst_check sub domain fvSub) s)
+        XorPolarityList (p, Metis_Set.transform (subst_check sub domain fvSub) s)
       | Exists fv_c_n_f => subst_quant Exists sub domain fvSub fv_c_n_f
       | Forall fv_c_n_f => subst_quant Forall sub domain fvSub fv_c_n_f
       | _ => raise Bug "subst_domain"
@@ -14600,12 +14042,12 @@
   and subst_quant quant sub domain fvSub (fv,c,bv,fm) =
       let
         val sub_fv = subst_fv fvSub domain
-        val fv = NameSet.union sub_fv (NameSet.difference fv domain)
-        val captured = NameSet.intersect bv sub_fv
-        val bv = NameSet.difference bv captured
-        val avoid = NameSet.union fv bv
+        val fv = Metis_NameSet.union sub_fv (Metis_NameSet.difference fv domain)
+        val captured = Metis_NameSet.intersect bv sub_fv
+        val bv = Metis_NameSet.difference bv captured
+        val avoid = Metis_NameSet.union fv bv
         val (_,bv,sub,domain,fvSub) =
-            NameSet.foldl subst_rename (avoid,bv,sub,domain,fvSub) captured
+            Metis_NameSet.foldl subst_rename (avoid,bv,sub,domain,fvSub) captured
         val fm = subst_domain sub domain fvSub fm
       in
         quant (fv,c,bv,fm)
@@ -14614,10 +14056,10 @@
   fun subst sub =
       let
         fun mk_dom (v,tm,(d,fv)) =
-            (NameSet.add d v, NameMap.insert fv (v, Term.freeVars tm))
-
-        val domain_fvSub = (NameSet.empty, NameMap.new ())
-        val (domain,fvSub) = Subst.foldl mk_dom domain_fvSub sub
+            (Metis_NameSet.add d v, Metis_NameMap.insert fv (v, Metis_Term.freeVars tm))
+
+        val domain_fvSub = (Metis_NameSet.empty, Metis_NameMap.new ())
+        val (domain,fvSub) = Metis_Subst.foldl mk_dom domain_fvSub sub
       in
         subst_check sub domain fvSub
       end;
@@ -14625,33 +14067,33 @@
 
 fun fromFormula fm =
     case fm of
-      Formula.True => True
-    | Formula.False => False
-    | Formula.Atom atm => Literal (Atom.freeVars atm, (true,atm))
-    | Formula.Not p => negateFromFormula p
-    | Formula.And (p,q) => And2 (fromFormula p, fromFormula q)
-    | Formula.Or (p,q) => Or2 (fromFormula p, fromFormula q)
-    | Formula.Imp (p,q) => Or2 (negateFromFormula p, fromFormula q)
-    | Formula.Iff (p,q) => Xor2 (negateFromFormula p, fromFormula q)
-    | Formula.Forall (v,p) => Forall1 (v, fromFormula p)
-    | Formula.Exists (v,p) => Exists1 (v, fromFormula p)
+      Metis_Formula.True => True
+    | Metis_Formula.False => False
+    | Metis_Formula.Metis_Atom atm => Metis_Literal (Metis_Atom.freeVars atm, (true,atm))
+    | Metis_Formula.Not p => negateFromFormula p
+    | Metis_Formula.And (p,q) => And2 (fromFormula p, fromFormula q)
+    | Metis_Formula.Or (p,q) => Or2 (fromFormula p, fromFormula q)
+    | Metis_Formula.Imp (p,q) => Or2 (negateFromFormula p, fromFormula q)
+    | Metis_Formula.Iff (p,q) => Xor2 (negateFromFormula p, fromFormula q)
+    | Metis_Formula.Forall (v,p) => Forall1 (v, fromFormula p)
+    | Metis_Formula.Exists (v,p) => Exists1 (v, fromFormula p)
 
 and negateFromFormula fm =
     case fm of
-      Formula.True => False
-    | Formula.False => True
-    | Formula.Atom atm => Literal (Atom.freeVars atm, (false,atm))
-    | Formula.Not p => fromFormula p
-    | Formula.And (p,q) => Or2 (negateFromFormula p, negateFromFormula q)
-    | Formula.Or (p,q) => And2 (negateFromFormula p, negateFromFormula q)
-    | Formula.Imp (p,q) => And2 (fromFormula p, negateFromFormula q)
-    | Formula.Iff (p,q) => Xor2 (fromFormula p, fromFormula q)
-    | Formula.Forall (v,p) => Exists1 (v, negateFromFormula p)
-    | Formula.Exists (v,p) => Forall1 (v, negateFromFormula p);
-
-local
-  fun lastElt (s : formula Set.set) =
-      case Set.findr (K true) s of
+      Metis_Formula.True => False
+    | Metis_Formula.False => True
+    | Metis_Formula.Metis_Atom atm => Metis_Literal (Metis_Atom.freeVars atm, (false,atm))
+    | Metis_Formula.Not p => fromFormula p
+    | Metis_Formula.And (p,q) => Or2 (negateFromFormula p, negateFromFormula q)
+    | Metis_Formula.Or (p,q) => And2 (negateFromFormula p, negateFromFormula q)
+    | Metis_Formula.Imp (p,q) => And2 (fromFormula p, negateFromFormula q)
+    | Metis_Formula.Iff (p,q) => Xor2 (fromFormula p, fromFormula q)
+    | Metis_Formula.Forall (v,p) => Exists1 (v, negateFromFormula p)
+    | Metis_Formula.Exists (v,p) => Forall1 (v, negateFromFormula p);
+
+local
+  fun lastElt (s : formula Metis_Set.set) =
+      case Metis_Set.findr (K true) s of
         NONE => raise Bug "lastElt: empty set"
       | SOME fm => fm;
 
@@ -14659,42 +14101,42 @@
       let
         val fm = lastElt s
       in
-        Set.add (Set.delete s fm) (negate fm)
+        Metis_Set.add (Metis_Set.delete s fm) (negate fm)
       end;
 
   fun form fm =
       case fm of
-        True => Formula.True
-      | False => Formula.False
-      | Literal (_,lit) => Literal.toFormula lit
-      | And (_,_,s) => Formula.listMkConj (Set.transform form s)
-      | Or (_,_,s) => Formula.listMkDisj (Set.transform form s)
+        True => Metis_Formula.True
+      | False => Metis_Formula.False
+      | Metis_Literal (_,lit) => Metis_Literal.toFormula lit
+      | And (_,_,s) => Metis_Formula.listMkConj (Metis_Set.transform form s)
+      | Or (_,_,s) => Metis_Formula.listMkDisj (Metis_Set.transform form s)
       | Xor (_,_,p,s) =>
         let
           val s = if p then negateLastElt s else s
         in
-          Formula.listMkEquiv (Set.transform form s)
+          Metis_Formula.listMkEquiv (Metis_Set.transform form s)
         end
-      | Exists (_,_,n,f) => Formula.listMkExists (NameSet.toList n, form f)
-      | Forall (_,_,n,f) => Formula.listMkForall (NameSet.toList n, form f);
+      | Exists (_,_,n,f) => Metis_Formula.listMkExists (Metis_NameSet.toList n, form f)
+      | Forall (_,_,n,f) => Metis_Formula.listMkForall (Metis_NameSet.toList n, form f);
 in
   val toFormula = form;
 end;
 
-fun toLiteral (Literal (_,lit)) = lit
-  | toLiteral _ = raise Error "Normalize.toLiteral";
-
-local
-  fun addLiteral (l,s) = LiteralSet.add s (toLiteral l);
-in
-  fun toClause False = LiteralSet.empty
-    | toClause (Or (_,_,s)) = Set.foldl addLiteral LiteralSet.empty s
-    | toClause l = LiteralSet.singleton (toLiteral l);
-end;
-
-val pp = Print.ppMap toFormula Formula.pp;
-
-val toString = Print.toString pp;
+fun toLiteral (Metis_Literal (_,lit)) = lit
+  | toLiteral _ = raise Error "Metis_Normalize.toLiteral";
+
+local
+  fun addLiteral (l,s) = Metis_LiteralSet.add s (toLiteral l);
+in
+  fun toClause False = Metis_LiteralSet.empty
+    | toClause (Or (_,_,s)) = Metis_Set.foldl addLiteral Metis_LiteralSet.empty s
+    | toClause l = Metis_LiteralSet.singleton (toLiteral l);
+end;
+
+val pp = Metis_Print.ppMap toFormula Metis_Formula.pp;
+
+val toString = Metis_Print.toString pp;
 
 (* ------------------------------------------------------------------------- *)
 (* Negation normal form.                                                     *)
@@ -14708,48 +14150,49 @@
 
 val newSkolemFunction =
     let
-      val counter : int StringMap.map Unsynchronized.ref = Unsynchronized.ref (StringMap.new ())
-    in
-      fn n =>
+      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 = Name.toString n
-           val i = Option.getOpt (StringMap.peek m s, 0)
-           val () = counter := StringMap.insert m (s, i + 1)
+           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
-           Name.fromString s
-         end
+           Metis_Name.fromString s
+         end)
     end;
 
 fun skolemize fv bv fm =
     let
-      val fv = NameSet.transform Term.Var fv
-
-      fun mk (v,s) = Subst.insert s (v, Term.Fn (newSkolemFunction v, fv))
-    in
-      subst (NameSet.foldl mk Subst.empty bv) fm
+      val fv = Metis_NameSet.transform Metis_Term.Var fv
+
+      fun mk (v,s) = Metis_Subst.insert s (v, Metis_Term.Fn (newSkolemFunction v, fv))
+    in
+      subst (Metis_NameSet.foldl mk Metis_Subst.empty bv) fm
     end;
 
 local
   fun rename avoid fv bv fm =
       let
-        val captured = NameSet.intersect avoid bv
-      in
-        if NameSet.null captured then fm
+        val captured = Metis_NameSet.intersect avoid bv
+      in
+        if Metis_NameSet.null captured then fm
         else
           let
             fun ren (v,(a,s)) =
                 let
-                  val v' = Term.variantPrime a v
+                  val v' = Metis_Term.variantPrime a v
                 in
-                  (NameSet.add a v', Subst.insert s (v, Term.Var v'))
+                  (Metis_NameSet.add a v', Metis_Subst.insert s (v, Metis_Term.Var v'))
                 end
 
-            val avoid = NameSet.union (NameSet.union avoid fv) bv
-
-            val (_,sub) = NameSet.foldl ren (avoid,Subst.empty) captured
+            val avoid = Metis_NameSet.union (Metis_NameSet.union avoid fv) bv
+
+            val (_,sub) = Metis_NameSet.foldl ren (avoid,Metis_Subst.empty) captured
           in
             subst sub fm
           end
@@ -14759,8 +14202,8 @@
 (*MetisTrace5
       let
         val fm' = cnfFm' avoid fm
-        val () = Print.trace pp "Normalize.cnfFm: fm" fm
-        val () = Print.trace pp "Normalize.cnfFm: fm'" fm'
+        val () = Metis_Print.trace pp "Metis_Normalize.cnfFm: fm" fm
+        val () = Metis_Print.trace pp "Metis_Normalize.cnfFm: fm'" fm'
       in
         fm'
       end
@@ -14769,12 +14212,12 @@
       case fm of
         True => True
       | False => False
-      | Literal _ => fm
-      | And (_,_,s) => AndList (Set.transform (cnfFm avoid) s)
+      | Metis_Literal _ => fm
+      | And (_,_,s) => AndList (Metis_Set.transform (cnfFm avoid) s)
       | Or (fv,_,s) =>
         let
-          val avoid = NameSet.union avoid fv
-          val (fms,_) = Set.foldl cnfOr ([],avoid) s
+          val avoid = Metis_NameSet.union avoid fv
+          val (fms,_) = Metis_Set.foldl cnfOr ([],avoid) s
         in
           pushOrList fms
         end
@@ -14786,12 +14229,12 @@
       let
         val fm = cnfFm avoid fm
         val fms = fm :: fms
-        val avoid = NameSet.union avoid (freeVars fm)
+        val avoid = Metis_NameSet.union avoid (freeVars fm)
       in
         (fms,avoid)
       end;
 in
-  val basicCnf = cnfFm NameSet.empty;
+  val basicCnf = cnfFm Metis_NameSet.empty;
 end;
 
 (* ------------------------------------------------------------------------- *)
@@ -14805,7 +14248,7 @@
       case fm of
         True => best
       | False => best
-      | Literal _ => best
+      | Metis_Literal _ => best
       | And (_,_,s) =>
         minBreakSet countClauses countAnd2 countTrue AndSet s best
       | Or (_,_,s) =>
@@ -14822,7 +14265,7 @@
               fun fwd (fm,(c1,s1,l)) =
                   let
                     val c1' = count2 (count fm, c1)
-                    and s1' = Set.add s1 fm
+                    and s1' = Metis_Set.add s1 fm
                   in
                     (c1', s1', (c1,s1,fm) :: l)
                   end
@@ -14830,7 +14273,7 @@
               fun bwd ((c1,s1,fm),(c2,s2,l)) =
                   let
                     val c2' = count2 (count fm, c2)
-                    and s2' = Set.add s2 fm
+                    and s2' = Metis_Set.add s2 fm
                   in
                     (c2', s2', (c1,s1,fm,c2,s2) :: l)
                   end
@@ -14859,7 +14302,7 @@
         val breakSet1 =
             let
               fun break c1 s1 fm c2 (best as (bcl,_)) =
-                  if Set.null s1 then best
+                  if Metis_Set.null s1 then best
                   else
                     let
                       val cDef = countDefinition (countXor2 (c1, count fm))
@@ -14868,14 +14311,14 @@
                       val noBetter = countLeqish bcl cl
                     in
                       if noBetter then best
-                      else (cl, SOME (mkSet (Set.add s1 fm)))
+                      else (cl, SOME (mkSet (Metis_Set.add s1 fm)))
                     end
             in
               fn ((c1,s1,fm,c2,s2),best) =>
                  break c1 s1 fm c2 (break c2 s2 fm c1 best)
             end
 
-        val fms = Set.toList fmSet
+        val fms = Metis_Set.toList fmSet
 
         fun breakSet measure best =
             let
@@ -14905,7 +14348,7 @@
                 case def of
                   NONE => ()
                 | SOME d =>
-                  Print.trace pp ("defCNF: before = " ^ countToString cl ^
+                  Metis_Print.trace pp ("defCNF: before = " ^ countToString cl ^
                                   ", after = " ^ countToString cl' ^
                                   ", definition") d
 *)
@@ -14919,11 +14362,11 @@
 (* Conjunctive normal form derivations.                                      *)
 (* ------------------------------------------------------------------------- *)
 
-datatype thm = Thm of formula * inference
+datatype thm = Metis_Thm of formula * inference
 
 and inference =
-    Axiom of Formula.formula
-  | Definition of string * Formula.formula
+    Axiom of Metis_Formula.formula
+  | Definition of string * Metis_Formula.formula
   | Simplify of thm * thm list
   | Conjunct of thm
   | Specialize of thm
@@ -14940,25 +14383,25 @@
     | Skolemize th => [th]
     | Clausify th => [th];
 
-fun compareThm (Thm (fm1,_), Thm (fm2,_)) = compare (fm1,fm2);
-
-fun parentsThm (Thm (_,inf)) = parentsInference inf;
-
-fun mkAxiom fm = Thm (fromFormula fm, Axiom fm);
-
-fun destThm (Thm (fm,inf)) = (toFormula fm, inf);
-
-local
-  val emptyProved : (thm,Formula.formula) Map.map = Map.new compareThm;
-
-  fun isProved proved th = Map.inDomain th proved;
+fun compareThm (Metis_Thm (fm1,_), Metis_Thm (fm2,_)) = compare (fm1,fm2);
+
+fun parentsThm (Metis_Thm (_,inf)) = parentsInference inf;
+
+fun mkAxiom fm = Metis_Thm (fromFormula fm, Axiom fm);
+
+fun destThm (Metis_Thm (fm,inf)) = (toFormula fm, inf);
+
+local
+  val emptyProved : (thm,Metis_Formula.formula) Metis_Map.map = Metis_Map.new compareThm;
+
+  fun isProved proved th = Metis_Map.inDomain th proved;
 
   fun isUnproved proved th = not (isProved proved th);
 
   fun lookupProved proved th =
-      case Map.peek proved th of
+      case Metis_Map.peek proved th of
         SOME fm => fm
-      | NONE => raise Bug "Normalize.lookupProved";
+      | NONE => raise Bug "Metis_Normalize.lookupProved";
 
   fun prove acc proved ths =
       case ths of
@@ -14979,7 +14422,7 @@
 
                 val acc = (fm,inf,fms) :: acc
 
-                val proved = Map.insert proved (th,fm)
+                val proved = Metis_Map.insert proved (th,fm)
               in
                 prove acc proved ths'
               end
@@ -15004,7 +14447,7 @@
     | Skolemize _ => "Skolemize"
     | Clausify _ => "Clausify";
 
-val ppInference = Print.ppMap toStringInference Print.ppString;
+val ppInference = Metis_Print.ppMap toStringInference Metis_Print.ppString;
 
 (* ------------------------------------------------------------------------- *)
 (* Simplifying with definitions.                                             *)
@@ -15012,30 +14455,30 @@
 
 datatype simplify =
     Simp of
-      {formula : (formula, formula * thm) Map.map,
-       andSet : (formula Set.set * formula * thm) list,
-       orSet : (formula Set.set * formula * thm) list,
-       xorSet : (formula Set.set * formula * thm) list};
+      {formula : (formula, formula * thm) Metis_Map.map,
+       andSet : (formula Metis_Set.set * formula * thm) list,
+       orSet : (formula Metis_Set.set * formula * thm) list,
+       xorSet : (formula Metis_Set.set * formula * thm) list};
 
 val simplifyEmpty =
     Simp
-      {formula = Map.new compare,
+      {formula = Metis_Map.new compare,
        andSet = [],
        orSet = [],
        xorSet = []};
 
 local
   fun simpler fm s =
-      Set.size s <> 1 orelse
-      case Set.pick s of
+      Metis_Set.size s <> 1 orelse
+      case Metis_Set.pick s of
         True => false
       | False => false
-      | Literal _ => false
+      | Metis_Literal _ => false
       | _ => true;
 
   fun addSet set_defs body_def =
       let
-        fun def_body_size (body,_,_) = Set.size body
+        fun def_body_size (body,_,_) = Metis_Set.size body
 
         val body_size = def_body_size body_def
 
@@ -15046,7 +14489,7 @@
             case Int.compare (def_body_size bd, body_size) of
               LESS => List.revAppend (acc, body_def :: l)
             | EQUAL =>
-              if Set.equal b body then List.revAppend (acc,l)
+              if Metis_Set.equal b body then List.revAppend (acc,l)
               else add (bd :: acc) bds
             | GREATER => add (bd :: acc) bds
       in
@@ -15084,25 +14527,25 @@
         case def of
           True =>
           let
-            fun addXorLiteral (fm as Literal _, simp) =
+            fun addXorLiteral (fm as Metis_Literal _, simp) =
                 let
-                  val s = Set.delete s fm
+                  val s = Metis_Set.delete s fm
                 in
                   if not (simpler fm s) then simp
                   else addXorSet simp (s, applyPolarity (not p) fm, th)
                 end
               | addXorLiteral (_,simp) = simp
           in
-            Set.foldl addXorLiteral simp s
+            Metis_Set.foldl addXorLiteral simp s
           end
         | _ => simp
       end
     | add (simp as Simp {formula,andSet,orSet,xorSet}) (body,def,th) =
-      if Map.inDomain body formula then simp
+      if Metis_Map.inDomain body formula then simp
       else
         let
-          val formula = Map.insert formula (body,(def,th))
-          val formula = Map.insert formula (negate body, (negate def, th))
+          val formula = Metis_Map.insert formula (body,(def,th))
+          val formula = Metis_Map.insert formula (negate body, (negate def, th))
         in
           Simp
             {formula = formula,
@@ -15112,7 +14555,7 @@
         end
 
   and addXorSet (simp as Simp {formula,andSet,orSet,xorSet}) (s,def,th) =
-      if Set.size s = 1 then add simp (Set.pick s, def, th)
+      if Metis_Set.size s = 1 then add simp (Metis_Set.pick s, def, th)
       else
         let
           val xorSet = addSet xorSet (s,def,th)
@@ -15124,19 +14567,19 @@
              xorSet = xorSet}
         end;
 in
-  fun simplifyAdd simp (th as Thm (fm,_)) = add simp (fm,True,th);
+  fun simplifyAdd simp (th as Metis_Thm (fm,_)) = add simp (fm,True,th);
 end;
 
 local
   fun simplifySet set_defs set =
       let
-        fun pred (s,_,_) = Set.subset s set
+        fun pred (s,_,_) = Metis_Set.subset s set
       in
         case List.find pred set_defs of
           NONE => NONE
         | SOME (s,f,th) =>
           let
-            val set = Set.add (Set.difference set s) f
+            val set = Metis_Set.add (Metis_Set.difference set s) f
           in
             SOME (set,th)
           end
@@ -15187,7 +14630,7 @@
                    try_simp_top fm inf
                  end)
             | _ =>
-              (case Map.peek formula fm of
+              (case Metis_Map.peek formula fm of
                  NONE => NONE
                | SOME (fm,th) =>
                  let
@@ -15222,7 +14665,7 @@
 
         and simp_set s inf =
             let
-              val (changed,l,inf) = Set.foldr simp_set_elt (false,[],inf) s
+              val (changed,l,inf) = Metis_Set.foldr simp_set_elt (false,[],inf) s
             in
               if changed then SOME (l,inf) else NONE
             end
@@ -15232,25 +14675,25 @@
               NONE => (changed, fm :: l, inf)
             | SOME (fm,inf) => (true, fm :: l, inf)
       in
-        fn th as Thm (fm,_) =>
+        fn th as Metis_Thm (fm,_) =>
            case simp fm [] of
              SOME (fm,ths) =>
              let
                val inf = Simplify (th,ths)
              in
-               Thm (fm,inf)
+               Metis_Thm (fm,inf)
              end
            | NONE => th
       end;
 end;
 
 (*MetisTrace2
-val simplify = fn simp => fn th as Thm (fm,_) =>
-    let
-      val th' as Thm (fm',_) = simplify simp th
+val simplify = fn simp => fn th as Metis_Thm (fm,_) =>
+    let
+      val th' as Metis_Thm (fm',_) = simplify simp th
       val () = if compare (fm,fm') = EQUAL then ()
-               else (Print.trace pp "Normalize.simplify: fm" fm;
-                     Print.trace pp "Normalize.simplify: fm'" fm')
+               else (Metis_Print.trace pp "Metis_Normalize.simplify: fm" fm;
+                     Metis_Print.trace pp "Metis_Normalize.simplify: fm'" fm')
     in
       th'
     end;
@@ -15264,27 +14707,28 @@
     let
       val counter : int Unsynchronized.ref = Unsynchronized.ref 0
     in
-      fn () =>
+      (* MODIFIED by Jasmin Blanchette *)
+      fn () => CRITICAL (fn () =>
          let
            val Unsynchronized.ref i = counter
            val () = counter := i + 1
          in
            definitionPrefix ^ "_" ^ Int.toString i
-         end
+         end)
     end;
 
 fun newDefinition def =
     let
       val fv = freeVars def
       val rel = newDefinitionRelation ()
-      val atm = (Name.fromString rel, NameSet.transform Term.Var fv)
-      val fm = Formula.Iff (Formula.Atom atm, toFormula def)
-      val fm = Formula.setMkForall (fv,fm)
+      val atm = (Metis_Name.fromString rel, Metis_NameSet.transform Metis_Term.Var fv)
+      val fm = Metis_Formula.Iff (Metis_Formula.Metis_Atom atm, toFormula def)
+      val fm = Metis_Formula.setMkForall (fv,fm)
       val inf = Definition (rel,fm)
-      val lit = Literal (fv,(false,atm))
+      val lit = Metis_Literal (fv,(false,atm))
       val fm = Xor2 (lit,def)
     in
-      Thm (fm,inf)
+      Metis_Thm (fm,inf)
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -15300,7 +14744,7 @@
 local
   fun def_cnf_inconsistent th =
       let
-        val cls = [(LiteralSet.empty,th)]
+        val cls = [(Metis_LiteralSet.empty,th)]
       in
         (cls,InconsistentCnf)
       end;
@@ -15308,14 +14752,14 @@
   fun def_cnf_clause inf (fm,acc) =
       let
         val cl = toClause fm
-        val th = Thm (fm,inf)
+        val th = Metis_Thm (fm,inf)
       in
         (cl,th) :: acc
       end
 (*MetisDebug
       handle Error err =>
-        (Print.trace pp "Normalize.addCnf.def_cnf_clause: fm" fm;
-         raise Bug ("Normalize.addCnf.def_cnf_clause: " ^ err));
+        (Metis_Print.trace pp "Metis_Normalize.addCnf.def_cnf_clause: fm" fm;
+         raise Bug ("Metis_Normalize.addCnf.def_cnf_clause: " ^ err));
 *)
 
   fun def_cnf cls simp ths =
@@ -15323,25 +14767,25 @@
         [] => (cls, ConsistentCnf simp)
       | th :: ths => def_cnf_formula cls simp (simplify simp th) ths
 
-  and def_cnf_formula cls simp (th as Thm (fm,_)) ths =
+  and def_cnf_formula cls simp (th as Metis_Thm (fm,_)) ths =
       case fm of
         True => def_cnf cls simp ths
       | False => def_cnf_inconsistent th
       | And (_,_,s) =>
         let
-          fun add (f,z) = Thm (f, Conjunct th) :: z
-        in
-          def_cnf cls simp (Set.foldr add ths s)
+          fun add (f,z) = Metis_Thm (f, Conjunct th) :: z
+        in
+          def_cnf cls simp (Metis_Set.foldr add ths s)
         end
       | Exists (fv,_,n,f) =>
         let
-          val th = Thm (skolemize fv n f, Skolemize th)
+          val th = Metis_Thm (skolemize fv n f, Skolemize th)
         in
           def_cnf_formula cls simp th ths
         end
       | Forall (_,_,_,f) =>
         let
-          val th = Thm (f, Specialize th)
+          val th = Metis_Thm (f, Specialize th)
         in
           def_cnf_formula cls simp th ths
         end
@@ -15364,11 +14808,11 @@
           in
             case fm of
               True => def_cnf cls simp ths
-            | False => def_cnf_inconsistent (Thm (fm,inf))
+            | False => def_cnf_inconsistent (Metis_Thm (fm,inf))
             | And (_,_,s) =>
               let
-                val inf = Conjunct (Thm (fm,inf))
-                val cls = Set.foldl (def_cnf_clause inf) cls s
+                val inf = Conjunct (Metis_Thm (fm,inf))
+                val cls = Metis_Set.foldl (def_cnf_clause inf) cls s
               in
                 def_cnf cls simp ths
               end
@@ -15405,7 +14849,6 @@
     end;
 
 end
-end;
 
 (**** Original file: Model.sig ****)
 
@@ -15414,11 +14857,11 @@
 (* Copyright (c) 2003 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-signature Model =
+signature Metis_Model =
 sig
 
 (* ------------------------------------------------------------------------- *)
-(* Model size.                                                               *)
+(* Metis_Model size.                                                               *)
 (* ------------------------------------------------------------------------- *)
 
 type size = {size : int}
@@ -15443,20 +14886,20 @@
 
 datatype fixed =
     Fixed of
-      {functions : fixedFunction Metis.NameArityMap.map,
-       relations : fixedRelation Metis.NameArityMap.map}
+      {functions : fixedFunction Metis_NameArityMap.map,
+       relations : fixedRelation Metis_NameArityMap.map}
 
 val emptyFixed : fixed
 
 val unionFixed : fixed -> fixed -> fixed
 
-val getFunctionFixed : fixed -> Metis.NameArity.nameArity -> fixedFunction
-
-val getRelationFixed : fixed -> Metis.NameArity.nameArity -> fixedRelation
-
-val insertFunctionFixed : fixed -> Metis.NameArity.nameArity * fixedFunction -> fixed
-
-val insertRelationFixed : fixed -> Metis.NameArity.nameArity * fixedRelation -> fixed
+val getFunctionFixed : fixed -> Metis_NameArity.nameArity -> fixedFunction
+
+val getRelationFixed : fixed -> Metis_NameArity.nameArity -> fixedRelation
+
+val insertFunctionFixed : fixed -> Metis_NameArity.nameArity * fixedFunction -> fixed
+
+val insertRelationFixed : fixed -> Metis_NameArity.nameArity * fixedRelation -> fixed
 
 val unionListFixed : fixed list -> fixed
 
@@ -15467,12 +14910,12 @@
 (* ------------------------------------------------------------------------- *)
 
 type fixedMap =
-     {functionMap : Metis.Name.name Metis.NameArityMap.map,
-      relationMap : Metis.Name.name Metis.NameArityMap.map}
+     {functionMap : Metis_Name.name Metis_NameArityMap.map,
+      relationMap : Metis_Name.name Metis_NameArityMap.map}
 
 val mapFixed : fixedMap -> fixed -> fixed
 
-val ppFixedMap : fixedMap Metis.Print.pp
+val ppFixedMap : fixedMap Metis_Print.pp
 
 (* ------------------------------------------------------------------------- *)
 (* Standard fixed model parts.                                               *)
@@ -15484,7 +14927,7 @@
 
 val projectionMax : int
 
-val projectionName : int -> Metis.Name.name
+val projectionName : int -> Metis_Name.name
 
 val projectionFixed : fixed
 
@@ -15494,41 +14937,41 @@
 
 val numeralMax : int
 
-val numeralName : int -> Metis.Name.name
-
-val addName : Metis.Name.name
-
-val divName : Metis.Name.name
-
-val dividesName : Metis.Name.name
-
-val evenName : Metis.Name.name
-
-val expName : Metis.Name.name
-
-val geName : Metis.Name.name
-
-val gtName : Metis.Name.name
-
-val isZeroName : Metis.Name.name
-
-val leName : Metis.Name.name
-
-val ltName : Metis.Name.name
-
-val modName : Metis.Name.name
-
-val multName : Metis.Name.name
-
-val negName : Metis.Name.name
-
-val oddName : Metis.Name.name
-
-val preName : Metis.Name.name
-
-val subName : Metis.Name.name
-
-val sucName : Metis.Name.name
+val numeralName : int -> Metis_Name.name
+
+val addName : Metis_Name.name
+
+val divName : Metis_Name.name
+
+val dividesName : Metis_Name.name
+
+val evenName : Metis_Name.name
+
+val expName : Metis_Name.name
+
+val geName : Metis_Name.name
+
+val gtName : Metis_Name.name
+
+val isZeroName : Metis_Name.name
+
+val leName : Metis_Name.name
+
+val ltName : Metis_Name.name
+
+val modName : Metis_Name.name
+
+val multName : Metis_Name.name
+
+val negName : Metis_Name.name
+
+val oddName : Metis_Name.name
+
+val preName : Metis_Name.name
+
+val subName : Metis_Name.name
+
+val sucName : Metis_Name.name
 
 val modularFixed : fixed
 
@@ -15536,45 +14979,45 @@
 
 (* Sets *)
 
-val cardName : Metis.Name.name
-
-val complementName : Metis.Name.name
-
-val differenceName : Metis.Name.name
-
-val emptyName : Metis.Name.name
-
-val memberName : Metis.Name.name
-
-val insertName : Metis.Name.name
-
-val intersectName : Metis.Name.name
-
-val singletonName : Metis.Name.name
-
-val subsetName : Metis.Name.name
-
-val symmetricDifferenceName : Metis.Name.name
-
-val unionName : Metis.Name.name
-
-val universeName : Metis.Name.name
+val cardName : Metis_Name.name
+
+val complementName : Metis_Name.name
+
+val differenceName : Metis_Name.name
+
+val emptyName : Metis_Name.name
+
+val memberName : Metis_Name.name
+
+val insertName : Metis_Name.name
+
+val intersectName : Metis_Name.name
+
+val singletonName : Metis_Name.name
+
+val subsetName : Metis_Name.name
+
+val symmetricDifferenceName : Metis_Name.name
+
+val unionName : Metis_Name.name
+
+val universeName : Metis_Name.name
 
 val setFixed : fixed
 
 (* Lists *)
 
-val appendName : Metis.Name.name
-
-val consName : Metis.Name.name
-
-val lengthName : Metis.Name.name
-
-val nilName : Metis.Name.name
-
-val nullName : Metis.Name.name
-
-val tailName : Metis.Name.name
+val appendName : Metis_Name.name
+
+val consName : Metis_Name.name
+
+val lengthName : Metis_Name.name
+
+val nilName : Metis_Name.name
+
+val nullName : Metis_Name.name
+
+val tailName : Metis_Name.name
 
 val listFixed : fixed
 
@@ -15586,23 +15029,23 @@
 
 val emptyValuation : valuation
 
-val zeroValuation : Metis.NameSet.set -> valuation
-
-val constantValuation : element -> Metis.NameSet.set -> valuation
-
-val peekValuation : valuation -> Metis.Name.name -> element option
-
-val getValuation : valuation -> Metis.Name.name -> element
-
-val insertValuation : valuation -> Metis.Name.name * element -> valuation
-
-val randomValuation : {size : int} -> Metis.NameSet.set -> valuation
+val zeroValuation : Metis_NameSet.set -> valuation
+
+val constantValuation : element -> Metis_NameSet.set -> valuation
+
+val peekValuation : valuation -> Metis_Name.name -> element option
+
+val getValuation : valuation -> Metis_Name.name -> element
+
+val insertValuation : valuation -> Metis_Name.name * element -> valuation
+
+val randomValuation : {size : int} -> Metis_NameSet.set -> valuation
 
 val incrementValuation :
-    {size : int} -> Metis.NameSet.set -> valuation -> valuation option
+    {size : int} -> Metis_NameSet.set -> valuation -> valuation option
 
 val foldValuation :
-    {size : int} -> Metis.NameSet.set -> (valuation * 'a -> 'a) -> 'a -> 'a
+    {size : int} -> Metis_NameSet.set -> (valuation * 'a -> 'a) -> 'a -> 'a
 
 (* ------------------------------------------------------------------------- *)
 (* A type of random finite models.                                           *)
@@ -15622,19 +15065,19 @@
 (* Interpreting terms and formulas in the model.                             *)
 (* ------------------------------------------------------------------------- *)
 
-val interpretFunction : model -> Metis.Term.functionName * element list -> element
-
-val interpretRelation : model -> Metis.Atom.relationName * element list -> bool
-
-val interpretTerm : model -> valuation -> Metis.Term.term -> element
-
-val interpretAtom : model -> valuation -> Metis.Atom.atom -> bool
-
-val interpretFormula : model -> valuation -> Metis.Formula.formula -> bool
-
-val interpretLiteral : model -> valuation -> Metis.Literal.literal -> bool
-
-val interpretClause : model -> valuation -> Metis.Thm.clause -> bool
+val interpretFunction : model -> Metis_Term.functionName * element list -> element
+
+val interpretRelation : model -> Metis_Atom.relationName * element list -> bool
+
+val interpretTerm : model -> valuation -> Metis_Term.term -> element
+
+val interpretAtom : model -> valuation -> Metis_Atom.atom -> bool
+
+val interpretFormula : model -> valuation -> Metis_Formula.formula -> bool
+
+val interpretLiteral : model -> valuation -> Metis_Literal.literal -> bool
+
+val interpretClause : model -> valuation -> Metis_Thm.clause -> bool
 
 (* ------------------------------------------------------------------------- *)
 (* Check whether random groundings of a formula are true in the model.       *)
@@ -15643,70 +15086,61 @@
 
 val check :
     (model -> valuation -> 'a -> bool) -> {maxChecks : int option} -> model ->
-    Metis.NameSet.set -> 'a -> {T : int, F : int}
+    Metis_NameSet.set -> 'a -> {T : int, F : int}
 
 val checkAtom :
-    {maxChecks : int option} -> model -> Metis.Atom.atom -> {T : int, F : int}
+    {maxChecks : int option} -> model -> Metis_Atom.atom -> {T : int, F : int}
 
 val checkFormula :
-    {maxChecks : int option} -> model -> Metis.Formula.formula -> {T : int, F : int}
+    {maxChecks : int option} -> model -> Metis_Formula.formula -> {T : int, F : int}
 
 val checkLiteral :
-    {maxChecks : int option} -> model -> Metis.Literal.literal -> {T : int, F : int}
+    {maxChecks : int option} -> model -> Metis_Literal.literal -> {T : int, F : int}
 
 val checkClause :
-    {maxChecks : int option} -> model -> Metis.Thm.clause -> {T : int, F : int}
+    {maxChecks : int option} -> model -> Metis_Thm.clause -> {T : int, F : int}
 
 (* ------------------------------------------------------------------------- *)
 (* Updating the model.                                                       *)
 (* ------------------------------------------------------------------------- *)
 
 val updateFunction :
-    model -> (Metis.Term.functionName * element list) * element -> unit
+    model -> (Metis_Term.functionName * element list) * element -> unit
 
 val updateRelation :
-    model -> (Metis.Atom.relationName * element list) * bool -> unit
+    model -> (Metis_Atom.relationName * element list) * bool -> unit
 
 (* ------------------------------------------------------------------------- *)
 (* Choosing a random perturbation to make a formula true in the model.       *)
 (* ------------------------------------------------------------------------- *)
 
-val perturbTerm : model -> valuation -> Metis.Term.term * element list -> unit
-
-val perturbAtom : model -> valuation -> Metis.Atom.atom * bool -> unit
-
-val perturbLiteral : model -> valuation -> Metis.Literal.literal -> unit
-
-val perturbClause : model -> valuation -> Metis.Thm.clause -> unit
+val perturbTerm : model -> valuation -> Metis_Term.term * element list -> unit
+
+val perturbAtom : model -> valuation -> Metis_Atom.atom * bool -> unit
+
+val perturbLiteral : model -> valuation -> Metis_Literal.literal -> unit
+
+val perturbClause : model -> valuation -> Metis_Thm.clause -> unit
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty printing.                                                          *)
 (* ------------------------------------------------------------------------- *)
 
-val pp : model Metis.Print.pp
+val pp : model Metis_Print.pp
 
 end
 
 (**** Original file: Model.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* RANDOM FINITE MODELS                                                      *)
 (* Copyright (c) 2003 Joe Hurd, distributed under the BSD License            *)
 (* ========================================================================= *)
 
-structure Model :> Model =
+structure Metis_Model :> Metis_Model =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Constants.                                                                *)
@@ -15763,12 +15197,12 @@
 
 fun intToBool 1 = true
   | intToBool 0 = false
-  | intToBool _ = raise Bug "Model.intToBool";
+  | intToBool _ = raise Bug "Metis_Model.intToBool";
 
 fun minMaxInterval i j = interval i (1 + j - i);
 
 (* ------------------------------------------------------------------------- *)
-(* Model size.                                                               *)
+(* Metis_Model size.                                                               *)
 (* ------------------------------------------------------------------------- *)
 
 type size = {size : int};
@@ -15813,31 +15247,31 @@
 
 datatype fixed =
     Fixed of
-      {functions : fixedFunction NameArityMap.map,
-       relations : fixedRelation NameArityMap.map};
+      {functions : fixedFunction Metis_NameArityMap.map,
+       relations : fixedRelation Metis_NameArityMap.map};
 
 val uselessFixedFunction : fixedFunction = K (K NONE);
 
 val uselessFixedRelation : fixedRelation = K (K NONE);
 
-val emptyFunctions : fixedFunction NameArityMap.map = NameArityMap.new ();
-
-val emptyRelations : fixedRelation NameArityMap.map = NameArityMap.new ();
+val emptyFunctions : fixedFunction Metis_NameArityMap.map = Metis_NameArityMap.new ();
+
+val emptyRelations : fixedRelation Metis_NameArityMap.map = Metis_NameArityMap.new ();
 
 fun fixed0 f sz elts =
     case elts of
       [] => f sz
-    | _ => raise Bug "Model.fixed0: wrong arity";
+    | _ => raise Bug "Metis_Model.fixed0: wrong arity";
 
 fun fixed1 f sz elts =
     case elts of
       [x] => f sz x
-    | _ => raise Bug "Model.fixed1: wrong arity";
+    | _ => raise Bug "Metis_Model.fixed1: wrong arity";
 
 fun fixed2 f sz elts =
     case elts of
       [x,y] => f sz x y
-    | _ => raise Bug "Model.fixed2: wrong arity";
+    | _ => raise Bug "Metis_Model.fixed2: wrong arity";
 
 val emptyFixed =
     let
@@ -15853,14 +15287,14 @@
     let
       val Fixed {functions = fns, ...} = fix
     in
-      NameArityMap.peek fns name_arity
+      Metis_NameArityMap.peek fns name_arity
     end;
 
 fun peekRelationFixed fix name_arity =
     let
       val Fixed {relations = rels, ...} = fix
     in
-      NameArityMap.peek rels name_arity
+      Metis_NameArityMap.peek rels name_arity
     end;
 
 fun getFunctionFixed fix name_arity =
@@ -15877,7 +15311,7 @@
     let
       val Fixed {functions = fns, relations = rels} = fix
 
-      val fns = NameArityMap.insert fns name_arity_fn
+      val fns = Metis_NameArityMap.insert fns name_arity_fn
     in
       Fixed
         {functions = fns,
@@ -15888,7 +15322,7 @@
     let
       val Fixed {functions = fns, relations = rels} = fix
 
-      val rels = NameArityMap.insert rels name_arity_rel
+      val rels = Metis_NameArityMap.insert rels name_arity_rel
     in
       Fixed
         {functions = fns,
@@ -15896,16 +15330,16 @@
     end;
 
 local
-  fun union _ = raise Bug "Model.unionFixed: nameArity clash";
+  fun union _ = raise Bug "Metis_Model.unionFixed: nameArity clash";
 in
   fun unionFixed fix1 fix2 =
       let
         val Fixed {functions = fns1, relations = rels1} = fix1
         and Fixed {functions = fns2, relations = rels2} = fix2
 
-        val fns = NameArityMap.union union fns1 fns2
-
-        val rels = NameArityMap.union union rels1 rels2
+        val fns = Metis_NameArityMap.union union fns1 fns2
+
+        val rels = Metis_NameArityMap.union union rels1 rels2
       in
         Fixed
           {functions = fns,
@@ -15924,18 +15358,18 @@
   fun hasTypeFn _ elts =
       case elts of
         [x,_] => SOME x
-      | _ => raise Bug "Model.hasTypeFn: wrong arity";
+      | _ => raise Bug "Metis_Model.hasTypeFn: wrong arity";
 
   fun eqRel _ elts =
       case elts of
         [x,y] => SOME (x = y)
-      | _ => raise Bug "Model.eqRel: wrong arity";
+      | _ => raise Bug "Metis_Model.eqRel: wrong arity";
 in
   val basicFixed =
       let
-        val fns = NameArityMap.singleton (Term.hasTypeFunction,hasTypeFn)
-
-        val rels = NameArityMap.singleton (Atom.eqRelation,eqRel)
+        val fns = Metis_NameArityMap.singleton (Metis_Term.hasTypeFunction,hasTypeFn)
+
+        val rels = Metis_NameArityMap.singleton (Metis_Atom.eqRelation,eqRel)
       in
         Fixed
           {functions = fns,
@@ -15948,17 +15382,17 @@
 (* ------------------------------------------------------------------------- *)
 
 type fixedMap =
-     {functionMap : Name.name NameArityMap.map,
-      relationMap : Name.name NameArityMap.map};
+     {functionMap : Metis_Name.name Metis_NameArityMap.map,
+      relationMap : Metis_Name.name Metis_NameArityMap.map};
 
 fun mapFixed fixMap fix =
     let
       val {functionMap = fnMap, relationMap = relMap} = fixMap
       and Fixed {functions = fns, relations = rels} = fix
 
-      val fns = NameArityMap.compose fnMap fns
-
-      val rels = NameArityMap.compose relMap rels
+      val fns = Metis_NameArityMap.compose fnMap fns
+
+      val rels = Metis_NameArityMap.compose relMap rels
     in
       Fixed
         {functions = fns,
@@ -15968,27 +15402,27 @@
 local
   fun mkEntry tag (na,n) = (tag,na,n);
 
-  fun mkList tag m = map (mkEntry tag) (NameArityMap.toList m);
+  fun mkList tag m = map (mkEntry tag) (Metis_NameArityMap.toList m);
 
   fun ppEntry (tag,source_arity,target) =
-      Print.blockProgram Print.Inconsistent 2
-        [Print.addString tag,
-         Print.addBreak 1,
-         NameArity.pp source_arity,
-         Print.addString " ->",
-         Print.addBreak 1,
-         Name.pp target];
+      Metis_Print.blockProgram Metis_Print.Inconsistent 2
+        [Metis_Print.addString tag,
+         Metis_Print.addBreak 1,
+         Metis_NameArity.pp source_arity,
+         Metis_Print.addString " ->",
+         Metis_Print.addBreak 1,
+         Metis_Name.pp target];
 in
   fun ppFixedMap fixMap =
       let
         val {functionMap = fnMap, relationMap = relMap} = fixMap
       in
         case mkList "function" fnMap @ mkList "relation" relMap of
-          [] => Print.skip
+          [] => Metis_Print.skip
         | entry :: entries =>
-          Print.blockProgram Print.Consistent 0
+          Metis_Print.blockProgram Metis_Print.Consistent 0
             (ppEntry entry ::
-             map (Print.sequence Print.addNewline o ppEntry) entries)
+             map (Metis_Print.sequence Metis_Print.addNewline o ppEntry) entries)
       end;
 end;
 
@@ -16006,12 +15440,12 @@
 fun projectionName i =
     let
       val _ = projectionMin <= i orelse
-              raise Bug "Model.projectionName: less than projectionMin"
+              raise Bug "Metis_Model.projectionName: less than projectionMin"
 
       val _ = i <= projectionMax orelse
-              raise Bug "Model.projectionName: greater than projectionMax"
-    in
-      Name.fromString ("project" ^ Int.toString i)
+              raise Bug "Metis_Model.projectionName: greater than projectionMax"
+    in
+      Metis_Name.fromString ("project" ^ Int.toString i)
     end;
 
 fun projectionFn i _ elts = SOME (List.nth (elts, i - 1));
@@ -16022,7 +15456,7 @@
 
       fun addProj i acc =
           if i > arity then acc
-          else addProj (i + 1) (NameArityMap.insert acc (mkProj i))
+          else addProj (i + 1) (Metis_NameArityMap.insert acc (mkProj i))
 
       val fns = addProj projectionMin emptyFunctions
 
@@ -16046,33 +15480,33 @@
 fun numeralName i =
     let
       val _ = numeralMin <= i orelse
-              raise Bug "Model.numeralName: less than numeralMin"
+              raise Bug "Metis_Model.numeralName: less than numeralMin"
 
       val _ = i <= numeralMax orelse
-              raise Bug "Model.numeralName: greater than numeralMax"
+              raise Bug "Metis_Model.numeralName: greater than numeralMax"
 
       val s = if i < 0 then "negative" ^ Int.toString (~i) else Int.toString i
     in
-      Name.fromString s
-    end;
-
-val addName = Name.fromString "+"
-and divName = Name.fromString "div"
-and dividesName = Name.fromString "divides"
-and evenName = Name.fromString "even"
-and expName = Name.fromString "exp"
-and geName = Name.fromString ">="
-and gtName = Name.fromString ">"
-and isZeroName = Name.fromString "isZero"
-and leName = Name.fromString "<="
-and ltName = Name.fromString "<"
-and modName = Name.fromString "mod"
-and multName = Name.fromString "*"
-and negName = Name.fromString "~"
-and oddName = Name.fromString "odd"
-and preName = Name.fromString "pre"
-and subName = Name.fromString "-"
-and sucName = Name.fromString "suc";
+      Metis_Name.fromString s
+    end;
+
+val addName = Metis_Name.fromString "+"
+and divName = Metis_Name.fromString "div"
+and dividesName = Metis_Name.fromString "divides"
+and evenName = Metis_Name.fromString "even"
+and expName = Metis_Name.fromString "exp"
+and geName = Metis_Name.fromString ">="
+and gtName = Metis_Name.fromString ">"
+and isZeroName = Metis_Name.fromString "isZero"
+and leName = Metis_Name.fromString "<="
+and ltName = Metis_Name.fromString "<"
+and modName = Metis_Name.fromString "mod"
+and multName = Metis_Name.fromString "*"
+and negName = Metis_Name.fromString "~"
+and oddName = Metis_Name.fromString "odd"
+and preName = Metis_Name.fromString "pre"
+and subName = Metis_Name.fromString "-"
+and sucName = Metis_Name.fromString "suc";
 
 local
   (* Support *)
@@ -16136,7 +15570,7 @@
   val modularFixed =
       let
         val fns =
-            NameArityMap.fromList
+            Metis_NameArityMap.fromList
               (map (fn i => ((numeralName i,0), fixed0 (numeralFn i)))
                  numeralList @
                [((addName,2), fixed2 addFn),
@@ -16150,7 +15584,7 @@
                 ((sucName,1), fixed1 sucFn)])
 
         val rels =
-            NameArityMap.fromList
+            Metis_NameArityMap.fromList
               [((dividesName,2), fixed2 dividesRel),
                ((evenName,1), fixed1 evenRel),
                ((geName,2), fixed2 geRel),
@@ -16236,7 +15670,7 @@
   val overflowFixed =
       let
         val fns =
-            NameArityMap.fromList
+            Metis_NameArityMap.fromList
               (map (fn i => ((numeralName i,0), fixed0 (numeralFn i)))
                  numeralList @
                [((addName,2), fixed2 addFn),
@@ -16250,7 +15684,7 @@
                 ((sucName,1), fixed1 sucFn)])
 
         val rels =
-            NameArityMap.fromList
+            Metis_NameArityMap.fromList
               [((dividesName,2), fixed2 dividesRel),
                ((evenName,1), fixed1 evenRel),
                ((geName,2), fixed2 geRel),
@@ -16268,18 +15702,18 @@
 
 (* Sets *)
 
-val cardName = Name.fromString "card"
-and complementName = Name.fromString "complement"
-and differenceName = Name.fromString "difference"
-and emptyName = Name.fromString "empty"
-and memberName = Name.fromString "member"
-and insertName = Name.fromString "insert"
-and intersectName = Name.fromString "intersect"
-and singletonName = Name.fromString "singleton"
-and subsetName = Name.fromString "subset"
-and symmetricDifferenceName = Name.fromString "symmetricDifference"
-and unionName = Name.fromString "union"
-and universeName = Name.fromString "universe";
+val cardName = Metis_Name.fromString "card"
+and complementName = Metis_Name.fromString "complement"
+and differenceName = Metis_Name.fromString "difference"
+and emptyName = Metis_Name.fromString "empty"
+and memberName = Metis_Name.fromString "member"
+and insertName = Metis_Name.fromString "insert"
+and intersectName = Metis_Name.fromString "intersect"
+and singletonName = Metis_Name.fromString "singleton"
+and subsetName = Metis_Name.fromString "subset"
+and symmetricDifferenceName = Metis_Name.fromString "symmetricDifference"
+and unionName = Metis_Name.fromString "union"
+and universeName = Metis_Name.fromString "universe";
 
 local
   (* Support *)
@@ -16377,7 +15811,7 @@
   val setFixed =
       let
         val fns =
-            NameArityMap.fromList
+            Metis_NameArityMap.fromList
               [((cardName,1), fixed1 cardFn),
                ((complementName,1), fixed1 complementFn),
                ((differenceName,2), fixed2 differenceFn),
@@ -16390,7 +15824,7 @@
                ((universeName,0), fixed0 universeFn)]
 
         val rels =
-            NameArityMap.fromList
+            Metis_NameArityMap.fromList
               [((memberName,2), fixed2 memberRel),
                ((subsetName,2), fixed2 subsetRel)]
       in
@@ -16402,12 +15836,12 @@
 
 (* Lists *)
 
-val appendName = Name.fromString "@"
-and consName = Name.fromString "::"
-and lengthName = Name.fromString "length"
-and nilName = Name.fromString "nil"
-and nullName = Name.fromString "null"
-and tailName = Name.fromString "tail";
+val appendName = Metis_Name.fromString "@"
+and consName = Metis_Name.fromString "::"
+and lengthName = Metis_Name.fromString "length"
+and nilName = Metis_Name.fromString "nil"
+and nullName = Metis_Name.fromString "null"
+and tailName = Metis_Name.fromString "tail";
 
 local
   val baseFix =
@@ -16422,13 +15856,13 @@
       end;
 
   val fixMap =
-      {functionMap = NameArityMap.fromList
+      {functionMap = Metis_NameArityMap.fromList
                        [((appendName,2),addName),
                         ((consName,2),sucName),
                         ((lengthName,1), projectionName 1),
                         ((nilName,0), numeralName 0),
                         ((tailName,1),preName)],
-       relationMap = NameArityMap.fromList
+       relationMap = Metis_NameArityMap.fromList
                        [((nullName,1),isZeroName)]};
 
 in
@@ -16439,19 +15873,19 @@
 (* Valuations.                                                               *)
 (* ------------------------------------------------------------------------- *)
 
-datatype valuation = Valuation of element NameMap.map;
-
-val emptyValuation = Valuation (NameMap.new ());
-
-fun insertValuation (Valuation m) v_i = Valuation (NameMap.insert m v_i);
-
-fun peekValuation (Valuation m) v = NameMap.peek m v;
+datatype valuation = Valuation of element Metis_NameMap.map;
+
+val emptyValuation = Valuation (Metis_NameMap.new ());
+
+fun insertValuation (Valuation m) v_i = Valuation (Metis_NameMap.insert m v_i);
+
+fun peekValuation (Valuation m) v = Metis_NameMap.peek m v;
 
 fun constantValuation i =
     let
       fun add (v,V) = insertValuation V (v,i)
     in
-      NameSet.foldl add emptyValuation
+      Metis_NameSet.foldl add emptyValuation
     end;
 
 val zeroValuation = constantValuation zeroElement;
@@ -16459,13 +15893,13 @@
 fun getValuation V v =
     case peekValuation V v of
       SOME i => i
-    | NONE => raise Error "Model.getValuation: incomplete valuation";
+    | NONE => raise Error "Metis_Model.getValuation: incomplete valuation";
 
 fun randomValuation {size = N} vs =
     let
-      fun f (v,V) = insertValuation V (v, Portable.randomInt N)
-    in
-      NameSet.foldl f emptyValuation vs
+      fun f (v,V) = insertValuation V (v, Metis_Portable.randomInt N)
+    in
+      Metis_NameSet.foldl f emptyValuation vs
     end;
 
 fun incrementValuation N vars =
@@ -16485,7 +15919,7 @@
               if carry then inc vs V else SOME V
             end
     in
-      inc (NameSet.toList vars)
+      inc (Metis_NameSet.toList vars)
     end;
 
 fun foldValuation N vars f =
@@ -16522,7 +15956,7 @@
     | SOME space => ArrayTable (Array.array (space,UNKNOWN));
 
 local
-  fun randomResult R = Portable.randomInt R;
+  fun randomResult R = Metis_Portable.randomInt R;
 in
   fun lookupTable N R table elts =
       case table of
@@ -16565,13 +15999,13 @@
     Tables of
       {domainSize : int,
        rangeSize : int,
-       tableMap : table NameArityMap.map Unsynchronized.ref};
+       tableMap : table Metis_NameArityMap.map Unsynchronized.ref};
 
 fun newTables N R =
     Tables
       {domainSize = N,
        rangeSize = R,
-       tableMap = Unsynchronized.ref (NameArityMap.new ())};
+       tableMap = Unsynchronized.ref (Metis_NameArityMap.new ())};
 
 fun getTables tables n_a =
     let
@@ -16579,7 +16013,7 @@
 
       val Unsynchronized.ref m = tm
     in
-      case NameArityMap.peek m n_a of
+      case Metis_NameArityMap.peek m n_a of
         SOME t => t
       | NONE =>
         let
@@ -16587,7 +16021,7 @@
 
           val t = newTable N a
 
-          val m = NameArityMap.insert m (n_a,t)
+          val m = Metis_NameArityMap.insert m (n_a,t)
 
           val () = tm := m
         in
@@ -16624,10 +16058,10 @@
 type parameters = {size : int, fixed : fixed};
 
 datatype model =
-    Model of
+    Metis_Model of
       {size : int,
-       fixedFunctions : (element list -> element option) NameArityMap.map,
-       fixedRelations : (element list -> bool option) NameArityMap.map,
+       fixedFunctions : (element list -> element option) Metis_NameArityMap.map,
+       fixedRelations : (element list -> bool option) Metis_NameArityMap.map,
        randomFunctions : tables,
        randomRelations : tables};
 
@@ -16635,13 +16069,13 @@
     let
       val Fixed {functions = fns, relations = rels} = fixed
 
-      val fixFns = NameArityMap.transform (fn f => f {size = N}) fns
-      and fixRels = NameArityMap.transform (fn r => r {size = N}) rels
+      val fixFns = Metis_NameArityMap.transform (fn f => f {size = N}) fns
+      and fixRels = Metis_NameArityMap.transform (fn r => r {size = N}) rels
 
       val rndFns = newTables N N
       and rndRels = newTables N 2
     in
-      Model
+      Metis_Model
         {size = N,
          fixedFunctions = fixFns,
          fixedRelations = fixRels,
@@ -16649,13 +16083,13 @@
          randomRelations = rndRels}
     end;
 
-fun size (Model {size = N, ...}) = N;
+fun size (Metis_Model {size = N, ...}) = N;
 
 fun peekFixedFunction M (n,elts) =
     let
-      val Model {fixedFunctions = fixFns, ...} = M
-    in
-      case NameArityMap.peek fixFns (n, length elts) of
+      val Metis_Model {fixedFunctions = fixFns, ...} = M
+    in
+      case Metis_NameArityMap.peek fixFns (n, length elts) of
         NONE => NONE
       | SOME fixFn => fixFn elts
     end;
@@ -16664,9 +16098,9 @@
 
 fun peekFixedRelation M (n,elts) =
     let
-      val Model {fixedRelations = fixRels, ...} = M
-    in
-      case NameArityMap.peek fixRels (n, length elts) of
+      val Metis_Model {fixedRelations = fixRels, ...} = M
+    in
+      case Metis_NameArityMap.peek fixRels (n, length elts) of
         NONE => NONE
       | SOME fixRel => fixRel elts
     end;
@@ -16693,12 +16127,12 @@
 
 fun destTerm tm =
     case tm of
-      Term.Var _ => tm
-    | Term.Fn f_tms =>
-      case Term.stripApp tm of
+      Metis_Term.Var _ => tm
+    | Metis_Term.Fn f_tms =>
+      case Metis_Term.stripApp tm of
         (_,[]) => tm
-      | (v as Term.Var _, tms) => Term.Fn (Term.appName, v :: tms)
-      | (Term.Fn (f,tms), tms') => Term.Fn (f, tms @ tms');
+      | (v as Metis_Term.Var _, tms) => Metis_Term.Fn (Metis_Term.appName, v :: tms)
+      | (Metis_Term.Fn (f,tms), tms') => Metis_Term.Fn (f, tms @ tms');
 
 (* ------------------------------------------------------------------------- *)
 (* Interpreting terms and formulas in the model.                             *)
@@ -16709,7 +16143,7 @@
       SOME r => r
     | NONE =>
       let
-        val Model {randomFunctions = rndFns, ...} = M
+        val Metis_Model {randomFunctions = rndFns, ...} = M
       in
         lookupTables rndFns n_elts
       end;
@@ -16719,7 +16153,7 @@
       SOME r => r
     | NONE =>
       let
-        val Model {randomRelations = rndRels, ...} = M
+        val Metis_Model {randomRelations = rndRels, ...} = M
       in
         intToBool (lookupTables rndRels n_elts)
       end;
@@ -16728,8 +16162,8 @@
     let
       fun interpret tm =
           case destTerm tm of
-            Term.Var v => getValuation V v
-          | Term.Fn (f,tms) => interpretFunction M (f, map interpret tms)
+            Metis_Term.Var v => getValuation V v
+          | Metis_Term.Fn (f,tms) => interpretFunction M (f, map interpret tms)
     in
       interpret
     end;
@@ -16743,17 +16177,17 @@
 
       fun interpret V fm =
           case fm of
-            Formula.True => true
-          | Formula.False => false
-          | Formula.Atom atm => interpretAtom M V atm
-          | Formula.Not p => not (interpret V p)
-          | Formula.Or (p,q) => interpret V p orelse interpret V q
-          | Formula.And (p,q) => interpret V p andalso interpret V q
-          | Formula.Imp (p,q) => interpret V (Formula.Or (Formula.Not p, q))
-          | Formula.Iff (p,q) => interpret V p = interpret V q
-          | Formula.Forall (v,p) => interpret' V p v N
-          | Formula.Exists (v,p) =>
-            interpret V (Formula.Not (Formula.Forall (v, Formula.Not p)))
+            Metis_Formula.True => true
+          | Metis_Formula.False => false
+          | Metis_Formula.Metis_Atom atm => interpretAtom M V atm
+          | Metis_Formula.Not p => not (interpret V p)
+          | Metis_Formula.Or (p,q) => interpret V p orelse interpret V q
+          | Metis_Formula.And (p,q) => interpret V p andalso interpret V q
+          | Metis_Formula.Imp (p,q) => interpret V (Metis_Formula.Or (Metis_Formula.Not p, q))
+          | Metis_Formula.Iff (p,q) => interpret V p = interpret V q
+          | Metis_Formula.Forall (v,p) => interpret' V p v N
+          | Metis_Formula.Exists (v,p) =>
+            interpret V (Metis_Formula.Not (Metis_Formula.Forall (v, Metis_Formula.Not p)))
 
       and interpret' V fm v i =
           i = 0 orelse
@@ -16774,7 +16208,7 @@
       if pol then b else not b
     end;
 
-fun interpretClause M V cl = LiteralSet.exists (interpretLiteral M V) cl;
+fun interpretClause M V cl = Metis_LiteralSet.exists (interpretLiteral M V) cl;
 
 (* ------------------------------------------------------------------------- *)
 (* Check whether random groundings of a formula are true in the model.       *)
@@ -16794,7 +16228,7 @@
           case maxChecks of
             NONE => maxChecks
           | SOME m =>
-            case expInt N (NameSet.size fv) of
+            case expInt N (Metis_NameSet.size fv) of
               SOME n => if n <= m then NONE else maxChecks
             | NONE => maxChecks
     in
@@ -16804,16 +16238,16 @@
     end;
 
 fun checkAtom maxChecks M atm =
-    check interpretAtom maxChecks M (Atom.freeVars atm) atm;
+    check interpretAtom maxChecks M (Metis_Atom.freeVars atm) atm;
 
 fun checkFormula maxChecks M fm =
-    check interpretFormula maxChecks M (Formula.freeVars fm) fm;
+    check interpretFormula maxChecks M (Metis_Formula.freeVars fm) fm;
 
 fun checkLiteral maxChecks M lit =
-    check interpretLiteral maxChecks M (Literal.freeVars lit) lit;
+    check interpretLiteral maxChecks M (Metis_Literal.freeVars lit) lit;
 
 fun checkClause maxChecks M cl =
-    check interpretClause maxChecks M (LiteralSet.freeVars cl) cl;
+    check interpretClause maxChecks M (Metis_LiteralSet.freeVars cl) cl;
 
 (* ------------------------------------------------------------------------- *)
 (* Updating the model.                                                       *)
@@ -16821,7 +16255,7 @@
 
 fun updateFunction M func_elts_elt =
     let
-      val Model {randomFunctions = rndFns, ...} = M
+      val Metis_Model {randomFunctions = rndFns, ...} = M
 
       val () = updateTables rndFns func_elts_elt
     in
@@ -16830,7 +16264,7 @@
 
 fun updateRelation M (rel_elts,pol) =
     let
-      val Model {randomRelations = rndRels, ...} = M
+      val Metis_Model {randomRelations = rndRels, ...} = M
 
       val () = updateTables rndRels (rel_elts, boolToInt pol)
     in
@@ -16843,14 +16277,14 @@
 
 datatype modelTerm =
     ModelVar
-  | ModelFn of Term.functionName * modelTerm list * int list;
+  | ModelFn of Metis_Term.functionName * modelTerm list * int list;
 
 fun modelTerm M V =
     let
       fun modelTm tm =
           case destTerm tm of
-            Term.Var v => (ModelVar, getValuation V v)
-          | Term.Fn (f,tms) =>
+            Metis_Term.Var v => (ModelVar, getValuation V v)
+          | Metis_Term.Fn (f,tms) =>
             let
               val (tms,xs) = unzip (map modelTm tms)
             in
@@ -16865,8 +16299,8 @@
 (* ------------------------------------------------------------------------- *)
 
 datatype perturbation =
-    FunctionPerturbation of (Term.functionName * element list) * element
-  | RelationPerturbation of (Atom.relationName * element list) * bool;
+    FunctionPerturbation of (Metis_Term.functionName * element list) * element
+  | RelationPerturbation of (Metis_Atom.relationName * element list) * bool;
 
 fun perturb M pert =
     case pert of
@@ -16926,7 +16360,7 @@
             in
               pert (x :: ys) tms xs acc
             end
-          | pert _ _ _ _ = raise Bug "Model.pertTerms.pert"
+          | pert _ _ _ _ = raise Bug "Metis_Model.pertTerms.pert"
       in
         pert []
       end;
@@ -16948,11 +16382,11 @@
 
   fun pertLiteral M V ((pol,atm),acc) = pertAtom M V pol atm acc;
 
-  fun pertClause M V cl acc = LiteralSet.foldl (pertLiteral M V) acc cl;
+  fun pertClause M V cl acc = Metis_LiteralSet.foldl (pertLiteral M V) acc cl;
 
   fun pickPerturb M perts =
       if null perts then ()
-      else perturb M (List.nth (perts, Portable.randomInt (length perts)));
+      else perturb M (List.nth (perts, Metis_Portable.randomInt (length perts)));
 in
   fun perturbTerm M V (tm,target) =
       pickPerturb M (pertTerm M target (fst (modelTerm M V tm)) []);
@@ -16970,13 +16404,12 @@
 (* ------------------------------------------------------------------------- *)
 
 fun pp M =
-    Print.program
-      [Print.addString "Model{",
-       Print.ppInt (size M),
-       Print.addString "}"];
-
-end
-end;
+    Metis_Print.program
+      [Metis_Print.addString "Metis_Model{",
+       Metis_Print.ppInt (size M),
+       Metis_Print.addString "}"];
+
+end
 
 (**** Original file: Problem.sig ****)
 
@@ -16985,7 +16418,7 @@
 (* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Problem =
+signature Metis_Problem =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -16993,21 +16426,21 @@
 (* ------------------------------------------------------------------------- *)
 
 type problem =
-     {axioms : Metis.Thm.clause list,
-      conjecture : Metis.Thm.clause list}
+     {axioms : Metis_Thm.clause list,
+      conjecture : Metis_Thm.clause list}
 
 val size : problem -> {clauses : int,
                        literals : int,
                        symbols : int,
                        typedSymbols : int}
 
-val freeVars : problem -> Metis.NameSet.set
-
-val toClauses : problem -> Metis.Thm.clause list
-
-val toFormula : problem -> Metis.Formula.formula
-
-val toGoal : problem -> Metis.Formula.formula
+val freeVars : problem -> Metis_NameSet.set
+
+val toClauses : problem -> Metis_Thm.clause list
+
+val toFormula : problem -> Metis_Formula.formula
+
+val toGoal : problem -> Metis_Formula.formula
 
 val toString : problem -> string
 
@@ -17046,42 +16479,33 @@
 
 (**** Original file: Problem.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* CNF PROBLEMS                                                              *)
 (* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Problem :> Problem =
+structure Metis_Problem :> Metis_Problem =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Problems.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
 type problem =
-     {axioms : Thm.clause list,
-      conjecture : Thm.clause list};
+     {axioms : Metis_Thm.clause list,
+      conjecture : Metis_Thm.clause list};
 
 fun toClauses {axioms,conjecture} = axioms @ conjecture;
 
 fun size prob =
     let
-      fun lits (cl,n) = n + LiteralSet.size cl
-
-      fun syms (cl,n) = n + LiteralSet.symbols cl
-
-      fun typedSyms (cl,n) = n + LiteralSet.typedSymbols cl
+      fun lits (cl,n) = n + Metis_LiteralSet.size cl
+
+      fun syms (cl,n) = n + Metis_LiteralSet.symbols cl
+
+      fun typedSyms (cl,n) = n + Metis_LiteralSet.typedSymbols cl
 
       val cls = toClauses prob
     in
@@ -17092,33 +16516,33 @@
     end;
 
 fun freeVars {axioms,conjecture} =
-    NameSet.union
-      (LiteralSet.freeVarsList axioms)
-      (LiteralSet.freeVarsList conjecture);
+    Metis_NameSet.union
+      (Metis_LiteralSet.freeVarsList axioms)
+      (Metis_LiteralSet.freeVarsList conjecture);
 
 local
   fun clauseToFormula cl =
-      Formula.listMkDisj (LiteralSet.transform Literal.toFormula cl);
+      Metis_Formula.listMkDisj (Metis_LiteralSet.transform Metis_Literal.toFormula cl);
 in
   fun toFormula prob =
-      Formula.listMkConj (map clauseToFormula (toClauses prob));
+      Metis_Formula.listMkConj (map clauseToFormula (toClauses prob));
 
   fun toGoal {axioms,conjecture} =
       let
-        val clToFm = Formula.generalize o clauseToFormula
-        val clsToFm = Formula.listMkConj o map clToFm
-
-        val fm = Formula.False
+        val clToFm = Metis_Formula.generalize o clauseToFormula
+        val clsToFm = Metis_Formula.listMkConj o map clToFm
+
+        val fm = Metis_Formula.False
         val fm =
             if null conjecture then fm
-            else Formula.Imp (clsToFm conjecture, fm)
-        val fm = Formula.Imp (clsToFm axioms, fm)
+            else Metis_Formula.Imp (clsToFm conjecture, fm)
+        val fm = Metis_Formula.Imp (clsToFm axioms, fm)
       in
         fm
       end;
 end;
 
-fun toString prob = Formula.toString (toFormula prob);
+fun toString prob = Metis_Formula.toString (toFormula prob);
 
 (* ------------------------------------------------------------------------- *)
 (* Categorizing problems.                                                    *)
@@ -17153,35 +16577,35 @@
 
       val rels =
           let
-            fun f (cl,set) = NameAritySet.union set (LiteralSet.relations cl)
-          in
-            List.foldl f NameAritySet.empty cls
+            fun f (cl,set) = Metis_NameAritySet.union set (Metis_LiteralSet.relations cl)
+          in
+            List.foldl f Metis_NameAritySet.empty cls
           end
 
       val funs =
           let
-            fun f (cl,set) = NameAritySet.union set (LiteralSet.functions cl)
-          in
-            List.foldl f NameAritySet.empty cls
+            fun f (cl,set) = Metis_NameAritySet.union set (Metis_LiteralSet.functions cl)
+          in
+            List.foldl f Metis_NameAritySet.empty cls
           end
 
       val propositional =
-          if NameAritySet.allNullary rels then Propositional
-          else if NameAritySet.allNullary funs then EffectivelyPropositional
+          if Metis_NameAritySet.allNullary rels then Propositional
+          else if Metis_NameAritySet.allNullary funs then EffectivelyPropositional
           else NonPropositional
 
       val equality =
-          if not (NameAritySet.member Atom.eqRelation rels) then NonEquality
-          else if NameAritySet.size rels = 1 then PureEquality
+          if not (Metis_NameAritySet.member Metis_Atom.eqRelation rels) then NonEquality
+          else if Metis_NameAritySet.size rels = 1 then PureEquality
           else Equality
 
       val horn =
-          if List.exists LiteralSet.null cls then Trivial
-          else if List.all (fn cl => LiteralSet.size cl = 1) cls then Unit
+          if List.exists Metis_LiteralSet.null cls then Trivial
+          else if List.all (fn cl => Metis_LiteralSet.size cl = 1) cls then Unit
           else 
             let
-              fun pos cl = LiteralSet.count Literal.positive cl <= 1
-              fun neg cl = LiteralSet.count Literal.negative cl <= 1
+              fun pos cl = Metis_LiteralSet.count Metis_Literal.positive cl <= 1
+              fun neg cl = Metis_LiteralSet.count Metis_Literal.negative cl <= 1
             in
               case (List.all pos cls, List.all neg cls) of
                 (true,true) => DoubleHorn
@@ -17215,7 +16639,6 @@
      | NonHorn => "non-horn");
 
 end
-end;
 
 (**** Original file: TermNet.sig ****)
 
@@ -17224,7 +16647,7 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature TermNet =
+signature Metis_TermNet =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -17245,15 +16668,15 @@
 
 val size : 'a termNet -> int
 
-val insert : 'a termNet -> Metis.Term.term * 'a -> 'a termNet
-
-val fromList : parameters -> (Metis.Term.term * 'a) list -> 'a termNet
+val insert : 'a termNet -> Metis_Term.term * 'a -> 'a termNet
+
+val fromList : parameters -> (Metis_Term.term * 'a) list -> 'a termNet
 
 val filter : ('a -> bool) -> 'a termNet -> 'a termNet
 
 val toString : 'a termNet -> string
 
-val pp : 'a Metis.Print.pp -> 'a termNet Metis.Print.pp
+val pp : 'a Metis_Print.pp -> 'a termNet Metis_Print.pp
 
 (* ------------------------------------------------------------------------- *)
 (* Matching and unification queries.                                         *)
@@ -17262,41 +16685,32 @@
 (* Filter afterwards to get the precise set of satisfying values.            *)
 (* ------------------------------------------------------------------------- *)
 
-val match : 'a termNet -> Metis.Term.term -> 'a list
-
-val matched : 'a termNet -> Metis.Term.term -> 'a list
-
-val unify : 'a termNet -> Metis.Term.term -> 'a list
+val match : 'a termNet -> Metis_Term.term -> 'a list
+
+val matched : 'a termNet -> Metis_Term.term -> 'a list
+
+val unify : 'a termNet -> Metis_Term.term -> 'a list
 
 end
 
 (**** Original file: TermNet.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS              *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure TermNet :> TermNet =
+structure Metis_TermNet :> Metis_TermNet =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Anonymous variables.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
-val anonymousName = Name.fromString "_";
-val anonymousVar = Term.Var anonymousName;
+val anonymousName = Metis_Name.fromString "_";
+val anonymousVar = Metis_Term.Var anonymousName;
 
 (* ------------------------------------------------------------------------- *)
 (* Quotient terms.                                                           *)
@@ -17304,12 +16718,12 @@
 
 datatype qterm =
     Var
-  | Fn of NameArity.nameArity * qterm list;
+  | Fn of Metis_NameArity.nameArity * qterm list;
 
 local
   fun cmp [] = EQUAL
     | cmp (q1_q2 :: qs) =
-      if Portable.pointerEqual q1_q2 then cmp qs
+      if Metis_Portable.pointerEqual q1_q2 then cmp qs
       else
         case q1_q2 of
           (Var,Var) => EQUAL
@@ -17318,7 +16732,7 @@
         | (Fn f1, Fn f2) => fnCmp f1 f2 qs
 
   and fnCmp (n1,q1) (n2,q2) qs =
-    case NameArity.compare (n1,n2) of
+    case Metis_NameArity.compare (n1,n2) of
       LESS => LESS
     | EQUAL => cmp (zip q1 q2 @ qs)
     | GREATER => GREATER;
@@ -17332,15 +16746,15 @@
 
 fun equalFnQterm f1 f2 = compareFnQterm (f1,f2) = EQUAL;
 
-fun termToQterm (Term.Var _) = Var
-  | termToQterm (Term.Fn (f,l)) = Fn ((f, length l), map termToQterm l);
+fun termToQterm (Metis_Term.Var _) = Var
+  | termToQterm (Metis_Term.Fn (f,l)) = Fn ((f, length l), map termToQterm l);
 
 local
   fun qm [] = true
     | qm ((Var,_) :: rest) = qm rest
     | qm ((Fn _, Var) :: _) = false
     | qm ((Fn (f,a), Fn (g,b)) :: rest) =
-      NameArity.equal f g andalso qm (zip a b @ rest);
+      Metis_NameArity.equal f g andalso qm (zip a b @ rest);
 in
   fun matchQtermQterm qtm qtm' = qm [(qtm,qtm')];
 end;
@@ -17348,22 +16762,22 @@
 local
   fun qm [] = true
     | qm ((Var,_) :: rest) = qm rest
-    | qm ((Fn _, Term.Var _) :: _) = false
-    | qm ((Fn ((f,n),a), Term.Fn (g,b)) :: rest) =
-      Name.equal f g andalso n = length b andalso qm (zip a b @ rest);
+    | qm ((Fn _, Metis_Term.Var _) :: _) = false
+    | qm ((Fn ((f,n),a), Metis_Term.Fn (g,b)) :: rest) =
+      Metis_Name.equal f g andalso n = length b andalso qm (zip a b @ rest);
 in
   fun matchQtermTerm qtm tm = qm [(qtm,tm)];
 end;
 
 local
   fun qn qsub [] = SOME qsub
-    | qn qsub ((Term.Var v, qtm) :: rest) =
-      (case NameMap.peek qsub v of
-         NONE => qn (NameMap.insert qsub (v,qtm)) rest
+    | qn qsub ((Metis_Term.Var v, qtm) :: rest) =
+      (case Metis_NameMap.peek qsub v of
+         NONE => qn (Metis_NameMap.insert qsub (v,qtm)) rest
        | SOME qtm' => if equalQterm qtm qtm' then qn qsub rest else NONE)
-    | qn _ ((Term.Fn _, Var) :: _) = NONE
-    | qn qsub ((Term.Fn (f,a), Fn ((g,n),b)) :: rest) =
-      if Name.equal f g andalso length a = n then qn qsub (zip a b @ rest)
+    | qn _ ((Metis_Term.Fn _, Var) :: _) = NONE
+    | qn qsub ((Metis_Term.Fn (f,a), Fn ((g,n),b)) :: rest) =
+      if Metis_Name.equal f g andalso length a = n then qn qsub (zip a b @ rest)
       else NONE;
 in
   fun matchTermQterm qsub tm qtm = qn qsub [(tm,qtm)];
@@ -17374,23 +16788,23 @@
     | qv x Var = x
     | qv (Fn (f,a)) (Fn (g,b)) =
       let
-        val _ = NameArity.equal f g orelse raise Error "TermNet.qv"
+        val _ = Metis_NameArity.equal f g orelse raise Error "Metis_TermNet.qv"
       in
         Fn (f, zipWith qv a b)
       end;
 
   fun qu qsub [] = qsub
     | qu qsub ((Var, _) :: rest) = qu qsub rest
-    | qu qsub ((qtm, Term.Var v) :: rest) =
+    | qu qsub ((qtm, Metis_Term.Var v) :: rest) =
       let
         val qtm =
-            case NameMap.peek qsub v of NONE => qtm | SOME qtm' => qv qtm qtm'
-      in
-        qu (NameMap.insert qsub (v,qtm)) rest
+            case Metis_NameMap.peek qsub v of NONE => qtm | SOME qtm' => qv qtm qtm'
+      in
+        qu (Metis_NameMap.insert qsub (v,qtm)) rest
       end
-    | qu qsub ((Fn ((f,n),a), Term.Fn (g,b)) :: rest) =
-      if Name.equal f g andalso n = length b then qu qsub (zip a b @ rest)
-      else raise Error "TermNet.qu";
+    | qu qsub ((Fn ((f,n),a), Metis_Term.Fn (g,b)) :: rest) =
+      if Metis_Name.equal f g andalso n = length b then qu qsub (zip a b @ rest)
+      else raise Error "Metis_TermNet.qu";
 in
   fun unifyQtermQterm qtm qtm' = total (qv qtm) qtm';
 
@@ -17399,9 +16813,9 @@
 
 local
   fun qtermToTerm Var = anonymousVar
-    | qtermToTerm (Fn ((f,_),l)) = Term.Fn (f, map qtermToTerm l);
-in
-  val ppQterm = Print.ppMap qtermToTerm Term.pp;
+    | qtermToTerm (Fn ((f,_),l)) = Metis_Term.Fn (f, map qtermToTerm l);
+in
+  val ppQterm = Metis_Print.ppMap qtermToTerm Metis_Term.pp;
 end;
 
 (* ------------------------------------------------------------------------- *)
@@ -17413,7 +16827,7 @@
 datatype 'a net =
     Result of 'a list
   | Single of qterm * 'a net
-  | Multiple of 'a net option * 'a net NameArityMap.map;
+  | Multiple of 'a net option * 'a net Metis_NameArityMap.map;
 
 datatype 'a termNet = Net of parameters * int * (int * (int * 'a) net) option;
 
@@ -17427,7 +16841,7 @@
   fun computeSize (Result l) = length l
     | computeSize (Single (_,n)) = computeSize n
     | computeSize (Multiple (vs,fs)) =
-      NameArityMap.foldl
+      Metis_NameArityMap.foldl
         (fn (_,n,acc) => acc + computeSize n)
         (case vs of SOME n => computeSize n | NONE => 0)
         fs;
@@ -17450,16 +16864,16 @@
   fun add (Result l) [] (Result l') = Result (l @ l')
     | add a (input1 as qtm :: qtms) (Single (qtm',n)) =
       if equalQterm qtm qtm' then Single (qtm, add a qtms n)
-      else add a input1 (add n [qtm'] (Multiple (NONE, NameArityMap.new ())))
+      else add a input1 (add n [qtm'] (Multiple (NONE, Metis_NameArityMap.new ())))
     | add a (Var :: qtms) (Multiple (vs,fs)) =
       Multiple (SOME (oadd a qtms vs), fs)
     | add a (Fn (f,l) :: qtms) (Multiple (vs,fs)) =
       let
-        val n = NameArityMap.peek fs f
-      in
-        Multiple (vs, NameArityMap.insert fs (f, oadd a (l @ qtms) n))
+        val n = Metis_NameArityMap.peek fs f
+      in
+        Multiple (vs, Metis_NameArityMap.insert fs (f, oadd a (l @ qtms) n))
       end
-    | add _ _ _ = raise Bug "TermNet.insert: Match"
+    | add _ _ _ = raise Bug "Metis_TermNet.insert: Match"
 
   and oadd a qtms NONE = singles qtms a
     | oadd a qtms (SOME n) = add a qtms n;
@@ -17468,7 +16882,7 @@
 in
   fun insert (Net (p,k,n)) (tm,a) =
       Net (p, k + 1, ins (k,a) (termToQterm tm) (pre n))
-      handle Error _ => raise Bug "TermNet.insert: should never fail";
+      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;
@@ -17487,18 +16901,18 @@
           let
             val vs = Option.mapPartial filt vs
 
-            val fs = NameArityMap.mapPartial (fn (_,n) => filt n) fs
-          in
-            if not (Option.isSome vs) andalso NameArityMap.null fs then NONE
+            val fs = Metis_NameArityMap.mapPartial (fn (_,n) => filt n) fs
+          in
+            if not (Option.isSome vs) andalso Metis_NameArityMap.null fs then NONE
             else SOME (Multiple (vs,fs))
           end
     in
       fn net as Net (_,_,NONE) => net
        | Net (p, k, SOME (_,n)) => Net (p, k, netSize (filt n))
     end
-    handle Error _ => raise Bug "TermNet.filter: should never fail";
-
-fun toString net = "TermNet[" ^ Int.toString (size net) ^ "]";
+    handle Error _ => raise Bug "Metis_TermNet.filter: should never fail";
+
+fun toString net = "Metis_TermNet[" ^ Int.toString (size net) ^ "]";
 
 (* ------------------------------------------------------------------------- *)
 (* Specialized fold operations to support matching and unification.          *)
@@ -17529,7 +16943,7 @@
   val stackAddFn = addFn;
 
   fun stackValue ([],[],[qtm]) = qtm
-    | stackValue _ = raise Bug "TermNet.stackValue";
+    | stackValue _ = raise Bug "Metis_TermNet.stackValue";
 end;
 
 local
@@ -17550,9 +16964,9 @@
         fun getFns (f as (_,k), net, x) =
             (k + n, stackAddFn f stack, net) :: x
       in
-        fold inc acc (NameArityMap.foldr getFns rest fns)
+        fold inc acc (Metis_NameArityMap.foldr getFns rest fns)
       end
-    | fold _ _ _ = raise Bug "TermNet.foldTerms.fold";
+    | fold _ _ _ = raise Bug "Metis_TermNet.foldTerms.fold";
 in
   fun foldTerms inc acc net = fold inc acc [(1,stackEmpty,net)];
 end;
@@ -17565,10 +16979,10 @@
         | fold (Var :: pats, Multiple (v,_)) =
           (case v of NONE => acc | SOME net => fold (pats,net))
         | fold (Fn (f,a) :: pats, Multiple (_,fns)) =
-          (case NameArityMap.peek fns f of
+          (case Metis_NameArityMap.peek fns f of
              NONE => acc
            | SOME net => fold (a @ pats, net))
-        | fold _ = raise Bug "TermNet.foldEqualTerms.fold";
+        | fold _ = raise Bug "Metis_TermNet.foldEqualTerms.fold";
     in
       fn net => fold ([pat],net)
     end;
@@ -17598,13 +17012,13 @@
             | SOME net => (pats, stackAddQterm pat stack, net) :: rest
 
         val rest =
-            case NameArityMap.peek fns f of
+            case Metis_NameArityMap.peek fns f of
               NONE => rest
             | SOME net => (a @ pats, stackAddFn f stack, net) :: rest
       in
         fold inc acc rest
       end
-    | fold _ _ _ = raise Bug "TermNet.foldUnifiableTerms.fold";
+    | fold _ _ _ = raise Bug "Metis_TermNet.foldUnifiableTerms.fold";
 in
   fun foldUnifiableTerms pat inc acc net =
       fold inc acc [([pat],stackEmpty,net)];
@@ -17636,25 +17050,25 @@
 
         val rest =
             case tm of
-              Term.Var _ => rest
-            | Term.Fn (f,l) =>
-              case NameArityMap.peek fs (f, length l) of
+              Metis_Term.Var _ => rest
+            | Metis_Term.Fn (f,l) =>
+              case Metis_NameArityMap.peek fs (f, length l) of
                 NONE => rest
               | SOME n => (n, l @ tms) :: rest
       in
         mat acc rest
       end
-    | mat _ _ = raise Bug "TermNet.match: Match";
+    | mat _ _ = raise Bug "Metis_TermNet.match: Match";
 in
   fun match (Net (_,_,NONE)) _ = []
     | match (Net (p, _, SOME (_,n))) tm =
       finally p (mat [] [(n,[tm])])
-      handle Error _ => raise Bug "TermNet.match: should never fail";
+      handle Error _ => raise Bug "Metis_TermNet.match: should never fail";
 end;
 
 local
   fun unseenInc qsub v tms (qtm,net,rest) =
-      (NameMap.insert qsub (v,qtm), net, tms) :: rest;
+      (Metis_NameMap.insert qsub (v,qtm), net, tms) :: rest;
 
   fun seenInc qsub tms (_,net,rest) = (qsub,net,tms) :: rest;
 
@@ -17664,30 +17078,30 @@
       (case matchTermQterm qsub tm qtm of
          NONE => mat acc rest
        | SOME qsub => mat acc ((qsub,net,tms) :: rest))
-    | mat acc ((qsub, net as Multiple _, Term.Var v :: tms) :: rest) =
-      (case NameMap.peek qsub v of
+    | mat acc ((qsub, net as Multiple _, Metis_Term.Var v :: tms) :: rest) =
+      (case Metis_NameMap.peek qsub v of
          NONE => mat acc (foldTerms (unseenInc qsub v tms) rest net)
        | SOME qtm => mat acc (foldEqualTerms qtm (seenInc qsub tms) rest net))
-    | mat acc ((qsub, Multiple (_,fns), Term.Fn (f,a) :: tms) :: rest) =
+    | mat acc ((qsub, Multiple (_,fns), Metis_Term.Fn (f,a) :: tms) :: rest) =
       let
         val rest =
-            case NameArityMap.peek fns (f, length a) of
+            case Metis_NameArityMap.peek fns (f, length a) of
               NONE => rest
             | SOME net => (qsub, net, a @ tms) :: rest
       in
         mat acc rest
       end
-    | mat _ _ = raise Bug "TermNet.matched.mat";
+    | mat _ _ = raise Bug "Metis_TermNet.matched.mat";
 in
   fun matched (Net (_,_,NONE)) _ = []
     | matched (Net (parm, _, SOME (_,net))) tm =
-      finally parm (mat [] [(NameMap.new (), net, [tm])])
-      handle Error _ => raise Bug "TermNet.matched: should never fail";
+      finally parm (mat [] [(Metis_NameMap.new (), net, [tm])])
+      handle Error _ => raise Bug "Metis_TermNet.matched: should never fail";
 end;
 
 local
   fun inc qsub v tms (qtm,net,rest) =
-      (NameMap.insert qsub (v,qtm), net, tms) :: rest;
+      (Metis_NameMap.insert qsub (v,qtm), net, tms) :: rest;
 
   fun mat acc [] = acc
     | mat acc ((_, Result l, []) :: rest) = mat (l @ acc) rest
@@ -17695,27 +17109,27 @@
       (case unifyQtermTerm qsub qtm tm of
          NONE => mat acc rest
        | SOME qsub => mat acc ((qsub,net,tms) :: rest))
-    | mat acc ((qsub, net as Multiple _, Term.Var v :: tms) :: rest) =
-      (case NameMap.peek qsub v of
+    | mat acc ((qsub, net as Multiple _, Metis_Term.Var v :: tms) :: rest) =
+      (case Metis_NameMap.peek qsub v of
          NONE => mat acc (foldTerms (inc qsub v tms) rest net)
        | SOME qtm => mat acc (foldUnifiableTerms qtm (inc qsub v tms) rest net))
-    | mat acc ((qsub, Multiple (v,fns), Term.Fn (f,a) :: tms) :: rest) =
+    | mat acc ((qsub, Multiple (v,fns), Metis_Term.Fn (f,a) :: tms) :: rest) =
       let
         val rest = case v of NONE => rest | SOME net => (qsub,net,tms) :: rest
 
         val rest =
-            case NameArityMap.peek fns (f, length a) of
+            case Metis_NameArityMap.peek fns (f, length a) of
               NONE => rest
             | SOME net => (qsub, net, a @ tms) :: rest
       in
         mat acc rest
       end
-    | mat _ _ = raise Bug "TermNet.unify.mat";
+    | mat _ _ = raise Bug "Metis_TermNet.unify.mat";
 in
   fun unify (Net (_,_,NONE)) _ = []
     | unify (Net (parm, _, SOME (_,net))) tm =
-      finally parm (mat [] [(NameMap.new (), net, [tm])])
-      handle Error _ => raise Bug "TermNet.unify: should never fail";
+      finally parm (mat [] [(Metis_NameMap.new (), net, [tm])])
+      handle Error _ => raise Bug "Metis_TermNet.unify: should never fail";
 end;
 
 (* ------------------------------------------------------------------------- *)
@@ -17725,18 +17139,17 @@
 local
   fun inc (qtm, Result l, acc) =
       foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l
-    | inc _ = raise Bug "TermNet.pp.inc";
+    | inc _ = raise Bug "Metis_TermNet.pp.inc";
 
   fun toList (Net (_,_,NONE)) = []
     | toList (Net (parm, _, SOME (_,net))) =
       finally parm (foldTerms inc [] net);
 in
   fun pp ppA =
-      Print.ppMap toList (Print.ppList (Print.ppOp2 " |->" ppQterm ppA));
-end;
-
-end
-end;
+      Metis_Print.ppMap toList (Metis_Print.ppList (Metis_Print.ppOp2 " |->" ppQterm ppA));
+end;
+
+end
 
 (**** Original file: AtomNet.sig ****)
 
@@ -17745,7 +17158,7 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature AtomNet =
+signature Metis_AtomNet =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -17764,15 +17177,15 @@
 
 val size : 'a atomNet -> int
 
-val insert : 'a atomNet -> Metis.Atom.atom * 'a -> 'a atomNet
-
-val fromList : parameters -> (Metis.Atom.atom * 'a) list -> 'a atomNet
+val insert : 'a atomNet -> Metis_Atom.atom * 'a -> 'a atomNet
+
+val fromList : parameters -> (Metis_Atom.atom * 'a) list -> 'a atomNet
 
 val filter : ('a -> bool) -> 'a atomNet -> 'a atomNet
 
 val toString : 'a atomNet -> string
 
-val pp : 'a Metis.Print.pp -> 'a atomNet Metis.Print.pp
+val pp : 'a Metis_Print.pp -> 'a atomNet Metis_Print.pp
 
 (* ------------------------------------------------------------------------- *)
 (* Matching and unification queries.                                         *)
@@ -17781,69 +17194,60 @@
 (* Filter afterwards to get the precise set of satisfying values.            *)
 (* ------------------------------------------------------------------------- *)
 
-val match : 'a atomNet -> Metis.Atom.atom -> 'a list
-
-val matched : 'a atomNet -> Metis.Atom.atom -> 'a list
-
-val unify : 'a atomNet -> Metis.Atom.atom -> 'a list
+val match : 'a atomNet -> Metis_Atom.atom -> 'a list
+
+val matched : 'a atomNet -> Metis_Atom.atom -> 'a list
+
+val unify : 'a atomNet -> Metis_Atom.atom -> 'a list
 
 end
 
 (**** Original file: AtomNet.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS              *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure AtomNet :> AtomNet =
+structure Metis_AtomNet :> Metis_AtomNet =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Helper functions.                                                         *)
 (* ------------------------------------------------------------------------- *)
 
-fun atomToTerm atom = Term.Fn atom;
-
-fun termToAtom (Term.Var _) = raise Bug "AtomNet.termToAtom"
-  | termToAtom (Term.Fn atom) = atom;
+fun atomToTerm atom = Metis_Term.Fn atom;
+
+fun termToAtom (Metis_Term.Var _) = raise Bug "Metis_AtomNet.termToAtom"
+  | termToAtom (Metis_Term.Fn atom) = atom;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of atom sets that can be efficiently matched and unified.          *)
 (* ------------------------------------------------------------------------- *)
 
-type parameters = TermNet.parameters;
-
-type 'a atomNet = 'a TermNet.termNet;
+type parameters = Metis_TermNet.parameters;
+
+type 'a atomNet = 'a Metis_TermNet.termNet;
 
 (* ------------------------------------------------------------------------- *)
 (* Basic operations.                                                         *)
 (* ------------------------------------------------------------------------- *)
 
-val new = TermNet.new;
-
-val size = TermNet.size;
-
-fun insert net (atm,a) = TermNet.insert net (atomToTerm atm, a);
+val new = Metis_TermNet.new;
+
+val size = Metis_TermNet.size;
+
+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;
 
-val filter = TermNet.filter;
-
-fun toString net = "AtomNet[" ^ Int.toString (size net) ^ "]";
-
-val pp = TermNet.pp;
+val filter = Metis_TermNet.filter;
+
+fun toString net = "Metis_AtomNet[" ^ Int.toString (size net) ^ "]";
+
+val pp = Metis_TermNet.pp;
 
 (* ------------------------------------------------------------------------- *)
 (* Matching and unification queries.                                         *)
@@ -17852,14 +17256,13 @@
 (* Filter afterwards to get the precise set of satisfying values.            *)
 (* ------------------------------------------------------------------------- *)
 
-fun match net atm = TermNet.match net (atomToTerm atm);
-
-fun matched net atm = TermNet.matched net (atomToTerm atm);
-
-fun unify net atm = TermNet.unify net (atomToTerm atm);
-
-end
-end;
+fun match net atm = Metis_TermNet.match net (atomToTerm atm);
+
+fun matched net atm = Metis_TermNet.matched net (atomToTerm atm);
+
+fun unify net atm = Metis_TermNet.unify net (atomToTerm atm);
+
+end
 
 (**** Original file: LiteralNet.sig ****)
 
@@ -17868,7 +17271,7 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature LiteralNet =
+signature Metis_LiteralNet =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -17889,15 +17292,15 @@
 
 val profile : 'a literalNet -> {positive : int, negative : int}
 
-val insert : 'a literalNet -> Metis.Literal.literal * 'a -> 'a literalNet
-
-val fromList : parameters -> (Metis.Literal.literal * 'a) list -> 'a literalNet
+val insert : 'a literalNet -> Metis_Literal.literal * 'a -> 'a literalNet
+
+val fromList : parameters -> (Metis_Literal.literal * 'a) list -> 'a literalNet
 
 val filter : ('a -> bool) -> 'a literalNet -> 'a literalNet
 
 val toString : 'a literalNet -> string
 
-val pp : 'a Metis.Print.pp -> 'a literalNet Metis.Print.pp
+val pp : 'a Metis_Print.pp -> 'a literalNet Metis_Print.pp
 
 (* ------------------------------------------------------------------------- *)
 (* Matching and unification queries.                                         *)
@@ -17906,55 +17309,46 @@
 (* Filter afterwards to get the precise set of satisfying values.            *)
 (* ------------------------------------------------------------------------- *)
 
-val match : 'a literalNet -> Metis.Literal.literal -> 'a list
-
-val matched : 'a literalNet -> Metis.Literal.literal -> 'a list
-
-val unify : 'a literalNet -> Metis.Literal.literal -> 'a list
+val match : 'a literalNet -> Metis_Literal.literal -> 'a list
+
+val matched : 'a literalNet -> Metis_Literal.literal -> 'a list
+
+val unify : 'a literalNet -> Metis_Literal.literal -> 'a list
 
 end
 
 (**** Original file: LiteralNet.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS           *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure LiteralNet :> LiteralNet =
+structure Metis_LiteralNet :> Metis_LiteralNet =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of literal sets that can be efficiently matched and unified.       *)
 (* ------------------------------------------------------------------------- *)
 
-type parameters = AtomNet.parameters;
+type parameters = Metis_AtomNet.parameters;
 
 type 'a literalNet =
-    {positive : 'a AtomNet.atomNet,
-     negative : 'a AtomNet.atomNet};
+    {positive : 'a Metis_AtomNet.atomNet,
+     negative : 'a Metis_AtomNet.atomNet};
 
 (* ------------------------------------------------------------------------- *)
 (* Basic operations.                                                         *)
 (* ------------------------------------------------------------------------- *)
 
-fun new parm = {positive = AtomNet.new parm, negative = AtomNet.new parm};
-
-local
-  fun pos ({positive,...} : 'a literalNet) = AtomNet.size positive;
-
-  fun neg ({negative,...} : 'a literalNet) = AtomNet.size negative;
+fun new parm = {positive = Metis_AtomNet.new parm, negative = Metis_AtomNet.new parm};
+
+local
+  fun pos ({positive,...} : 'a literalNet) = Metis_AtomNet.size positive;
+
+  fun neg ({negative,...} : 'a literalNet) = Metis_AtomNet.size negative;
 in
   fun size net = pos net + neg net;
 
@@ -17962,22 +17356,22 @@
 end;
 
 fun insert {positive,negative} ((true,atm),a) =
-    {positive = AtomNet.insert positive (atm,a), negative = negative}
+    {positive = Metis_AtomNet.insert positive (atm,a), negative = negative}
   | insert {positive,negative} ((false,atm),a) =
-    {positive = positive, negative = AtomNet.insert negative (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 filter pred {positive,negative} =
-    {positive = AtomNet.filter pred positive,
-     negative = AtomNet.filter pred negative};
-
-fun toString net = "LiteralNet[" ^ Int.toString (size net) ^ "]";
+    {positive = Metis_AtomNet.filter pred positive,
+     negative = Metis_AtomNet.filter pred negative};
+
+fun toString net = "Metis_LiteralNet[" ^ Int.toString (size net) ^ "]";
 
 fun pp ppA =
-    Print.ppMap
+    Metis_Print.ppMap
       (fn {positive,negative} => (positive,negative))
-      (Print.ppOp2 " + NEGATIVE" (AtomNet.pp ppA) (AtomNet.pp ppA));
+      (Metis_Print.ppOp2 " + NEGATIVE" (Metis_AtomNet.pp ppA) (Metis_AtomNet.pp ppA));
 
 (* ------------------------------------------------------------------------- *)
 (* Matching and unification queries.                                         *)
@@ -17987,19 +17381,18 @@
 (* ------------------------------------------------------------------------- *)
 
 fun match ({positive,...} : 'a literalNet) (true,atm) =
-    AtomNet.match positive atm
-  | match {negative,...} (false,atm) = AtomNet.match negative atm;
+    Metis_AtomNet.match positive atm
+  | match {negative,...} (false,atm) = Metis_AtomNet.match negative atm;
 
 fun matched ({positive,...} : 'a literalNet) (true,atm) =
-    AtomNet.matched positive atm
-  | matched {negative,...} (false,atm) = AtomNet.matched negative atm;
+    Metis_AtomNet.matched positive atm
+  | matched {negative,...} (false,atm) = Metis_AtomNet.matched negative atm;
 
 fun unify ({positive,...} : 'a literalNet) (true,atm) =
-    AtomNet.unify positive atm
-  | unify {negative,...} (false,atm) = AtomNet.unify negative atm;
-
-end
-end;
+    Metis_AtomNet.unify positive atm
+  | unify {negative,...} (false,atm) = Metis_AtomNet.unify negative atm;
+
+end
 
 (**** Original file: Subsume.sig ****)
 
@@ -18008,7 +17401,7 @@
 (* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Subsume =
+signature Metis_Subsume =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -18021,11 +17414,11 @@
 
 val size : 'a subsume -> int
 
-val insert : 'a subsume -> Metis.Thm.clause * 'a -> 'a subsume
+val insert : 'a subsume -> Metis_Thm.clause * 'a -> 'a subsume
 
 val filter : ('a -> bool) -> 'a subsume -> 'a subsume
 
-val pp : 'a subsume Metis.Print.pp
+val pp : 'a subsume Metis_Print.pp
 
 val toString : 'a subsume -> string
 
@@ -18034,47 +17427,38 @@
 (* ------------------------------------------------------------------------- *)
 
 val subsumes :
-    (Metis.Thm.clause * Metis.Subst.subst * 'a -> bool) -> 'a subsume -> Metis.Thm.clause ->
-    (Metis.Thm.clause * Metis.Subst.subst * 'a) option
-
-val isSubsumed : 'a subsume -> Metis.Thm.clause -> bool
+    (Metis_Thm.clause * Metis_Subst.subst * 'a -> bool) -> 'a subsume -> Metis_Thm.clause ->
+    (Metis_Thm.clause * Metis_Subst.subst * 'a) option
+
+val isSubsumed : 'a subsume -> Metis_Thm.clause -> bool
 
 val strictlySubsumes :  (* exclude subsuming clauses with more literals *)
-    (Metis.Thm.clause * Metis.Subst.subst * 'a -> bool) -> 'a subsume -> Metis.Thm.clause ->
-    (Metis.Thm.clause * Metis.Subst.subst * 'a) option
-
-val isStrictlySubsumed : 'a subsume -> Metis.Thm.clause -> bool
+    (Metis_Thm.clause * Metis_Subst.subst * 'a -> bool) -> 'a subsume -> Metis_Thm.clause ->
+    (Metis_Thm.clause * Metis_Subst.subst * 'a) option
+
+val isStrictlySubsumed : 'a subsume -> Metis_Thm.clause -> bool
 
 (* ------------------------------------------------------------------------- *)
 (* Single clause versions.                                                   *)
 (* ------------------------------------------------------------------------- *)
 
-val clauseSubsumes : Metis.Thm.clause -> Metis.Thm.clause -> Metis.Subst.subst option
-
-val clauseStrictlySubsumes : Metis.Thm.clause -> Metis.Thm.clause -> Metis.Subst.subst option
+val clauseSubsumes : Metis_Thm.clause -> Metis_Thm.clause -> Metis_Subst.subst option
+
+val clauseStrictlySubsumes : Metis_Thm.clause -> Metis_Thm.clause -> Metis_Subst.subst option
 
 end
 
 (**** Original file: Subsume.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES                        *)
 (* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Subsume :> Subsume =
+structure Metis_Subsume :> Metis_Subsume =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Helper functions.                                                         *)
@@ -18091,7 +17475,7 @@
 
 local
   fun addSym (lit,acc) =
-      case total Literal.sym lit of
+      case total Metis_Literal.sym lit of
         NONE => acc
       | SOME lit => lit :: acc
 in
@@ -18100,20 +17484,20 @@
 
 fun sortClause cl =
     let
-      val lits = LiteralSet.toList cl
-    in
-      sortMap Literal.typedSymbols (revCompare Int.compare) lits
+      val lits = Metis_LiteralSet.toList cl
+    in
+      sortMap Metis_Literal.typedSymbols (revCompare Int.compare) lits
     end;
 
 fun incompatible lit =
     let
       val lits = clauseSym [lit]
     in
-      fn lit' => not (List.exists (can (Literal.unify Subst.empty lit')) lits)
-    end;
-
-(* ------------------------------------------------------------------------- *)
-(* Clause ids and lengths.                                                   *)
+      fn lit' => not (List.exists (can (Metis_Literal.unify Metis_Subst.empty lit')) lits)
+    end;
+
+(* ------------------------------------------------------------------------- *)
+(* Metis_Clause ids and lengths.                                                   *)
 (* ------------------------------------------------------------------------- *)
 
 type clauseId = int;
@@ -18121,7 +17505,7 @@
 type clauseLength = int;
 
 local
-  type idSet = (clauseId * clauseLength) Set.set;
+  type idSet = (clauseId * clauseLength) Metis_Set.set;
 
   fun idCompare ((id1,len1),(id2,len2)) =
       case Int.compare (len1,len2) of
@@ -18129,14 +17513,14 @@
       | EQUAL => Int.compare (id1,id2)
       | GREATER => GREATER;
 in
-  val idSetEmpty : idSet = Set.empty idCompare;
-
-  fun idSetAdd (id_len,set) : idSet = Set.add set id_len;
+  val idSetEmpty : idSet = Metis_Set.empty idCompare;
+
+  fun idSetAdd (id_len,set) : idSet = Metis_Set.add set id_len;
 
   fun idSetAddMax max (id_len as (_,len), set) : idSet =
-      if len <= max then Set.add set id_len else set;
-
-  fun idSetIntersect set1 set2 : idSet = Set.intersect set1 set2;
+      if len <= max then Metis_Set.add set id_len else set;
+
+  fun idSetIntersect set1 set2 : idSet = Metis_Set.intersect set1 set2;
 end;
 
 (* ------------------------------------------------------------------------- *)
@@ -18144,90 +17528,90 @@
 (* ------------------------------------------------------------------------- *)
 
 datatype 'a subsume =
-    Subsume of
-      {empty : (Thm.clause * Subst.subst * 'a) list,
-       unit : (Literal.literal * Thm.clause * 'a)  LiteralNet.literalNet,
+    Metis_Subsume of
+      {empty : (Metis_Thm.clause * Metis_Subst.subst * 'a) list,
+       unit : (Metis_Literal.literal * Metis_Thm.clause * 'a)  Metis_LiteralNet.literalNet,
        nonunit :
          {nextId : clauseId,
-          clauses : (Literal.literal list * Thm.clause * 'a) IntMap.map,
-          fstLits : (clauseId * clauseLength) LiteralNet.literalNet,
-          sndLits : (clauseId * clauseLength) LiteralNet.literalNet}};
+          clauses : (Metis_Literal.literal list * Metis_Thm.clause * 'a) Metis_IntMap.map,
+          fstLits : (clauseId * clauseLength) Metis_LiteralNet.literalNet,
+          sndLits : (clauseId * clauseLength) Metis_LiteralNet.literalNet}};
 
 fun new () =
-    Subsume
+    Metis_Subsume
       {empty = [],
-       unit = LiteralNet.new {fifo = false},
+       unit = Metis_LiteralNet.new {fifo = false},
        nonunit =
          {nextId = 0,
-          clauses = IntMap.new (),
-          fstLits = LiteralNet.new {fifo = false},
-          sndLits = LiteralNet.new {fifo = false}}};
-
-fun size (Subsume {empty, unit, nonunit = {clauses,...}}) =
-    length empty + LiteralNet.size unit + IntMap.size clauses;
+          clauses = Metis_IntMap.new (),
+          fstLits = Metis_LiteralNet.new {fifo = false},
+          sndLits = Metis_LiteralNet.new {fifo = false}}};
+
+fun size (Metis_Subsume {empty, unit, nonunit = {clauses,...}}) =
+    length empty + Metis_LiteralNet.size unit + Metis_IntMap.size clauses;
       
-fun insert (Subsume {empty,unit,nonunit}) (cl',a) =
+fun insert (Metis_Subsume {empty,unit,nonunit}) (cl',a) =
     case sortClause cl' of
       [] =>
       let
-        val empty = (cl',Subst.empty,a) :: empty
-      in
-        Subsume {empty = empty, unit = unit, nonunit = nonunit}
+        val empty = (cl',Metis_Subst.empty,a) :: empty
+      in
+        Metis_Subsume {empty = empty, unit = unit, nonunit = nonunit}
       end
     | [lit] =>
       let
-        val unit = LiteralNet.insert unit (lit,(lit,cl',a))
-      in
-        Subsume {empty = empty, unit = unit, nonunit = nonunit}
+        val unit = Metis_LiteralNet.insert unit (lit,(lit,cl',a))
+      in
+        Metis_Subsume {empty = empty, unit = unit, nonunit = nonunit}
       end
     | fstLit :: (nonFstLits as sndLit :: otherLits) =>
       let
         val {nextId,clauses,fstLits,sndLits} = nonunit
-        val id_length = (nextId, LiteralSet.size cl')
-        val fstLits = LiteralNet.insert fstLits (fstLit,id_length)
+        val id_length = (nextId, Metis_LiteralSet.size cl')
+        val fstLits = Metis_LiteralNet.insert fstLits (fstLit,id_length)
         val (sndLit,otherLits) =
             case findRest (incompatible fstLit) nonFstLits of
               SOME sndLit_otherLits => sndLit_otherLits
             | NONE => (sndLit,otherLits)
-        val sndLits = LiteralNet.insert sndLits (sndLit,id_length)
+        val sndLits = Metis_LiteralNet.insert sndLits (sndLit,id_length)
         val lits' = otherLits @ [fstLit,sndLit]
-        val clauses = IntMap.insert clauses (nextId,(lits',cl',a))
+        val clauses = Metis_IntMap.insert clauses (nextId,(lits',cl',a))
         val nextId = nextId + 1
         val nonunit = {nextId = nextId, clauses = clauses,
                        fstLits = fstLits, sndLits = sndLits}
       in
-        Subsume {empty = empty, unit = unit, nonunit = nonunit}
-      end;
-
-fun filter pred (Subsume {empty,unit,nonunit}) =
+        Metis_Subsume {empty = empty, unit = unit, nonunit = nonunit}
+      end;
+
+fun filter pred (Metis_Subsume {empty,unit,nonunit}) =
     let
       val empty = List.filter (pred o #3) empty
 
-      val unit = LiteralNet.filter (pred o #3) unit
+      val unit = Metis_LiteralNet.filter (pred o #3) unit
 
       val nonunit =
           let
             val {nextId,clauses,fstLits,sndLits} = nonunit
-            val clauses' = IntMap.filter (pred o #3 o snd) clauses
-          in
-            if IntMap.size clauses = IntMap.size clauses' then nonunit
+            val clauses' = Metis_IntMap.filter (pred o #3 o snd) clauses
+          in
+            if Metis_IntMap.size clauses = Metis_IntMap.size clauses' then nonunit
             else
               let
-                fun predId (id,_) = IntMap.inDomain id clauses'
-                val fstLits = LiteralNet.filter predId fstLits
-                and sndLits = LiteralNet.filter predId sndLits
+                fun predId (id,_) = Metis_IntMap.inDomain id clauses'
+                val fstLits = Metis_LiteralNet.filter predId fstLits
+                and sndLits = Metis_LiteralNet.filter predId sndLits
               in
                 {nextId = nextId, clauses = clauses',
                  fstLits = fstLits, sndLits = sndLits}
               end
           end
     in
-      Subsume {empty = empty, unit = unit, nonunit = nonunit}
-    end;
-
-fun toString subsume = "Subsume{" ^ Int.toString (size subsume) ^ "}";
-
-fun pp subsume = Print.ppMap toString Print.ppString subsume;
+      Metis_Subsume {empty = empty, unit = unit, nonunit = nonunit}
+    end;
+
+fun toString subsume = "Metis_Subsume{" ^ Int.toString (size subsume) ^ "}";
+
+fun pp subsume = Metis_Print.ppMap toString Metis_Print.ppString subsume;
 
 (* ------------------------------------------------------------------------- *)
 (* Subsumption checking.                                                     *)
@@ -18235,7 +17619,7 @@
 
 local
   fun matchLit lit' (lit,acc) =
-      case total (Literal.match Subst.empty lit') lit of
+      case total (Metis_Literal.match Metis_Subst.empty lit') lit of
         SOME sub => sub :: acc
       | NONE => acc;
 in
@@ -18246,7 +17630,7 @@
             case List.foldl (matchLit lit') [] cl of
               [] => NONE
             | [sub'] =>
-              (case total (Subst.union sub) sub' of
+              (case total (Metis_Subst.union sub) sub' of
                  NONE => NONE
                | SOME sub => mkSubsl acc sub lits')
             | subs => mkSubsl (subs :: acc) sub lits'
@@ -18263,12 +17647,12 @@
             let
               val others = (sub, subs :: subsl) :: others
             in
-              case total (Subst.union sub) sub' of
+              case total (Metis_Subst.union sub) sub' of
                 NONE => search others
               | SOME sub => search ((sub,subsl) :: others)
             end
       in
-        case mkSubsl [] Subst.empty lits' of
+        case mkSubsl [] Metis_Subst.empty lits' of
           NONE => NONE
         | SOME sub_subsl => search [sub_subsl]
       end;
@@ -18282,7 +17666,7 @@
         fun subLit lit =
             let
               fun subUnit (lit',cl',a) =
-                  case total (Literal.match Subst.empty lit') lit of
+                  case total (Metis_Literal.match Metis_Subst.empty lit') lit of
                     NONE => NONE
                   | SOME sub =>
                     let
@@ -18291,7 +17675,7 @@
                       if pred x then SOME x else NONE
                     end
             in
-              first subUnit (LiteralNet.match unit lit)
+              first subUnit (Metis_LiteralNet.match unit lit)
             end
       in
         first subLit
@@ -18302,13 +17686,13 @@
         val addId = case max of NONE => idSetAdd | SOME n => idSetAddMax n
 
         fun subLit lits (lit,acc) =
-            List.foldl addId acc (LiteralNet.match lits lit)
+            List.foldl addId acc (Metis_LiteralNet.match lits lit)
 
         val {nextId = _, clauses, fstLits, sndLits} = nonunit
 
         fun subCl' (id,_) =
             let
-              val (lits',cl',a) = IntMap.get clauses id
+              val (lits',cl',a) = Metis_IntMap.get clauses id
             in
               genClauseSubsumes pred cl' lits' cl a
             end
@@ -18317,17 +17701,17 @@
         val sndCands = List.foldl (subLit sndLits) idSetEmpty cl
         val cands = idSetIntersect fstCands sndCands
       in
-        Set.firstl subCl' cands
-      end;
-
-  fun genSubsumes pred (Subsume {empty,unit,nonunit}) max cl =
+        Metis_Set.firstl subCl' cands
+      end;
+
+  fun genSubsumes pred (Metis_Subsume {empty,unit,nonunit}) max cl =
       case emptySubsumes pred empty of
         s as SOME _ => s
       | NONE =>
         if max = SOME 0 then NONE
         else
           let
-            val cl = clauseSym (LiteralSet.toList cl)
+            val cl = clauseSym (Metis_LiteralSet.toList cl)
           in
             case unitSubsumes pred unit cl of
               s as SOME _ => s
@@ -18339,38 +17723,38 @@
   fun subsumes pred subsume cl = genSubsumes pred subsume NONE cl;
 
   fun strictlySubsumes pred subsume cl =
-      genSubsumes pred subsume (SOME (LiteralSet.size cl)) cl;
+      genSubsumes pred subsume (SOME (Metis_LiteralSet.size cl)) cl;
 end;
 
 (*MetisTrace4
 val subsumes = fn pred => fn subsume => fn cl =>
     let
-      val ppCl = LiteralSet.pp
-      val ppSub = Subst.pp
-      val () = Print.trace ppCl "Subsume.subsumes: cl" cl
+      val ppCl = Metis_LiteralSet.pp
+      val ppSub = Metis_Subst.pp
+      val () = Metis_Print.trace ppCl "Metis_Subsume.subsumes: cl" cl
       val result = subsumes pred subsume cl
       val () =
           case result of
-            NONE => trace "Subsume.subsumes: not subsumed\n"
+            NONE => trace "Metis_Subsume.subsumes: not subsumed\n"
           | SOME (cl,sub,_) =>
-            (Print.trace ppCl "Subsume.subsumes: subsuming cl" cl;
-             Print.trace ppSub "Subsume.subsumes: subsuming sub" sub)
+            (Metis_Print.trace ppCl "Metis_Subsume.subsumes: subsuming cl" cl;
+             Metis_Print.trace ppSub "Metis_Subsume.subsumes: subsuming sub" sub)
     in
       result
     end;
 
 val strictlySubsumes = fn pred => fn subsume => fn cl =>
     let
-      val ppCl = LiteralSet.pp
-      val ppSub = Subst.pp
-      val () = Print.trace ppCl "Subsume.strictlySubsumes: cl" cl
+      val ppCl = Metis_LiteralSet.pp
+      val ppSub = Metis_Subst.pp
+      val () = Metis_Print.trace ppCl "Metis_Subsume.strictlySubsumes: cl" cl
       val result = strictlySubsumes pred subsume cl
       val () =
           case result of
-            NONE => trace "Subsume.subsumes: not subsumed\n"
+            NONE => trace "Metis_Subsume.subsumes: not subsumed\n"
           | SOME (cl,sub,_) =>
-            (Print.trace ppCl "Subsume.subsumes: subsuming cl" cl;
-             Print.trace ppSub "Subsume.subsumes: subsuming sub" sub)
+            (Metis_Print.trace ppCl "Metis_Subsume.subsumes: subsuming cl" cl;
+             Metis_Print.trace ppSub "Metis_Subsume.subsumes: subsuming sub" sub)
     in
       result
     end;
@@ -18388,7 +17772,7 @@
 fun clauseSubsumes cl' cl =
     let
       val lits' = sortClause cl'
-      and lits = clauseSym (LiteralSet.toList cl)
+      and lits = clauseSym (Metis_LiteralSet.toList cl)
     in
       case genClauseSubsumes (K true) cl' lits' lits () of
         SOME (_,sub,()) => SOME sub
@@ -18396,11 +17780,10 @@
     end;
 
 fun clauseStrictlySubsumes cl' cl =
-    if LiteralSet.size cl' > LiteralSet.size cl then NONE
+    if Metis_LiteralSet.size cl' > Metis_LiteralSet.size cl then NONE
     else clauseSubsumes cl' cl;
 
 end
-end;
 
 (**** Original file: KnuthBendixOrder.sig ****)
 
@@ -18409,7 +17792,7 @@
 (* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature KnuthBendixOrder =
+signature Metis_KnuthBendixOrder =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -18418,41 +17801,32 @@
 (* ------------------------------------------------------------------------- *)
 
 type kbo =
-     {weight : Metis.Term.function -> int,
-      precedence : Metis.Term.function * Metis.Term.function -> order}
+     {weight : Metis_Term.function -> int,
+      precedence : Metis_Term.function * Metis_Term.function -> order}
 
 val default : kbo
 
-val compare : kbo -> Metis.Term.term * Metis.Term.term -> order option
+val compare : kbo -> Metis_Term.term * Metis_Term.term -> order option
 
 end
 
 (**** Original file: KnuthBendixOrder.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* KNUTH-BENDIX TERM ORDERING CONSTRAINTS                                    *)
 (* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure KnuthBendixOrder :> KnuthBendixOrder =
+structure Metis_KnuthBendixOrder :> Metis_KnuthBendixOrder =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Helper functions.                                                         *)
 (* ------------------------------------------------------------------------- *)
 
-fun notEqualTerm (x,y) = not (Term.equal x y);
+fun notEqualTerm (x,y) = not (Metis_Term.equal x y);
 
 fun firstNotEqualTerm f l =
     case List.find notEqualTerm l of
@@ -18465,20 +17839,20 @@
 (* ------------------------------------------------------------------------- *)
 
 type kbo =
-     {weight : Term.function -> int,
-      precedence : Term.function * Term.function -> order};
+     {weight : Metis_Term.function -> int,
+      precedence : Metis_Term.function * Metis_Term.function -> order};
 
 (* Default weight = uniform *)
 
-val uniformWeight : Term.function -> int = K 1;
+val uniformWeight : Metis_Term.function -> int = K 1;
 
 (* Default precedence = by arity *)
 
-val arityPrecedence : Term.function * Term.function -> order =
+val arityPrecedence : Metis_Term.function * Metis_Term.function -> order =
     fn ((f1,n1),(f2,n2)) =>
        case Int.compare (n1,n2) of
          LESS => LESS
-       | EQUAL => Name.compare (f1,f2)
+       | EQUAL => Metis_Name.compare (f1,f2)
        | GREATER => GREATER;
 
 (* The default order *)
@@ -18486,22 +17860,22 @@
 val default = {weight = uniformWeight, precedence = arityPrecedence};
 
 (* ------------------------------------------------------------------------- *)
-(* Term weight-1 represented as a linear function of the weight-1 of the     *)
+(* Metis_Term weight-1 represented as a linear function of the weight-1 of the     *)
 (* variables in the term (plus a constant).                                  *)
 (*                                                                           *)
 (* Note that the conditions on weight functions ensure that all weights are  *)
 (* at least 1, so all weight-1s are at least 0.                              *)
 (* ------------------------------------------------------------------------- *)
 
-datatype weight = Weight of int NameMap.map * int;
-
-val weightEmpty : int NameMap.map = NameMap.new ();
+datatype weight = Weight of int Metis_NameMap.map * int;
+
+val weightEmpty : int Metis_NameMap.map = Metis_NameMap.new ();
 
 val weightZero = Weight (weightEmpty,0);
 
-fun weightIsZero (Weight (m,c)) = c = 0 andalso NameMap.null m;
-
-fun weightNeg (Weight (m,c)) = Weight (NameMap.transform ~ m, ~c);
+fun weightIsZero (Weight (m,c)) = c = 0 andalso Metis_NameMap.null m;
+
+fun weightNeg (Weight (m,c)) = Weight (Metis_NameMap.transform ~ m, ~c);
 
 local
   fun add ((_,n1),(_,n2)) =
@@ -18512,7 +17886,7 @@
       end;
 in
   fun weightAdd (Weight (m1,c1)) (Weight (m2,c2)) =
-      Weight (NameMap.union add m1 m2, c1 + c2);
+      Weight (Metis_NameMap.union add m1 m2, c1 + c2);
 end;
 
 fun weightSubtract w1 w2 = weightAdd w1 (weightNeg w2);
@@ -18520,46 +17894,46 @@
 fun weightTerm weight =
     let
       fun wt m c [] = Weight (m,c)
-        | wt m c (Term.Var v :: tms) =
-          let
-            val n = Option.getOpt (NameMap.peek m v, 0)
-          in
-            wt (NameMap.insert m (v, n + 1)) (c + 1) tms
-          end
-        | wt m c (Term.Fn (f,a) :: tms) =
+        | wt m c (Metis_Term.Var v :: tms) =
+          let
+            val n = Option.getOpt (Metis_NameMap.peek m v, 0)
+          in
+            wt (Metis_NameMap.insert m (v, n + 1)) (c + 1) tms
+          end
+        | wt m c (Metis_Term.Fn (f,a) :: tms) =
           wt m (c + weight (f, length a)) (a @ tms)
     in
       fn tm => wt weightEmpty ~1 [tm]
     end;
 
 fun weightLowerBound (w as Weight (m,c)) =
-    if NameMap.exists (fn (_,n) => n < 0) m then NONE else SOME c;
+    if Metis_NameMap.exists (fn (_,n) => n < 0) m then NONE else SOME c;
 
 (*MetisDebug
 val ppWeightList =
     let
       fun ppCoeff n =
-          if n < 0 then Print.sequence (Print.addString "~") (ppCoeff (~n))
-          else if n = 1 then Print.skip
-          else Print.ppInt n
-
-      fun pp_tm (NONE,n) = Print.ppInt n
-        | pp_tm (SOME v, n) = Print.sequence (ppCoeff n) (Name.pp v)
-    in
-      fn [] => Print.ppInt 0
-       | tms => Print.ppOpList " +" pp_tm tms
+          if n < 0 then Metis_Print.sequence (Metis_Print.addString "~") (ppCoeff (~n))
+          else if n = 1 then Metis_Print.skip
+          else Metis_Print.ppInt n
+
+      fun pp_tm (NONE,n) = Metis_Print.ppInt n
+        | pp_tm (SOME v, n) = Metis_Print.sequence (ppCoeff n) (Metis_Name.pp v)
+    in
+      fn [] => Metis_Print.ppInt 0
+       | tms => Metis_Print.ppOpList " +" pp_tm tms
     end;
 
 fun ppWeight (Weight (m,c)) =
     let
-      val l = NameMap.toList m
+      val l = Metis_NameMap.toList m
       val l = map (fn (v,n) => (SOME v, n)) l
       val l = if c = 0 then l else l @ [(NONE,c)]
     in
       ppWeightList l
     end;
 
-val weightToString = Print.toString ppWeight;
+val weightToString = Metis_Print.toString ppWeight;
 *)
 
 (* ------------------------------------------------------------------------- *)
@@ -18590,7 +17964,7 @@
           | SOME 0 => precedenceLess tm1 tm2
           | SOME n => n > 0
 
-      and precedenceLess (Term.Fn (f1,a1)) (Term.Fn (f2,a2)) =
+      and precedenceLess (Metis_Term.Fn (f1,a1)) (Metis_Term.Fn (f2,a2)) =
           (case precedence ((f1, length a1), (f2, length a2)) of
              LESS => true
            | EQUAL => firstNotEqualTerm weightLess (zip a1 a2)
@@ -18609,7 +17983,7 @@
             else NONE
           end
 
-      and precedenceCmp (Term.Fn (f1,a1)) (Term.Fn (f2,a2)) =
+      and precedenceCmp (Metis_Term.Fn (f1,a1)) (Metis_Term.Fn (f2,a2)) =
           (case precedence ((f1, length a1), (f2, length a2)) of
              LESS => SOME LESS
            | EQUAL => firstNotEqualTerm weightCmp (zip a1 a2)
@@ -18617,27 +17991,26 @@
         | precedenceCmp _ _ = raise Bug "kboOrder.precendenceCmp"
     in
       fn (tm1,tm2) =>
-         if Term.equal tm1 tm2 then SOME EQUAL else weightCmp tm1 tm2
+         if Metis_Term.equal tm1 tm2 then SOME EQUAL else weightCmp tm1 tm2
     end;
 
 (*MetisTrace7
 val compare = fn kbo => fn (tm1,tm2) =>
     let
-      val () = Print.trace Term.pp "KnuthBendixOrder.compare: tm1" tm1
-      val () = Print.trace Term.pp "KnuthBendixOrder.compare: tm2" tm2
+      val () = Metis_Print.trace Metis_Term.pp "Metis_KnuthBendixOrder.compare: tm1" tm1
+      val () = Metis_Print.trace Metis_Term.pp "Metis_KnuthBendixOrder.compare: tm2" tm2
       val result = compare kbo (tm1,tm2)
       val () =
           case result of
-            NONE => trace "KnuthBendixOrder.compare: result = Incomparable\n"
+            NONE => trace "Metis_KnuthBendixOrder.compare: result = Incomparable\n"
           | SOME x =>
-            Print.trace Print.ppOrder "KnuthBendixOrder.compare: result" x
+            Metis_Print.trace Metis_Print.ppOrder "Metis_KnuthBendixOrder.compare: result" x
     in
       result
     end;
 *)
 
 end
-end;
 
 (**** Original file: Rewrite.sig ****)
 
@@ -18646,7 +18019,7 @@
 (* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Rewrite =
+signature Metis_Rewrite =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -18657,21 +18030,21 @@
 
 val toStringOrient : orient -> string
 
-val ppOrient : orient Metis.Print.pp
+val ppOrient : orient Metis_Print.pp
 
 val toStringOrientOption : orient option -> string
 
-val ppOrientOption : orient option Metis.Print.pp
+val ppOrientOption : orient option Metis_Print.pp
 
 (* ------------------------------------------------------------------------- *)
 (* A type of rewrite systems.                                                *)
 (* ------------------------------------------------------------------------- *)
 
-type reductionOrder = Metis.Term.term * Metis.Term.term -> order option
+type reductionOrder = Metis_Term.term * Metis_Term.term -> order option
 
 type equationId = int
 
-type equation = Metis.Rule.equation
+type equation = Metis_Rule.equation
 
 type rewrite
 
@@ -18689,7 +18062,7 @@
 
 val toString : rewrite -> string
 
-val pp : rewrite Metis.Print.pp
+val pp : rewrite Metis_Print.pp
 
 (* ------------------------------------------------------------------------- *)
 (* Add equations into the system.                                            *)
@@ -18703,23 +18076,23 @@
 (* Rewriting (the order must be a refinement of the rewrite order).          *)
 (* ------------------------------------------------------------------------- *)
 
-val rewrConv : rewrite -> reductionOrder -> Metis.Rule.conv
-
-val rewriteConv : rewrite -> reductionOrder -> Metis.Rule.conv
+val rewrConv : rewrite -> reductionOrder -> Metis_Rule.conv
+
+val rewriteConv : rewrite -> reductionOrder -> Metis_Rule.conv
 
 val rewriteLiteralsRule :
-    rewrite -> reductionOrder -> Metis.LiteralSet.set -> Metis.Rule.rule
-
-val rewriteRule : rewrite -> reductionOrder -> Metis.Rule.rule
-
-val rewrIdConv : rewrite -> reductionOrder -> equationId -> Metis.Rule.conv
-
-val rewriteIdConv : rewrite -> reductionOrder -> equationId -> Metis.Rule.conv
+    rewrite -> reductionOrder -> Metis_LiteralSet.set -> Metis_Rule.rule
+
+val rewriteRule : rewrite -> reductionOrder -> Metis_Rule.rule
+
+val rewrIdConv : rewrite -> reductionOrder -> equationId -> Metis_Rule.conv
+
+val rewriteIdConv : rewrite -> reductionOrder -> equationId -> Metis_Rule.conv
 
 val rewriteIdLiteralsRule :
-    rewrite -> reductionOrder -> equationId -> Metis.LiteralSet.set -> Metis.Rule.rule
-
-val rewriteIdRule : rewrite -> reductionOrder -> equationId -> Metis.Rule.rule
+    rewrite -> reductionOrder -> equationId -> Metis_LiteralSet.set -> Metis_Rule.rule
+
+val rewriteIdRule : rewrite -> reductionOrder -> equationId -> Metis_Rule.rule
 
 (* ------------------------------------------------------------------------- *)
 (* Inter-reduce the equations in the system.                                 *)
@@ -18735,32 +18108,23 @@
 (* Rewriting as a derived rule.                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val rewrite : equation list -> Metis.Thm.thm -> Metis.Thm.thm
-
-val orderedRewrite : reductionOrder -> equation list -> Metis.Thm.thm -> Metis.Thm.thm
+val rewrite : equation list -> Metis_Thm.thm -> Metis_Thm.thm
+
+val orderedRewrite : reductionOrder -> equation list -> Metis_Thm.thm -> Metis_Thm.thm
 
 end
 
 (**** Original file: Rewrite.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* ORDERED REWRITING FOR FIRST ORDER TERMS                                   *)
 (* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Rewrite :> Rewrite =
+structure Metis_Rewrite :> Metis_Rewrite =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Orientations of equations.                                                *)
@@ -18773,121 +18137,121 @@
       LeftToRight => "-->"
     | RightToLeft => "<--";
 
-val ppOrient = Print.ppMap toStringOrient Print.ppString;
+val ppOrient = Metis_Print.ppMap toStringOrient Metis_Print.ppString;
 
 fun toStringOrientOption orto =
     case orto of
       SOME ort => toStringOrient ort
     | NONE => "<->";
 
-val ppOrientOption = Print.ppMap toStringOrientOption Print.ppString;
+val ppOrientOption = Metis_Print.ppMap toStringOrientOption Metis_Print.ppString;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of rewrite systems.                                                *)
 (* ------------------------------------------------------------------------- *)
 
-type reductionOrder = Term.term * Term.term -> order option;
+type reductionOrder = Metis_Term.term * Metis_Term.term -> order option;
 
 type equationId = int;
 
-type equation = Rule.equation;
+type equation = Metis_Rule.equation;
 
 datatype rewrite =
-    Rewrite of
+    Metis_Rewrite of
       {order : reductionOrder,
-       known : (equation * orient option) IntMap.map,
-       redexes : (equationId * orient) TermNet.termNet,
-       subterms : (equationId * bool * Term.path) TermNet.termNet,
-       waiting : IntSet.set};
+       known : (equation * orient option) Metis_IntMap.map,
+       redexes : (equationId * orient) Metis_TermNet.termNet,
+       subterms : (equationId * bool * Metis_Term.path) Metis_TermNet.termNet,
+       waiting : Metis_IntSet.set};
 
 fun updateWaiting rw waiting =
     let
-      val Rewrite {order, known, redexes, subterms, waiting = _} = rw
-    in
-      Rewrite
+      val Metis_Rewrite {order, known, redexes, subterms, waiting = _} = rw
+    in
+      Metis_Rewrite
         {order = order, known = known, redexes = redexes,
          subterms = subterms, waiting = waiting}
     end;
 
-fun deleteWaiting (rw as Rewrite {waiting,...}) id =
-    updateWaiting rw (IntSet.delete waiting id);
+fun deleteWaiting (rw as Metis_Rewrite {waiting,...}) id =
+    updateWaiting rw (Metis_IntSet.delete waiting id);
 
 (* ------------------------------------------------------------------------- *)
 (* Basic operations                                                          *)
 (* ------------------------------------------------------------------------- *)
 
 fun new order =
-    Rewrite
+    Metis_Rewrite
       {order = order,
-       known = IntMap.new (),
-       redexes = TermNet.new {fifo = false},
-       subterms = TermNet.new {fifo = false},
-       waiting = IntSet.empty};
-
-fun peek (Rewrite {known,...}) id = IntMap.peek known id;
-
-fun size (Rewrite {known,...}) = IntMap.size known;
-
-fun equations (Rewrite {known,...}) =
-    IntMap.foldr (fn (_,(eqn,_),eqns) => eqn :: eqns) [] known;
-
-val pp = Print.ppMap equations (Print.ppList Rule.ppEquation);
+       known = Metis_IntMap.new (),
+       redexes = Metis_TermNet.new {fifo = false},
+       subterms = Metis_TermNet.new {fifo = false},
+       waiting = Metis_IntSet.empty};
+
+fun peek (Metis_Rewrite {known,...}) id = Metis_IntMap.peek known id;
+
+fun size (Metis_Rewrite {known,...}) = Metis_IntMap.size known;
+
+fun equations (Metis_Rewrite {known,...}) =
+    Metis_IntMap.foldr (fn (_,(eqn,_),eqns) => eqn :: eqns) [] known;
+
+val pp = Metis_Print.ppMap equations (Metis_Print.ppList Metis_Rule.ppEquation);
 
 (*MetisTrace1
 local
   fun ppEq ((x_y,_),ort) =
-      Print.ppOp2 (" " ^ toStringOrientOption ort) Term.pp Term.pp x_y;
+      Metis_Print.ppOp2 (" " ^ toStringOrientOption ort) Metis_Term.pp Metis_Term.pp x_y;
 
   fun ppField f ppA a =
-      Print.blockProgram Print.Inconsistent 2
-        [Print.addString (f ^ " ="),
-         Print.addBreak 1,
+      Metis_Print.blockProgram Metis_Print.Inconsistent 2
+        [Metis_Print.addString (f ^ " ="),
+         Metis_Print.addBreak 1,
          ppA a];
 
   val ppKnown =
       ppField "known"
-        (Print.ppMap IntMap.toList
-           (Print.ppList (Print.ppPair Print.ppInt ppEq)));
+        (Metis_Print.ppMap Metis_IntMap.toList
+           (Metis_Print.ppList (Metis_Print.ppPair Metis_Print.ppInt ppEq)));
 
   val ppRedexes =
       ppField "redexes"
-        (TermNet.pp (Print.ppPair Print.ppInt ppOrient));
+        (Metis_TermNet.pp (Metis_Print.ppPair Metis_Print.ppInt ppOrient));
 
   val ppSubterms =
       ppField "subterms"
-        (TermNet.pp
-           (Print.ppMap
+        (Metis_TermNet.pp
+           (Metis_Print.ppMap
               (fn (i,l,p) => (i, (if l then 0 else 1) :: p))
-              (Print.ppPair Print.ppInt Term.ppPath)));
+              (Metis_Print.ppPair Metis_Print.ppInt Metis_Term.ppPath)));
 
   val ppWaiting =
       ppField "waiting"
-        (Print.ppMap (IntSet.toList) (Print.ppList Print.ppInt));
-in
-  fun pp (Rewrite {known,redexes,subterms,waiting,...}) =
-      Print.blockProgram Print.Inconsistent 2
-        [Print.addString "Rewrite",
-         Print.addBreak 1,
-         Print.blockProgram Print.Inconsistent 1
-           [Print.addString "{",
+        (Metis_Print.ppMap (Metis_IntSet.toList) (Metis_Print.ppList Metis_Print.ppInt));
+in
+  fun pp (Metis_Rewrite {known,redexes,subterms,waiting,...}) =
+      Metis_Print.blockProgram Metis_Print.Inconsistent 2
+        [Metis_Print.addString "Metis_Rewrite",
+         Metis_Print.addBreak 1,
+         Metis_Print.blockProgram Metis_Print.Inconsistent 1
+           [Metis_Print.addString "{",
             ppKnown known,
 (*MetisTrace5
-            Print.addString ",",
-            Print.addBreak 1,
+            Metis_Print.addString ",",
+            Metis_Print.addBreak 1,
             ppRedexes redexes,
-            Print.addString ",",
-            Print.addBreak 1,
+            Metis_Print.addString ",",
+            Metis_Print.addBreak 1,
             ppSubterms subterms,
-            Print.addString ",",
-            Print.addBreak 1,
+            Metis_Print.addString ",",
+            Metis_Print.addBreak 1,
             ppWaiting waiting,
 *)
-            Print.skip],
-         Print.addString "}"]
-end;
-*)
-
-val toString = Print.toString pp;
+            Metis_Print.skip],
+         Metis_Print.addString "}"]
+end;
+*)
+
+val toString = Metis_Print.toString pp;
 
 (* ------------------------------------------------------------------------- *)
 (* Debug functions.                                                          *)
@@ -18896,43 +18260,43 @@
 fun termReducible order known id =
     let
       fun eqnRed ((l,r),_) tm =
-          case total (Subst.match Subst.empty l) tm of
+          case total (Metis_Subst.match Metis_Subst.empty l) tm of
             NONE => false
           | SOME sub =>
-            order (tm, Subst.subst (Subst.normalize sub) r) = SOME GREATER
+            order (tm, Metis_Subst.subst (Metis_Subst.normalize sub) r) = SOME GREATER
 
       fun knownRed tm (eqnId,(eqn,ort)) =
           eqnId <> id andalso
           ((ort <> SOME RightToLeft andalso eqnRed eqn tm) orelse
-           (ort <> SOME LeftToRight andalso eqnRed (Rule.symEqn eqn) tm))
-
-      fun termRed tm = IntMap.exists (knownRed tm) known orelse subtermRed tm
-      and subtermRed (Term.Var _) = false
-        | subtermRed (Term.Fn (_,tms)) = List.exists termRed tms
+           (ort <> SOME LeftToRight andalso eqnRed (Metis_Rule.symEqn eqn) tm))
+
+      fun termRed tm = Metis_IntMap.exists (knownRed tm) known orelse subtermRed tm
+      and subtermRed (Metis_Term.Var _) = false
+        | subtermRed (Metis_Term.Fn (_,tms)) = List.exists termRed tms
     in
       termRed
     end;
 
 fun literalReducible order known id lit =
-    List.exists (termReducible order known id) (Literal.arguments lit);
+    List.exists (termReducible order known id) (Metis_Literal.arguments lit);
 
 fun literalsReducible order known id lits =
-    LiteralSet.exists (literalReducible order known id) lits;
+    Metis_LiteralSet.exists (literalReducible order known id) lits;
 
 fun thmReducible order known id th =
-    literalsReducible order known id (Thm.clause th);
+    literalsReducible order known id (Metis_Thm.clause th);
 
 (* ------------------------------------------------------------------------- *)
 (* Add equations into the system.                                            *)
 (* ------------------------------------------------------------------------- *)
 
-fun orderToOrient (SOME EQUAL) = raise Error "Rewrite.orient: reflexive"
+fun orderToOrient (SOME EQUAL) = raise Error "Metis_Rewrite.orient: reflexive"
   | orderToOrient (SOME GREATER) = SOME LeftToRight
   | orderToOrient (SOME LESS) = SOME RightToLeft
   | orderToOrient NONE = NONE;
 
 local
-  fun ins redexes redex id ort = TermNet.insert redexes (redex,(id,ort));
+  fun ins redexes redex id ort = Metis_TermNet.insert redexes (redex,(id,ort));
 in
   fun addRedexes id (((l,r),_),ort) redexes =
       case ort of
@@ -18941,21 +18305,21 @@
       | NONE => ins (ins redexes l id LeftToRight) r id RightToLeft;
 end;
 
-fun add (rw as Rewrite {known,...}) (id,eqn) =
-    if IntMap.inDomain id known then rw
+fun add (rw as Metis_Rewrite {known,...}) (id,eqn) =
+    if Metis_IntMap.inDomain id known then rw
     else
       let
-        val Rewrite {order,redexes,subterms,waiting, ...} = rw
+        val Metis_Rewrite {order,redexes,subterms,waiting, ...} = rw
         val ort = orderToOrient (order (fst eqn))
-        val known = IntMap.insert known (id,(eqn,ort))
+        val known = Metis_IntMap.insert known (id,(eqn,ort))
         val redexes = addRedexes id (eqn,ort) redexes
-        val waiting = IntSet.add waiting id
+        val waiting = Metis_IntSet.add waiting id
         val rw =
-            Rewrite
+            Metis_Rewrite
               {order = order, known = known, redexes = redexes,
                subterms = subterms, waiting = waiting}
 (*MetisTrace5
-        val () = Print.trace pp "Rewrite.add: result" rw
+        val () = Metis_Print.trace pp "Metis_Rewrite.add: result" rw
 *)
       in
         rw
@@ -18970,7 +18334,7 @@
 local
   fun reorder ((i,_),(j,_)) = Int.compare (j,i);
 in
-  fun matchingRedexes redexes tm = sort reorder (TermNet.match redexes tm);
+  fun matchingRedexes redexes tm = sort reorder (Metis_TermNet.match redexes tm);
 end;
 
 fun wellOriented NONE _ = true
@@ -18982,118 +18346,118 @@
   | redexResidue RightToLeft ((l,r),_) = (r,l);
 
 fun orientedEquation LeftToRight eqn = eqn
-  | orientedEquation RightToLeft eqn = Rule.symEqn eqn;
+  | orientedEquation RightToLeft eqn = Metis_Rule.symEqn eqn;
 
 fun rewrIdConv' order known redexes id tm =
     let
       fun rewr (id',lr) =
           let
             val _ = id <> id' orelse raise Error "same theorem"
-            val (eqn,ort) = IntMap.get known id'
+            val (eqn,ort) = Metis_IntMap.get known id'
             val _ = wellOriented ort lr orelse raise Error "orientation"
             val (l,r) = redexResidue lr eqn
-            val sub = Subst.normalize (Subst.match Subst.empty l tm)
-            val tm' = Subst.subst sub r
+            val sub = Metis_Subst.normalize (Metis_Subst.match Metis_Subst.empty l tm)
+            val tm' = Metis_Subst.subst sub r
             val _ = Option.isSome ort orelse
                     order (tm,tm') = SOME GREATER orelse
                     raise Error "order"
             val (_,th) = orientedEquation lr eqn
           in
-            (tm', Thm.subst sub th)
+            (tm', Metis_Thm.subst sub th)
           end
     in
       case first (total rewr) (matchingRedexes redexes tm) of
-        NONE => raise Error "Rewrite.rewrIdConv: no matching rewrites"
+        NONE => raise Error "Metis_Rewrite.rewrIdConv: no matching rewrites"
       | SOME res => res
     end;
 
 fun rewriteIdConv' order known redexes id =
-    if IntMap.null known then Rule.allConv
-    else Rule.repeatTopDownConv (rewrIdConv' order known redexes id);
+    if Metis_IntMap.null known then Metis_Rule.allConv
+    else Metis_Rule.repeatTopDownConv (rewrIdConv' order known redexes id);
 
 fun mkNeqConv order lit =
     let
-      val (l,r) = Literal.destNeq lit
+      val (l,r) = Metis_Literal.destNeq lit
     in
       case order (l,r) of
         NONE => raise Error "incomparable"
       | SOME LESS =>
         let
-          val th = Rule.symmetryRule l r
+          val th = Metis_Rule.symmetryRule l r
         in
           fn tm =>
-             if Term.equal tm r then (l,th) else raise Error "mkNeqConv: RL"
+             if Metis_Term.equal tm r then (l,th) else raise Error "mkNeqConv: Metis_RL"
         end
       | SOME EQUAL => raise Error "irreflexive"
       | SOME GREATER =>
         let
-          val th = Thm.assume lit
+          val th = Metis_Thm.assume lit
         in
           fn tm =>
-             if Term.equal tm l then (r,th) else raise Error "mkNeqConv: LR"
+             if Metis_Term.equal tm l then (r,th) else raise Error "mkNeqConv: LR"
         end
     end;
 
-datatype neqConvs = NeqConvs of Rule.conv LiteralMap.map;
-
-val neqConvsEmpty = NeqConvs (LiteralMap.new ());
-
-fun neqConvsNull (NeqConvs m) = LiteralMap.null m;
+datatype neqConvs = NeqConvs of Metis_Rule.conv Metis_LiteralMap.map;
+
+val neqConvsEmpty = NeqConvs (Metis_LiteralMap.new ());
+
+fun neqConvsNull (NeqConvs m) = Metis_LiteralMap.null m;
 
 fun neqConvsAdd order (neq as NeqConvs m) lit =
     case total (mkNeqConv order) lit of
       NONE => NONE
-    | SOME conv => SOME (NeqConvs (LiteralMap.insert m (lit,conv)));
+    | SOME conv => SOME (NeqConvs (Metis_LiteralMap.insert m (lit,conv)));
 
 fun mkNeqConvs order =
     let
       fun add (lit,(neq,lits)) =
           case neqConvsAdd order neq lit of
             SOME neq => (neq,lits)
-          | NONE => (neq, LiteralSet.add lits lit)
-    in
-      LiteralSet.foldl add (neqConvsEmpty,LiteralSet.empty)
-    end;
-
-fun neqConvsDelete (NeqConvs m) lit = NeqConvs (LiteralMap.delete m lit);
+          | NONE => (neq, Metis_LiteralSet.add lits lit)
+    in
+      Metis_LiteralSet.foldl add (neqConvsEmpty,Metis_LiteralSet.empty)
+    end;
+
+fun neqConvsDelete (NeqConvs m) lit = NeqConvs (Metis_LiteralMap.delete m lit);
 
 fun neqConvsToConv (NeqConvs m) =
-    Rule.firstConv (LiteralMap.foldr (fn (_,c,l) => c :: l) [] m);
+    Metis_Rule.firstConv (Metis_LiteralMap.foldr (fn (_,c,l) => c :: l) [] m);
 
 fun neqConvsFoldl f b (NeqConvs m) =
-    LiteralMap.foldl (fn (l,_,z) => f (l,z)) b m;
+    Metis_LiteralMap.foldl (fn (l,_,z) => f (l,z)) b m;
 
 fun neqConvsRewrIdLiterule order known redexes id neq =
-    if IntMap.null known andalso neqConvsNull neq then Rule.allLiterule
+    if Metis_IntMap.null known andalso neqConvsNull neq then Metis_Rule.allLiterule
     else
       let
         val neq_conv = neqConvsToConv neq
         val rewr_conv = rewrIdConv' order known redexes id
-        val conv = Rule.orelseConv neq_conv rewr_conv
-        val conv = Rule.repeatTopDownConv conv
-      in
-        Rule.allArgumentsLiterule conv
+        val conv = Metis_Rule.orelseConv neq_conv rewr_conv
+        val conv = Metis_Rule.repeatTopDownConv conv
+      in
+        Metis_Rule.allArgumentsLiterule conv
       end;
 
 fun rewriteIdEqn' order known redexes id (eqn as (l_r,th)) =
     let
-      val (neq,_) = mkNeqConvs order (Thm.clause th)
+      val (neq,_) = mkNeqConvs order (Metis_Thm.clause th)
       val literule = neqConvsRewrIdLiterule order known redexes id neq
       val (strongEqn,lit) =
-          case Rule.equationLiteral eqn of
-            NONE => (true, Literal.mkEq l_r)
+          case Metis_Rule.equationLiteral eqn of
+            NONE => (true, Metis_Literal.mkEq l_r)
           | SOME lit => (false,lit)
       val (lit',litTh) = literule lit
     in
-      if Literal.equal lit lit' then eqn
+      if Metis_Literal.equal lit lit' then eqn
       else
-        (Literal.destEq lit',
+        (Metis_Literal.destEq lit',
          if strongEqn then th
-         else if not (Thm.negateMember lit litTh) then litTh
-         else Thm.resolve lit th litTh)
+         else if not (Metis_Thm.negateMember lit litTh) then litTh
+         else Metis_Thm.resolve lit th litTh)
     end
 (*MetisDebug
-    handle Error err => raise Error ("Rewrite.rewriteIdEqn':\n" ^ err);
+    handle Error err => raise Error ("Metis_Rewrite.rewriteIdEqn':\n" ^ err);
 *)
 
 fun rewriteIdLiteralsRule' order known redexes id lits th =
@@ -19105,14 +18469,14 @@
             val neq = neqConvsDelete neq lit
             val (lit',litTh) = mk_literule neq lit
           in
-            if Literal.equal lit lit' then acc
+            if Metis_Literal.equal lit lit' then acc
             else
               let
-                val th = Thm.resolve lit th litTh
+                val th = Metis_Thm.resolve lit th litTh
               in
                 case neqConvsAdd order neq lit' of
                   SOME neq => (true,neq,lits,th)
-                | NONE => (changed, neq, LiteralSet.add lits lit', th)
+                | NONE => (changed, neq, Metis_LiteralSet.add lits lit', th)
               end
           end
 
@@ -19132,50 +18496,50 @@
       val rewr_literule = mk_literule neq
 
       fun rewr_lit (lit,th) =
-          if Thm.member lit th then Rule.literalRule rewr_literule lit th
+          if Metis_Thm.member lit th then Metis_Rule.literalRule rewr_literule lit th
           else th
     in
-      LiteralSet.foldl rewr_lit th lits
+      Metis_LiteralSet.foldl rewr_lit th lits
     end;
 
 fun rewriteIdRule' order known redexes id th =
-    rewriteIdLiteralsRule' order known redexes id (Thm.clause th) th;
+    rewriteIdLiteralsRule' order known redexes id (Metis_Thm.clause th) th;
 
 (*MetisDebug
 val rewriteIdRule' = fn order => fn known => fn redexes => fn id => fn th =>
     let
 (*MetisTrace6
-      val () = Print.trace Thm.pp "Rewrite.rewriteIdRule': th" th
+      val () = Metis_Print.trace Metis_Thm.pp "Metis_Rewrite.rewriteIdRule': th" th
 *)
       val result = rewriteIdRule' order known redexes id th
 (*MetisTrace6
-      val () = Print.trace Thm.pp "Rewrite.rewriteIdRule': result" result
+      val () = Metis_Print.trace Metis_Thm.pp "Metis_Rewrite.rewriteIdRule': result" result
 *)
       val _ = not (thmReducible order known id result) orelse
               raise Bug "rewriteIdRule: should be normalized"
     in
       result
     end
-    handle Error err => raise Error ("Rewrite.rewriteIdRule:\n" ^ err);
-*)
-
-fun rewrIdConv (Rewrite {known,redexes,...}) order =
+    handle Error err => raise Error ("Metis_Rewrite.rewriteIdRule:\n" ^ err);
+*)
+
+fun rewrIdConv (Metis_Rewrite {known,redexes,...}) order =
     rewrIdConv' order known redexes;
 
 fun rewrConv rewrite order = rewrIdConv rewrite order ~1;
 
-fun rewriteIdConv (Rewrite {known,redexes,...}) order =
+fun rewriteIdConv (Metis_Rewrite {known,redexes,...}) order =
     rewriteIdConv' order known redexes;
 
 fun rewriteConv rewrite order = rewriteIdConv rewrite order ~1;
 
-fun rewriteIdLiteralsRule (Rewrite {known,redexes,...}) order =
+fun rewriteIdLiteralsRule (Metis_Rewrite {known,redexes,...}) order =
     rewriteIdLiteralsRule' order known redexes;
 
 fun rewriteLiteralsRule rewrite order =
     rewriteIdLiteralsRule rewrite order ~1;
 
-fun rewriteIdRule (Rewrite {known,redexes,...}) order =
+fun rewriteIdRule (Metis_Rewrite {known,redexes,...}) order =
     rewriteIdRule' order known redexes;
 
 fun rewriteRule rewrite order = rewriteIdRule rewrite order ~1;
@@ -19186,17 +18550,17 @@
 
 fun addSubterms id (((l,r),_) : equation) subterms =
     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)
+      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)
     in
       subterms
     end;
 
 fun sameRedexes NONE _ _ = false
-  | sameRedexes (SOME LeftToRight) (l0,_) (l,_) = Term.equal l0 l
-  | sameRedexes (SOME RightToLeft) (_,r0) (_,r) = Term.equal r0 r;
+  | sameRedexes (SOME LeftToRight) (l0,_) (l,_) = Metis_Term.equal l0 l
+  | sameRedexes (SOME RightToLeft) (_,r0) (_,r) = Metis_Term.equal r0 r;
 
 fun redexResidues NONE (l,r) = [(l,r,false),(r,l,false)]
   | redexResidues (SOME LeftToRight) (l,r) = [(l,r,true)]
@@ -19206,14 +18570,14 @@
     let
       fun checkValidRewr (l,r,ord) id' left path =
           let
-            val (((x,y),_),_) = IntMap.get known id'
-            val tm = Term.subterm (if left then x else y) path
-            val sub = Subst.match Subst.empty l tm
+            val (((x,y),_),_) = Metis_IntMap.get known id'
+            val tm = Metis_Term.subterm (if left then x else y) path
+            val sub = Metis_Subst.match Metis_Subst.empty l tm
           in
             if ord then ()
             else
               let
-                val tm' = Subst.subst (Subst.normalize sub) r
+                val tm' = Metis_Subst.subst (Metis_Subst.normalize sub) r
               in
                 if order (tm,tm') = SOME GREATER then ()
                 else raise Error "order"
@@ -19221,13 +18585,13 @@
           end
 
       fun addRed lr ((id',left,path),todo) =
-          if id <> id' andalso not (IntSet.member id' todo) andalso
+          if id <> id' andalso not (Metis_IntSet.member id' todo) andalso
              can (checkValidRewr lr id' left) path
-          then IntSet.add todo id'
+          then Metis_IntSet.add todo id'
           else todo
 
       fun findRed (lr as (l,_,_), todo) =
-          List.foldl (addRed lr) todo (TermNet.matched subterms l)
+          List.foldl (addRed lr) todo (Metis_TermNet.matched subterms l)
     in
       List.foldl findRed
     end;
@@ -19235,29 +18599,29 @@
 fun reduce1 new id (eqn0,ort0) (rpl,spl,todo,rw,changed) =
     let
       val (eq0,_) = eqn0
-      val Rewrite {order,known,redexes,subterms,waiting} = rw
+      val Metis_Rewrite {order,known,redexes,subterms,waiting} = rw
       val eqn as (eq,_) = rewriteIdEqn' order known redexes id eqn0
       val identical =
           let
             val (l0,r0) = eq0
             and (l,r) = eq
           in
-            Term.equal l l0 andalso Term.equal r r0
+            Metis_Term.equal l l0 andalso Metis_Term.equal r r0
           end
       val same_redexes = identical orelse sameRedexes ort0 eq0 eq
-      val rpl = if same_redexes then rpl else IntSet.add rpl id
-      val spl = if new orelse identical then spl else IntSet.add spl id
+      val rpl = if same_redexes then rpl else Metis_IntSet.add rpl id
+      val spl = if new orelse identical then spl else Metis_IntSet.add spl id
       val changed =
-          if not new andalso identical then changed else IntSet.add changed id
+          if not new andalso identical then changed else Metis_IntSet.add changed id
       val ort =
           if same_redexes then SOME ort0 else total orderToOrient (order eq)
     in
       case ort of
         NONE =>
         let
-          val known = IntMap.delete known id
+          val known = Metis_IntMap.delete known id
           val rw =
-              Rewrite
+              Metis_Rewrite
                 {order = order, known = known, redexes = redexes,
                  subterms = subterms, waiting = waiting}
         in
@@ -19271,7 +18635,7 @@
                 findReducibles
                   order known subterms id todo (redexResidues ort eq)
           val known =
-              if identical then known else IntMap.insert known (id,(eqn,ort))
+              if identical then known else Metis_IntMap.insert known (id,(eqn,ort))
           val redexes =
               if same_redexes then redexes
               else addRedexes id (eqn,ort) redexes
@@ -19279,7 +18643,7 @@
               if new orelse not identical then addSubterms id eqn subterms
               else subterms
           val rw =
-              Rewrite
+              Metis_Rewrite
                 {order = order, known = known, redexes = redexes,
                  subterms = subterms, waiting = waiting}
         in
@@ -19290,49 +18654,49 @@
 fun pick known set =
     let
       fun oriented id =
-          case IntMap.peek known id of
+          case Metis_IntMap.peek known id of
             SOME (x as (_, SOME _)) => SOME (id,x)
           | _ => NONE
 
       fun any id =
-          case IntMap.peek known id of SOME x => SOME (id,x) | _ => NONE
-    in
-      case IntSet.firstl oriented set of
+          case Metis_IntMap.peek known id of SOME x => SOME (id,x) | _ => NONE
+    in
+      case Metis_IntSet.firstl oriented set of
         x as SOME _ => x
-      | NONE => IntSet.firstl any set
+      | NONE => Metis_IntSet.firstl any set
     end;
 
 local
   fun cleanRedexes known redexes rpl =
-      if IntSet.null rpl then redexes
+      if Metis_IntSet.null rpl then redexes
       else
         let
-          fun filt (id,_) = not (IntSet.member id rpl)
+          fun filt (id,_) = not (Metis_IntSet.member id rpl)
 
           fun addReds (id,reds) =
-              case IntMap.peek known id of
+              case Metis_IntMap.peek known id of
                 NONE => reds
               | SOME eqn_ort => addRedexes id eqn_ort reds
 
-          val redexes = TermNet.filter filt redexes
-          val redexes = IntSet.foldl addReds redexes rpl
+          val redexes = Metis_TermNet.filter filt redexes
+          val redexes = Metis_IntSet.foldl addReds redexes rpl
         in
           redexes
         end;
 
   fun cleanSubterms known subterms spl =
-      if IntSet.null spl then subterms
+      if Metis_IntSet.null spl then subterms
       else
         let
-          fun filt (id,_,_) = not (IntSet.member id spl)
+          fun filt (id,_,_) = not (Metis_IntSet.member id spl)
 
           fun addSubtms (id,subtms) =
-              case IntMap.peek known id of
+              case Metis_IntMap.peek known id of
                 NONE => subtms
               | SOME (eqn,_) => addSubterms id eqn subtms
 
-          val subterms = TermNet.filter filt subterms
-          val subterms = IntSet.foldl addSubtms subterms spl
+          val subterms = Metis_TermNet.filter filt subterms
+          val subterms = Metis_IntSet.foldl addSubtms subterms spl
         in
           subterms
         end;
@@ -19340,15 +18704,15 @@
   fun rebuild rpl spl rw =
       let
 (*MetisTrace5
-        val ppPl = Print.ppMap IntSet.toList (Print.ppList Print.ppInt)
-        val () = Print.trace ppPl "Rewrite.rebuild: rpl" rpl
-        val () = Print.trace ppPl "Rewrite.rebuild: spl" spl
-*)
-        val Rewrite {order,known,redexes,subterms,waiting} = rw
+        val ppPl = Metis_Print.ppMap Metis_IntSet.toList (Metis_Print.ppList Metis_Print.ppInt)
+        val () = Metis_Print.trace ppPl "Metis_Rewrite.rebuild: rpl" rpl
+        val () = Metis_Print.trace ppPl "Metis_Rewrite.rebuild: spl" spl
+*)
+        val Metis_Rewrite {order,known,redexes,subterms,waiting} = rw
         val redexes = cleanRedexes known redexes rpl
         val subterms = cleanSubterms known subterms spl
       in
-        Rewrite
+        Metis_Rewrite
           {order = order,
            known = known,
            redexes = redexes,
@@ -19357,11 +18721,11 @@
       end;
 end;
 
-fun reduceAcc (rpl, spl, todo, rw as Rewrite {known,waiting,...}, changed) =
+fun reduceAcc (rpl, spl, todo, rw as Metis_Rewrite {known,waiting,...}, changed) =
     case pick known todo of
       SOME (id,eqn_ort) =>
       let
-        val todo = IntSet.delete todo id
+        val todo = Metis_IntSet.delete todo id
       in
         reduceAcc (reduce1 false id eqn_ort (rpl,spl,todo,rw,changed))
       end
@@ -19373,34 +18737,34 @@
         in
           reduceAcc (reduce1 true id eqn_ort (rpl,spl,todo,rw,changed))
         end
-      | NONE => (rebuild rpl spl rw, IntSet.toList changed);
-
-fun isReduced (Rewrite {waiting,...}) = IntSet.null waiting;
+      | NONE => (rebuild rpl spl rw, Metis_IntSet.toList changed);
+
+fun isReduced (Metis_Rewrite {waiting,...}) = Metis_IntSet.null waiting;
 
 fun reduce' rw =
     if isReduced rw then (rw,[])
-    else reduceAcc (IntSet.empty,IntSet.empty,IntSet.empty,rw,IntSet.empty);
+    else reduceAcc (Metis_IntSet.empty,Metis_IntSet.empty,Metis_IntSet.empty,rw,Metis_IntSet.empty);
 
 (*MetisDebug
 val reduce' = fn rw =>
     let
 (*MetisTrace4
-      val () = Print.trace pp "Rewrite.reduce': rw" rw
-*)
-      val Rewrite {known,order,...} = rw
-      val result as (Rewrite {known = known', ...}, _) = reduce' rw
+      val () = Metis_Print.trace pp "Metis_Rewrite.reduce': rw" rw
+*)
+      val Metis_Rewrite {known,order,...} = rw
+      val result as (Metis_Rewrite {known = known', ...}, _) = reduce' rw
 (*MetisTrace4
-      val ppResult = Print.ppPair pp (Print.ppList Print.ppInt)
-      val () = Print.trace ppResult "Rewrite.reduce': result" result
-*)
-      val ths = map (fn (id,((_,th),_)) => (id,th)) (IntMap.toList known')
+      val ppResult = Metis_Print.ppPair pp (Metis_Print.ppList Metis_Print.ppInt)
+      val () = Metis_Print.trace ppResult "Metis_Rewrite.reduce': result" result
+*)
+      val ths = map (fn (id,((_,th),_)) => (id,th)) (Metis_IntMap.toList known')
       val _ =
           not (List.exists (uncurry (thmReducible order known')) ths) orelse
-          raise Bug "Rewrite.reduce': not fully reduced"
+          raise Bug "Metis_Rewrite.reduce': not fully reduced"
     in
       result
     end
-    handle Error err => raise Bug ("Rewrite.reduce': shouldn't fail\n" ^ err);
+    handle Error err => raise Bug ("Metis_Rewrite.reduce': shouldn't fail\n" ^ err);
 *)
 
 fun reduce rw = fst (reduce' rw);
@@ -19423,7 +18787,6 @@
 val rewrite = orderedRewrite (K (SOME GREATER));
 
 end
-end;
 
 (**** Original file: Units.sig ****)
 
@@ -19432,14 +18795,14 @@
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Units =
+signature Metis_Units =
 sig
 
 (* ------------------------------------------------------------------------- *)
 (* A type of unit store.                                                     *)
 (* ------------------------------------------------------------------------- *)
 
-type unitThm = Metis.Literal.literal * Metis.Thm.thm
+type unitThm = Metis_Literal.literal * Metis_Thm.thm
 
 type units
 
@@ -19453,7 +18816,7 @@
 
 val toString : units -> string
 
-val pp : units Metis.Print.pp
+val pp : units Metis_Print.pp
 
 (* ------------------------------------------------------------------------- *)
 (* Add units into the store.                                                 *)
@@ -19467,73 +18830,64 @@
 (* Matching.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-val match : units -> Metis.Literal.literal -> (unitThm * Metis.Subst.subst) option
+val match : units -> Metis_Literal.literal -> (unitThm * Metis_Subst.subst) option
 
 (* ------------------------------------------------------------------------- *)
 (* Reducing by repeated matching and resolution.                             *)
 (* ------------------------------------------------------------------------- *)
 
-val reduce : units -> Metis.Rule.rule
+val reduce : units -> Metis_Rule.rule
 
 end
 
 (**** Original file: Units.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* A STORE FOR UNIT THEOREMS                                                 *)
 (* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Units :> Units =
+structure Metis_Units :> Metis_Units =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of unit store.                                                     *)
 (* ------------------------------------------------------------------------- *)
 
-type unitThm = Literal.literal * Thm.thm;
-
-datatype units = Units of unitThm LiteralNet.literalNet;
+type unitThm = Metis_Literal.literal * Metis_Thm.thm;
+
+datatype units = Metis_Units of unitThm Metis_LiteralNet.literalNet;
 
 (* ------------------------------------------------------------------------- *)
 (* Basic operations.                                                         *)
 (* ------------------------------------------------------------------------- *)
 
-val empty = Units (LiteralNet.new {fifo = false});
-
-fun size (Units net) = LiteralNet.size net;
+val empty = Metis_Units (Metis_LiteralNet.new {fifo = false});
+
+fun size (Metis_Units net) = Metis_LiteralNet.size net;
 
 fun toString units = "U{" ^ Int.toString (size units) ^ "}";
 
-val pp = Print.ppMap toString Print.ppString;
+val pp = Metis_Print.ppMap toString Metis_Print.ppString;
 
 (* ------------------------------------------------------------------------- *)
 (* Add units into the store.                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun add (units as Units net) (uTh as (lit,th)) =
-    let
-      val net = LiteralNet.insert net (lit,uTh)
-    in
-      case total Literal.sym lit of
-        NONE => Units net
+fun add (units as Metis_Units net) (uTh as (lit,th)) =
+    let
+      val net = Metis_LiteralNet.insert net (lit,uTh)
+    in
+      case total Metis_Literal.sym lit of
+        NONE => Metis_Units net
       | SOME (lit' as (pol,_)) =>
         let
-          val th' = (if pol then Rule.symEq else Rule.symNeq) lit th
-          val net = LiteralNet.insert net (lit',(lit',th'))
-        in
-          Units net
+          val th' = (if pol then Metis_Rule.symEq else Metis_Rule.symNeq) lit th
+          val net = Metis_LiteralNet.insert net (lit',(lit',th'))
+        in
+          Metis_Units net
         end
     end;
 
@@ -19543,14 +18897,14 @@
 (* Matching.                                                                 *)
 (* ------------------------------------------------------------------------- *)
 
-fun match (Units net) lit =
+fun match (Metis_Units net) lit =
     let
       fun check (uTh as (lit',_)) =
-          case total (Literal.match Subst.empty lit') lit of
+          case total (Metis_Literal.match Metis_Subst.empty lit') lit of
             NONE => NONE
           | SOME sub => SOME (uTh,sub)
     in
-      first check (LiteralNet.match net lit)
+      first check (Metis_LiteralNet.match net lit)
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -19560,41 +18914,40 @@
 fun reduce units =
     let
       fun red1 (lit,news_th) =
-          case total Literal.destIrrefl lit of
+          case total Metis_Literal.destIrrefl lit of
             SOME tm =>
             let
               val (news,th) = news_th
-              val th = Thm.resolve lit th (Thm.refl tm)
+              val th = Metis_Thm.resolve lit th (Metis_Thm.refl tm)
             in
               (news,th)
             end
           | NONE =>
             let
-              val lit' = Literal.negate lit
+              val lit' = Metis_Literal.negate lit
             in
               case match units lit' of
                 NONE => news_th
               | SOME ((_,rth),sub) =>
                 let
                   val (news,th) = news_th
-                  val rth = Thm.subst sub rth
-                  val th = Thm.resolve lit th rth
-                  val new = LiteralSet.delete (Thm.clause rth) lit'
-                  val news = LiteralSet.union new news
+                  val rth = Metis_Thm.subst sub rth
+                  val th = Metis_Thm.resolve lit th rth
+                  val new = Metis_LiteralSet.delete (Metis_Thm.clause rth) lit'
+                  val news = Metis_LiteralSet.union new news
                 in
                   (news,th)
                 end
             end
 
       fun red (news,th) =
-          if LiteralSet.null news then th
-          else red (LiteralSet.foldl red1 (LiteralSet.empty,th) news)
-    in
-      fn th => Rule.removeSym (red (Thm.clause th, th))
-    end;
-
-end
-end;
+          if Metis_LiteralSet.null news then th
+          else red (Metis_LiteralSet.foldl red1 (Metis_LiteralSet.empty,th) news)
+    in
+      fn th => Metis_Rule.removeSym (red (Metis_Thm.clause th, th))
+    end;
+
+end
 
 (**** Original file: Clause.sig ****)
 
@@ -19603,7 +18956,7 @@
 (* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Clause =
+signature Metis_Clause =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -19616,13 +18969,13 @@
   | PositiveLiteralOrder
 
 type parameters =
-     {ordering : Metis.KnuthBendixOrder.kbo,
+     {ordering : Metis_KnuthBendixOrder.kbo,
       orderLiterals : literalOrder,
       orderTerms : bool}
 
 type clauseId = int
 
-type clauseInfo = {parameters : parameters, id : clauseId, thm : Metis.Thm.thm}
+type clauseInfo = {parameters : parameters, id : clauseId, thm : Metis_Thm.thm}
 
 type clause
 
@@ -19640,11 +18993,11 @@
 
 val id : clause -> clauseId
 
-val thm : clause -> Metis.Thm.thm
+val thm : clause -> Metis_Thm.thm
 
 val equalThms : clause -> clause -> bool
 
-val literals : clause -> Metis.Thm.clause
+val literals : clause -> Metis_Thm.clause
 
 val isTautology : clause -> bool
 
@@ -19654,21 +19007,21 @@
 (* The term ordering is used to cut down inferences.                         *)
 (* ------------------------------------------------------------------------- *)
 
-val largestLiterals : clause -> Metis.LiteralSet.set
+val largestLiterals : clause -> Metis_LiteralSet.set
 
 val largestEquations :
-    clause -> (Metis.Literal.literal * Metis.Rewrite.orient * Metis.Term.term) list
+    clause -> (Metis_Literal.literal * Metis_Rewrite.orient * Metis_Term.term) list
 
 val largestSubterms :
-    clause -> (Metis.Literal.literal * Metis.Term.path * Metis.Term.term) list
-
-val allSubterms : clause -> (Metis.Literal.literal * Metis.Term.path * Metis.Term.term) list
+    clause -> (Metis_Literal.literal * Metis_Term.path * Metis_Term.term) list
+
+val allSubterms : clause -> (Metis_Literal.literal * Metis_Term.path * Metis_Term.term) list
 
 (* ------------------------------------------------------------------------- *)
 (* Subsumption.                                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val subsumes : clause Metis.Subsume.subsume -> clause -> bool
+val subsumes : clause Metis_Subsume.subsume -> clause -> bool
 
 (* ------------------------------------------------------------------------- *)
 (* Simplifying rules: these preserve the clause id.                          *)
@@ -19678,9 +19031,9 @@
 
 val simplify : clause -> clause option
 
-val reduce : Metis.Units.units -> clause -> clause
-
-val rewrite : Metis.Rewrite.rewrite -> clause -> clause
+val reduce : Metis_Units.units -> clause -> clause
+
+val rewrite : Metis_Rewrite.rewrite -> clause -> clause
 
 (* ------------------------------------------------------------------------- *)
 (* Inference rules: these generate new clause ids.                           *)
@@ -19688,11 +19041,11 @@
 
 val factor : clause -> clause list
 
-val resolve : clause * Metis.Literal.literal -> clause * Metis.Literal.literal -> clause
+val resolve : clause * Metis_Literal.literal -> clause * Metis_Literal.literal -> clause
 
 val paramodulate :
-    clause * Metis.Literal.literal * Metis.Rewrite.orient * Metis.Term.term ->
-    clause * Metis.Literal.literal * Metis.Term.path * Metis.Term.term -> clause
+    clause * Metis_Literal.literal * Metis_Rewrite.orient * Metis_Term.term ->
+    clause * Metis_Literal.literal * Metis_Term.path * Metis_Term.term -> clause
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty printing.                                                          *)
@@ -19700,7 +19053,7 @@
 
 val showId : bool Unsynchronized.ref
 
-val pp : clause Metis.Print.pp
+val pp : clause Metis_Print.pp
 
 val toString : clause -> string
 
@@ -19708,24 +19061,15 @@
 
 (**** Original file: Clause.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* CLAUSE = ID + THEOREM                                                     *)
 (* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Clause :> Clause =
+structure Metis_Clause :> Metis_Clause =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Helper functions.                                                         *)
@@ -19735,7 +19079,9 @@
     let
       val r = Unsynchronized.ref 0
     in
-      fn () => case r of Unsynchronized.ref n => let val () = r := n + 1 in n end
+      (* MODIFIED by Jasmin Blanchette *)
+      fn () => CRITICAL (fn () =>
+        case r of Unsynchronized.ref n => let val () = r := n + 1 in n end)
     end;
 
 (* ------------------------------------------------------------------------- *)
@@ -19748,15 +19094,15 @@
   | PositiveLiteralOrder;
 
 type parameters =
-     {ordering : KnuthBendixOrder.kbo,
+     {ordering : Metis_KnuthBendixOrder.kbo,
       orderLiterals : literalOrder,
       orderTerms : bool};
 
 type clauseId = int;
 
-type clauseInfo = {parameters : parameters, id : clauseId, thm : Thm.thm};
-
-datatype clause = Clause of clauseInfo;
+type clauseInfo = {parameters : parameters, id : clauseId, thm : Metis_Thm.thm};
+
+datatype clause = Metis_Clause of clauseInfo;
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty printing.                                                          *)
@@ -19765,48 +19111,48 @@
 val showId = Unsynchronized.ref false;
 
 local
-  val ppIdThm = Print.ppPair Print.ppInt Thm.pp;
-in
-  fun pp (Clause {id,thm,...}) =
-      if !showId then ppIdThm (id,thm) else Thm.pp thm;
-end;
-
-fun toString cl = Print.toString pp cl;
+  val ppIdThm = Metis_Print.ppPair Metis_Print.ppInt Metis_Thm.pp;
+in
+  fun pp (Metis_Clause {id,thm,...}) =
+      if !showId then ppIdThm (id,thm) else Metis_Thm.pp thm;
+end;
+
+fun toString cl = Metis_Print.toString pp cl;
 
 (* ------------------------------------------------------------------------- *)
 (* Basic operations.                                                         *)
 (* ------------------------------------------------------------------------- *)
 
 val default : parameters =
-    {ordering = KnuthBendixOrder.default,
+    {ordering = Metis_KnuthBendixOrder.default,
      orderLiterals = UnsignedLiteralOrder (* PositiveLiteralOrder *) (* MODIFIED by Jasmin Blanchette *),
      orderTerms = true};
 
-fun mk info = Clause info
-
-fun dest (Clause info) = info;
-
-fun id (Clause {id = i, ...}) = i;
-
-fun thm (Clause {thm = th, ...}) = th;
-
-fun equalThms cl cl' = Thm.equal (thm cl) (thm cl');
+fun mk info = Metis_Clause info
+
+fun dest (Metis_Clause info) = info;
+
+fun id (Metis_Clause {id = i, ...}) = i;
+
+fun thm (Metis_Clause {thm = th, ...}) = th;
+
+fun equalThms cl cl' = Metis_Thm.equal (thm cl) (thm cl');
 
 fun new parameters thm =
-    Clause {parameters = parameters, id = newId (), thm = thm};
-
-fun literals cl = Thm.clause (thm cl);
-
-fun isTautology (Clause {thm,...}) = Thm.isTautology thm;
-
-fun isContradiction (Clause {thm,...}) = Thm.isContradiction thm;
+    Metis_Clause {parameters = parameters, id = newId (), thm = thm};
+
+fun literals cl = Metis_Thm.clause (thm cl);
+
+fun isTautology (Metis_Clause {thm,...}) = Metis_Thm.isTautology thm;
+
+fun isContradiction (Metis_Clause {thm,...}) = Metis_Thm.isContradiction thm;
 
 (* ------------------------------------------------------------------------- *)
 (* The term ordering is used to cut down inferences.                         *)
 (* ------------------------------------------------------------------------- *)
 
 fun strictlyLess ordering x_y =
-    case KnuthBendixOrder.compare ordering x_y of
+    case Metis_KnuthBendixOrder.compare ordering x_y of
       SOME LESS => true
     | _ => false;
 
@@ -19815,8 +19161,8 @@
 
 local
   fun atomToTerms atm =
-      case total Atom.destEq atm of
-        NONE => [Term.Fn atm]
+      case total Metis_Atom.destEq atm of
+        NONE => [Metis_Term.Fn atm]
       | SOME (l,r) => [l,r];
 
   fun notStrictlyLess ordering (xs,ys) =
@@ -19833,19 +19179,19 @@
         let
           fun addLit ((_,atm),acc) = atomToTerms atm @ acc
 
-          val tms = LiteralSet.foldl addLit [] lits
+          val tms = Metis_LiteralSet.foldl addLit [] lits
         in
           fn (_,atm') => notStrictlyLess ordering (atomToTerms atm', tms)
         end
       | PositiveLiteralOrder =>
-        case LiteralSet.findl (K true) lits of
+        case Metis_LiteralSet.findl (K true) lits of
           NONE => K true
         | SOME (pol,_) =>
           let
             fun addLit ((p,atm),acc) =
                 if p = pol then atomToTerms atm @ acc else acc
 
-            val tms = LiteralSet.foldl addLit [] lits
+            val tms = Metis_LiteralSet.foldl addLit [] lits
           in
             fn (pol',atm') =>
                if pol <> pol' then pol
@@ -19853,44 +19199,44 @@
           end;
 end;
 
-fun largestLiterals (Clause {parameters,thm,...}) =
-    let
-      val litSet = Thm.clause thm
+fun largestLiterals (Metis_Clause {parameters,thm,...}) =
+    let
+      val litSet = Metis_Thm.clause thm
       val isLarger = isLargerLiteral parameters litSet
-      fun addLit (lit,s) = if isLarger lit then LiteralSet.add s lit else s
-    in
-      LiteralSet.foldr addLit LiteralSet.empty litSet
+      fun addLit (lit,s) = if isLarger lit then Metis_LiteralSet.add s lit else s
+    in
+      Metis_LiteralSet.foldr addLit Metis_LiteralSet.empty litSet
     end;
 
 (*MetisTrace6
 val largestLiterals = fn cl =>
     let
-      val ppResult = LiteralSet.pp
-      val () = Print.trace pp "Clause.largestLiterals: cl" cl
+      val ppResult = Metis_LiteralSet.pp
+      val () = Metis_Print.trace pp "Metis_Clause.largestLiterals: cl" cl
       val result = largestLiterals cl
-      val () = Print.trace ppResult "Clause.largestLiterals: result" result
+      val () = Metis_Print.trace ppResult "Metis_Clause.largestLiterals: result" result
     in
       result
     end;
 *)
 
-fun largestEquations (cl as Clause {parameters,...}) =
+fun largestEquations (cl as Metis_Clause {parameters,...}) =
     let
       fun addEq lit ort (l_r as (l,_)) acc =
           if isLargerTerm parameters l_r then (lit,ort,l) :: acc else acc
 
       fun addLit (lit,acc) =
-          case total Literal.destEq lit of
+          case total Metis_Literal.destEq lit of
             NONE => acc
           | SOME (l,r) =>
             let
-              val acc = addEq lit Rewrite.RightToLeft (r,l) acc
-              val acc = addEq lit Rewrite.LeftToRight (l,r) acc
+              val acc = addEq lit Metis_Rewrite.RightToLeft (r,l) acc
+              val acc = addEq lit Metis_Rewrite.LeftToRight (l,r) acc
             in
               acc
             end
     in
-      LiteralSet.foldr addLit [] (largestLiterals cl)
+      Metis_LiteralSet.foldr addLit [] (largestLiterals cl)
     end;
 
 local
@@ -19898,88 +19244,88 @@
       let
         fun addTm ((path,tm),acc) = (lit,path,tm) :: acc
       in
-        foldl addTm acc (Literal.nonVarTypedSubterms lit)
-      end;
-in
-  fun largestSubterms cl = LiteralSet.foldl addLit [] (largestLiterals cl);
-
-  fun allSubterms cl = LiteralSet.foldl addLit [] (literals cl);
+        foldl addTm acc (Metis_Literal.nonVarTypedSubterms lit)
+      end;
+in
+  fun largestSubterms cl = Metis_LiteralSet.foldl addLit [] (largestLiterals cl);
+
+  fun allSubterms cl = Metis_LiteralSet.foldl addLit [] (literals cl);
 end;
 
 (* ------------------------------------------------------------------------- *)
 (* Subsumption.                                                              *)
 (* ------------------------------------------------------------------------- *)
 
-fun subsumes (subs : clause Subsume.subsume) cl =
-    Subsume.isStrictlySubsumed subs (literals cl);
+fun subsumes (subs : clause Metis_Subsume.subsume) cl =
+    Metis_Subsume.isStrictlySubsumed subs (literals cl);
 
 (* ------------------------------------------------------------------------- *)
 (* Simplifying rules: these preserve the clause id.                          *)
 (* ------------------------------------------------------------------------- *)
 
-fun freshVars (Clause {parameters,id,thm}) =
-    Clause {parameters = parameters, id = id, thm = Rule.freshVars thm};
-
-fun simplify (Clause {parameters,id,thm}) =
-    case Rule.simplify thm of
+fun freshVars (Metis_Clause {parameters,id,thm}) =
+    Metis_Clause {parameters = parameters, id = id, thm = Metis_Rule.freshVars thm};
+
+fun simplify (Metis_Clause {parameters,id,thm}) =
+    case Metis_Rule.simplify thm of
       NONE => NONE
-    | SOME thm => SOME (Clause {parameters = parameters, id = id, thm = thm});
-
-fun reduce units (Clause {parameters,id,thm}) =
-    Clause {parameters = parameters, id = id, thm = Units.reduce units thm};
-
-fun rewrite rewr (cl as Clause {parameters,id,thm}) =
+    | SOME thm => SOME (Metis_Clause {parameters = parameters, id = id, thm = thm});
+
+fun reduce units (Metis_Clause {parameters,id,thm}) =
+    Metis_Clause {parameters = parameters, id = id, thm = Metis_Units.reduce units thm};
+
+fun rewrite rewr (cl as Metis_Clause {parameters,id,thm}) =
     let
       fun simp th =
           let
             val {ordering,...} = parameters
-            val cmp = KnuthBendixOrder.compare ordering
-          in
-            Rewrite.rewriteIdRule rewr cmp id th
+            val cmp = Metis_KnuthBendixOrder.compare ordering
+          in
+            Metis_Rewrite.rewriteIdRule rewr cmp id th
           end
 
 (*MetisTrace4
-      val () = Print.trace Rewrite.pp "Clause.rewrite: rewr" rewr
-      val () = Print.trace Print.ppInt "Clause.rewrite: id" id
-      val () = Print.trace pp "Clause.rewrite: cl" cl
+      val () = Metis_Print.trace Metis_Rewrite.pp "Metis_Clause.rewrite: rewr" rewr
+      val () = Metis_Print.trace Metis_Print.ppInt "Metis_Clause.rewrite: id" id
+      val () = Metis_Print.trace pp "Metis_Clause.rewrite: cl" cl
 *)
 
       val thm =
-          case Rewrite.peek rewr id of
+          case Metis_Rewrite.peek rewr id of
             NONE => simp thm
-          | SOME ((_,thm),_) => if Rewrite.isReduced rewr then thm else simp thm
-
-      val result = Clause {parameters = parameters, id = id, thm = thm}
+          | SOME ((_,thm),_) => if Metis_Rewrite.isReduced rewr then thm else simp thm
+
+      val result = Metis_Clause {parameters = parameters, id = id, thm = thm}
 
 (*MetisTrace4
-      val () = Print.trace pp "Clause.rewrite: result" result
+      val () = Metis_Print.trace pp "Metis_Clause.rewrite: result" result
 *)
     in
       result
     end
 (*MetisDebug
-    handle Error err => raise Error ("Clause.rewrite:\n" ^ err);
+    handle Error err => raise Error ("Metis_Clause.rewrite:\n" ^ err);
 *)
 
 (* ------------------------------------------------------------------------- *)
 (* Inference rules: these generate new clause ids.                           *)
 (* ------------------------------------------------------------------------- *)
 
-fun factor (cl as Clause {parameters,thm,...}) =
+fun factor (cl as Metis_Clause {parameters,thm,...}) =
     let
       val lits = largestLiterals cl
 
-      fun apply sub = new parameters (Thm.subst sub thm)
-    in
-      map apply (Rule.factor' lits)
+      fun apply sub = new parameters (Metis_Thm.subst sub thm)
+    in
+      map apply (Metis_Rule.factor' lits)
     end;
 
 (*MetisTrace5
 val factor = fn cl =>
     let
-      val () = Print.trace pp "Clause.factor: cl" cl
+      val () = Metis_Print.trace pp "Metis_Clause.factor: cl" cl
       val result = factor cl
-      val () = Print.trace (Print.ppList pp) "Clause.factor: result" result
+      val () = Metis_Print.trace (Metis_Print.ppList pp) "Metis_Clause.factor: result" result
     in
       result
     end;
@@ -19988,38 +19334,38 @@
 fun resolve (cl1,lit1) (cl2,lit2) =
     let
 (*MetisTrace5
-      val () = Print.trace pp "Clause.resolve: cl1" cl1
-      val () = Print.trace Literal.pp "Clause.resolve: lit1" lit1
-      val () = Print.trace pp "Clause.resolve: cl2" cl2
-      val () = Print.trace Literal.pp "Clause.resolve: lit2" lit2
-*)
-      val Clause {parameters, thm = th1, ...} = cl1
-      and Clause {thm = th2, ...} = cl2
-      val sub = Literal.unify Subst.empty lit1 (Literal.negate lit2)
+      val () = Metis_Print.trace pp "Metis_Clause.resolve: cl1" cl1
+      val () = Metis_Print.trace Metis_Literal.pp "Metis_Clause.resolve: lit1" lit1
+      val () = Metis_Print.trace pp "Metis_Clause.resolve: cl2" cl2
+      val () = Metis_Print.trace Metis_Literal.pp "Metis_Clause.resolve: lit2" lit2
+*)
+      val Metis_Clause {parameters, thm = th1, ...} = cl1
+      and Metis_Clause {thm = th2, ...} = cl2
+      val sub = Metis_Literal.unify Metis_Subst.empty lit1 (Metis_Literal.negate lit2)
 (*MetisTrace5
-      val () = Print.trace Subst.pp "Clause.resolve: sub" sub
-*)
-      val lit1 = Literal.subst sub lit1
-      val lit2 = Literal.negate lit1
-      val th1 = Thm.subst sub th1
-      and th2 = Thm.subst sub th2
-      val _ = isLargerLiteral parameters (Thm.clause th1) lit1 orelse
+      val () = Metis_Print.trace Metis_Subst.pp "Metis_Clause.resolve: sub" sub
+*)
+      val lit1 = Metis_Literal.subst sub lit1
+      val lit2 = Metis_Literal.negate lit1
+      val th1 = Metis_Thm.subst sub th1
+      and th2 = Metis_Thm.subst sub th2
+      val _ = isLargerLiteral parameters (Metis_Thm.clause th1) lit1 orelse
 (*MetisTrace5
-              (trace "Clause.resolve: th1 violates ordering\n"; false) orelse
+              (trace "Metis_Clause.resolve: th1 violates ordering\n"; false) orelse
 *)
               raise Error "resolve: clause1: ordering constraints"
-      val _ = isLargerLiteral parameters (Thm.clause th2) lit2 orelse
+      val _ = isLargerLiteral parameters (Metis_Thm.clause th2) lit2 orelse
 (*MetisTrace5
-              (trace "Clause.resolve: th2 violates ordering\n"; false) orelse
+              (trace "Metis_Clause.resolve: th2 violates ordering\n"; false) orelse
 *)
               raise Error "resolve: clause2: ordering constraints"
-      val th = Thm.resolve lit1 th1 th2
+      val th = Metis_Thm.resolve lit1 th1 th2
 (*MetisTrace5
-      val () = Print.trace Thm.pp "Clause.resolve: th" th
-*)
-      val cl = Clause {parameters = parameters, id = newId (), thm = th}
+      val () = Metis_Print.trace Metis_Thm.pp "Metis_Clause.resolve: th" th
+*)
+      val cl = Metis_Clause {parameters = parameters, id = newId (), thm = th}
 (*MetisTrace5
-      val () = Print.trace pp "Clause.resolve: cl" cl
+      val () = Metis_Print.trace pp "Metis_Clause.resolve: cl" cl
 *)
     in
       cl
@@ -20028,56 +19374,55 @@
 fun paramodulate (cl1,lit1,ort1,tm1) (cl2,lit2,path2,tm2) =
     let
 (*MetisTrace5
-      val () = Print.trace pp "Clause.paramodulate: cl1" cl1
-      val () = Print.trace Literal.pp "Clause.paramodulate: lit1" lit1
-      val () = Print.trace Rewrite.ppOrient "Clause.paramodulate: ort1" ort1
-      val () = Print.trace Term.pp "Clause.paramodulate: tm1" tm1
-      val () = Print.trace pp "Clause.paramodulate: cl2" cl2
-      val () = Print.trace Literal.pp "Clause.paramodulate: lit2" lit2
-      val () = Print.trace Term.ppPath "Clause.paramodulate: path2" path2
-      val () = Print.trace Term.pp "Clause.paramodulate: tm2" tm2
-*)
-      val Clause {parameters, thm = th1, ...} = cl1
-      and Clause {thm = th2, ...} = cl2
-      val sub = Subst.unify Subst.empty tm1 tm2
-      val lit1 = Literal.subst sub lit1
-      and lit2 = Literal.subst sub lit2
-      and th1 = Thm.subst sub th1
-      and th2 = Thm.subst sub th2
-
-      val _ = isLargerLiteral parameters (Thm.clause th1) lit1 orelse
-              raise Error "Clause.paramodulate: with clause: ordering"
-      val _ = isLargerLiteral parameters (Thm.clause th2) lit2 orelse
-              raise Error "Clause.paramodulate: into clause: ordering"
-
-      val eqn = (Literal.destEq lit1, th1)
+      val () = Metis_Print.trace pp "Metis_Clause.paramodulate: cl1" cl1
+      val () = Metis_Print.trace Metis_Literal.pp "Metis_Clause.paramodulate: lit1" lit1
+      val () = Metis_Print.trace Metis_Rewrite.ppOrient "Metis_Clause.paramodulate: ort1" ort1
+      val () = Metis_Print.trace Metis_Term.pp "Metis_Clause.paramodulate: tm1" tm1
+      val () = Metis_Print.trace pp "Metis_Clause.paramodulate: cl2" cl2
+      val () = Metis_Print.trace Metis_Literal.pp "Metis_Clause.paramodulate: lit2" lit2
+      val () = Metis_Print.trace Metis_Term.ppPath "Metis_Clause.paramodulate: path2" path2
+      val () = Metis_Print.trace Metis_Term.pp "Metis_Clause.paramodulate: tm2" tm2
+*)
+      val Metis_Clause {parameters, thm = th1, ...} = cl1
+      and Metis_Clause {thm = th2, ...} = cl2
+      val sub = Metis_Subst.unify Metis_Subst.empty tm1 tm2
+      val lit1 = Metis_Literal.subst sub lit1
+      and lit2 = Metis_Literal.subst sub lit2
+      and th1 = Metis_Thm.subst sub th1
+      and th2 = Metis_Thm.subst sub th2
+
+      val _ = isLargerLiteral parameters (Metis_Thm.clause th1) lit1 orelse
+              raise Error "Metis_Clause.paramodulate: with clause: ordering"
+      val _ = isLargerLiteral parameters (Metis_Thm.clause th2) lit2 orelse
+              raise Error "Metis_Clause.paramodulate: into clause: ordering"
+
+      val eqn = (Metis_Literal.destEq lit1, th1)
       val eqn as (l_r,_) =
           case ort1 of
-            Rewrite.LeftToRight => eqn
-          | Rewrite.RightToLeft => Rule.symEqn eqn
+            Metis_Rewrite.LeftToRight => eqn
+          | Metis_Rewrite.RightToLeft => Metis_Rule.symEqn eqn
 (*MetisTrace6
-      val () = Print.trace Rule.ppEquation "Clause.paramodulate: eqn" eqn
+      val () = Metis_Print.trace Metis_Rule.ppEquation "Metis_Clause.paramodulate: eqn" eqn
 *)
       val _ = isLargerTerm parameters l_r orelse
-              raise Error "Clause.paramodulate: equation: ordering constraints"
-      val th = Rule.rewrRule eqn lit2 path2 th2
+              raise Error "Metis_Clause.paramodulate: equation: ordering constraints"
+      val th = Metis_Rule.rewrRule eqn lit2 path2 th2
 (*MetisTrace5
-      val () = Print.trace Thm.pp "Clause.paramodulate: th" th
-*)
-    in
-      Clause {parameters = parameters, id = newId (), thm = th}
+      val () = Metis_Print.trace Metis_Thm.pp "Metis_Clause.paramodulate: th" th
+*)
+    in
+      Metis_Clause {parameters = parameters, id = newId (), thm = th}
     end
 (*MetisTrace5
     handle Error err =>
       let
-        val () = trace ("Clause.paramodulate: failed: " ^ err ^ "\n")
+        val () = trace ("Metis_Clause.paramodulate: failed: " ^ err ^ "\n")
       in
         raise Error err
       end;
 *)
 
 end
-end;
 
 (**** Original file: Active.sig ****)
 
@@ -20086,7 +19431,7 @@
 (* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Active =
+signature Metis_Active =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -20099,7 +19444,7 @@
       rewrite : bool}
 
 type parameters =
-     {clause : Metis.Clause.parameters,
+     {clause : Metis_Clause.parameters,
       prefactor : simplify,
       postfactor : simplify}
 
@@ -20113,50 +19458,41 @@
 
 val size : active -> int
 
-val saturation : active -> Metis.Clause.clause list
+val saturation : active -> Metis_Clause.clause list
 
 (* ------------------------------------------------------------------------- *)
 (* Create a new active clause set and initialize clauses.                    *)
 (* ------------------------------------------------------------------------- *)
 
 val new :
-    parameters -> {axioms : Metis.Thm.thm list, conjecture : Metis.Thm.thm list} ->
-    active * {axioms : Metis.Clause.clause list, conjecture : Metis.Clause.clause list}
+    parameters -> {axioms : Metis_Thm.thm list, conjecture : Metis_Thm.thm list} ->
+    active * {axioms : Metis_Clause.clause list, conjecture : Metis_Clause.clause list}
 
 (* ------------------------------------------------------------------------- *)
 (* Add a clause into the active set and deduce all consequences.             *)
 (* ------------------------------------------------------------------------- *)
 
-val add : active -> Metis.Clause.clause -> active * Metis.Clause.clause list
+val add : active -> Metis_Clause.clause -> active * Metis_Clause.clause list
 
 (* ------------------------------------------------------------------------- *)
 (* Pretty printing.                                                          *)
 (* ------------------------------------------------------------------------- *)
 
-val pp : active Metis.Print.pp
+val pp : active Metis_Print.pp
 
 end
 
 (**** Original file: Active.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* THE ACTIVE SET OF CLAUSES                                                 *)
 (* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Active :> Active =
+structure Metis_Active :> Metis_Active =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* Helper functions.                                                         *)
@@ -20168,23 +19504,23 @@
       let
         fun add (cl,rw) =
             let
-              val {id, thm = th, ...} = Clause.dest cl
-            in
-              case total Thm.destUnitEq th of
-                SOME l_r => Rewrite.add rw (id,(l_r,th))
+              val {id, thm = th, ...} = Metis_Clause.dest cl
+            in
+              case total Metis_Thm.destUnitEq th of
+                SOME l_r => Metis_Rewrite.add rw (id,(l_r,th))
               | NONE => rw
             end
       in
-        foldl add (Rewrite.new (KnuthBendixOrder.compare ordering))
+        foldl add (Metis_Rewrite.new (Metis_KnuthBendixOrder.compare ordering))
       end;
 
   fun allFactors red =
       let
         fun allClause cl =
-            List.all red (cl :: Clause.factor cl) orelse
-            let
-              val () = Print.trace Clause.pp
-                         "Active.isSaturated.allFactors: cl" cl
+            List.all red (cl :: Metis_Clause.factor cl) orelse
+            let
+              val () = Metis_Print.trace Metis_Clause.pp
+                         "Metis_Active.isSaturated.allFactors: cl" cl
             in
               false
             end
@@ -20197,30 +19533,30 @@
         fun allClause2 cl_lit cl =
             let
               fun allLiteral2 lit =
-                  case total (Clause.resolve cl_lit) (cl,lit) of
+                  case total (Metis_Clause.resolve cl_lit) (cl,lit) of
                     NONE => true
                   | SOME cl => allFactors red [cl]
             in
-              LiteralSet.all allLiteral2 (Clause.literals cl)
+              Metis_LiteralSet.all allLiteral2 (Metis_Clause.literals cl)
             end orelse
             let
-              val () = Print.trace Clause.pp
-                         "Active.isSaturated.allResolutions: cl2" cl
+              val () = Metis_Print.trace Metis_Clause.pp
+                         "Metis_Active.isSaturated.allResolutions: cl2" cl
             in
               false
             end
 
         fun allClause1 allCls cl =
             let
-              val cl = Clause.freshVars cl
+              val cl = Metis_Clause.freshVars cl
 
               fun allLiteral1 lit = List.all (allClause2 (cl,lit)) allCls
             in
-              LiteralSet.all allLiteral1 (Clause.literals cl)
+              Metis_LiteralSet.all allLiteral1 (Metis_Clause.literals cl)
             end orelse
             let
-              val () = Print.trace Clause.pp
-                         "Active.isSaturated.allResolutions: cl1" cl
+              val () = Metis_Print.trace Metis_Clause.pp
+                         "Metis_Active.isSaturated.allResolutions: cl1" cl
             in
               false
             end
@@ -20237,60 +19573,60 @@
             let
               fun allLiteral2 lit =
                   let
-                    val para = Clause.paramodulate cl_lit_ort_tm
+                    val para = Metis_Clause.paramodulate cl_lit_ort_tm
 
                     fun allSubterms (path,tm) =
                         case total para (cl,lit,path,tm) of
                           NONE => true
                         | SOME cl => allFactors red [cl]
                   in
-                    List.all allSubterms (Literal.nonVarTypedSubterms lit)
+                    List.all allSubterms (Metis_Literal.nonVarTypedSubterms lit)
                   end orelse
                   let
-                    val () = Print.trace Literal.pp
-                               "Active.isSaturated.allParamodulations: lit2" lit
+                    val () = Metis_Print.trace Metis_Literal.pp
+                               "Metis_Active.isSaturated.allParamodulations: lit2" lit
                   in
                     false
                   end
             in
-              LiteralSet.all allLiteral2 (Clause.literals cl)
+              Metis_LiteralSet.all allLiteral2 (Metis_Clause.literals cl)
             end orelse
             let
-              val () = Print.trace Clause.pp
-                         "Active.isSaturated.allParamodulations: cl2" cl
+              val () = Metis_Print.trace Metis_Clause.pp
+                         "Metis_Active.isSaturated.allParamodulations: cl2" cl
               val (_,_,ort,_) = cl_lit_ort_tm
-              val () = Print.trace Rewrite.ppOrient
-                         "Active.isSaturated.allParamodulations: ort1" ort
+              val () = Metis_Print.trace Metis_Rewrite.ppOrient
+                         "Metis_Active.isSaturated.allParamodulations: ort1" ort
             in
               false
             end
 
         fun allClause1 cl =
             let
-              val cl = Clause.freshVars cl
+              val cl = Metis_Clause.freshVars cl
 
               fun allLiteral1 lit =
                   let
                     fun allCl2 x = List.all (allClause2 x) cls
                   in
-                    case total Literal.destEq lit of
+                    case total Metis_Literal.destEq lit of
                       NONE => true
                     | SOME (l,r) =>
-                      allCl2 (cl,lit,Rewrite.LeftToRight,l) andalso
-                      allCl2 (cl,lit,Rewrite.RightToLeft,r)
+                      allCl2 (cl,lit,Metis_Rewrite.LeftToRight,l) andalso
+                      allCl2 (cl,lit,Metis_Rewrite.RightToLeft,r)
                   end orelse
                   let
-                    val () = Print.trace Literal.pp
-                               "Active.isSaturated.allParamodulations: lit1" lit
+                    val () = Metis_Print.trace Metis_Literal.pp
+                               "Metis_Active.isSaturated.allParamodulations: lit1" lit
                   in
                     false
                   end
             in
-              LiteralSet.all allLiteral1 (Clause.literals cl)
+              Metis_LiteralSet.all allLiteral1 (Metis_Clause.literals cl)
             end orelse
             let
-              val () = Print.trace Clause.pp
-                         "Active.isSaturated.allParamodulations: cl1" cl
+              val () = Metis_Print.trace Metis_Clause.pp
+                         "Metis_Active.isSaturated.allParamodulations: cl1" cl
             in
               false
             end
@@ -20301,20 +19637,20 @@
   fun redundant {subsume,reduce,rewrite} =
       let
         fun simp cl =
-            case Clause.simplify cl of
+            case Metis_Clause.simplify cl of
               NONE => true
             | SOME cl =>
-              Subsume.isStrictlySubsumed subsume (Clause.literals cl) orelse
+              Metis_Subsume.isStrictlySubsumed subsume (Metis_Clause.literals cl) orelse
               let
                 val cl' = cl
-                val cl' = Clause.reduce reduce cl'
-                val cl' = Clause.rewrite rewrite cl'
+                val cl' = Metis_Clause.reduce reduce cl'
+                val cl' = Metis_Clause.rewrite rewrite cl'
               in
-                not (Clause.equalThms cl cl') andalso
+                not (Metis_Clause.equalThms cl cl') andalso
                 (simp cl' orelse
                  let
-                   val () = Print.trace Clause.pp
-                              "Active.isSaturated.redundant: cl'" cl'
+                   val () = Metis_Print.trace Metis_Clause.pp
+                              "Metis_Active.isSaturated.redundant: cl'" cl'
                  in
                    false
                  end)
@@ -20323,8 +19659,8 @@
         fn cl =>
            simp cl orelse
            let
-             val () = Print.trace Clause.pp
-                        "Active.isSaturated.redundant: cl" cl
+             val () = Metis_Print.trace Metis_Clause.pp
+                        "Metis_Active.isSaturated.redundant: cl" cl
            in
              false
            end
@@ -20332,7 +19668,7 @@
 in
   fun isSaturated ordering subs cls =
       let
-        val rd = Units.empty
+        val rd = Metis_Units.empty
         val rw = mkRewrite ordering cls
         val red = redundant {subsume = subs, reduce = rd, rewrite = rw}
       in
@@ -20340,9 +19676,9 @@
          allResolutions red cls andalso
          allParamodulations red cls) orelse
         let
-          val () = Print.trace Rewrite.pp "Active.isSaturated: rw" rw
-          val () = Print.trace (Print.ppList Clause.pp)
-                     "Active.isSaturated: clauses" cls
+          val () = Metis_Print.trace Metis_Rewrite.pp "Metis_Active.isSaturated: rw" rw
+          val () = Metis_Print.trace (Metis_Print.ppList Metis_Clause.pp)
+                     "Metis_Active.isSaturated: clauses" cls
         in
           false
         end
@@ -20351,7 +19687,7 @@
 
 fun checkSaturated ordering subs cls =
     if isSaturated ordering subs cls then ()
-    else raise Bug "Active.checkSaturated";
+    else raise Bug "Metis_Active.checkSaturated";
 *)
 
 (* ------------------------------------------------------------------------- *)
@@ -20361,35 +19697,35 @@
 type simplify = {subsume : bool, reduce : bool, rewrite : bool};
 
 type parameters =
-     {clause : Clause.parameters,
+     {clause : Metis_Clause.parameters,
       prefactor : simplify,
       postfactor : simplify};
 
 datatype active =
-    Active of
+    Metis_Active of
       {parameters : parameters,
-       clauses : Clause.clause IntMap.map,
-       units : Units.units,
-       rewrite : Rewrite.rewrite,
-       subsume : Clause.clause Subsume.subsume,
-       literals : (Clause.clause * Literal.literal) LiteralNet.literalNet,
+       clauses : Metis_Clause.clause Metis_IntMap.map,
+       units : Metis_Units.units,
+       rewrite : Metis_Rewrite.rewrite,
+       subsume : Metis_Clause.clause Metis_Subsume.subsume,
+       literals : (Metis_Clause.clause * Metis_Literal.literal) Metis_LiteralNet.literalNet,
        equations :
-         (Clause.clause * Literal.literal * Rewrite.orient * Term.term)
-         TermNet.termNet,
+         (Metis_Clause.clause * Metis_Literal.literal * Metis_Rewrite.orient * Metis_Term.term)
+         Metis_TermNet.termNet,
        subterms :
-         (Clause.clause * Literal.literal * Term.path * Term.term)
-         TermNet.termNet,
-       allSubterms : (Clause.clause * Term.term) TermNet.termNet};
-
-fun getSubsume (Active {subsume = s, ...}) = s;
+         (Metis_Clause.clause * Metis_Literal.literal * Metis_Term.path * Metis_Term.term)
+         Metis_TermNet.termNet,
+       allSubterms : (Metis_Clause.clause * Metis_Term.term) Metis_TermNet.termNet};
+
+fun getSubsume (Metis_Active {subsume = s, ...}) = s;
 
 fun setRewrite active rewrite =
     let
-      val Active
+      val Metis_Active
             {parameters,clauses,units,subsume,literals,equations,
              subterms,allSubterms,...} = active
     in
-      Active
+      Metis_Active
         {parameters = parameters, clauses = clauses, units = units,
          rewrite = rewrite, subsume = subsume, literals = literals,
          equations = equations, subterms = subterms, allSubterms = allSubterms}
@@ -20402,7 +19738,7 @@
 val maxSimplify : simplify = {subsume = true, reduce = true, rewrite = true};
 
 val default : parameters =
-    {clause = Clause.default,
+    {clause = Metis_Clause.default,
      prefactor = maxSimplify,
      postfactor = maxSimplify};
 
@@ -20411,43 +19747,43 @@
       val {clause,...} = parameters
       val {ordering,...} = clause
     in
-      Active
+      Metis_Active
         {parameters = parameters,
-         clauses = IntMap.new (),
-         units = Units.empty,
-         rewrite = Rewrite.new (KnuthBendixOrder.compare ordering),
-         subsume = Subsume.new (),
-         literals = LiteralNet.new {fifo = false},
-         equations = TermNet.new {fifo = false},
-         subterms = TermNet.new {fifo = false},
-         allSubterms = TermNet.new {fifo = false}}
-    end;
-
-fun size (Active {clauses,...}) = IntMap.size clauses;
-
-fun clauses (Active {clauses = cls, ...}) =
+         clauses = Metis_IntMap.new (),
+         units = Metis_Units.empty,
+         rewrite = Metis_Rewrite.new (Metis_KnuthBendixOrder.compare ordering),
+         subsume = Metis_Subsume.new (),
+         literals = Metis_LiteralNet.new {fifo = false},
+         equations = Metis_TermNet.new {fifo = false},
+         subterms = Metis_TermNet.new {fifo = false},
+         allSubterms = Metis_TermNet.new {fifo = false}}
+    end;
+
+fun size (Metis_Active {clauses,...}) = Metis_IntMap.size clauses;
+
+fun clauses (Metis_Active {clauses = cls, ...}) =
     let
       fun add (_,cl,acc) = cl :: acc
     in
-      IntMap.foldr add [] cls
+      Metis_IntMap.foldr add [] cls
     end;
 
 fun saturation active =
     let
       fun remove (cl,(cls,subs)) =
           let
-            val lits = Clause.literals cl
-          in
-            if Subsume.isStrictlySubsumed subs lits then (cls,subs)
-            else (cl :: cls, Subsume.insert subs (lits,()))
+            val lits = Metis_Clause.literals cl
+          in
+            if Metis_Subsume.isStrictlySubsumed subs lits then (cls,subs)
+            else (cl :: cls, Metis_Subsume.insert subs (lits,()))
           end
 
       val cls = clauses active
-      val (cls,_) = foldl remove ([], Subsume.new ()) cls
-      val (cls,subs) = foldl remove ([], Subsume.new ()) cls
+      val (cls,_) = foldl remove ([], Metis_Subsume.new ()) cls
+      val (cls,subs) = foldl remove ([], Metis_Subsume.new ()) cls
 
 (*MetisDebug
-      val Active {parameters,...} = active
+      val Metis_Active {parameters,...} = active
       val {clause,...} = parameters
       val {ordering,...} = clause
       val () = checkSaturated ordering subs cls
@@ -20462,55 +19798,55 @@
 
 val pp =
     let
-      fun toStr active = "Active{" ^ Int.toString (size active) ^ "}"
-    in
-      Print.ppMap toStr Print.ppString
+      fun toStr active = "Metis_Active{" ^ Int.toString (size active) ^ "}"
+    in
+      Metis_Print.ppMap toStr Metis_Print.ppString
     end;
 
 (*MetisDebug
 local
   fun ppField f ppA a =
-      Print.blockProgram Print.Inconsistent 2
-        [Print.addString (f ^ " ="),
-         Print.addBreak 1,
+      Metis_Print.blockProgram Metis_Print.Inconsistent 2
+        [Metis_Print.addString (f ^ " ="),
+         Metis_Print.addBreak 1,
          ppA a];
 
   val ppClauses =
       ppField "clauses"
-        (Print.ppMap IntMap.toList
-           (Print.ppList (Print.ppPair Print.ppInt Clause.pp)));
-
-  val ppRewrite = ppField "rewrite" Rewrite.pp;
+        (Metis_Print.ppMap Metis_IntMap.toList
+           (Metis_Print.ppList (Metis_Print.ppPair Metis_Print.ppInt Metis_Clause.pp)));
+
+  val ppRewrite = ppField "rewrite" Metis_Rewrite.pp;
 
   val ppSubterms =
       ppField "subterms"
-        (TermNet.pp
-           (Print.ppMap (fn (c,l,p,t) => ((Clause.id c, l, p), t))
-              (Print.ppPair
-                 (Print.ppTriple Print.ppInt Literal.pp Term.ppPath)
-                 Term.pp)));
-in
-  fun pp (Active {clauses,rewrite,subterms,...}) =
-      Print.blockProgram Print.Inconsistent 2
-        [Print.addString "Active",
-         Print.addBreak 1,
-         Print.blockProgram Print.Inconsistent 1
-           [Print.addString "{",
+        (Metis_TermNet.pp
+           (Metis_Print.ppMap (fn (c,l,p,t) => ((Metis_Clause.id c, l, p), t))
+              (Metis_Print.ppPair
+                 (Metis_Print.ppTriple Metis_Print.ppInt Metis_Literal.pp Metis_Term.ppPath)
+                 Metis_Term.pp)));
+in
+  fun pp (Metis_Active {clauses,rewrite,subterms,...}) =
+      Metis_Print.blockProgram Metis_Print.Inconsistent 2
+        [Metis_Print.addString "Metis_Active",
+         Metis_Print.addBreak 1,
+         Metis_Print.blockProgram Metis_Print.Inconsistent 1
+           [Metis_Print.addString "{",
             ppClauses clauses,
-            Print.addString ",",
-            Print.addBreak 1,
+            Metis_Print.addString ",",
+            Metis_Print.addBreak 1,
             ppRewrite rewrite,
 (*MetisTrace5
-            Print.addString ",",
-            Print.addBreak 1,
+            Metis_Print.addString ",",
+            Metis_Print.addBreak 1,
             ppSubterms subterms,
 *)
-            Print.skip],
-         Print.addString "}"];
-end;
-*)
-
-val toString = Print.toString pp;
+            Metis_Print.skip],
+         Metis_Print.addString "}"];
+end;
+*)
+
+val toString = Metis_Print.toString pp;
 
 (* ------------------------------------------------------------------------- *)
 (* Simplify clauses.                                                         *)
@@ -20522,36 +19858,36 @@
 
       fun rewrite cl =
           let
-            val cl' = Clause.rewrite rewr cl
-          in
-            if Clause.equalThms cl cl' then SOME cl else Clause.simplify cl'
+            val cl' = Metis_Clause.rewrite rewr cl
+          in
+            if Metis_Clause.equalThms cl cl' then SOME cl else Metis_Clause.simplify cl'
           end
     in
       fn cl =>
-         case Clause.simplify cl of
+         case Metis_Clause.simplify cl of
            NONE => NONE
          | SOME cl =>
            case (if w then rewrite cl else SOME cl) of
              NONE => NONE
            | SOME cl =>
              let
-               val cl = if r then Clause.reduce units cl else cl
+               val cl = if r then Metis_Clause.reduce units cl else cl
              in
-               if s andalso Clause.subsumes subs cl then NONE else SOME cl
+               if s andalso Metis_Clause.subsumes subs cl then NONE else SOME cl
              end
     end;
 
 (*MetisDebug
 val simplify = fn simp => fn units => fn rewr => fn subs => fn cl =>
     let
-      fun traceCl s = Print.trace Clause.pp ("Active.simplify: " ^ s)
+      fun traceCl s = Metis_Print.trace Metis_Clause.pp ("Metis_Active.simplify: " ^ s)
 (*MetisTrace4
-      val ppClOpt = Print.ppOption Clause.pp
+      val ppClOpt = Metis_Print.ppOption Metis_Clause.pp
       val () = traceCl "cl" cl
 *)
       val cl' = simplify simp units rewr subs cl
 (*MetisTrace4
-      val () = Print.trace ppClOpt "Active.simplify: cl'" cl'
+      val () = Metis_Print.trace ppClOpt "Metis_Active.simplify: cl'" cl'
 *)
       val () =
           case cl' of
@@ -20561,7 +19897,7 @@
               (case simplify simp units rewr subs cl' of
                  NONE => SOME ("away", K ())
                | SOME cl'' =>
-                 if Clause.equalThms cl' cl'' then NONE
+                 if Metis_Clause.equalThms cl' cl'' then NONE
                  else SOME ("further", fn () => traceCl "cl''" cl'')) of
               NONE => ()
             | SOME (e,f) =>
@@ -20572,7 +19908,7 @@
               in
                 raise
                   Bug
-                    ("Active.simplify: clause should have been simplified "^e)
+                    ("Metis_Active.simplify: clause should have been simplified "^e)
               end
     in
       cl'
@@ -20581,7 +19917,7 @@
 
 fun simplifyActive simp active =
     let
-      val Active {units,rewrite,subsume,...} = active
+      val Metis_Active {units,rewrite,subsume,...} = active
     in
       simplify simp units rewrite subsume
     end;
@@ -20592,70 +19928,70 @@
 
 fun addUnit units cl =
     let
-      val th = Clause.thm cl
-    in
-      case total Thm.destUnit th of
-        SOME lit => Units.add units (lit,th)
+      val th = Metis_Clause.thm cl
+    in
+      case total Metis_Thm.destUnit th of
+        SOME lit => Metis_Units.add units (lit,th)
       | NONE => units
     end;
 
 fun addRewrite rewrite cl =
     let
-      val th = Clause.thm cl
-    in
-      case total Thm.destUnitEq th of
-        SOME l_r => Rewrite.add rewrite (Clause.id cl, (l_r,th))
+      val th = Metis_Clause.thm cl
+    in
+      case total Metis_Thm.destUnitEq th of
+        SOME l_r => Metis_Rewrite.add rewrite (Metis_Clause.id cl, (l_r,th))
       | NONE => rewrite
     end;
 
-fun addSubsume subsume cl = Subsume.insert subsume (Clause.literals cl, cl);
+fun addSubsume subsume cl = Metis_Subsume.insert subsume (Metis_Clause.literals cl, cl);
 
 fun addLiterals literals cl =
     let
       fun add (lit as (_,atm), literals) =
-          if Atom.isEq atm then literals
-          else LiteralNet.insert literals (lit,(cl,lit))
-    in
-      LiteralSet.foldl add literals (Clause.largestLiterals cl)
+          if Metis_Atom.isEq atm then literals
+          else Metis_LiteralNet.insert literals (lit,(cl,lit))
+    in
+      Metis_LiteralSet.foldl add literals (Metis_Clause.largestLiterals cl)
     end;
 
 fun addEquations equations cl =
     let
       fun add ((lit,ort,tm),equations) =
-          TermNet.insert equations (tm,(cl,lit,ort,tm))
-    in
-      foldl add equations (Clause.largestEquations cl)
+          Metis_TermNet.insert equations (tm,(cl,lit,ort,tm))
+    in
+      foldl add equations (Metis_Clause.largestEquations cl)
     end;
 
 fun addSubterms subterms cl =
     let
       fun add ((lit,path,tm),subterms) =
-          TermNet.insert subterms (tm,(cl,lit,path,tm))
-    in
-      foldl add subterms (Clause.largestSubterms cl)
+          Metis_TermNet.insert subterms (tm,(cl,lit,path,tm))
+    in
+      foldl add subterms (Metis_Clause.largestSubterms cl)
     end;
 
 fun addAllSubterms allSubterms cl =
     let
       fun add ((_,_,tm),allSubterms) =
-          TermNet.insert allSubterms (tm,(cl,tm))
-    in
-      foldl add allSubterms (Clause.allSubterms cl)
+          Metis_TermNet.insert allSubterms (tm,(cl,tm))
+    in
+      foldl add allSubterms (Metis_Clause.allSubterms cl)
     end;
 
 fun addClause active cl =
     let
-      val Active
+      val Metis_Active
             {parameters,clauses,units,rewrite,subsume,literals,
              equations,subterms,allSubterms} = active
-      val clauses = IntMap.insert clauses (Clause.id cl, cl)
+      val clauses = Metis_IntMap.insert clauses (Metis_Clause.id cl, cl)
       and subsume = addSubsume subsume cl
       and literals = addLiterals literals cl
       and equations = addEquations equations cl
       and subterms = addSubterms subterms cl
       and allSubterms = addAllSubterms allSubterms cl
     in
-      Active
+      Metis_Active
         {parameters = parameters, clauses = clauses, units = units,
          rewrite = rewrite, subsume = subsume, literals = literals,
          equations = equations, subterms = subterms,
@@ -20664,13 +20000,13 @@
 
 fun addFactorClause active cl =
     let
-      val Active
+      val Metis_Active
             {parameters,clauses,units,rewrite,subsume,literals,
              equations,subterms,allSubterms} = active
       val units = addUnit units cl
       and rewrite = addRewrite rewrite cl
     in
-      Active
+      Metis_Active
         {parameters = parameters, clauses = clauses, units = units,
          rewrite = rewrite, subsume = subsume, literals = literals,
          equations = equations, subterms = subterms, allSubterms = allSubterms}
@@ -20683,66 +20019,66 @@
 fun deduceResolution literals cl (lit as (_,atm), acc) =
     let
       fun resolve (cl_lit,acc) =
-          case total (Clause.resolve cl_lit) (cl,lit) of
+          case total (Metis_Clause.resolve cl_lit) (cl,lit) of
             SOME cl' => cl' :: acc
           | NONE => acc
 (*MetisTrace4
-      val () = Print.trace Literal.pp "Active.deduceResolution: lit" lit
-*)
-    in
-      if Atom.isEq atm then acc
-      else foldl resolve acc (LiteralNet.unify literals (Literal.negate lit))
+      val () = Metis_Print.trace Metis_Literal.pp "Metis_Active.deduceResolution: lit" lit
+*)
+    in
+      if Metis_Atom.isEq atm then acc
+      else foldl resolve acc (Metis_LiteralNet.unify literals (Metis_Literal.negate lit))
     end;
 
 fun deduceParamodulationWith subterms cl ((lit,ort,tm),acc) =
     let
       fun para (cl_lit_path_tm,acc) =
-          case total (Clause.paramodulate (cl,lit,ort,tm)) cl_lit_path_tm of
+          case total (Metis_Clause.paramodulate (cl,lit,ort,tm)) cl_lit_path_tm of
             SOME cl' => cl' :: acc
           | NONE => acc
     in
-      foldl para acc (TermNet.unify subterms tm)
+      foldl para acc (Metis_TermNet.unify subterms tm)
     end;
 
 fun deduceParamodulationInto equations cl ((lit,path,tm),acc) =
     let
       fun para (cl_lit_ort_tm,acc) =
-          case total (Clause.paramodulate cl_lit_ort_tm) (cl,lit,path,tm) of
+          case total (Metis_Clause.paramodulate cl_lit_ort_tm) (cl,lit,path,tm) of
             SOME cl' => cl' :: acc
           | NONE => acc
     in
-      foldl para acc (TermNet.unify equations tm)
+      foldl para acc (Metis_TermNet.unify equations tm)
     end;
 
 fun deduce active cl =
     let
-      val Active {parameters,literals,equations,subterms,...} = active
-
-      val lits = Clause.largestLiterals cl
-      val eqns = Clause.largestEquations cl
+      val Metis_Active {parameters,literals,equations,subterms,...} = active
+
+      val lits = Metis_Clause.largestLiterals cl
+      val eqns = Metis_Clause.largestEquations cl
       val subtms =
-          if TermNet.null equations then [] else Clause.largestSubterms cl
+          if Metis_TermNet.null equations then [] else Metis_Clause.largestSubterms cl
 (*MetisTrace5
-      val () = Print.trace LiteralSet.pp "Active.deduce: lits" lits
-      val () = Print.trace
-                 (Print.ppList
-                    (Print.ppMap (fn (lit,ort,_) => (lit,ort))
-                      (Print.ppPair Literal.pp Rewrite.ppOrient)))
-                 "Active.deduce: eqns" eqns
-      val () = Print.trace
-                 (Print.ppList
-                    (Print.ppTriple Literal.pp Term.ppPath Term.pp))
-                 "Active.deduce: subtms" subtms
+      val () = Metis_Print.trace Metis_LiteralSet.pp "Metis_Active.deduce: lits" lits
+      val () = Metis_Print.trace
+                 (Metis_Print.ppList
+                    (Metis_Print.ppMap (fn (lit,ort,_) => (lit,ort))
+                      (Metis_Print.ppPair Metis_Literal.pp Metis_Rewrite.ppOrient)))
+                 "Metis_Active.deduce: eqns" eqns
+      val () = Metis_Print.trace
+                 (Metis_Print.ppList
+                    (Metis_Print.ppTriple Metis_Literal.pp Metis_Term.ppPath Metis_Term.pp))
+                 "Metis_Active.deduce: subtms" subtms
 *)
 
       val acc = []
-      val acc = LiteralSet.foldl (deduceResolution literals cl) acc lits
+      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 = rev acc
 
 (*MetisTrace5
-      val () = Print.trace (Print.ppList Clause.pp) "Active.deduce: acc" acc
+      val () = Metis_Print.trace (Metis_Print.ppList Metis_Clause.pp) "Metis_Active.deduce: acc" acc
 *)
     in
       acc
@@ -20755,23 +20091,23 @@
 local
   fun clause_rewritables active =
       let
-        val Active {clauses,rewrite,...} = active
+        val Metis_Active {clauses,rewrite,...} = active
 
         fun rewr (id,cl,ids) =
             let
-              val cl' = Clause.rewrite rewrite cl
-            in
-              if Clause.equalThms cl cl' then ids else IntSet.add ids id
-            end
-      in
-        IntMap.foldr rewr IntSet.empty clauses
+              val cl' = Metis_Clause.rewrite rewrite cl
+            in
+              if Metis_Clause.equalThms cl cl' then ids else Metis_IntSet.add ids id
+            end
+      in
+        Metis_IntMap.foldr rewr Metis_IntSet.empty clauses
       end;
 
   fun orderedRedexResidues (((l,r),_),ort) =
       case ort of
         NONE => []
-      | SOME Rewrite.LeftToRight => [(l,r,true)]
-      | SOME Rewrite.RightToLeft => [(r,l,true)];
+      | SOME Metis_Rewrite.LeftToRight => [(l,r,true)]
+      | SOME Metis_Rewrite.RightToLeft => [(r,l,true)];
 
   fun unorderedRedexResidues (((l,r),_),ort) =
       case ort of
@@ -20780,22 +20116,22 @@
 
   fun rewrite_rewritables active rewr_ids =
       let
-        val Active {parameters,rewrite,clauses,allSubterms,...} = active
+        val Metis_Active {parameters,rewrite,clauses,allSubterms,...} = active
         val {clause = {ordering,...}, ...} = parameters
-        val order = KnuthBendixOrder.compare ordering
+        val order = Metis_KnuthBendixOrder.compare ordering
 
         fun addRewr (id,acc) =
-            if IntMap.inDomain id clauses then IntSet.add acc id else acc
+            if Metis_IntMap.inDomain id clauses then Metis_IntSet.add acc id else acc
 
         fun addReduce ((l,r,ord),acc) =
             let
               fun isValidRewr tm =
-                  case total (Subst.match Subst.empty l) tm of
+                  case total (Metis_Subst.match Metis_Subst.empty l) tm of
                     NONE => false
                   | SOME sub =>
                     ord orelse
                     let
-                      val tm' = Subst.subst (Subst.normalize sub) r
+                      val tm' = Metis_Subst.subst (Metis_Subst.normalize sub) r
                     in
                       order (tm,tm') = SOME GREATER
                     end
@@ -20803,27 +20139,27 @@
               fun addRed ((cl,tm),acc) =
                   let
 (*MetisTrace5
-                    val () = Print.trace Clause.pp "Active.addRed: cl" cl
-                    val () = Print.trace Term.pp "Active.addRed: tm" tm
-*)
-                    val id = Clause.id cl
+                    val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.addRed: cl" cl
+                    val () = Metis_Print.trace Metis_Term.pp "Metis_Active.addRed: tm" tm
+*)
+                    val id = Metis_Clause.id cl
                   in
-                    if IntSet.member id acc then acc
+                    if Metis_IntSet.member id acc then acc
                     else if not (isValidRewr tm) then acc
-                    else IntSet.add acc id
+                    else Metis_IntSet.add acc id
                   end
 
 (*MetisTrace5
-              val () = Print.trace Term.pp "Active.addReduce: l" l
-              val () = Print.trace Term.pp "Active.addReduce: r" r
-              val () = Print.trace Print.ppBool "Active.addReduce: ord" ord
-*)
-            in
-              List.foldl addRed acc (TermNet.matched allSubterms l)
+              val () = Metis_Print.trace Metis_Term.pp "Metis_Active.addReduce: l" l
+              val () = Metis_Print.trace Metis_Term.pp "Metis_Active.addReduce: r" r
+              val () = Metis_Print.trace Metis_Print.ppBool "Metis_Active.addReduce: ord" ord
+*)
+            in
+              List.foldl addRed acc (Metis_TermNet.matched allSubterms l)
             end
 
         fun addEquation redexResidues (id,acc) =
-            case Rewrite.peek rewrite id of
+            case Metis_Rewrite.peek rewrite id of
               NONE => acc
             | SOME eqn_ort => List.foldl addReduce acc (redexResidues eqn_ort)
 
@@ -20831,7 +20167,7 @@
 
         val addUnordered = addEquation unorderedRedexResidues
 
-        val ids = IntSet.empty
+        val ids = Metis_IntSet.empty
         val ids = List.foldl addRewr ids rewr_ids
         val ids = List.foldl addOrdered ids rewr_ids
         val ids = List.foldl addUnordered ids rewr_ids
@@ -20852,19 +20188,19 @@
         val rewrite_ids = rewrite_rewritables active ids
 
         val () =
-            if IntSet.equal rewrite_ids clause_ids then ()
+            if Metis_IntSet.equal rewrite_ids clause_ids then ()
             else
               let
-                val ppIdl = Print.ppList Print.ppInt
-                val ppIds = Print.ppMap IntSet.toList ppIdl
-                val () = Print.trace pp "Active.rewritables: active" active
-                val () = Print.trace ppIdl "Active.rewritables: ids" ids
-                val () = Print.trace ppIds
-                           "Active.rewritables: clause_ids" clause_ids
-                val () = Print.trace ppIds
-                           "Active.rewritables: rewrite_ids" rewrite_ids
+                val ppIdl = Metis_Print.ppList Metis_Print.ppInt
+                val ppIds = Metis_Print.ppMap Metis_IntSet.toList ppIdl
+                val () = Metis_Print.trace pp "Metis_Active.rewritables: active" active
+                val () = Metis_Print.trace ppIdl "Metis_Active.rewritables: ids" ids
+                val () = Metis_Print.trace ppIds
+                           "Metis_Active.rewritables: clause_ids" clause_ids
+                val () = Metis_Print.trace ppIds
+                           "Metis_Active.rewritables: rewrite_ids" rewrite_ids
               in
-                raise Bug "Active.rewritables: ~(rewrite_ids SUBSET clause_ids)"
+                raise Bug "Metis_Active.rewritables: ~(rewrite_ids SUBSET clause_ids)"
               end
       in
         if choose_clause_rewritables active ids then clause_ids else rewrite_ids
@@ -20872,14 +20208,14 @@
 *)
 
   fun delete active ids =
-      if IntSet.null ids then active
+      if Metis_IntSet.null ids then active
       else
         let
-          fun idPred id = not (IntSet.member id ids)
-
-          fun clausePred cl = idPred (Clause.id cl)
-
-          val Active
+          fun idPred id = not (Metis_IntSet.member id ids)
+
+          fun clausePred cl = idPred (Metis_Clause.id cl)
+
+          val Metis_Active
                 {parameters,
                  clauses,
                  units,
@@ -20890,14 +20226,14 @@
                  subterms,
                  allSubterms} = active
 
-          val clauses = IntMap.filter (idPred o fst) clauses
-          and subsume = Subsume.filter clausePred subsume
-          and literals = LiteralNet.filter (clausePred o #1) literals
-          and equations = TermNet.filter (clausePred o #1) equations
-          and subterms = TermNet.filter (clausePred o #1) subterms
-          and allSubterms = TermNet.filter (clausePred o fst) allSubterms
-        in
-          Active
+          val clauses = Metis_IntMap.filter (idPred o fst) clauses
+          and subsume = Metis_Subsume.filter clausePred subsume
+          and literals = Metis_LiteralNet.filter (clausePred o #1) literals
+          and equations = Metis_TermNet.filter (clausePred o #1) equations
+          and subterms = Metis_TermNet.filter (clausePred o #1) subterms
+          and allSubterms = Metis_TermNet.filter (clausePred o fst) allSubterms
+        in
+          Metis_Active
             {parameters = parameters,
              clauses = clauses,
              units = units,
@@ -20909,27 +20245,27 @@
              allSubterms = allSubterms}
         end;
 in
-  fun extract_rewritables (active as Active {clauses,rewrite,...}) =
-      if Rewrite.isReduced rewrite then (active,[])
+  fun extract_rewritables (active as Metis_Active {clauses,rewrite,...}) =
+      if Metis_Rewrite.isReduced rewrite then (active,[])
       else
         let
 (*MetisTrace3
-          val () = trace "Active.extract_rewritables: inter-reducing\n"
-*)
-          val (rewrite,ids) = Rewrite.reduce' rewrite
+          val () = trace "Metis_Active.extract_rewritables: inter-reducing\n"
+*)
+          val (rewrite,ids) = Metis_Rewrite.reduce' rewrite
           val active = setRewrite active rewrite
           val ids = rewritables active ids
-          val cls = IntSet.transform (IntMap.get clauses) ids
+          val cls = Metis_IntSet.transform (Metis_IntMap.get clauses) ids
 (*MetisTrace3
-          val ppCls = Print.ppList Clause.pp
-          val () = Print.trace ppCls "Active.extract_rewritables: cls" cls
+          val ppCls = Metis_Print.ppList Metis_Clause.pp
+          val () = Metis_Print.trace ppCls "Metis_Active.extract_rewritables: cls" cls
 *)
         in
           (delete active ids, cls)
         end
 (*MetisDebug
         handle Error err =>
-          raise Bug ("Active.extract_rewritables: shouldn't fail\n" ^ err);
+          raise Bug ("Metis_Active.extract_rewritables: shouldn't fail\n" ^ err);
 *)
 end;
 
@@ -20940,7 +20276,7 @@
 local
   fun prefactor_simplify active subsume =
       let
-        val Active {parameters,units,rewrite,...} = active
+        val Metis_Active {parameters,units,rewrite,...} = active
         val {prefactor,...} = parameters
       in
         simplify prefactor units rewrite subsume
@@ -20948,7 +20284,7 @@
 
   fun postfactor_simplify active subsume =
       let
-        val Active {parameters,units,rewrite,...} = active
+        val Metis_Active {parameters,units,rewrite,...} = active
         val {postfactor,...} = parameters
       in
         simplify postfactor units rewrite subsume
@@ -20957,9 +20293,9 @@
   val sort_utilitywise =
       let
         fun utility cl =
-            case LiteralSet.size (Clause.literals cl) of
+            case Metis_LiteralSet.size (Metis_Clause.literals cl) of
               0 => ~1
-            | 1 => if Thm.isUnitEq (Clause.thm cl) then 0 else 1
+            | 1 => if Metis_Thm.isUnitEq (Metis_Clause.thm cl) then 0 else 1
             | n => n
       in
         sortMap utility Int.compare
@@ -20982,7 +20318,7 @@
         NONE => active_subsume_acc
       | SOME cl =>
         let
-          val cls = sort_utilitywise (cl :: Clause.factor cl)
+          val cls = sort_utilitywise (cl :: Metis_Clause.factor cl)
         in
           foldl factor_add active_subsume_acc cls
         end;
@@ -21004,10 +20340,10 @@
 (*MetisTrace4
 val factor = fn active => fn cls =>
     let
-      val ppCls = Print.ppList Clause.pp
-      val () = Print.trace ppCls "Active.factor: cls" cls
+      val ppCls = Metis_Print.ppList Metis_Clause.pp
+      val () = Metis_Print.trace ppCls "Metis_Active.factor: cls" cls
       val (active,cls') = factor active cls
-      val () = Print.trace ppCls "Active.factor: cls'" cls'
+      val () = Metis_Print.trace ppCls "Metis_Active.factor: cls'" cls'
     in
       (active,cls')
     end;
@@ -21022,7 +20358,7 @@
       val {clause,...} = parameters
 
       fun mk_clause th =
-          Clause.mk {parameters = clause, id = Clause.newId (), thm = th}
+          Metis_Clause.mk {parameters = clause, id = Metis_Clause.newId (), thm = th}
 
       val active = empty parameters
       val (active,axioms) = factor active (map mk_clause axioms)
@@ -21039,27 +20375,26 @@
     case simplifyActive maxSimplify active cl of
       NONE => (active,[])
     | SOME cl' =>
-      if Clause.isContradiction cl' then (active,[cl'])
-      else if not (Clause.equalThms cl cl') then factor active [cl']
+      if Metis_Clause.isContradiction cl' then (active,[cl'])
+      else if not (Metis_Clause.equalThms cl cl') then factor active [cl']
       else
         let
 (*MetisTrace2
-          val () = Print.trace Clause.pp "Active.add: cl" cl
+          val () = Metis_Print.trace Metis_Clause.pp "Metis_Active.add: cl" cl
 *)
           val active = addClause active cl
-          val cl = Clause.freshVars cl
+          val cl = Metis_Clause.freshVars cl
           val cls = deduce active cl
           val (active,cls) = factor active cls
 (*MetisTrace2
-          val ppCls = Print.ppList Clause.pp
-          val () = Print.trace ppCls "Active.add: cls" cls
+          val ppCls = Metis_Print.ppList Metis_Clause.pp
+          val () = Metis_Print.trace ppCls "Metis_Active.add: cls" cls
 *)
         in
           (active,cls)
         end;
 
 end
-end;
 
 (**** Original file: Waiting.sig ****)
 
@@ -21068,7 +20403,7 @@
 (* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Waiting =
+signature Metis_Waiting =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -21092,7 +20427,7 @@
 type weight = real
 
 type modelParameters =
-     {model : Metis.Model.parameters,
+     {model : Metis_Model.parameters,
       initialPerturbations : int,
       maxChecks : int option,
       perturbations : int,
@@ -21120,47 +20455,38 @@
 
 val new :
     parameters ->
-    {axioms : Metis.Clause.clause list,
-     conjecture : Metis.Clause.clause list} -> waiting
+    {axioms : Metis_Clause.clause list,
+     conjecture : Metis_Clause.clause list} -> waiting
 
 val size : waiting -> int
 
-val pp : waiting Metis.Print.pp
+val pp : waiting Metis_Print.pp
 
 (* ------------------------------------------------------------------------- *)
 (* Adding new clauses.                                                       *)
 (* ------------------------------------------------------------------------- *)
 
-val add : waiting -> distance * Metis.Clause.clause list -> waiting
+val add : waiting -> distance * Metis_Clause.clause list -> waiting
 
 (* ------------------------------------------------------------------------- *)
 (* Removing the lightest clause.                                             *)
 (* ------------------------------------------------------------------------- *)
 
-val remove : waiting -> ((distance * Metis.Clause.clause) * waiting) option
+val remove : waiting -> ((distance * Metis_Clause.clause) * waiting) option
 
 end
 
 (**** Original file: Waiting.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* THE WAITING SET OF CLAUSES                                                *)
 (* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Waiting :> Waiting =
+structure Metis_Waiting :> Metis_Waiting =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of waiting sets of clauses.                                        *)
@@ -21169,7 +20495,7 @@
 type weight = real;
 
 type modelParameters =
-     {model : Model.parameters,
+     {model : Metis_Model.parameters,
       initialPerturbations : int,
       maxChecks : int option,
       perturbations : int,
@@ -21184,10 +20510,10 @@
 type distance = real;
 
 datatype waiting =
-    Waiting of
+    Metis_Waiting of
       {parameters : parameters,
-       clauses : (weight * (distance * Clause.clause)) Heap.heap,
-       models : Model.model list};
+       clauses : (weight * (distance * Metis_Clause.clause)) Metis_Heap.heap,
+       models : Metis_Model.model list};
 
 (* ------------------------------------------------------------------------- *)
 (* Basic operations.                                                         *)
@@ -21195,7 +20521,7 @@
 
 val defaultModels : modelParameters list =
     [(* MODIFIED by Jasmin Blanchette
-      {model = Model.default,
+      {model = Metis_Model.default,
       initialPerturbations = 100,
       maxChecks = SOME 20,
       perturbations = 0,
@@ -21207,31 +20533,31 @@
       variablesWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *)
       models = defaultModels};
 
-fun size (Waiting {clauses,...}) = Heap.size clauses;
+fun size (Metis_Waiting {clauses,...}) = Metis_Heap.size clauses;
 
 val pp =
-    Print.ppMap
-      (fn w => "Waiting{" ^ Int.toString (size w) ^ "}")
-      Print.ppString;
+    Metis_Print.ppMap
+      (fn w => "Metis_Waiting{" ^ Int.toString (size w) ^ "}")
+      Metis_Print.ppString;
 
 (*MetisDebug
 val pp =
-    Print.ppMap
-      (fn Waiting {clauses,...} =>
-          map (fn (w,(_,cl)) => (w, Clause.id cl, cl)) (Heap.toList clauses))
-      (Print.ppList (Print.ppTriple Print.ppReal Print.ppInt Clause.pp));
+    Metis_Print.ppMap
+      (fn Metis_Waiting {clauses,...} =>
+          map (fn (w,(_,cl)) => (w, Metis_Clause.id cl, cl)) (Metis_Heap.toList clauses))
+      (Metis_Print.ppList (Metis_Print.ppTriple Metis_Print.ppReal Metis_Print.ppInt Metis_Clause.pp));
 *)
 
 (* ------------------------------------------------------------------------- *)
 (* Perturbing the models.                                                    *)
 (* ------------------------------------------------------------------------- *)
 
-type modelClause = NameSet.set * Thm.clause;
+type modelClause = Metis_NameSet.set * Metis_Thm.clause;
 
 fun mkModelClause cl =
     let
-      val lits = Clause.literals cl
-      val fvs = LiteralSet.freeVars lits
+      val lits = Metis_Clause.literals cl
+      val fvs = Metis_LiteralSet.freeVars lits
     in
       (fvs,lits)
     end;
@@ -21242,14 +20568,14 @@
     if null cls then K ()
     else
       let
-        val N = {size = Model.size M}
+        val N = {size = Metis_Model.size M}
 
         fun perturbClause (fv,cl) =
             let
-              val V = Model.randomValuation N fv
-            in
-              if Model.interpretClause M V cl then ()
-              else Model.perturbClause M V cl
+              val V = Metis_Model.randomValuation N fv
+            in
+              if Metis_Model.interpretClause M V cl then ()
+              else Metis_Model.perturbClause M V cl
             end
 
         fun perturbClauses () = app perturbClause cls
@@ -21260,7 +20586,7 @@
 fun initialModel axioms conjecture parm =
     let
       val {model,initialPerturbations,...} : modelParameters = parm
-      val m = Model.new model
+      val m = Metis_Model.new model
       val () = perturbModel m conjecture initialPerturbations
       val () = perturbModel m axioms initialPerturbations
     in
@@ -21273,7 +20599,7 @@
           let
             val {maxChecks,weight,...} : modelParameters = parm
             val n = {maxChecks = maxChecks}
-            val {T,F} = Model.check Model.interpretClause n model fv cl
+            val {T,F} = Metis_Model.check Metis_Model.interpretClause n model fv cl
           in
             Math.pow (1.0 + Real.fromInt T / Real.fromInt (T + F), weight) * z
           end
@@ -21294,46 +20620,46 @@
     end;
 
 (* ------------------------------------------------------------------------- *)
-(* Clause weights.                                                           *)
-(* ------------------------------------------------------------------------- *)
-
-local
-  fun clauseSymbols cl = Real.fromInt (LiteralSet.typedSymbols cl);
+(* Metis_Clause weights.                                                           *)
+(* ------------------------------------------------------------------------- *)
+
+local
+  fun clauseSymbols cl = Real.fromInt (Metis_LiteralSet.typedSymbols cl);
 
   fun clauseVariables cl =
-      Real.fromInt (NameSet.size (LiteralSet.freeVars cl) + 1);
-
-  fun clauseLiterals cl = Real.fromInt (LiteralSet.size cl);
-
-  fun clausePriority cl = 1e~12 * Real.fromInt (Clause.id cl);
+      Real.fromInt (Metis_NameSet.size (Metis_LiteralSet.freeVars cl) + 1);
+
+  fun clauseLiterals cl = Real.fromInt (Metis_LiteralSet.size cl);
+
+  fun clausePriority cl = 1e~12 * Real.fromInt (Metis_Clause.id cl);
 in
   fun clauseWeight (parm : parameters) mods dist mcl cl =
       let
 (*MetisTrace3
-        val () = Print.trace Clause.pp "Waiting.clauseWeight: cl" cl
+        val () = Metis_Print.trace Metis_Clause.pp "Metis_Waiting.clauseWeight: cl" cl
 *)
         val {symbolsWeight,variablesWeight,literalsWeight,models,...} = parm
-        val lits = Clause.literals cl
+        val lits = Metis_Clause.literals cl
         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 *)
 (*MetisTrace4
-        val () = trace ("Waiting.clauseWeight: dist = " ^
+        val () = trace ("Metis_Waiting.clauseWeight: dist = " ^
                         Real.toString dist ^ "\n")
-        val () = trace ("Waiting.clauseWeight: symbolsW = " ^
+        val () = trace ("Metis_Waiting.clauseWeight: symbolsW = " ^
                         Real.toString symbolsW ^ "\n")
-        val () = trace ("Waiting.clauseWeight: variablesW = " ^
+        val () = trace ("Metis_Waiting.clauseWeight: variablesW = " ^
                         Real.toString variablesW ^ "\n")
-        val () = trace ("Waiting.clauseWeight: literalsW = " ^
+        val () = trace ("Metis_Waiting.clauseWeight: literalsW = " ^
                         Real.toString literalsW ^ "\n")
-        val () = trace ("Waiting.clauseWeight: modelsW = " ^
+        val () = trace ("Metis_Waiting.clauseWeight: modelsW = " ^
                         Real.toString modelsW ^ "\n")
 *)
         val weight = dist * symbolsW * variablesW * literalsW * modelsW
         val weight = weight + clausePriority cl
 (*MetisTrace3
-        val () = trace ("Waiting.clauseWeight: weight = " ^
+        val () = trace ("Metis_Waiting.clauseWeight: weight = " ^
                         Real.toString weight ^ "\n")
 *)
       in
@@ -21347,7 +20673,7 @@
 
 fun add' waiting dist mcls cls =
     let
-      val Waiting {parameters,clauses,models} = waiting
+      val Metis_Waiting {parameters,clauses,models} = waiting
       val {models = modelParameters, ...} = parameters
 
       val dist = dist + Math.ln (Real.fromInt (length cls))
@@ -21356,28 +20682,28 @@
           let
             val weight = clauseWeight parameters models dist mcl cl
           in
-            Heap.add acc (weight,(dist,cl))
+            Metis_Heap.add acc (weight,(dist,cl))
           end
 
       val clauses = List.foldl addCl clauses (zip mcls cls)
 
       val () = perturbModels modelParameters models mcls
     in
-      Waiting {parameters = parameters, clauses = clauses, models = models}
+      Metis_Waiting {parameters = parameters, clauses = clauses, models = models}
     end;
 
 fun add waiting (_,[]) = waiting
   | add waiting (dist,cls) =
     let
 (*MetisTrace3
-      val () = Print.trace pp "Waiting.add: waiting" waiting
-      val () = Print.trace (Print.ppList Clause.pp) "Waiting.add: cls" cls
+      val () = Metis_Print.trace pp "Metis_Waiting.add: waiting" waiting
+      val () = Metis_Print.trace (Metis_Print.ppList Metis_Clause.pp) "Metis_Waiting.add: cls" cls
 *)
 
       val waiting = add' waiting dist (mkModelClauses cls) cls
 
 (*MetisTrace3
-      val () = Print.trace pp "Waiting.add: waiting" waiting
+      val () = Metis_Print.trace pp "Metis_Waiting.add: waiting" waiting
 *)
     in
       waiting
@@ -21389,10 +20715,10 @@
   fun empty parameters axioms conjecture =
       let
         val {models = modelParameters, ...} = parameters
-        val clauses = Heap.new cmp
+        val clauses = Metis_Heap.new cmp
         and models = map (initialModel axioms conjecture) modelParameters
       in
-        Waiting {parameters = parameters, clauses = clauses, models = models}
+        Metis_Waiting {parameters = parameters, clauses = clauses, models = models}
       end;
 in
   fun new parameters {axioms,conjecture} =
@@ -21410,20 +20736,19 @@
 (* Removing the lightest clause.                                             *)
 (* ------------------------------------------------------------------------- *)
 
-fun remove (Waiting {parameters,clauses,models}) =
-    if Heap.null clauses then NONE
+fun remove (Metis_Waiting {parameters,clauses,models}) =
+    if Metis_Heap.null clauses then NONE
     else
       let
-        val ((_,dcl),clauses) = Heap.remove clauses
+        val ((_,dcl),clauses) = Metis_Heap.remove clauses
         val waiting =
-            Waiting
+            Metis_Waiting
               {parameters = parameters, clauses = clauses, models = models}
       in
         SOME (dcl,waiting)
       end;
 
 end
-end;
 
 (**** Original file: Resolution.sig ****)
 
@@ -21432,7 +20757,7 @@
 (* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-signature Resolution =
+signature Metis_Resolution =
 sig
 
 (* ------------------------------------------------------------------------- *)
@@ -21440,8 +20765,8 @@
 (* ------------------------------------------------------------------------- *)
 
 type parameters =
-     {active : Metis.Active.parameters,
-      waiting : Metis.Waiting.parameters}
+     {active : Metis_Active.parameters,
+      waiting : Metis_Waiting.parameters}
 
 type resolution
 
@@ -21452,22 +20777,22 @@
 val default : parameters
 
 val new :
-    parameters -> {axioms : Metis.Thm.thm list, conjecture : Metis.Thm.thm list} ->
+    parameters -> {axioms : Metis_Thm.thm list, conjecture : Metis_Thm.thm list} ->
     resolution
 
-val active : resolution -> Metis.Active.active
-
-val waiting : resolution -> Metis.Waiting.waiting
-
-val pp : resolution Metis.Print.pp
+val active : resolution -> Metis_Active.active
+
+val waiting : resolution -> Metis_Waiting.waiting
+
+val pp : resolution Metis_Print.pp
 
 (* ------------------------------------------------------------------------- *)
 (* The main proof loop.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
 datatype decision =
-    Contradiction of Metis.Thm.thm
-  | Satisfiable of Metis.Thm.thm list
+    Contradiction of Metis_Thm.thm
+  | Satisfiable of Metis_Thm.thm list
 
 datatype state =
     Decided of decision
@@ -21481,74 +20806,65 @@
 
 (**** Original file: Resolution.sml ****)
 
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
 (* ========================================================================= *)
 (* THE RESOLUTION PROOF PROCEDURE                                            *)
 (* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
 (* ========================================================================= *)
 
-structure Resolution :> Resolution =
+structure Metis_Resolution :> Metis_Resolution =
 struct
 
-open Useful;
+open Metis_Useful;
 
 (* ------------------------------------------------------------------------- *)
 (* A type of resolution proof procedures.                                    *)
 (* ------------------------------------------------------------------------- *)
 
 type parameters =
-     {active : Active.parameters,
-      waiting : Waiting.parameters};
+     {active : Metis_Active.parameters,
+      waiting : Metis_Waiting.parameters};
 
 datatype resolution =
-    Resolution of
+    Metis_Resolution of
       {parameters : parameters,
-       active : Active.active,
-       waiting : Waiting.waiting};
+       active : Metis_Active.active,
+       waiting : Metis_Waiting.waiting};
 
 (* ------------------------------------------------------------------------- *)
 (* Basic operations.                                                         *)
 (* ------------------------------------------------------------------------- *)
 
 val default : parameters =
-    {active = Active.default,
-     waiting = Waiting.default};
+    {active = Metis_Active.default,
+     waiting = Metis_Waiting.default};
 
 fun new parameters ths =
     let
       val {active = activeParm, waiting = waitingParm} = parameters
-      val (active,cls) = Active.new activeParm ths  (* cls = factored ths *)
-      val waiting = Waiting.new waitingParm cls
-    in
-      Resolution {parameters = parameters, active = active, waiting = waiting}
-    end;
-
-fun active (Resolution {active = a, ...}) = a;
-
-fun waiting (Resolution {waiting = w, ...}) = w;
+      val (active,cls) = Metis_Active.new activeParm ths  (* cls = factored ths *)
+      val waiting = Metis_Waiting.new waitingParm cls
+    in
+      Metis_Resolution {parameters = parameters, active = active, waiting = waiting}
+    end;
+
+fun active (Metis_Resolution {active = a, ...}) = a;
+
+fun waiting (Metis_Resolution {waiting = w, ...}) = w;
 
 val pp =
-    Print.ppMap
-      (fn Resolution {active,waiting,...} =>
-          "Resolution(" ^ Int.toString (Active.size active) ^
-          "<-" ^ Int.toString (Waiting.size waiting) ^ ")")
-      Print.ppString;
+    Metis_Print.ppMap
+      (fn Metis_Resolution {active,waiting,...} =>
+          "Metis_Resolution(" ^ Int.toString (Metis_Active.size active) ^
+          "<-" ^ Int.toString (Metis_Waiting.size waiting) ^ ")")
+      Metis_Print.ppString;
 
 (* ------------------------------------------------------------------------- *)
 (* The main proof loop.                                                      *)
 (* ------------------------------------------------------------------------- *)
 
 datatype decision =
-    Contradiction of Thm.thm
-  | Satisfiable of Thm.thm list;
+    Contradiction of Metis_Thm.thm
+  | Satisfiable of Metis_Thm.thm list;
 
 datatype state =
     Decided of decision
@@ -21556,28 +20872,28 @@
 
 fun iterate resolution =
     let
-      val Resolution {parameters,active,waiting} = resolution
+      val Metis_Resolution {parameters,active,waiting} = resolution
 (*MetisTrace2
-      val () = Print.trace Active.pp "Resolution.iterate: active" active
-      val () = Print.trace Waiting.pp "Resolution.iterate: waiting" waiting
-*)
-    in
-      case Waiting.remove waiting of
+      val () = Metis_Print.trace Metis_Active.pp "Metis_Resolution.iterate: active" active
+      val () = Metis_Print.trace Metis_Waiting.pp "Metis_Resolution.iterate: waiting" waiting
+*)
+    in
+      case Metis_Waiting.remove waiting of
         NONE =>
-        Decided (Satisfiable (map Clause.thm (Active.saturation active)))
+        Decided (Satisfiable (map Metis_Clause.thm (Metis_Active.saturation active)))
       | SOME ((d,cl),waiting) =>
-        if Clause.isContradiction cl then
-          Decided (Contradiction (Clause.thm cl))
+        if Metis_Clause.isContradiction cl then
+          Decided (Contradiction (Metis_Clause.thm cl))
         else
           let
 (*MetisTrace1
-            val () = Print.trace Clause.pp "Resolution.iterate: cl" cl
-*)
-            val (active,cls) = Active.add active cl
-            val waiting = Waiting.add waiting (d,cls)
+            val () = Metis_Print.trace Metis_Clause.pp "Metis_Resolution.iterate: cl" cl
+*)
+            val (active,cls) = Metis_Active.add active cl
+            val waiting = Metis_Waiting.add waiting (d,cls)
           in
             Undecided
-              (Resolution
+              (Metis_Resolution
                  {parameters = parameters, active = active, waiting = waiting})
           end
     end;
@@ -21588,2808 +20904,5 @@
     | Undecided resolution => loop resolution;
 
 end
-end;
-
-(**** Original file: Tptp.sig ****)
-
-(* ========================================================================= *)
-(* THE TPTP PROBLEM FILE FORMAT                                              *)
-(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
-(* ========================================================================= *)
-
-signature Tptp =
-sig
-
-(* ------------------------------------------------------------------------- *)
-(* Mapping to and from TPTP variable, function and relation names.           *)
-(* ------------------------------------------------------------------------- *)
-
-type mapping
-
-val defaultMapping : mapping
-
-val mkMapping :
-    {functionMapping : {name : Metis.Name.name, arity : int, tptp : string} list,
-     relationMapping : {name : Metis.Name.name, arity : int, tptp : string} list} ->
-    mapping
-
-val addVarSetMapping : mapping -> Metis.NameSet.set -> mapping
-
-(* ------------------------------------------------------------------------- *)
-(* Interpreting TPTP functions and relations in a finite model.              *)
-(* ------------------------------------------------------------------------- *)
-
-val defaultFixedMap : Metis.Model.fixedMap
-
-val defaultModel : Metis.Model.parameters
-
-val ppFixedMap : mapping -> Metis.Model.fixedMap Metis.Print.pp
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP roles.                                                               *)
-(* ------------------------------------------------------------------------- *)
-
-datatype role =
-    AxiomRole
-  | ConjectureRole
-  | DefinitionRole
-  | NegatedConjectureRole
-  | PlainRole
-  | TheoremRole
-  | OtherRole of string;
-
-val isCnfConjectureRole : role -> bool
-
-val isFofConjectureRole : role -> bool
-
-val toStringRole : role -> string
-
-val fromStringRole : string -> role
-
-val ppRole : role Metis.Print.pp
-
-(* ------------------------------------------------------------------------- *)
-(* SZS statuses.                                                             *)
-(* ------------------------------------------------------------------------- *)
-
-datatype status =
-    CounterSatisfiableStatus
-  | TheoremStatus
-  | SatisfiableStatus
-  | UnknownStatus
-  | UnsatisfiableStatus
-
-val toStringStatus : status -> string
-
-val ppStatus : status Metis.Print.pp
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP literals.                                                            *)
-(* ------------------------------------------------------------------------- *)
-
-datatype literal =
-    Boolean of bool
-  | Literal of Metis.Literal.literal
-
-val negateLiteral : literal -> literal
-
-val functionsLiteral : literal -> Metis.NameAritySet.set
-
-val relationLiteral : literal -> Metis.Atom.relation option
-
-val freeVarsLiteral : literal -> Metis.NameSet.set
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP formula names.                                                       *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formulaName =
-    FormulaName of string
-
-val ppFormulaName : formulaName Metis.Print.pp
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP formula bodies.                                                      *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formulaBody =
-    CnfFormulaBody of literal list
-  | FofFormulaBody of Metis.Formula.formula
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP formula sources.                                                     *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formulaSource =
-    NoFormulaSource
-  | StripFormulaSource of
-      {inference : string,
-       parents : formulaName list}
-  | NormalizeFormulaSource of
-      {inference : Metis.Normalize.inference,
-       parents : formulaName list}
-  | ProofFormulaSource of
-      {inference : Metis.Proof.inference,
-       parents : formulaName list}
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP formulas.                                                            *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formula =
-    Formula of
-      {name : formulaName,
-       role : role,
-       body : formulaBody,
-       source : formulaSource}
-
-val nameFormula : formula -> formulaName
-
-val roleFormula : formula -> role
-
-val bodyFormula : formula -> formulaBody
-
-val sourceFormula : formula -> formulaSource
-
-val functionsFormula : formula -> Metis.NameAritySet.set
-
-val relationsFormula : formula -> Metis.NameAritySet.set
-
-val freeVarsFormula : formula -> Metis.NameSet.set
-
-val freeVarsListFormula : formula list -> Metis.NameSet.set
-
-val isCnfConjectureFormula : formula -> bool
-val isFofConjectureFormula : formula -> bool
-val isConjectureFormula : formula -> bool
-
-(* ------------------------------------------------------------------------- *)
-(* Clause information.                                                       *)
-(* ------------------------------------------------------------------------- *)
-
-datatype clauseSource =
-    CnfClauseSource of formulaName * literal list
-  | FofClauseSource of Metis.Normalize.thm
-
-type 'a clauseInfo = 'a Metis.LiteralSetMap.map
-
-type clauseNames = formulaName clauseInfo
-
-type clauseRoles = role clauseInfo
-
-type clauseSources = clauseSource clauseInfo
-
-val noClauseNames : clauseNames
-
-val noClauseRoles : clauseRoles
-
-val noClauseSources : clauseSources
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP problems.                                                            *)
-(* ------------------------------------------------------------------------- *)
-
-type comments = string list
-
-type includes = string list
-
-datatype problem =
-    Problem of
-      {comments : comments,
-       includes : includes,
-       formulas : formula list}
-
-val hasCnfConjecture : problem -> bool
-val hasFofConjecture : problem -> bool
-val hasConjecture : problem -> bool
-
-val freeVars : problem -> Metis.NameSet.set
-
-val mkProblem :
-    {comments : comments,
-     includes : includes,
-     names : clauseNames,
-     roles : clauseRoles,
-     problem : Metis.Problem.problem} -> problem
-
-val normalize :
-    problem ->
-    {subgoal : Metis.Formula.formula * formulaName list,
-     problem : Metis.Problem.problem,
-     sources : clauseSources} list
-
-val goal : problem -> Metis.Formula.formula
-
-val read : {mapping : mapping, filename : string} -> problem
-
-val write :
-    {problem : problem,
-     mapping : mapping,
-     filename : string} -> unit
-
-val prove : {filename : string, mapping : mapping} -> bool
-
-(* ------------------------------------------------------------------------- *)
-(* TSTP proofs.                                                              *)
-(* ------------------------------------------------------------------------- *)
-
-val fromProof :
-    {problem : problem,
-     proofs : {subgoal : Metis.Formula.formula * formulaName list,
-               sources : clauseSources,
-               refutation : Metis.Thm.thm} list} -> formula list
-
-end
-
-(**** Original file: Tptp.sml ****)
-
-structure Metis = struct open Metis
-(* Metis-specific ML environment *)
-nonfix ++ -- RL;
-val explode = String.explode;
-val implode = String.implode;
-val print = TextIO.print;
-val foldl = List.foldl;
-val foldr = List.foldr;
-
-(* ========================================================================= *)
-(* THE TPTP PROBLEM FILE FORMAT                                              *)
-(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License            *)
-(* ========================================================================= *)
-
-structure Tptp :> Tptp =
-struct
-
-open Useful;
-
-(* ------------------------------------------------------------------------- *)
-(* Default TPTP function and relation name mapping.                          *)
-(* ------------------------------------------------------------------------- *)
-
-val defaultFunctionMapping =
-    [(* Mapping TPTP functions to infix symbols *)
-     {name = "~", arity = 1, tptp = "negate"},
-     {name = "*", arity = 2, tptp = "multiply"},
-     {name = "/", arity = 2, tptp = "divide"},
-     {name = "+", arity = 2, tptp = "add"},
-     {name = "-", arity = 2, tptp = "subtract"},
-     {name = "::", arity = 2, tptp = "cons"},
-     {name = "@", arity = 2, tptp = "append"},
-     {name = ",", arity = 2, tptp = "pair"},
-     (* Expanding HOL symbols to TPTP alphanumerics *)
-     {name = ":", arity = 2, tptp = "has_type"},
-     {name = ".", arity = 2, tptp = "apply"}];
-
-val defaultRelationMapping =
-    [(* Mapping TPTP relations to infix symbols *)
-     {name = "=", arity = 2, tptp = "="},  (* this preserves the = symbol *)
-     {name = "==", arity = 2, tptp = "equalish"},
-     {name = "<=", arity = 2, tptp = "less_equal"},
-     {name = "<", arity = 2, tptp = "less_than"},
-     {name = ">=", arity = 2, tptp = "greater_equal"},
-     {name = ">", arity = 2, tptp = "greater_than"},
-     (* Expanding HOL symbols to TPTP alphanumerics *)
-     {name = "{}", arity = 1, tptp = "bool"}];
-
-(* ------------------------------------------------------------------------- *)
-(* Interpreting TPTP functions and relations in a finite model.              *)
-(* ------------------------------------------------------------------------- *)
-
-val defaultFunctionModel =
-    [{name = "~", arity = 1, model = Model.negName},
-     {name = "*", arity = 2, model = Model.multName},
-     {name = "/", arity = 2, model = Model.divName},
-     {name = "+", arity = 2, model = Model.addName},
-     {name = "-", arity = 2, model = Model.subName},
-     {name = "::", arity = 2, model = Model.consName},
-     {name = "@", arity = 2, model = Model.appendName},
-     {name = ":", arity = 2, model = Term.hasTypeFunctionName},
-     {name = "additive_identity", arity = 0, model = Model.numeralName 0},
-     {name = "app", arity = 2, model = Model.appendName},
-     {name = "complement", arity = 1, model = Model.complementName},
-     {name = "difference", arity = 2, model = Model.differenceName},
-     {name = "divide", arity = 2, model = Model.divName},
-     {name = "empty_set", arity = 0, model = Model.emptyName},
-     {name = "identity", arity = 0, model = Model.numeralName 1},
-     {name = "identity_map", arity = 1, model = Model.projectionName 1},
-     {name = "intersection", arity = 2, model = Model.intersectName},
-     {name = "minus", arity = 1, model = Model.negName},
-     {name = "multiplicative_identity", arity = 0, model = Model.numeralName 1},
-     {name = "n0", arity = 0, model = Model.numeralName 0},
-     {name = "n1", arity = 0, model = Model.numeralName 1},
-     {name = "n2", arity = 0, model = Model.numeralName 2},
-     {name = "n3", arity = 0, model = Model.numeralName 3},
-     {name = "n4", arity = 0, model = Model.numeralName 4},
-     {name = "n5", arity = 0, model = Model.numeralName 5},
-     {name = "n6", arity = 0, model = Model.numeralName 6},
-     {name = "n7", arity = 0, model = Model.numeralName 7},
-     {name = "n8", arity = 0, model = Model.numeralName 8},
-     {name = "n9", arity = 0, model = Model.numeralName 9},
-     {name = "nil", arity = 0, model = Model.nilName},
-     {name = "null_class", arity = 0, model = Model.emptyName},
-     {name = "singleton", arity = 1, model = Model.singletonName},
-     {name = "successor", arity = 1, model = Model.sucName},
-     {name = "symmetric_difference", arity = 2,
-      model = Model.symmetricDifferenceName},
-     {name = "union", arity = 2, model = Model.unionName},
-     {name = "universal_class", arity = 0, model = Model.universeName}];
-
-val defaultRelationModel =
-    [{name = "=", arity = 2, model = Atom.eqRelationName},
-     {name = "==", arity = 2, model = Atom.eqRelationName},
-     {name = "<=", arity = 2, model = Model.leName},
-     {name = "<", arity = 2, model = Model.ltName},
-     {name = ">=", arity = 2, model = Model.geName},
-     {name = ">", arity = 2, model = Model.gtName},
-     {name = "divides", arity = 2, model = Model.dividesName},
-     {name = "element_of_set", arity = 2, model = Model.memberName},
-     {name = "equal", arity = 2, model = Atom.eqRelationName},
-     {name = "equal_elements", arity = 2, model = Atom.eqRelationName},
-     {name = "equal_sets", arity = 2, model = Atom.eqRelationName},
-     {name = "equivalent", arity = 2, model = Atom.eqRelationName},
-     {name = "less", arity = 2, model = Model.ltName},
-     {name = "less_or_equal", arity = 2, model = Model.leName},
-     {name = "member", arity = 2, model = Model.memberName},
-     {name = "subclass", arity = 2, model = Model.subsetName},
-     {name = "subset", arity = 2, model = Model.subsetName}];
-
-(* ------------------------------------------------------------------------- *)
-(* Helper functions.                                                         *)
-(* ------------------------------------------------------------------------- *)
-
-fun isHdTlString hp tp s =
-    let
-      fun ct 0 = true
-        | ct i = tp (String.sub (s,i)) andalso ct (i - 1)
-
-      val n = size s
-    in
-      n > 0 andalso hp (String.sub (s,0)) andalso ct (n - 1)
-    end;
-
-fun stripSuffix pred s =
-    let
-      fun f 0 = ""
-        | f n =
-          let
-            val n' = n - 1
-          in
-            if pred (String.sub (s,n')) then f n'
-            else String.substring (s,0,n)
-          end
-    in
-      f (size s)
-    end;
-
-fun variant avoid s =
-    if not (StringSet.member s avoid) then s
-    else
-      let
-        val s = stripSuffix Char.isDigit s
-
-        fun var i =
-            let
-              val s_i = s ^ Int.toString i
-            in
-              if not (StringSet.member s_i avoid) then s_i else var (i + 1)
-            end
-      in
-        var 0
-      end;
-
-(* ------------------------------------------------------------------------- *)
-(* Mapping to legal TPTP names.                                              *)
-(* ------------------------------------------------------------------------- *)
-
-local
-  fun nonEmptyPred p l =
-      case l of
-        [] => false
-      | c :: cs => p (c,cs);
-
-  fun existsPred l x = List.exists (fn p => p x) l;
-
-  fun isTptpChar #"_" = true
-    | isTptpChar c = Char.isAlphaNum c;
-
-  fun isTptpName l s = nonEmptyPred (existsPred l) (explode s);
-
-  fun isRegular (c,cs) =
-      Char.isLower c andalso List.all isTptpChar cs;
-
-  fun isNumber (c,cs) =
-      Char.isDigit c andalso List.all Char.isDigit cs;
-
-  fun isDefined (c,cs) =
-      c = #"$" andalso nonEmptyPred isRegular cs;
-
-  fun isSystem (c,cs) =
-      c = #"$" andalso nonEmptyPred isDefined cs;
-in
-  fun mkTptpVarName s =
-      let
-        val s =
-            case List.filter isTptpChar (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
-      end;
-
-  val isTptpConstName = isTptpName [isRegular,isNumber,isDefined,isSystem]
-  and isTptpFnName = isTptpName [isRegular,isDefined,isSystem]
-  and isTptpPropName = isTptpName [isRegular,isDefined,isSystem]
-  and isTptpRelName = isTptpName [isRegular,isDefined,isSystem];
-
-  val isTptpFormulaName = isTptpName [isRegular,isNumber];
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Mapping to legal TPTP variable names.                                     *)
-(* ------------------------------------------------------------------------- *)
-
-datatype varToTptp = VarToTptp of StringSet.set * string NameMap.map;
-
-val emptyVarToTptp = VarToTptp (StringSet.empty, NameMap.new ());
-
-fun addVarToTptp vm v =
-    let
-      val VarToTptp (avoid,mapping) = vm
-    in
-      if NameMap.inDomain v mapping then vm
-      else
-        let
-          val s = variant avoid (mkTptpVarName (Name.toString v))
-
-          val avoid = StringSet.add avoid s
-          and mapping = NameMap.insert mapping (v,s)
-        in
-          VarToTptp (avoid,mapping)
-        end
-    end;
-
-local
-  fun add (v,vm) = addVarToTptp vm v;
-in
-  val addListVarToTptp = List.foldl add;
-
-  val addSetVarToTptp = NameSet.foldl add;
-end;
-
-val fromListVarToTptp = addListVarToTptp emptyVarToTptp;
-
-val fromSetVarToTptp = addSetVarToTptp emptyVarToTptp;
-
-fun getVarToTptp vm v =
-    let
-      val VarToTptp (_,mapping) = vm
-    in
-      case NameMap.peek mapping v of
-        SOME s => s
-      | NONE => raise Bug "Tptp.getVarToTptp: unknown var"
-    end;
-
-(* ------------------------------------------------------------------------- *)
-(* Mapping from TPTP variable names.                                         *)
-(* ------------------------------------------------------------------------- *)
-
-fun getVarFromTptp s = Name.fromString s;
-
-(* ------------------------------------------------------------------------- *)
-(* Mapping to TPTP function and relation names.                              *)
-(* ------------------------------------------------------------------------- *)
-
-datatype nameToTptp = NameToTptp of string NameArityMap.map;
-
-local
-  val emptyNames : string NameArityMap.map = NameArityMap.new ();
-
-  fun addNames ({name,arity,tptp},mapping) =
-      NameArityMap.insert mapping ((name,arity),tptp);
-
-  val fromListNames = List.foldl addNames emptyNames;
-in
-  fun mkNameToTptp mapping = NameToTptp (fromListNames mapping);
-end;
-
-local
-  fun escapeChar c =
-      case c of
-        #"\\" => "\\\\"
-      | #"'" => "\\'"
-      | #"\n" => "\\n"
-      | #"\t" => "\\t"
-      | _ => str c;
-
-  val escapeString = String.translate escapeChar;
-in
-  fun singleQuote s = "'" ^ escapeString s ^ "'";
-end;
-
-fun getNameToTptp isTptp s = if isTptp s then s else singleQuote s;
-
-fun getNameArityToTptp isZeroTptp isPlusTptp (NameToTptp mapping) na =
-    case NameArityMap.peek mapping na of
-      SOME s => s
-    | NONE =>
-      let
-        val (n,a) = na
-        val isTptp = if a = 0 then isZeroTptp else isPlusTptp
-        val s = Name.toString n
-      in
-        getNameToTptp isTptp s
-      end;
-
-(* ------------------------------------------------------------------------- *)
-(* Mapping from TPTP function and relation names.                            *)
-(* ------------------------------------------------------------------------- *)
-
-datatype nameFromTptp = NameFromTptp of (string * int, Name.name) Map.map;
-
-local
-  val stringArityCompare = prodCompare String.compare Int.compare;
-
-  val emptyStringArityMap = Map.new stringArityCompare;
-
-  fun addStringArityMap ({name,arity,tptp},mapping) =
-      Map.insert mapping ((tptp,arity),name);
-
-  val fromListStringArityMap =
-      List.foldl addStringArityMap emptyStringArityMap;
-in
-  fun mkNameFromTptp mapping = NameFromTptp (fromListStringArityMap mapping);
-end;
-
-fun getNameFromTptp (NameFromTptp mapping) sa =
-    case Map.peek mapping sa of
-      SOME n => n
-    | NONE =>
-      let
-        val (s,_) = sa
-      in
-        Name.fromString s
-      end;
-
-(* ------------------------------------------------------------------------- *)
-(* Mapping to and from TPTP variable, function and relation names.           *)
-(* ------------------------------------------------------------------------- *)
-
-datatype mapping =
-    Mapping of
-      {varTo : varToTptp,
-       fnTo : nameToTptp,
-       relTo : nameToTptp,
-       fnFrom : nameFromTptp,
-       relFrom : nameFromTptp};
-
-fun mkMapping mapping =
-    let
-      val {functionMapping,relationMapping} = mapping
-
-      val varTo = emptyVarToTptp
-      val fnTo = mkNameToTptp functionMapping
-      val relTo = mkNameToTptp relationMapping
-
-      val fnFrom = mkNameFromTptp functionMapping
-      val relFrom = mkNameFromTptp relationMapping
-    in
-      Mapping
-        {varTo = varTo,
-         fnTo = fnTo,
-         relTo = relTo,
-         fnFrom = fnFrom,
-         relFrom = relFrom}
-    end;
-
-fun addVarListMapping mapping vs =
-    let
-      val Mapping
-            {varTo,
-             fnTo,
-             relTo,
-             fnFrom,
-             relFrom} = mapping
-
-      val varTo = addListVarToTptp varTo vs
-    in
-      Mapping
-        {varTo = varTo,
-         fnTo = fnTo,
-         relTo = relTo,
-         fnFrom = fnFrom,
-         relFrom = relFrom}
-    end;
-
-fun addVarSetMapping mapping vs =
-    let
-      val Mapping
-            {varTo,
-             fnTo,
-             relTo,
-             fnFrom,
-             relFrom} = mapping
-
-      val varTo = addSetVarToTptp varTo vs
-    in
-      Mapping
-        {varTo = varTo,
-         fnTo = fnTo,
-         relTo = relTo,
-         fnFrom = fnFrom,
-         relFrom = relFrom}
-    end;
-
-fun varToTptp mapping v =
-    let
-      val Mapping {varTo,...} = mapping
-    in
-      getVarToTptp varTo v
-    end;
-
-fun fnToTptp mapping fa =
-    let
-      val Mapping {fnTo,...} = mapping
-    in
-      getNameArityToTptp isTptpConstName isTptpFnName fnTo fa
-    end;
-
-fun relToTptp mapping ra =
-    let
-      val Mapping {relTo,...} = mapping
-    in
-      getNameArityToTptp isTptpPropName isTptpRelName relTo ra
-    end;
-
-fun varFromTptp (_ : mapping) v = getVarFromTptp v;
-
-fun fnFromTptp mapping fa =
-    let
-      val Mapping {fnFrom,...} = mapping
-    in
-      getNameFromTptp fnFrom fa
-    end;
-
-fun relFromTptp mapping ra =
-    let
-      val Mapping {relFrom,...} = mapping
-    in
-      getNameFromTptp relFrom ra
-    end;
-
-val defaultMapping =
-    let
-      fun lift {name,arity,tptp} =
-          {name = Name.fromString name, arity = arity, tptp = tptp}
-
-      val functionMapping = map lift defaultFunctionMapping
-      and relationMapping = map lift defaultRelationMapping
-
-      val mapping =
-          {functionMapping = functionMapping,
-           relationMapping = relationMapping}
-    in
-      mkMapping mapping
-    end;
-
-(* ------------------------------------------------------------------------- *)
-(* Interpreting TPTP functions and relations in a finite model.              *)
-(* ------------------------------------------------------------------------- *)
-
-fun mkFixedMap funcModel relModel =
-    let
-      fun mkEntry {name,arity,model} = ((Name.fromString name, arity), model)
-
-      fun mkMap l = NameArityMap.fromList (map mkEntry l)
-    in
-      {functionMap = mkMap funcModel,
-       relationMap = mkMap relModel}
-    end;
-
-val defaultFixedMap = mkFixedMap defaultFunctionModel defaultRelationModel;
-
-val defaultModel =
-    let
-      val {size = N, fixed = fix} = Model.default
-
-      val fix = Model.mapFixed defaultFixedMap fix
-    in
-      {size = N, fixed = fix}
-    end;
-
-local
-  fun toTptpMap toTptp =
-      let
-        fun add ((src,arity),dest,m) =
-            let
-              val src = Name.fromString (toTptp (src,arity))
-            in
-              NameArityMap.insert m ((src,arity),dest)
-            end
-      in
-        fn m => NameArityMap.foldl add (NameArityMap.new ()) m
-      end;
-
-  fun toTptpFixedMap mapping fixMap =
-      let
-        val {functionMap = fnMap, relationMap = relMap} = fixMap
-
-        val fnMap = toTptpMap (fnToTptp mapping) fnMap
-        and relMap = toTptpMap (relToTptp mapping) relMap
-      in
-        {functionMap = fnMap,
-         relationMap = relMap}
-      end;
-in
-  fun ppFixedMap mapping fixMap =
-      Model.ppFixedMap (toTptpFixedMap mapping fixMap);
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP roles.                                                               *)
-(* ------------------------------------------------------------------------- *)
-
-datatype role =
-    AxiomRole
-  | ConjectureRole
-  | DefinitionRole
-  | NegatedConjectureRole
-  | PlainRole
-  | TheoremRole
-  | OtherRole of string;
-
-fun isCnfConjectureRole role =
-    case role of
-      NegatedConjectureRole => true
-    | _ => false;
-
-fun isFofConjectureRole role =
-    case role of
-      ConjectureRole => true
-    | _ => false;
-
-fun toStringRole role =
-    case role of
-      AxiomRole => "axiom"
-    | ConjectureRole => "conjecture"
-    | DefinitionRole => "definition"
-    | NegatedConjectureRole => "negated_conjecture"
-    | PlainRole => "plain"
-    | TheoremRole => "theorem"
-    | OtherRole s => s;
-
-fun fromStringRole s =
-    case s of
-      "axiom" => AxiomRole
-    | "conjecture" => ConjectureRole
-    | "definition" => DefinitionRole
-    | "negated_conjecture" => NegatedConjectureRole
-    | "plain" => PlainRole
-    | "theorem" => TheoremRole
-    | _ => OtherRole s;
-
-val ppRole = Print.ppMap toStringRole Print.ppString;
-
-(* ------------------------------------------------------------------------- *)
-(* SZS statuses.                                                             *)
-(* ------------------------------------------------------------------------- *)
-
-datatype status =
-    CounterSatisfiableStatus
-  | TheoremStatus
-  | SatisfiableStatus
-  | UnknownStatus
-  | UnsatisfiableStatus;
-
-fun toStringStatus status =
-    case status of
-      CounterSatisfiableStatus => "CounterSatisfiable"
-    | TheoremStatus => "Theorem"
-    | SatisfiableStatus => "Satisfiable"
-    | UnknownStatus => "Unknown"
-    | UnsatisfiableStatus => "Unsatisfiable";
-
-val ppStatus = Print.ppMap toStringStatus Print.ppString;
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP literals.                                                            *)
-(* ------------------------------------------------------------------------- *)
-
-datatype literal =
-    Boolean of bool
-  | Literal of Literal.literal;
-
-fun destLiteral lit =
-    case lit of
-      Literal l => l
-    | _ => raise Error "Tptp.destLiteral";
-
-fun isBooleanLiteral lit =
-    case lit of
-      Boolean _ => true
-    | _ => false;
-
-fun equalBooleanLiteral b lit =
-    case lit of
-      Boolean b' => b = b'
-    | _ => false;
-
-fun negateLiteral (Boolean b) = (Boolean (not b))
-  | negateLiteral (Literal l) = (Literal (Literal.negate l));
-
-fun functionsLiteral (Boolean _) = NameAritySet.empty
-  | functionsLiteral (Literal lit) = Literal.functions lit;
-
-fun relationLiteral (Boolean _) = NONE
-  | relationLiteral (Literal lit) = SOME (Literal.relation lit);
-
-fun literalToFormula (Boolean true) = Formula.True
-  | literalToFormula (Boolean false) = Formula.False
-  | literalToFormula (Literal lit) = Literal.toFormula lit;
-
-fun literalFromFormula Formula.True = Boolean true
-  | literalFromFormula Formula.False = Boolean false
-  | literalFromFormula fm = Literal (Literal.fromFormula fm);
-
-fun freeVarsLiteral (Boolean _) = NameSet.empty
-  | freeVarsLiteral (Literal lit) = Literal.freeVars lit;
-
-fun literalSubst sub lit =
-    case lit of
-      Boolean _ => lit
-    | Literal l => Literal (Literal.subst sub l);
-
-(* ------------------------------------------------------------------------- *)
-(* Printing formulas using TPTP syntax.                                      *)
-(* ------------------------------------------------------------------------- *)
-
-fun ppVar mapping v =
-    let
-      val s = varToTptp mapping v
-    in
-      Print.addString s
-    end;
-
-fun ppFnName mapping fa = Print.addString (fnToTptp mapping fa);
-
-fun ppConst mapping c = ppFnName mapping (c,0);
-
-fun ppTerm mapping =
-    let
-      fun term tm =
-          case tm of
-            Term.Var v => ppVar mapping v
-          | Term.Fn (f,tms) =>
-            case length tms of
-              0 => ppConst mapping f
-            | a =>
-              Print.blockProgram Print.Inconsistent 2
-                [ppFnName mapping (f,a),
-                 Print.addString "(",
-                 Print.ppOpList "," term tms,
-                 Print.addString ")"]
-    in
-      Print.block Print.Inconsistent 0 o term
-    end;
-
-fun ppRelName mapping ra = Print.addString (relToTptp mapping ra);
-
-fun ppProp mapping p = ppRelName mapping (p,0);
-
-fun ppAtom mapping (r,tms) =
-    case length tms of
-      0 => ppProp mapping r
-    | a =>
-      Print.blockProgram Print.Inconsistent 2
-        [ppRelName mapping (r,a),
-         Print.addString "(",
-         Print.ppOpList "," (ppTerm mapping) tms,
-         Print.addString ")"];
-
-local
-  val neg = Print.sequence (Print.addString "~") (Print.addBreak 1);
-
-  fun fof mapping fm =
-      case fm of
-        Formula.And _ => assoc_binary mapping ("&", Formula.stripConj fm)
-      | Formula.Or _ => assoc_binary mapping ("|", Formula.stripDisj fm)
-      | Formula.Imp a_b => nonassoc_binary mapping ("=>",a_b)
-      | Formula.Iff a_b => nonassoc_binary mapping ("<=>",a_b)
-      | _ => unitary mapping fm
-
-  and nonassoc_binary mapping (s,a_b) =
-      Print.ppOp2 (" " ^ s) (unitary mapping) (unitary mapping) a_b
-
-  and assoc_binary mapping (s,l) = Print.ppOpList (" " ^ s) (unitary mapping) l
-
-  and unitary mapping fm =
-      case fm of
-        Formula.True => Print.addString "$true"
-      | Formula.False => Print.addString "$false"
-      | Formula.Forall _ => quantified mapping ("!", Formula.stripForall fm)
-      | Formula.Exists _ => quantified mapping ("?", Formula.stripExists fm)
-      | Formula.Not _ =>
-        (case total Formula.destNeq fm of
-           SOME a_b => Print.ppOp2 " !=" (ppTerm mapping) (ppTerm mapping) a_b
-         | NONE =>
-           let
-             val (n,fm) = Formula.stripNeg fm
-           in
-             Print.blockProgram Print.Inconsistent 2
-               [Print.duplicate n neg,
-                unitary mapping fm]
-           end)
-      | Formula.Atom atm =>
-        (case total Formula.destEq fm of
-           SOME a_b => Print.ppOp2 " =" (ppTerm mapping) (ppTerm mapping) a_b
-         | NONE => ppAtom mapping atm)
-      | _ =>
-        Print.blockProgram Print.Inconsistent 1
-          [Print.addString "(",
-           fof mapping fm,
-           Print.addString ")"]
-
-  and quantified mapping (q,(vs,fm)) =
-      let
-        val mapping = addVarListMapping mapping vs
-      in
-        Print.blockProgram Print.Inconsistent 2
-          [Print.addString q,
-           Print.addString " ",
-           Print.blockProgram Print.Inconsistent (String.size q)
-             [Print.addString "[",
-              Print.ppOpList "," (ppVar mapping) vs,
-              Print.addString "] :"],
-           Print.addBreak 1,
-           unitary mapping fm]
-      end;
-in
-  fun ppFof mapping fm = Print.block Print.Inconsistent 0 (fof mapping fm);
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Lexing TPTP files.                                                        *)
-(* ------------------------------------------------------------------------- *)
-
-datatype token =
-    AlphaNum of string
-  | Punct of char
-  | Quote of string;
-
-fun isAlphaNum #"_" = true
-  | isAlphaNum c = Char.isAlphaNum c;
-
-local
-  open Parse;
-
-  infixr 9 >>++
-  infixr 8 ++
-  infixr 7 >>
-  infixr 6 ||
-
-  val alphaNumToken = atLeastOne (some isAlphaNum) >> (AlphaNum o implode);
-
-  val punctToken =
-      let
-        val punctChars = "<>=-*+/\\?@|!$%&#^:;~()[]{}.,"
-      in
-        some (Char.contains punctChars) >> Punct
-      end;
-
-  val quoteToken =
-      let
-        val escapeParser =
-            some (equal #"'") >> singleton ||
-            some (equal #"\\") >> singleton
-
-        fun stopOn #"'" = true
-          | stopOn #"\n" = true
-          | stopOn _ = false
-
-        val quotedParser =
-            some (equal #"\\") ++ escapeParser >> op:: ||
-            some (not o stopOn) >> singleton
-      in
-        exactChar #"'" ++ many quotedParser ++ exactChar #"'" >>
-        (fn (_,(l,_)) => Quote (implode (List.concat l)))
-      end;
-
-  val lexToken = alphaNumToken || punctToken || quoteToken;
-
-  val space = many (some Char.isSpace) >> K ();
-in
-  val lexer = (space ++ lexToken ++ space) >> (fn ((),(tok,())) => tok);
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP clauses.                                                             *)
-(* ------------------------------------------------------------------------- *)
-
-type clause = literal list;
-
-val clauseFunctions =
-    let
-      fun funcs (lit,acc) = NameAritySet.union (functionsLiteral lit) acc
-    in
-      foldl funcs NameAritySet.empty
-    end;
-
-val clauseRelations =
-    let
-      fun rels (lit,acc) =
-          case relationLiteral lit of
-            NONE => acc
-          | SOME r => NameAritySet.add acc r
-    in
-      foldl rels NameAritySet.empty
-    end;
-
-val clauseFreeVars =
-    let
-      fun fvs (lit,acc) = NameSet.union (freeVarsLiteral lit) acc
-    in
-      foldl fvs NameSet.empty
-    end;
-
-fun clauseSubst sub lits = map (literalSubst sub) lits;
-
-fun clauseToFormula lits = Formula.listMkDisj (map literalToFormula lits);
-
-fun clauseFromFormula fm = map literalFromFormula (Formula.stripDisj fm);
-
-fun clauseFromLiteralSet cl =
-    clauseFromFormula
-      (Formula.listMkDisj (LiteralSet.transform Literal.toFormula cl));
-
-fun clauseFromThm th = clauseFromLiteralSet (Thm.clause th);
-
-fun ppClause mapping = Print.ppMap clauseToFormula (ppFof mapping);
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP formula names.                                                       *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formulaName = FormulaName of string;
-
-datatype formulaNameSet = FormulaNameSet of formulaName Set.set;
-
-fun compareFormulaName (FormulaName s1, FormulaName s2) =
-    String.compare (s1,s2);
-
-fun toTptpFormulaName (FormulaName s) =
-    getNameToTptp isTptpFormulaName s;
-
-val ppFormulaName = Print.ppMap toTptpFormulaName Print.ppString;
-
-val emptyFormulaNameSet = FormulaNameSet (Set.empty compareFormulaName);
-
-fun memberFormulaNameSet n (FormulaNameSet s) = Set.member n s;
-
-fun addFormulaNameSet (FormulaNameSet s) n = FormulaNameSet (Set.add s n);
-
-fun addListFormulaNameSet (FormulaNameSet s) l =
-    FormulaNameSet (Set.addList s l);
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP formula bodies.                                                      *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formulaBody =
-    CnfFormulaBody of literal list
-  | FofFormulaBody of Formula.formula;
-
-fun destCnfFormulaBody body =
-    case body of
-      CnfFormulaBody x => x
-    | _ => raise Error "destCnfFormulaBody";
-
-val isCnfFormulaBody = can destCnfFormulaBody;
-
-fun destFofFormulaBody body =
-    case body of
-      FofFormulaBody x => x
-    | _ => raise Error "destFofFormulaBody";
-
-val isFofFormulaBody = can destFofFormulaBody;
-
-fun formulaBodyFunctions body =
-    case body of
-      CnfFormulaBody cl => clauseFunctions cl
-    | FofFormulaBody fm => Formula.functions fm;
-
-fun formulaBodyRelations body =
-    case body of
-      CnfFormulaBody cl => clauseRelations cl
-    | FofFormulaBody fm => Formula.relations fm;
-
-fun formulaBodyFreeVars body =
-    case body of
-      CnfFormulaBody cl => clauseFreeVars cl
-    | FofFormulaBody fm => Formula.freeVars fm;
-
-fun ppFormulaBody mapping body =
-    case body of
-      CnfFormulaBody cl => ppClause mapping cl
-    | FofFormulaBody fm => ppFof mapping (Formula.generalize fm);
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP formula sources.                                                     *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formulaSource =
-    NoFormulaSource
-  | StripFormulaSource of
-      {inference : string,
-       parents : formulaName list}
-  | NormalizeFormulaSource of
-      {inference : Normalize.inference,
-       parents : formulaName list}
-  | ProofFormulaSource of
-      {inference : Proof.inference,
-       parents : formulaName list};
-
-fun isNoFormulaSource source =
-    case source of
-      NoFormulaSource => true
-    | _ => false;
-
-fun functionsFormulaSource source =
-    case source of
-      NoFormulaSource => NameAritySet.empty
-    | StripFormulaSource _ => NameAritySet.empty
-    | NormalizeFormulaSource data =>
-      let
-        val {inference = inf, parents = _} = data
-      in
-        case inf of
-          Normalize.Axiom fm => Formula.functions fm
-        | Normalize.Definition (_,fm) => Formula.functions fm
-        | _ => NameAritySet.empty
-      end
-    | ProofFormulaSource data =>
-      let
-        val {inference = inf, parents = _} = data
-      in
-        case inf of
-          Proof.Axiom cl => LiteralSet.functions cl
-        | Proof.Assume atm => Atom.functions atm
-        | Proof.Subst (sub,_) => Subst.functions sub
-        | Proof.Resolve (atm,_,_) => Atom.functions atm
-        | Proof.Refl tm => Term.functions tm
-        | Proof.Equality (lit,_,tm) =>
-          NameAritySet.union (Literal.functions lit) (Term.functions tm)
-      end;
-
-fun relationsFormulaSource source =
-    case source of
-      NoFormulaSource => NameAritySet.empty
-    | StripFormulaSource _ => NameAritySet.empty
-    | NormalizeFormulaSource data =>
-      let
-        val {inference = inf, parents = _} = data
-      in
-        case inf of
-          Normalize.Axiom fm => Formula.relations fm
-        | Normalize.Definition (_,fm) => Formula.relations fm
-        | _ => NameAritySet.empty
-      end
-    | ProofFormulaSource data =>
-      let
-        val {inference = inf, parents = _} = data
-      in
-        case inf of
-          Proof.Axiom cl => LiteralSet.relations cl
-        | Proof.Assume atm => NameAritySet.singleton (Atom.relation atm)
-        | Proof.Subst _ => NameAritySet.empty
-        | Proof.Resolve (atm,_,_) => NameAritySet.singleton (Atom.relation atm)
-        | Proof.Refl tm => NameAritySet.empty
-        | Proof.Equality (lit,_,_) =>
-          NameAritySet.singleton (Literal.relation lit)
-      end;
-
-fun freeVarsFormulaSource source =
-    case source of
-      NoFormulaSource => NameSet.empty
-    | StripFormulaSource _ => NameSet.empty
-    | NormalizeFormulaSource data => NameSet.empty
-    | ProofFormulaSource data =>
-      let
-        val {inference = inf, parents = _} = data
-      in
-        case inf of
-          Proof.Axiom cl => LiteralSet.freeVars cl
-        | Proof.Assume atm => Atom.freeVars atm
-        | Proof.Subst (sub,_) => Subst.freeVars sub
-        | Proof.Resolve (atm,_,_) => Atom.freeVars atm
-        | Proof.Refl tm => Term.freeVars tm
-        | Proof.Equality (lit,_,tm) =>
-          NameSet.union (Literal.freeVars lit) (Term.freeVars tm)
-      end;
-
-local
-  val GEN_INFERENCE = "inference"
-  and GEN_INTRODUCED = "introduced";
-
-  fun nameStrip inf = inf;
-
-  fun ppStrip mapping inf = Print.skip;
-
-  fun nameNormalize inf =
-      case inf of
-        Normalize.Axiom _ => "canonicalize"
-      | Normalize.Definition _ => "canonicalize"
-      | Normalize.Simplify _ => "simplify"
-      | Normalize.Conjunct _ => "conjunct"
-      | Normalize.Specialize _ => "specialize"
-      | Normalize.Skolemize _ => "skolemize"
-      | Normalize.Clausify _ => "clausify";
-
-  fun ppNormalize mapping inf = Print.skip;
-
-  fun nameProof inf =
-      case inf of
-        Proof.Axiom _ => "canonicalize"
-      | Proof.Assume _ => "assume"
-      | Proof.Subst _ => "subst"
-      | Proof.Resolve _ => "resolve"
-      | Proof.Refl _ => "refl"
-      | Proof.Equality _ => "equality";
-
-  local
-    fun ppTermInf mapping = ppTerm mapping;
-
-    fun ppAtomInf mapping atm =
-        case total Atom.destEq atm of
-          SOME (a,b) => ppAtom mapping (Name.fromString "$equal", [a,b])
-        | NONE => ppAtom mapping atm;
-
-    fun ppLiteralInf mapping (pol,atm) =
-        Print.sequence
-          (if pol then Print.skip else Print.addString "~ ")
-          (ppAtomInf mapping atm);
-  in
-    fun ppProofTerm mapping =
-        Print.ppBracket "$fot(" ")" (ppTermInf mapping);
-
-    fun ppProofAtom mapping =
-        Print.ppBracket "$cnf(" ")" (ppAtomInf mapping);
-
-    fun ppProofLiteral mapping =
-        Print.ppBracket "$cnf(" ")" (ppLiteralInf mapping);
-  end;
-
-  val ppProofVar = ppVar;
-
-  val ppProofPath = Term.ppPath;
-
-  fun ppProof mapping inf =
-      Print.blockProgram Print.Inconsistent 1
-        [Print.addString "[",
-         (case inf of
-            Proof.Axiom _ => Print.skip
-          | Proof.Assume atm => ppProofAtom mapping atm
-          | Proof.Subst _ => Print.skip
-          | Proof.Resolve (atm,_,_) => ppProofAtom mapping atm
-          | Proof.Refl tm => ppProofTerm mapping tm
-          | Proof.Equality (lit,path,tm) =>
-            Print.program
-              [ppProofLiteral mapping lit,
-               Print.addString ",",
-               Print.addBreak 1,
-               ppProofPath path,
-               Print.addString ",",
-               Print.addBreak 1,
-               ppProofTerm mapping tm]),
-         Print.addString "]"];
-
-  val ppParent = ppFormulaName;
-
-  fun ppProofSubst mapping =
-      Print.ppMap Subst.toList
-        (Print.ppList
-           (Print.ppBracket "bind(" ")"
-              (Print.ppOp2 "," (ppProofVar mapping)
-                 (ppProofTerm mapping))));
-
-  fun ppProofParent mapping (p,s) =
-      if Subst.null s then ppParent p
-      else Print.ppOp2 " :" ppParent (ppProofSubst mapping) (p,s);
-in
-  fun ppFormulaSource mapping source =
-      case source of
-        NoFormulaSource => Print.skip
-      | StripFormulaSource {inference,parents} =>
-        let
-          val gen = GEN_INFERENCE
-
-          val name = nameStrip inference
-        in
-          Print.blockProgram Print.Inconsistent (size gen + 1)
-            [Print.addString gen,
-             Print.addString "(",
-             Print.addString name,
-             Print.addString ",",
-             Print.addBreak 1,
-             Print.ppBracket "[" "]" (ppStrip mapping) inference,
-             Print.addString ",",
-             Print.addBreak 1,
-             Print.ppList ppParent parents,
-             Print.addString ")"]
-        end
-      | NormalizeFormulaSource {inference,parents} =>
-        let
-          val gen = GEN_INFERENCE
-
-          val name = nameNormalize inference
-        in
-          Print.blockProgram Print.Inconsistent (size gen + 1)
-            [Print.addString gen,
-             Print.addString "(",
-             Print.addString name,
-             Print.addString ",",
-             Print.addBreak 1,
-             Print.ppBracket "[" "]" (ppNormalize mapping) inference,
-             Print.addString ",",
-             Print.addBreak 1,
-             Print.ppList ppParent parents,
-             Print.addString ")"]
-        end
-      | ProofFormulaSource {inference,parents} =>
-        let
-          val isTaut = null parents
-
-          val gen = if isTaut then GEN_INTRODUCED else GEN_INFERENCE
-
-          val name = nameProof inference
-
-          val parents =
-              let
-                val sub =
-                    case inference of
-                      Proof.Subst (s,_) => s
-                    | _ => Subst.empty
-              in
-                map (fn parent => (parent,sub)) parents
-              end
-        in
-          Print.blockProgram Print.Inconsistent (size gen + 1)
-            ([Print.addString gen,
-              Print.addString "("] @
-             (if isTaut then
-                [Print.addString "tautology",
-                 Print.addString ",",
-                 Print.addBreak 1,
-                 Print.blockProgram Print.Inconsistent 1
-                   [Print.addString "[",
-                    Print.addString name,
-                    Print.addString ",",
-                    Print.addBreak 1,
-                    ppProof mapping inference,
-                    Print.addString "]"]]
-              else
-                [Print.addString name,
-                 Print.addString ",",
-                 Print.addBreak 1,
-                 ppProof mapping inference,
-                 Print.addString ",",
-                 Print.addBreak 1,
-                 Print.ppList (ppProofParent mapping) parents]) @
-             [Print.addString ")"])
-        end
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP formulas.                                                            *)
-(* ------------------------------------------------------------------------- *)
-
-datatype formula =
-    Formula of
-      {name : formulaName,
-       role : role,
-       body : formulaBody,
-       source : formulaSource};
-
-fun nameFormula (Formula {name,...}) = name;
-
-fun roleFormula (Formula {role,...}) = role;
-
-fun bodyFormula (Formula {body,...}) = body;
-
-fun sourceFormula (Formula {source,...}) = source;
-
-fun destCnfFormula fm = destCnfFormulaBody (bodyFormula fm);
-
-val isCnfFormula = can destCnfFormula;
-
-fun destFofFormula fm = destFofFormulaBody (bodyFormula fm);
-
-val isFofFormula = can destFofFormula;
-
-fun functionsFormula fm =
-    let
-      val bodyFns = formulaBodyFunctions (bodyFormula fm)
-      and sourceFns = functionsFormulaSource (sourceFormula fm)
-    in
-      NameAritySet.union bodyFns sourceFns
-    end;
-
-fun relationsFormula fm =
-    let
-      val bodyRels = formulaBodyRelations (bodyFormula fm)
-      and sourceRels = relationsFormulaSource (sourceFormula fm)
-    in
-      NameAritySet.union bodyRels sourceRels
-    end;
-
-fun freeVarsFormula fm =
-    let
-      val bodyFvs = formulaBodyFreeVars (bodyFormula fm)
-      and sourceFvs = freeVarsFormulaSource (sourceFormula fm)
-    in
-      NameSet.union bodyFvs sourceFvs
-    end;
-
-val freeVarsListFormula =
-    let
-      fun add (fm,vs) = NameSet.union vs (freeVarsFormula fm)
-    in
-      List.foldl add NameSet.empty
-    end;
-
-val formulasFunctions =
-    let
-      fun funcs (fm,acc) = NameAritySet.union (functionsFormula fm) acc
-    in
-      foldl funcs NameAritySet.empty
-    end;
-
-val formulasRelations =
-    let
-      fun rels (fm,acc) = NameAritySet.union (relationsFormula fm) acc
-    in
-      foldl rels NameAritySet.empty
-    end;
-
-fun isCnfConjectureFormula fm =
-    case fm of
-      Formula {role, body = CnfFormulaBody _, ...} => isCnfConjectureRole role
-    | _ => false;
-
-fun isFofConjectureFormula fm =
-    case fm of
-      Formula {role, body = FofFormulaBody _, ...} => isFofConjectureRole role
-    | _ => false;
-
-fun isConjectureFormula fm =
-    isCnfConjectureFormula fm orelse
-    isFofConjectureFormula fm;
-
-(* Parsing and pretty-printing *)
-
-fun ppFormula mapping fm =
-    let
-      val Formula {name,role,body,source} = fm
-
-      val gen =
-          case body of
-            CnfFormulaBody _ => "cnf"
-          | FofFormulaBody _ => "fof"
-    in
-      Print.blockProgram Print.Inconsistent (size gen + 1)
-        ([Print.addString gen,
-          Print.addString "(",
-          ppFormulaName name,
-          Print.addString ",",
-          Print.addBreak 1,
-          ppRole role,
-          Print.addString ",",
-          Print.addBreak 1,
-          Print.blockProgram Print.Consistent 1
-            [Print.addString "(",
-             ppFormulaBody mapping body,
-             Print.addString ")"]] @
-         (if isNoFormulaSource source then []
-          else
-            [Print.addString ",",
-             Print.addBreak 1,
-             ppFormulaSource mapping source]) @
-         [Print.addString ")."])
-    end;
-
-fun formulaToString mapping = Print.toString (ppFormula mapping);
-
-local
-  open Parse;
-
-  infixr 9 >>++
-  infixr 8 ++
-  infixr 7 >>
-  infixr 6 ||
-
-  fun someAlphaNum p =
-      maybe (fn AlphaNum s => if p s then SOME s else NONE | _ => NONE);
-
-  fun alphaNumParser s = someAlphaNum (equal s) >> K ();
-
-  val lowerParser = someAlphaNum (fn s => Char.isLower (String.sub (s,0)));
-
-  val upperParser = someAlphaNum (fn s => Char.isUpper (String.sub (s,0)));
-
-  val stringParser = lowerParser || upperParser;
-
-  val numberParser = someAlphaNum (List.all Char.isDigit o explode);
-
-  fun somePunct p =
-      maybe (fn Punct c => if p c then SOME c else NONE | _ => NONE);
-
-  fun punctParser c = somePunct (equal c) >> K ();
-
-  val quoteParser = maybe (fn Quote s => SOME s | _ => NONE);
-
-  local
-    fun f [] = raise Bug "symbolParser"
-      | f [x] = x
-      | f (h :: t) = (h ++ f t) >> K ();
-  in
-    fun symbolParser s = f (map punctParser (explode s));
-  end;
-
-  val definedParser =
-      punctParser #"$" ++ someAlphaNum (K true) >> (fn ((),s) => "$" ^ s);
-
-  val systemParser =
-      punctParser #"$" ++ punctParser #"$" ++ someAlphaNum (K true) >>
-      (fn ((),((),s)) => "$$" ^ s);
-
-  val nameParser =
-      (stringParser || numberParser || quoteParser) >> FormulaName;
-
-  val roleParser = lowerParser >> fromStringRole;
-
-  local
-    fun isProposition s = isHdTlString Char.isLower isAlphaNum s;
-  in
-    val propositionParser =
-        someAlphaNum isProposition ||
-        definedParser ||
-        systemParser ||
-        quoteParser;
-  end;
-
-  local
-    fun isFunction s = isHdTlString Char.isLower isAlphaNum s;
-  in
-    val functionParser =
-        someAlphaNum isFunction ||
-        definedParser ||
-        systemParser ||
-        quoteParser;
-  end;
-
-  local
-    fun isConstant s = isHdTlString Char.isLower isAlphaNum s;
-  in
-    val constantParser =
-        someAlphaNum isConstant ||
-        definedParser ||
-        numberParser ||
-        systemParser ||
-        quoteParser;
-  end;
-
-  val varParser = upperParser;
-
-  val varListParser =
-      (punctParser #"[" ++ varParser ++
-       many ((punctParser #"," ++ varParser) >> snd) ++
-       punctParser #"]") >>
-      (fn ((),(h,(t,()))) => h :: t);
-
-  fun mkVarName mapping v = varFromTptp mapping v;
-
-  fun mkVar mapping v =
-      let
-        val v = mkVarName mapping v
-      in
-        Term.Var v
-      end
-
-  fun mkFn mapping (f,tms) =
-      let
-        val f = fnFromTptp mapping (f, length tms)
-      in
-        Term.Fn (f,tms)
-      end;
-
-  fun mkConst mapping c = mkFn mapping (c,[]);
-
-  fun mkAtom mapping (r,tms) =
-      let
-        val r = relFromTptp mapping (r, length tms)
-      in
-        (r,tms)
-      end;
-
-  fun termParser mapping input =
-      let
-        val fnP = functionArgumentsParser mapping >> mkFn mapping
-        val nonFnP = nonFunctionArgumentsTermParser mapping
-      in
-        fnP || nonFnP
-      end input
-
-  and functionArgumentsParser mapping input =
-      let
-        val commaTmP = (punctParser #"," ++ termParser mapping) >> snd
-      in
-        (functionParser ++ punctParser #"(" ++ termParser mapping ++
-         many commaTmP ++ punctParser #")") >>
-        (fn (f,((),(t,(ts,())))) => (f, t :: ts))
-      end input
-
-  and nonFunctionArgumentsTermParser mapping input =
-      let
-        val varP = varParser >> mkVar mapping
-        val constP = constantParser >> mkConst mapping
-      in
-        varP || constP
-      end input;
-
-  fun binaryAtomParser mapping tm input =
-      let
-        val eqP =
-            (punctParser #"=" ++ termParser mapping) >>
-            (fn ((),r) => (true,("$equal",[tm,r])))
-
-        val neqP =
-            (symbolParser "!=" ++ termParser mapping) >>
-            (fn ((),r) => (false,("$equal",[tm,r])))
-      in
-        eqP || neqP
-      end input;
-
-  fun maybeBinaryAtomParser mapping (s,tms) input =
-      let
-        val tm = mkFn mapping (s,tms)
-      in
-        optional (binaryAtomParser mapping tm) >>
-        (fn SOME lit => lit
-          | NONE => (true,(s,tms)))
-      end input;
-
-  fun literalAtomParser mapping input =
-      let
-        val fnP =
-            functionArgumentsParser mapping >>++
-            maybeBinaryAtomParser mapping
-
-        val nonFnP =
-            nonFunctionArgumentsTermParser mapping >>++
-            binaryAtomParser mapping
-
-        val propP = propositionParser >> (fn s => (true,(s,[])))
-      in
-        fnP || nonFnP || propP
-      end input;
-
-  fun atomParser mapping input =
-      let
-        fun mk (pol,rel) =
-          case rel of
-            ("$true",[]) => Boolean pol
-          | ("$false",[]) => Boolean (not pol)
-          | ("$equal",[l,r]) => Literal (pol, Atom.mkEq (l,r))
-          | (r,tms) => Literal (pol, mkAtom mapping (r,tms))
-      in
-        literalAtomParser mapping >> mk
-      end input;
-
-  fun literalParser mapping input =
-      let
-        val negP =
-            (punctParser #"~" ++ atomParser mapping) >>
-            (negateLiteral o snd)
-
-        val posP = atomParser mapping
-      in
-        negP || posP
-      end input;
-
-  fun disjunctionParser mapping input =
-      let
-        val orLitP = (punctParser #"|" ++ literalParser mapping) >> snd
-      in
-        (literalParser mapping ++ many orLitP) >> (fn (h,t) => h :: t)
-      end input;
-
-  fun clauseParser mapping input =
-      let
-        val disjP = disjunctionParser mapping
-
-        val bracketDisjP =
-            (punctParser #"(" ++ disjP ++ punctParser #")") >>
-            (fn ((),(c,())) => c)
-      in
-        bracketDisjP || disjP
-      end input;
-
-  val binaryConnectiveParser =
-      (symbolParser "<=>" >> K Formula.Iff) ||
-      (symbolParser "=>" >> K Formula.Imp) ||
-      (symbolParser "<=" >> K (fn (f,g) => Formula.Imp (g,f))) ||
-      (symbolParser "<~>" >> K (Formula.Not o Formula.Iff)) ||
-      (symbolParser "~|" >> K (Formula.Not o Formula.Or)) ||
-      (symbolParser "~&" >> K (Formula.Not o Formula.And));
-
-  val quantifierParser =
-      (punctParser #"!" >> K Formula.listMkForall) ||
-      (punctParser #"?" >> K Formula.listMkExists);
-
-  fun fofFormulaParser mapping input =
-      let
-        fun mk (f,NONE) = f
-          | mk (f, SOME t) = t f
-      in
-        (unitaryFormulaParser mapping ++
-         optional (binaryFormulaParser mapping)) >> mk
-      end input
-
-  and binaryFormulaParser mapping input =
-      let
-        val nonAssocP = nonAssocBinaryFormulaParser mapping
-
-        val assocP = assocBinaryFormulaParser mapping
-      in
-        nonAssocP || assocP
-      end input
-
-  and nonAssocBinaryFormulaParser mapping input =
-      let
-        fun mk (c,g) f = c (f,g)
-      in
-        (binaryConnectiveParser ++ unitaryFormulaParser mapping) >> mk
-      end input
-
-  and assocBinaryFormulaParser mapping input =
-      let
-        val orP = orFormulaParser mapping
-
-        val andP = andFormulaParser mapping
-      in
-        orP || andP
-      end input
-
-  and orFormulaParser mapping input =
-      let
-        val orFmP = (punctParser #"|" ++ unitaryFormulaParser mapping) >> snd
-      in
-        atLeastOne orFmP >>
-        (fn fs => fn f => Formula.listMkDisj (f :: fs))
-      end input
-
-  and andFormulaParser mapping input =
-      let
-        val andFmP = (punctParser #"&" ++ unitaryFormulaParser mapping) >> snd
-      in
-        atLeastOne andFmP >>
-        (fn fs => fn f => Formula.listMkConj (f :: fs))
-      end input
-
-  and unitaryFormulaParser mapping input =
-      let
-        val quantP = quantifiedFormulaParser mapping
-
-        val unaryP = unaryFormulaParser mapping
-
-        val brackP =
-            (punctParser #"(" ++ fofFormulaParser mapping ++
-             punctParser #")") >>
-            (fn ((),(f,())) => f)
-
-        val atomP =
-            atomParser mapping >>
-            (fn Boolean b => Formula.mkBoolean b
-              | Literal l => Literal.toFormula l)
-      in
-        quantP ||
-        unaryP ||
-        brackP ||
-        atomP
-      end input
-
-  and quantifiedFormulaParser mapping input =
-      let
-        fun mk (q,(vs,((),f))) = q (map (mkVarName mapping) vs, f)
-      in
-        (quantifierParser ++ varListParser ++ punctParser #":" ++
-         unitaryFormulaParser mapping) >> mk
-      end input
-
-  and unaryFormulaParser mapping input =
-      let
-        fun mk (c,f) = c f
-      in
-        (unaryConnectiveParser ++ unitaryFormulaParser mapping) >> mk
-      end input
-
-  and unaryConnectiveParser input =
-      (punctParser #"~" >> K Formula.Not) input;
-
-  fun cnfParser mapping input =
-      let
-        fun mk ((),((),(name,((),(role,((),(cl,((),())))))))) =
-            let
-              val body = CnfFormulaBody cl
-              val source = NoFormulaSource
-            in
-              Formula
-                {name = name,
-                 role = role,
-                 body = body,
-                 source = source}
-            end
-      in
-        (alphaNumParser "cnf" ++ punctParser #"(" ++
-         nameParser ++ punctParser #"," ++
-         roleParser ++ punctParser #"," ++
-         clauseParser mapping ++ punctParser #")" ++
-         punctParser #".") >> mk
-      end input;
-
-  fun fofParser mapping input =
-      let
-        fun mk ((),((),(name,((),(role,((),(fm,((),())))))))) =
-            let
-              val body = FofFormulaBody fm
-              val source = NoFormulaSource
-            in
-              Formula
-                {name = name,
-                 role = role,
-                 body = body,
-                 source = source}
-            end
-      in
-        (alphaNumParser "fof" ++ punctParser #"(" ++
-         nameParser ++ punctParser #"," ++
-         roleParser ++ punctParser #"," ++
-         fofFormulaParser mapping ++ punctParser #")" ++
-         punctParser #".") >> mk
-      end input;
-in
-  fun formulaParser mapping input =
-      let
-        val cnfP = cnfParser mapping
-
-        val fofP = fofParser mapping
-      in
-        cnfP || fofP
-      end input;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Include declarations.                                                     *)
-(* ------------------------------------------------------------------------- *)
-
-fun ppInclude i =
-    Print.blockProgram Print.Inconsistent 2
-      [Print.addString "include('",
-       Print.addString i,
-       Print.addString "')."];
-
-val includeToString = Print.toString ppInclude;
-
-local
-  open Parse;
-
-  infixr 9 >>++
-  infixr 8 ++
-  infixr 7 >>
-  infixr 6 ||
-
-  val filenameParser = maybe (fn Quote s => SOME s | _ => NONE);
-in
-  val includeParser =
-      (some (equal (AlphaNum "include")) ++
-       some (equal (Punct #"(")) ++
-       filenameParser ++
-       some (equal (Punct #")")) ++
-       some (equal (Punct #"."))) >>
-      (fn (_,(_,(f,(_,_)))) => f);
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Parsing TPTP files.                                                       *)
-(* ------------------------------------------------------------------------- *)
-
-datatype declaration =
-    IncludeDeclaration of string
-  | FormulaDeclaration of formula;
-
-val partitionDeclarations =
-    let
-      fun part (d,(il,fl)) =
-          case d of
-            IncludeDeclaration i => (i :: il, fl)
-          | FormulaDeclaration f => (il, f :: fl)
-    in
-      fn l => List.foldl part ([],[]) (rev l)
-    end;
-
-local
-  open Parse;
-
-  infixr 9 >>++
-  infixr 8 ++
-  infixr 7 >>
-  infixr 6 ||
-
-  fun declarationParser mapping =
-      (includeParser >> IncludeDeclaration) ||
-      (formulaParser mapping >> FormulaDeclaration);
-
-  fun parseChars parser chars =
-      let
-        val tokens = Parse.everything (lexer >> singleton) chars
-      in
-        Parse.everything (parser >> singleton) tokens
-      end;
-in
-  fun parseDeclaration mapping = parseChars (declarationParser mapping);
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* Clause information.                                                       *)
-(* ------------------------------------------------------------------------- *)
-
-datatype clauseSource =
-    CnfClauseSource of formulaName * literal list
-  | FofClauseSource of Normalize.thm;
-
-type 'a clauseInfo = 'a LiteralSetMap.map;
-
-type clauseNames = formulaName clauseInfo;
-
-type clauseRoles = role clauseInfo;
-
-type clauseSources = clauseSource clauseInfo;
-
-val noClauseNames : clauseNames = LiteralSetMap.new ();
-
-val allClauseNames : clauseNames -> formulaNameSet =
-    let
-      fun add (_,n,s) = addFormulaNameSet s n
-    in
-      LiteralSetMap.foldl add emptyFormulaNameSet
-    end;
-
-val noClauseRoles : clauseRoles = LiteralSetMap.new ();
-
-val noClauseSources : clauseSources = LiteralSetMap.new ();
-
-(* ------------------------------------------------------------------------- *)
-(* Comments.                                                                 *)
-(* ------------------------------------------------------------------------- *)
-
-fun mkLineComment "" = "%"
-  | mkLineComment line = "% " ^ line;
-
-fun destLineComment cs =
-    case cs of
-      [] => ""
-    | #"%" :: #" " :: rest => implode rest
-    | #"%" :: rest => implode rest
-    | _ => raise Error "Tptp.destLineComment";
-
-val isLineComment = can destLineComment;
-
-(* ------------------------------------------------------------------------- *)
-(* TPTP problems.                                                            *)
-(* ------------------------------------------------------------------------- *)
-
-type comments = string list;
-
-type includes = string list;
-
-datatype problem =
-    Problem of
-      {comments : comments,
-       includes : includes,
-       formulas : formula list};
-
-fun hasCnfConjecture (Problem {formulas,...}) =
-    List.exists isCnfConjectureFormula formulas;
-
-fun hasFofConjecture (Problem {formulas,...}) =
-    List.exists isFofConjectureFormula formulas;
-
-fun hasConjecture (Problem {formulas,...}) =
-    List.exists isConjectureFormula formulas;
-
-fun freeVars (Problem {formulas,...}) = freeVarsListFormula formulas;
-
-local
-  fun bump n avoid =
-      let
-        val s = FormulaName (Int.toString n)
-      in
-        if memberFormulaNameSet s avoid then bump (n + 1) avoid
-        else (s, n, addFormulaNameSet avoid s)
-      end;
-
-  fun fromClause defaultRole names roles cl (n,avoid) =
-      let
-        val (name,n,avoid) =
-            case LiteralSetMap.peek names cl of
-              SOME name => (name,n,avoid)
-            | NONE => bump n avoid
-
-        val role = Option.getOpt (LiteralSetMap.peek roles cl, defaultRole)
-
-        val body = CnfFormulaBody (clauseFromLiteralSet cl)
-
-        val source = NoFormulaSource
-
-        val formula =
-            Formula
-              {name = name,
-               role = role,
-               body = body,
-               source = source}
-      in
-        (formula,(n,avoid))
-      end;
-in
-  fun mkProblem {comments,includes,names,roles,problem} =
-      let
-        fun fromCl defaultRole = fromClause defaultRole names roles
-
-        val {axioms,conjecture} = problem
-
-        val n_avoid = (0, allClauseNames names)
-
-        val (axiomFormulas,n_avoid) = maps (fromCl AxiomRole) axioms n_avoid
-
-        val (conjectureFormulas,_) =
-            maps (fromCl NegatedConjectureRole) conjecture n_avoid
-
-        val formulas = axiomFormulas @ conjectureFormulas
-      in
-        Problem
-          {comments = comments,
-           includes = includes,
-           formulas = formulas}
-      end;
-end;
-
-type normalization =
-     {problem : Problem.problem,
-      sources : clauseSources};
-
-val initialNormalization : normalization =
-    {problem = {axioms = [], conjecture = []},
-     sources = noClauseSources};
-
-datatype problemGoal =
-    NoGoal
-  | CnfGoal of (formulaName * clause) list
-  | FofGoal of (formulaName * Formula.formula) list;
-
-local
-  fun partitionFormula (formula,(cnfAxioms,fofAxioms,cnfGoals,fofGoals)) =
-      let
-        val Formula {name,role,body,...} = formula
-      in
-        case body of
-          CnfFormulaBody cl =>
-          if isCnfConjectureRole role then
-            let
-              val cnfGoals = (name,cl) :: cnfGoals
-            in
-              (cnfAxioms,fofAxioms,cnfGoals,fofGoals)
-            end
-          else
-            let
-              val cnfAxioms = (name,cl) :: cnfAxioms
-            in
-              (cnfAxioms,fofAxioms,cnfGoals,fofGoals)
-            end
-        | FofFormulaBody fm =>
-          if isFofConjectureRole role then
-            let
-              val fofGoals = (name,fm) :: fofGoals
-            in
-              (cnfAxioms,fofAxioms,cnfGoals,fofGoals)
-            end
-          else
-            let
-              val fofAxioms = (name,fm) :: fofAxioms
-            in
-              (cnfAxioms,fofAxioms,cnfGoals,fofGoals)
-            end
-      end;
-
-  fun partitionFormulas fms =
-      let
-        val (cnfAxioms,fofAxioms,cnfGoals,fofGoals) =
-            List.foldl partitionFormula ([],[],[],[]) fms
-
-        val goal =
-            case (rev cnfGoals, rev fofGoals) of
-              ([],[]) => NoGoal
-            | (cnfGoals,[]) => CnfGoal cnfGoals
-            | ([],fofGoals) => FofGoal fofGoals
-            | (_ :: _, _ :: _) =>
-              raise Error "TPTP problem has both cnf and fof conjecture formulas"
-      in
-        {cnfAxioms = rev cnfAxioms,
-         fofAxioms = rev fofAxioms,
-         goal = goal}
-      end;
-
-  fun addClauses role clauses acc : normalization =
-      let
-        fun addClause (cl_src,sources) =
-            LiteralSetMap.insert sources cl_src
-
-        val {problem,sources} : normalization = acc
-        val {axioms,conjecture} = problem
-
-        val cls = map fst clauses
-        val (axioms,conjecture) =
-            if isCnfConjectureRole role then (axioms, cls @ conjecture)
-            else (cls @ axioms, conjecture)
-
-        val problem = {axioms = axioms, conjecture = conjecture}
-        and sources = List.foldl addClause sources clauses
-      in
-        {problem = problem,
-         sources = sources}
-      end;
-
-  fun addCnf role ((name,clause),(norm,cnf)) =
-      if List.exists (equalBooleanLiteral true) clause then (norm,cnf)
-      else
-        let
-          val cl = List.mapPartial (total destLiteral) clause
-          val cl = LiteralSet.fromList cl
-
-          val src = CnfClauseSource (name,clause)
-
-          val norm = addClauses role [(cl,src)] norm
-        in
-          (norm,cnf)
-        end;
-
-  val addCnfAxiom = addCnf AxiomRole;
-
-  val addCnfGoal = addCnf NegatedConjectureRole;
-
-  fun addFof role (th,(norm,cnf)) =
-      let
-        fun sourcify (cl,inf) = (cl, FofClauseSource inf)
-
-        val (clauses,cnf) = Normalize.addCnf th cnf
-        val clauses = map sourcify clauses
-        val norm = addClauses role clauses norm
-      in
-        (norm,cnf)
-      end;
-
-  fun addFofAxiom ((_,fm),acc) =
-      addFof AxiomRole (Normalize.mkAxiom fm, acc);
-
-  fun normProblem subgoal (norm,_) =
-      let
-        val {problem,sources} = norm
-        val {axioms,conjecture} = problem
-        val problem = {axioms = rev axioms, conjecture = rev conjecture}
-      in
-        {subgoal = subgoal,
-         problem = problem,
-         sources = sources}
-      end;
-
-  val normProblemFalse = normProblem (Formula.False,[]);
-
-  fun splitProblem acc =
-      let
-        fun mk parents subgoal =
-            let
-              val subgoal = Formula.generalize subgoal
-
-              val th = Normalize.mkAxiom (Formula.Not subgoal)
-
-              val acc = addFof NegatedConjectureRole (th,acc)
-            in
-              normProblem (subgoal,parents) acc
-            end
-
-        fun split (name,goal) =
-            let
-              val subgoals = Formula.splitGoal goal
-              val subgoals =
-                  if null subgoals then [Formula.True] else subgoals
-
-              val parents = [name]
-            in
-              map (mk parents) subgoals
-            end
-      in
-        fn goals => List.concat (map split goals)
-      end;
-
-  fun clausesToGoal cls =
-      let
-        val cls = map (Formula.generalize o clauseToFormula o snd) cls
-      in
-        Formula.listMkConj cls
-      end;
-
-  fun formulasToGoal fms =
-      let
-        val fms = map (Formula.generalize o snd) fms
-      in
-        Formula.listMkConj fms
-      end;
-in
-  fun goal (Problem {formulas,...}) =
-      let
-        val {cnfAxioms,fofAxioms,goal} = partitionFormulas formulas
-
-        val fm =
-            case goal of
-              NoGoal => Formula.False
-            | CnfGoal cls => Formula.Imp (clausesToGoal cls, Formula.False)
-            | FofGoal goals => formulasToGoal goals
-
-        val fm =
-            if null fofAxioms then fm
-            else Formula.Imp (formulasToGoal fofAxioms, fm)
-
-        val fm =
-            if null cnfAxioms then fm
-            else Formula.Imp (clausesToGoal cnfAxioms, fm)
-      in
-        fm
-      end;
-
-  fun normalize (Problem {formulas,...}) =
-      let
-        val {cnfAxioms,fofAxioms,goal} = partitionFormulas formulas
-
-        val acc = (initialNormalization, Normalize.initialCnf)
-        val acc = List.foldl addCnfAxiom acc cnfAxioms
-        val acc = List.foldl addFofAxiom acc fofAxioms
-      in
-        case goal of
-          NoGoal => [normProblemFalse acc]
-        | CnfGoal cls => [normProblemFalse (List.foldl addCnfGoal acc cls)]
-        | FofGoal goals => splitProblem acc goals
-      end;
-end;
-
-local
-  datatype blockComment =
-      OutsideBlockComment
-    | EnteringBlockComment
-    | InsideBlockComment
-    | LeavingBlockComment;
-
-  fun stripLineComments acc strm =
-      case strm of
-        Stream.Nil => (rev acc, Stream.Nil)
-      | Stream.Cons (line,rest) =>
-        case total destLineComment line of
-          SOME s => stripLineComments (s :: acc) (rest ())
-        | NONE => (rev acc, Stream.filter (not o isLineComment) strm);
-
-  fun advanceBlockComment c state =
-      case state of
-        OutsideBlockComment =>
-        if c = #"/" then (Stream.Nil, EnteringBlockComment)
-        else (Stream.singleton c, OutsideBlockComment)
-      | EnteringBlockComment =>
-        if c = #"*" then (Stream.Nil, InsideBlockComment)
-        else if c = #"/" then (Stream.singleton #"/", EnteringBlockComment)
-        else (Stream.fromList [#"/",c], OutsideBlockComment)
-      | InsideBlockComment =>
-        if c = #"*" then (Stream.Nil, LeavingBlockComment)
-        else (Stream.Nil, InsideBlockComment)
-      | LeavingBlockComment =>
-        if c = #"/" then (Stream.Nil, OutsideBlockComment)
-        else if c = #"*" then (Stream.Nil, LeavingBlockComment)
-        else (Stream.Nil, InsideBlockComment);
-
-  fun eofBlockComment state =
-      case state of
-        OutsideBlockComment => Stream.Nil
-      | EnteringBlockComment => Stream.singleton #"/"
-      | _ => raise Error "EOF inside a block comment";
-
-  val stripBlockComments =
-      Stream.mapsConcat advanceBlockComment eofBlockComment
-        OutsideBlockComment;
-in
-  fun read {mapping,filename} =
-      let
-        (* Estimating parse error line numbers *)
-
-        val lines = Stream.fromTextFile {filename = filename}
-
-        val {chars,parseErrorLocation} = Parse.initialize {lines = lines}
-      in
-        (let
-           (* The character stream *)
-
-           val (comments,chars) = stripLineComments [] chars
-
-           val chars = Parse.everything Parse.any chars
-
-           val chars = stripBlockComments chars
-
-           (* The declaration stream *)
-
-           val declarations = Stream.toList (parseDeclaration mapping chars)
-
-           val (includes,formulas) = partitionDeclarations declarations
-         in
-           Problem
-             {comments = comments,
-              includes = includes,
-              formulas = formulas}
-         end
-         handle Parse.NoParse => raise Error "parse error")
-        handle Error err =>
-          raise Error ("error in TPTP file \"" ^ filename ^ "\" " ^
-                       parseErrorLocation () ^ "\n" ^ err)
-      end;
-end;
-
-local
-  val newline = Stream.singleton "\n";
-
-  fun spacer top = if top then Stream.Nil else newline;
-
-  fun mkComment comment = mkLineComment comment ^ "\n";
-
-  fun mkInclude inc = includeToString inc ^ "\n";
-
-  fun formulaStream _ _ [] = Stream.Nil
-    | formulaStream mapping top (h :: t) =
-      Stream.append
-        (Stream.concatList
-           [spacer top,
-            Stream.singleton (formulaToString mapping h),
-            newline])
-        (fn () => formulaStream mapping false t);
-in
-  fun write {problem,mapping,filename} =
-      let
-        val Problem {comments,includes,formulas} = problem
-
-        val includesTop = null comments
-        val formulasTop = includesTop andalso null includes
-      in
-        Stream.toTextFile
-          {filename = filename}
-          (Stream.concatList
-             [Stream.map mkComment (Stream.fromList comments),
-              spacer includesTop,
-              Stream.map mkInclude (Stream.fromList includes),
-              formulaStream mapping formulasTop formulas])
-      end;
-end;
-
-local
-  fun refute {axioms,conjecture} =
-      let
-        val axioms = map Thm.axiom axioms
-        and conjecture = map Thm.axiom conjecture
-        val problem = {axioms = axioms, conjecture = conjecture}
-        val resolution = Resolution.new Resolution.default problem
-      in
-        case Resolution.loop resolution of
-          Resolution.Contradiction _ => true
-        | Resolution.Satisfiable _ => false
-      end;
-in
-  fun prove filename =
-      let
-        val problem = read filename
-        val problems = map #problem (normalize problem)
-      in
-        List.all refute problems
-      end;
-end;
-
-(* ------------------------------------------------------------------------- *)
-(* TSTP proofs.                                                              *)
-(* ------------------------------------------------------------------------- *)
-
-local
-  fun newName avoid prefix =
-      let
-        fun bump i =
-            let
-              val name = FormulaName (prefix ^ Int.toString i)
-              val i = i + 1
-            in
-              if memberFormulaNameSet name avoid then bump i else (name,i)
-            end
-      in
-        bump
-      end;
-
-  fun lookupClauseSource sources cl =
-      case LiteralSetMap.peek sources cl of
-        SOME src => src
-      | NONE => raise Bug "Tptp.lookupClauseSource";
-
-  fun lookupFormulaName fmNames fm =
-      case FormulaMap.peek fmNames fm of
-        SOME name => name
-      | NONE => raise Bug "Tptp.lookupFormulaName";
-
-  fun lookupClauseName clNames cl =
-      case LiteralSetMap.peek clNames cl of
-        SOME name => name
-      | NONE => raise Bug "Tptp.lookupClauseName";
-
-  fun lookupClauseSourceName sources fmNames cl =
-      case lookupClauseSource sources cl of
-        CnfClauseSource (name,_) => name
-      | FofClauseSource th =>
-        let
-          val (fm,_) = Normalize.destThm th
-        in
-          lookupFormulaName fmNames fm
-        end;
-
-  fun collectProofDeps sources ((_,inf),names_ths) =
-      case inf of
-        Proof.Axiom cl =>
-        let
-          val (names,ths) = names_ths
-        in
-          case lookupClauseSource sources cl of
-            CnfClauseSource (name,_) =>
-            let
-              val names = addFormulaNameSet names name
-            in
-              (names,ths)
-            end
-          | FofClauseSource th =>
-            let
-              val ths = th :: ths
-            in
-              (names,ths)
-            end
-        end
-      | _ => names_ths;
-
-  fun collectNormalizeDeps ((_,inf,_),fofs_defs) =
-      case inf of
-        Normalize.Axiom fm =>
-        let
-          val (fofs,defs) = fofs_defs
-          val fofs = FormulaSet.add fofs fm
-        in
-          (fofs,defs)
-        end
-      | Normalize.Definition n_d =>
-        let
-          val (fofs,defs) = fofs_defs
-          val defs = StringMap.insert defs n_d
-        in
-          (fofs,defs)
-        end
-      | _ => fofs_defs;
-
-  fun collectSubgoalProofDeps subgoalProof (names,fofs,defs) =
-      let
-        val {subgoal,sources,refutation} = subgoalProof
-
-        val names = addListFormulaNameSet names (snd subgoal)
-
-        val proof = Proof.proof refutation
-
-        val (names,ths) =
-            List.foldl (collectProofDeps sources) (names,[]) proof
-
-        val normalization = Normalize.proveThms (rev ths)
-
-        val (fofs,defs) =
-            List.foldl collectNormalizeDeps (fofs,defs) normalization
-
-        val subgoalProof =
-            {subgoal = subgoal,
-             normalization = normalization,
-             sources = sources,
-             proof = proof}
-      in
-        (subgoalProof,(names,fofs,defs))
-      end;
-
-  fun addProblemFormula names fofs (formula,(avoid,formulas,fmNames)) =
-      let
-        val name = nameFormula formula
-
-        val avoid = addFormulaNameSet avoid name
-
-        val (formulas,fmNames) =
-            if memberFormulaNameSet name names then
-              (formula :: formulas, fmNames)
-            else
-              case bodyFormula formula of
-                CnfFormulaBody _ => (formulas,fmNames)
-              | FofFormulaBody fm =>
-                if not (FormulaSet.member fm fofs) then (formulas,fmNames)
-                else (formula :: formulas, FormulaMap.insert fmNames (fm,name))
-      in
-        (avoid,formulas,fmNames)
-      end;
-
-  fun addDefinitionFormula avoid (_,def,(formulas,i,fmNames)) =
-      let
-        val (name,i) = newName avoid "definition_" i
-
-        val role = DefinitionRole
-
-        val body = FofFormulaBody def
-
-        val source = NoFormulaSource
-
-        val formula =
-            Formula
-              {name = name,
-               role = role,
-               body = body,
-               source = source}
-
-        val formulas = formula :: formulas
-
-        val fmNames = FormulaMap.insert fmNames (def,name)
-      in
-        (formulas,i,fmNames)
-      end;
-
-  fun addSubgoalFormula avoid subgoalProof (formulas,i) =
-      let
-        val {subgoal,normalization,sources,proof} = subgoalProof
-
-        val (fm,pars) = subgoal
-
-        val (name,i) = newName avoid "subgoal_" i
-
-        val number = i - 1
-
-        val (subgoal,formulas) =
-            if null pars then (NONE,formulas)
-            else
-              let
-                val role = PlainRole
-
-                val body = FofFormulaBody fm
-
-                val source =
-                    StripFormulaSource
-                      {inference = "strip",
-                       parents = pars}
-
-                val formula =
-                    Formula
-                      {name = name,
-                       role = role,
-                       body = body,
-                       source = source}
-              in
-                (SOME (name,fm), formula :: formulas)
-              end
-
-        val subgoalProof =
-            {number = number,
-             subgoal = subgoal,
-             normalization = normalization,
-             sources = sources,
-             proof = proof}
-      in
-        (subgoalProof,(formulas,i))
-      end;
-
-  fun mkNormalizeFormulaSource fmNames inference fms =
-      let
-        val fms =
-            case inference of
-              Normalize.Axiom fm => fm :: fms
-            | Normalize.Definition (_,fm) => fm :: fms
-            | _ => fms
-
-        val parents = map (lookupFormulaName fmNames) fms
-      in
-        NormalizeFormulaSource
-          {inference = inference,
-           parents = parents}
-      end;
-
-  fun mkProofFormulaSource sources fmNames clNames inference =
-      let
-        val parents =
-            case inference of
-              Proof.Axiom cl => [lookupClauseSourceName sources fmNames cl]
-            | _ =>
-              let
-                val cls = map Thm.clause (Proof.parents inference)
-              in
-                map (lookupClauseName clNames) cls
-              end
-      in
-        ProofFormulaSource
-          {inference = inference,
-           parents = parents}
-      end;
-
-  fun addNormalizeFormula avoid prefix ((fm,inf,fms),acc) =
-      let
-        val (formulas,i,fmNames) = acc
-
-        val (name,i) = newName avoid prefix i
-
-        val role = PlainRole
-
-        val body = FofFormulaBody fm
-
-        val source = mkNormalizeFormulaSource fmNames inf fms
-
-        val formula =
-            Formula
-              {name = name,
-               role = role,
-               body = body,
-               source = source}
-
-        val formulas = formula :: formulas
-
-        val fmNames = FormulaMap.insert fmNames (fm,name)
-      in
-        (formulas,i,fmNames)
-      end;
-
-  fun isSameClause sources formulas inf =
-      case inf of
-        Proof.Axiom cl =>
-          (case lookupClauseSource sources cl of
-             CnfClauseSource (name,lits) =>
-             if List.exists isBooleanLiteral lits then NONE
-             else SOME name
-           | _ => NONE)
-      | _ => NONE;
-
-  fun addProofFormula avoid sources fmNames prefix ((th,inf),acc) =
-      let
-        val (formulas,i,clNames) = acc
-
-        val cl = Thm.clause th
-      in
-        case isSameClause sources formulas inf of
-          SOME name =>
-          let
-            val clNames = LiteralSetMap.insert clNames (cl,name)
-          in
-            (formulas,i,clNames)
-          end
-        | NONE =>
-          let
-            val (name,i) = newName avoid prefix i
-
-            val role = PlainRole
-
-            val body = CnfFormulaBody (clauseFromLiteralSet cl)
-
-            val source = mkProofFormulaSource sources fmNames clNames inf
-
-            val formula =
-                Formula
-                  {name = name,
-                   role = role,
-                   body = body,
-                   source = source}
-
-            val formulas = formula :: formulas
-
-            val clNames = LiteralSetMap.insert clNames (cl,name)
-          in
-            (formulas,i,clNames)
-          end
-      end;
-
-  fun addSubgoalProofFormulas avoid fmNames (subgoalProof,formulas) =
-      let
-        val {number,subgoal,normalization,sources,proof} = subgoalProof
-
-        val (formulas,fmNames) =
-            case subgoal of
-              NONE => (formulas,fmNames)
-            | SOME (name,fm) =>
-              let
-                val source =
-                    StripFormulaSource
-                      {inference = "negate",
-                       parents = [name]}
-
-                val prefix = "negate_" ^ Int.toString number ^ "_"
-
-                val (name,_) = newName avoid prefix 0
-
-                val role = PlainRole
-
-                val fm = Formula.Not fm
-
-                val body = FofFormulaBody fm
-
-                val formula =
-                    Formula
-                      {name = name,
-                       role = role,
-                       body = body,
-                       source = source}
-
-                val formulas = formula :: formulas
-
-                val fmNames = FormulaMap.insert fmNames (fm,name)
-              in
-                (formulas,fmNames)
-              end
-
-        val prefix = "normalize_" ^ Int.toString number ^ "_"
-        val (formulas,_,fmNames) =
-            List.foldl (addNormalizeFormula avoid prefix)
-              (formulas,0,fmNames) normalization
-
-        val prefix = "refute_" ^ Int.toString number ^ "_"
-        val clNames : formulaName LiteralSetMap.map = LiteralSetMap.new ()
-        val (formulas,_,_) =
-            List.foldl (addProofFormula avoid sources fmNames prefix)
-              (formulas,0,clNames) proof
-      in
-        formulas
-      end;
-in
-  fun fromProof {problem,proofs} =
-      let
-        val names = emptyFormulaNameSet
-        and fofs = FormulaSet.empty
-        and defs : Formula.formula StringMap.map = StringMap.new ()
-
-        val (proofs,(names,fofs,defs)) =
-            maps collectSubgoalProofDeps proofs (names,fofs,defs)
-
-        val Problem {formulas,...} = problem
-
-        val fmNames : formulaName FormulaMap.map = FormulaMap.new ()
-        val (avoid,formulas,fmNames) =
-            List.foldl (addProblemFormula names fofs)
-              (emptyFormulaNameSet,[],fmNames) formulas
-
-        val (formulas,_,fmNames) =
-            StringMap.foldl (addDefinitionFormula avoid)
-              (formulas,0,fmNames) defs
-
-        val (proofs,(formulas,_)) =
-            maps (addSubgoalFormula avoid) proofs (formulas,0)
-
-        val formulas =
-            List.foldl (addSubgoalProofFormulas avoid fmNames) formulas proofs
-      in
-        rev formulas
-      end
-(*MetisDebug
-      handle Error err => raise Bug ("Tptp.fromProof: shouldn't fail:\n" ^ err);
-*)
-end;
-
-end
-end;
+;
 print_depth 10;