merged
authorblanchet
Fri, 17 Sep 2010 16:38:11 +0200
changeset 39505 4301d70795d5
parent 39490 c3d0414ba6df (current diff)
parent 39504 99d6cad53c7e (diff)
child 39506 cf61ec140c4d
child 39547 5df45da44bfb
merged
src/HOL/IsaMakefile
src/HOL/Tools/Sledgehammer/metis_clauses.ML
--- a/src/HOL/IsaMakefile	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/IsaMakefile	Fri Sep 17 16:38:11 2010 +0200
@@ -317,7 +317,8 @@
   Tools/semiring_normalizer.ML \
   Tools/Sledgehammer/clausifier.ML \
   Tools/Sledgehammer/meson_tactic.ML \
-  Tools/Sledgehammer/metis_clauses.ML \
+  Tools/Sledgehammer/metis_reconstruct.ML \
+  Tools/Sledgehammer/metis_translate.ML \
   Tools/Sledgehammer/metis_tactics.ML \
   Tools/Sledgehammer/sledgehammer.ML \
   Tools/Sledgehammer/sledgehammer_filter.ML \
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -434,7 +434,7 @@
 
 fun run_minimize args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) =
   let
-    open Metis_Clauses
+    open Metis_Translate
     val thy = Proof.theory_of st
     val n0 = length (these (!named_thms))
     val (prover_name, _) = get_atp thy args
--- a/src/HOL/Sledgehammer.thy	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/Sledgehammer.thy	Fri Sep 17 16:38:11 2010 +0200
@@ -16,7 +16,8 @@
   ("~~/src/Tools/Metis/metis.ML")
   ("Tools/Sledgehammer/clausifier.ML")
   ("Tools/Sledgehammer/meson_tactic.ML")
-  ("Tools/Sledgehammer/metis_clauses.ML")
+  ("Tools/Sledgehammer/metis_translate.ML")
+  ("Tools/Sledgehammer/metis_reconstruct.ML")
   ("Tools/Sledgehammer/metis_tactics.ML")
   ("Tools/Sledgehammer/sledgehammer_util.ML")
   ("Tools/Sledgehammer/sledgehammer_filter.ML")
@@ -102,7 +103,8 @@
 use "Tools/Sledgehammer/meson_tactic.ML"
 setup Meson_Tactic.setup
 
-use "Tools/Sledgehammer/metis_clauses.ML"
+use "Tools/Sledgehammer/metis_translate.ML"
+use "Tools/Sledgehammer/metis_reconstruct.ML"
 use "Tools/Sledgehammer/metis_tactics.ML"
 setup Metis_Tactics.setup
 
--- a/src/HOL/Tools/ATP/atp_proof.ML	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_proof.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -11,6 +11,12 @@
   type 'a fo_term = 'a ATP_Problem.fo_term
   type 'a uniform_formula = 'a ATP_Problem.uniform_formula
 
+  datatype failure =
+    Unprovable | IncompleteUnprovable | CantConnect | TimedOut |
+    OutOfResources | SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl |
+    MalformedInput | MalformedOutput | Interrupted | InternalError |
+    UnknownError
+
   type step_name = string * string option
 
   datatype 'a step =
@@ -20,6 +26,13 @@
   type 'a proof = 'a uniform_formula step list
 
   val strip_spaces : (char -> bool) -> string -> string
+  val string_for_failure : failure -> string
+  val extract_important_message : string -> string
+  val extract_known_failure :
+    (failure * string) list -> string -> failure option
+  val extract_tstplike_proof_and_outcome :
+    bool -> int -> (string * string) list -> (failure * string) list -> string
+    -> string * failure option
   val is_same_step : step_name * step_name -> bool
   val atp_proof_from_tstplike_string : string -> string proof
   val map_term_names_in_atp_proof :
@@ -30,6 +43,13 @@
 structure ATP_Proof : ATP_PROOF =
 struct
 
+open ATP_Problem
+
+datatype failure =
+  Unprovable | IncompleteUnprovable | CantConnect | TimedOut | OutOfResources |
+  SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl | MalformedInput |
+  MalformedOutput | Interrupted | InternalError | UnknownError
+
 fun strip_spaces_in_list _ [] = []
   | strip_spaces_in_list _ [c1] = if Char.isSpace c1 then [] else [str c1]
   | strip_spaces_in_list is_evil [c1, c2] =
@@ -51,7 +71,85 @@
 fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
 val strip_spaces_except_between_ident_chars = strip_spaces is_ident_char
 
-open ATP_Problem
+val missing_message_tail =
+  " appears to be missing. You will need to install it if you want to run \
+  \ATPs remotely."
+
+fun string_for_failure Unprovable = "The ATP problem is unprovable."
+  | string_for_failure IncompleteUnprovable =
+    "The ATP cannot prove the problem."
+  | string_for_failure CantConnect = "Can't connect to remote server."
+  | string_for_failure TimedOut = "Timed out."
+  | string_for_failure OutOfResources = "The ATP ran out of resources."
+  | string_for_failure SpassTooOld =
+    "Isabelle requires a more recent version of SPASS with support for the \
+    \TPTP syntax. To install it, download and extract the package \
+    \\"http://isabelle.in.tum.de/dist/contrib/spass-3.7.tar.gz\" and add the \
+    \\"spass-3.7\" directory's absolute path to " ^
+    quote (Path.implode (Path.expand (Path.appends
+               (Path.variable "ISABELLE_HOME_USER" ::
+                map Path.basic ["etc", "components"])))) ^
+    " on a line of its own."
+  | string_for_failure VampireTooOld =
+    "Isabelle requires a more recent version of Vampire. To install it, follow \
+    \the instructions from the Sledgehammer manual (\"isabelle doc\
+    \ sledgehammer\")."
+  | string_for_failure NoPerl = "Perl" ^ missing_message_tail
+  | string_for_failure NoLibwwwPerl =
+    "The Perl module \"libwww-perl\"" ^ missing_message_tail
+  | string_for_failure MalformedInput =
+    "The ATP problem is malformed. Please report this to the Isabelle \
+    \developers."
+  | string_for_failure MalformedOutput = "The ATP output is malformed."
+  | string_for_failure Interrupted = "The ATP was interrupted."
+  | string_for_failure InternalError = "An internal ATP error occurred."
+  | string_for_failure UnknownError = "An unknown ATP error occurred."
+
+fun extract_delimited (begin_delim, end_delim) output =
+  output |> first_field begin_delim |> the |> snd
+         |> first_field end_delim |> the |> fst
+         |> first_field "\n" |> the |> snd
+  handle Option.Option => ""
+
+val tstp_important_message_delims =
+  ("% SZS start RequiredInformation", "% SZS end RequiredInformation")
+
+fun extract_important_message output =
+  case extract_delimited tstp_important_message_delims output of
+    "" => ""
+  | s => s |> space_explode "\n" |> filter_out (curry (op =) "")
+           |> map (perhaps (try (unprefix "%")))
+           |> map (perhaps (try (unprefix " ")))
+           |> space_implode "\n " |> quote
+
+(* Splits by the first possible of a list of delimiters. *)
+fun extract_tstplike_proof delims output =
+  case pairself (find_first (fn s => String.isSubstring s output))
+                (ListPair.unzip delims) of
+    (SOME begin_delim, SOME end_delim) =>
+    extract_delimited (begin_delim, end_delim) output
+  | _ => ""
+
+fun extract_known_failure known_failures output =
+  known_failures
+  |> find_first (fn (_, pattern) => String.isSubstring pattern output)
+  |> Option.map fst
+
+fun extract_tstplike_proof_and_outcome complete res_code proof_delims
+                                       known_failures output =
+  case extract_known_failure known_failures output of
+    NONE => (case extract_tstplike_proof proof_delims output of
+             "" => ("", SOME (if res_code = 0 andalso output = "" then
+                                Interrupted
+                              else
+                                UnknownError))
+           | tstplike_proof => if res_code = 0 then (tstplike_proof, NONE)
+                               else ("", SOME UnknownError))
+  | SOME failure =>
+    ("", SOME (if failure = IncompleteUnprovable andalso complete then
+                 Unprovable
+               else
+                 failure))
 
 fun mk_anot (AConn (ANot, [phi])) = phi
   | mk_anot phi = AConn (ANot, [phi])
--- a/src/HOL/Tools/ATP/atp_systems.ML	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_systems.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -7,11 +7,7 @@
 
 signature ATP_SYSTEMS =
 sig
-  datatype failure =
-    Unprovable | IncompleteUnprovable | CantConnect | TimedOut |
-    OutOfResources | SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl |
-    MalformedInput | MalformedOutput | Interrupted | InternalError |
-    UnknownError
+  type failure = ATP_Proof.failure
 
   type prover_config =
     {exec: string * string,
@@ -24,9 +20,6 @@
      explicit_forall: bool,
      use_conjecture_for_hypotheses: bool}
 
-  val string_for_failure : failure -> string
-  val known_failure_in_output :
-    string -> (failure * string) list -> failure option
   val add_prover: string * prover_config -> theory -> theory
   val get_prover: theory -> string -> prover_config
   val available_atps: theory -> unit
@@ -38,12 +31,9 @@
 structure ATP_Systems : ATP_SYSTEMS =
 struct
 
-(* prover configuration *)
+open ATP_Proof
 
-datatype failure =
-  Unprovable | IncompleteUnprovable | CantConnect | TimedOut | OutOfResources |
-  SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl | MalformedInput |
-  MalformedOutput | Interrupted | InternalError | UnknownError
+(* prover configuration *)
 
 type prover_config =
   {exec: string * string,
@@ -56,44 +46,6 @@
    explicit_forall: bool,
    use_conjecture_for_hypotheses: bool}
 
-val missing_message_tail =
-  " appears to be missing. You will need to install it if you want to run \
-  \ATPs remotely."
-
-fun string_for_failure Unprovable = "The ATP problem is unprovable."
-  | string_for_failure IncompleteUnprovable =
-    "The ATP cannot prove the problem."
-  | string_for_failure CantConnect = "Can't connect to remote server."
-  | string_for_failure TimedOut = "Timed out."
-  | string_for_failure OutOfResources = "The ATP ran out of resources."
-  | string_for_failure SpassTooOld =
-    "Isabelle requires a more recent version of SPASS with support for the \
-    \TPTP syntax. To install it, download and extract the package \
-    \\"http://isabelle.in.tum.de/dist/contrib/spass-3.7.tar.gz\" and add the \
-    \\"spass-3.7\" directory's absolute path to " ^
-    quote (Path.implode (Path.expand (Path.appends
-               (Path.variable "ISABELLE_HOME_USER" ::
-                map Path.basic ["etc", "components"])))) ^
-    " on a line of its own."
-  | string_for_failure VampireTooOld =
-    "Isabelle requires a more recent version of Vampire. To install it, follow \
-    \the instructions from the Sledgehammer manual (\"isabelle doc\
-    \ sledgehammer\")."
-  | string_for_failure NoPerl = "Perl" ^ missing_message_tail
-  | string_for_failure NoLibwwwPerl =
-    "The Perl module \"libwww-perl\"" ^ missing_message_tail
-  | string_for_failure MalformedInput =
-    "The ATP problem is malformed. Please report this to the Isabelle \
-    \developers."
-  | string_for_failure MalformedOutput = "The ATP output is malformed."
-  | string_for_failure Interrupted = "The ATP was interrupted."
-  | string_for_failure InternalError = "An internal ATP error occurred."
-  | string_for_failure UnknownError = "An unknown ATP error occurred."
-
-fun known_failure_in_output output =
-  find_first (fn (_, pattern) => String.isSubstring pattern output)
-  #> Option.map fst
-
 val known_perl_failures =
   [(CantConnect, "HTTP error"),
    (NoPerl, "env: perl"),
@@ -124,6 +76,7 @@
 
 fun to_secs bonus time = (Time.toMilliseconds time + bonus + 999) div 1000
 
+
 (* E prover *)
 
 (* Give older versions of E an extra second, because the "eproof" script wrongly
@@ -163,6 +116,8 @@
 val e = ("e", e_config)
 
 
+(* SPASS *)
+
 (* The "-VarWeight=3" option helps the higher-order problems, probably by
    counteracting the presence of "hAPP". *)
 val spass_config : prover_config =
@@ -226,11 +181,11 @@
 
 fun get_systems () =
   case bash_output "\"$ISABELLE_ATP/scripts/remote_atp\" -w 2>&1" of
-    (answer, 0) => split_lines answer
-  | (answer, _) =>
-    error (case known_failure_in_output answer known_perl_failures of
+    (output, 0) => split_lines output
+  | (output, _) =>
+    error (case extract_known_failure known_perl_failures output of
              SOME failure => string_for_failure failure
-           | NONE => perhaps (try (unsuffix "\n")) answer ^ ".")
+           | NONE => perhaps (try (unsuffix "\n")) output ^ ".")
 
 fun refresh_systems_on_tptp () =
   Synchronized.change systems (fn _ => get_systems ())
--- a/src/HOL/Tools/Sledgehammer/metis_clauses.ML	Fri Sep 17 16:15:45 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,524 +0,0 @@
-(*  Title:      HOL/Tools/Sledgehammer/metis_clauses.ML
-    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
-    Author:     Jasmin Blanchette, TU Muenchen
-
-Storing/printing FOL clauses and arity clauses.  Typed equality is
-treated differently.
-*)
-
-signature METIS_CLAUSES =
-sig
-  type name = string * string
-  datatype type_literal =
-    TyLitVar of name * name |
-    TyLitFree of name * name
-  datatype arLit =
-    TConsLit of name * name * name list |
-    TVarLit of name * name
-  datatype arity_clause =
-    ArityClause of {name: string, conclLit: arLit, premLits: arLit list}
-  datatype class_rel_clause =
-    ClassRelClause of {name: string, subclass: name, superclass: name}
-  datatype combtyp =
-    CombTVar of name |
-    CombTFree of name |
-    CombType of name * combtyp list
-  datatype combterm =
-    CombConst of name * combtyp * combtyp list (* Const and Free *) |
-    CombVar of name * combtyp |
-    CombApp of combterm * combterm
-  datatype fol_literal = FOLLiteral of bool * combterm
-
-  val type_wrapper_name : string
-  val bound_var_prefix : string
-  val schematic_var_prefix: string
-  val fixed_var_prefix: string
-  val tvar_prefix: string
-  val tfree_prefix: string
-  val const_prefix: string
-  val type_const_prefix: string
-  val class_prefix: string
-  val invert_const: string -> string
-  val ascii_of: string -> string
-  val unascii_of: string -> string
-  val strip_prefix_and_unascii: string -> string -> string option
-  val make_bound_var : string -> string
-  val make_schematic_var : string * int -> string
-  val make_fixed_var : string -> string
-  val make_schematic_type_var : string * int -> string
-  val make_fixed_type_var : string -> string
-  val make_fixed_const : string -> string
-  val make_fixed_type_const : string -> string
-  val make_type_class : string -> string
-  val skolem_theory_name: string
-  val skolem_prefix: string
-  val skolem_infix: string
-  val is_skolem_const_name: string -> bool
-  val num_type_args: theory -> string -> int
-  val type_literals_for_types : typ list -> type_literal list
-  val make_class_rel_clauses: theory -> class list -> class list -> class_rel_clause list
-  val make_arity_clauses: theory -> string list -> class list -> class list * arity_clause list
-  val combtyp_of : combterm -> combtyp
-  val strip_combterm_comb : combterm -> combterm * combterm list
-  val combterm_from_term :
-    theory -> (string * typ) list -> term -> combterm * typ list
-  val literals_of_term : theory -> term -> fol_literal list * typ list
-  val conceal_skolem_terms :
-    int -> (string * term) list -> term -> (string * term) list * term
-  val reveal_skolem_terms : (string * term) list -> term -> term
-  val tfree_classes_of_terms : term list -> string list
-  val tvar_classes_of_terms : term list -> string list
-  val type_consts_of_terms : theory -> term list -> string list
-end
-
-structure Metis_Clauses : METIS_CLAUSES =
-struct
-
-val type_wrapper_name = "ti"
-
-val bound_var_prefix = "B_"
-val schematic_var_prefix = "V_"
-val fixed_var_prefix = "v_"
-
-val tvar_prefix = "T_";
-val tfree_prefix = "t_";
-
-val const_prefix = "c_";
-val type_const_prefix = "tc_";
-val class_prefix = "class_";
-
-fun union_all xss = fold (union (op =)) xss []
-
-(* Readable names for the more common symbolic functions. Do not mess with the
-   last nine entries of the table unless you know what you are doing. *)
-val const_trans_table =
-  Symtab.make [(@{type_name Product_Type.prod}, "prod"),
-               (@{type_name Sum_Type.sum}, "sum"),
-               (@{const_name HOL.eq}, "equal"),
-               (@{const_name HOL.conj}, "and"),
-               (@{const_name HOL.disj}, "or"),
-               (@{const_name HOL.implies}, "implies"),
-               (@{const_name Set.member}, "member"),
-               (@{const_name fequal}, "fequal"),
-               (@{const_name COMBI}, "COMBI"),
-               (@{const_name COMBK}, "COMBK"),
-               (@{const_name COMBB}, "COMBB"),
-               (@{const_name COMBC}, "COMBC"),
-               (@{const_name COMBS}, "COMBS"),
-               (@{const_name True}, "True"),
-               (@{const_name False}, "False"),
-               (@{const_name If}, "If")]
-
-(* Invert the table of translations between Isabelle and ATPs. *)
-val const_trans_table_inv =
-  Symtab.update ("fequal", @{const_name HOL.eq})
-                (Symtab.make (map swap (Symtab.dest const_trans_table)))
-
-val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
-
-(*Escaping of special characters.
-  Alphanumeric characters are left unchanged.
-  The character _ goes to __
-  Characters in the range ASCII space to / go to _A to _P, respectively.
-  Other characters go to _nnn where nnn is the decimal ASCII code.*)
-val A_minus_space = Char.ord #"A" - Char.ord #" ";
-
-fun stringN_of_int 0 _ = ""
-  | stringN_of_int k n = stringN_of_int (k-1) (n div 10) ^ Int.toString (n mod 10);
-
-fun ascii_of_c c =
-  if Char.isAlphaNum c then String.str c
-  else if c = #"_" then "__"
-  else if #" " <= c andalso c <= #"/"
-       then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
-  else ("_" ^ stringN_of_int 3 (Char.ord c))  (*fixed width, in case more digits follow*)
-
-val ascii_of = String.translate ascii_of_c;
-
-(** Remove ASCII armouring from names in proof files **)
-
-(*We don't raise error exceptions because this code can run inside the watcher.
-  Also, the errors are "impossible" (hah!)*)
-fun unascii_aux rcs [] = String.implode(rev rcs)
-  | unascii_aux rcs [#"_"] = unascii_aux (#"_"::rcs) []  (*ERROR*)
-      (*Three types of _ escapes: __, _A to _P, _nnn*)
-  | unascii_aux rcs (#"_" :: #"_" :: cs) = unascii_aux (#"_"::rcs) cs
-  | unascii_aux rcs (#"_" :: c :: cs) =
-      if #"A" <= c andalso c<= #"P"  (*translation of #" " to #"/"*)
-      then unascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
-      else
-        let val digits = List.take (c::cs, 3) handle Subscript => []
-        in
-            case Int.fromString (String.implode digits) of
-                NONE => unascii_aux (c:: #"_"::rcs) cs  (*ERROR*)
-              | SOME n => unascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
-        end
-  | unascii_aux rcs (c::cs) = unascii_aux (c::rcs) cs
-val unascii_of = unascii_aux [] o String.explode
-
-(* If string s has the prefix s1, return the result of deleting it,
-   un-ASCII'd. *)
-fun strip_prefix_and_unascii s1 s =
-  if String.isPrefix s1 s then
-    SOME (unascii_of (String.extract (s, size s1, NONE)))
-  else
-    NONE
-
-(*Remove the initial ' character from a type variable, if it is present*)
-fun trim_type_var s =
-  if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
-  else error ("trim_type: Malformed type variable encountered: " ^ s);
-
-fun ascii_of_indexname (v,0) = ascii_of v
-  | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
-
-fun make_bound_var x = bound_var_prefix ^ ascii_of x
-fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v
-fun make_fixed_var x = fixed_var_prefix ^ ascii_of x
-
-fun make_schematic_type_var (x,i) =
-      tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
-fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
-
-fun lookup_const c =
-  case Symtab.lookup const_trans_table c of
-    SOME c' => c'
-  | NONE => ascii_of c
-
-(* HOL.eq MUST BE "equal" because it's built into ATPs. *)
-fun make_fixed_const @{const_name HOL.eq} = "equal"
-  | make_fixed_const c = const_prefix ^ lookup_const c
-
-fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
-
-fun make_type_class clas = class_prefix ^ ascii_of clas;
-
-val skolem_theory_name = "Sledgehammer" ^ Long_Name.separator ^ "Sko"
-val skolem_prefix = "sko_"
-val skolem_infix = "$"
-
-(* Hack: Could return false positives (e.g., a user happens to declare a
-   constant called "SomeTheory.sko_means_shoe_in_$wedish". *)
-val is_skolem_const_name =
-  Long_Name.base_name
-  #> String.isPrefix skolem_prefix andf String.isSubstring skolem_infix
-
-(* The number of type arguments of a constant, zero if it's monomorphic. For
-   (instances of) Skolem pseudoconstants, this information is encoded in the
-   constant name. *)
-fun num_type_args thy s =
-  if String.isPrefix skolem_theory_name s then
-    s |> unprefix skolem_theory_name
-      |> space_explode skolem_infix |> hd
-      |> space_explode "_" |> List.last |> Int.fromString |> the
-  else
-    (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
-
-(**** Definitions and functions for FOL clauses for TPTP format output ****)
-
-type name = string * string
-
-(**** Isabelle FOL clauses ****)
-
-(* The first component is the type class; the second is a TVar or TFree. *)
-datatype type_literal =
-  TyLitVar of name * name |
-  TyLitFree of name * name
-
-exception CLAUSE of string * term;
-
-(*Make literals for sorted type variables*)
-fun sorts_on_typs_aux (_, [])   = []
-  | sorts_on_typs_aux ((x,i),  s::ss) =
-      let val sorts = sorts_on_typs_aux ((x,i), ss)
-      in
-          if s = "HOL.type" then sorts
-          else if i = ~1 then TyLitFree (`make_type_class s, `make_fixed_type_var x) :: sorts
-          else TyLitVar (`make_type_class s, (make_schematic_type_var (x,i), x)) :: sorts
-      end;
-
-fun sorts_on_typs (TFree (a,s)) = sorts_on_typs_aux ((a,~1),s)
-  | sorts_on_typs (TVar (v,s))  = sorts_on_typs_aux (v,s);
-
-(*Given a list of sorted type variables, return a list of type literals.*)
-fun type_literals_for_types Ts =
-  fold (union (op =)) (map sorts_on_typs Ts) []
-
-(** make axiom and conjecture clauses. **)
-
-(**** Isabelle arities ****)
-
-datatype arLit =
-  TConsLit of name * name * name list |
-  TVarLit of name * name
-
-datatype arity_clause =
-  ArityClause of {name: string, conclLit: arLit, premLits: arLit list}
-
-
-fun gen_TVars 0 = []
-  | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
-
-fun pack_sort(_,[])  = []
-  | pack_sort(tvar, "HOL.type"::srt) = pack_sort (tvar, srt)   (*IGNORE sort "type"*)
-  | pack_sort(tvar, cls::srt) =
-    (`make_type_class cls, (tvar, tvar)) :: pack_sort (tvar, srt)
-
-(*Arity of type constructor tcon :: (arg1,...,argN)res*)
-fun make_axiom_arity_clause (tcons, name, (cls,args)) =
-  let
-    val tvars = gen_TVars (length args)
-    val tvars_srts = ListPair.zip (tvars, args)
-  in
-    ArityClause {name = name,
-                 conclLit = TConsLit (`make_type_class cls,
-                                      `make_fixed_type_const tcons,
-                                      tvars ~~ tvars),
-                 premLits = map TVarLit (union_all (map pack_sort tvars_srts))}
-  end
-
-
-(**** Isabelle class relations ****)
-
-datatype class_rel_clause =
-  ClassRelClause of {name: string, subclass: name, superclass: name}
-
-(*Generate all pairs (sub,super) such that sub is a proper subclass of super in theory thy.*)
-fun class_pairs _ [] _ = []
-  | class_pairs thy subs supers =
-      let
-        val class_less = Sorts.class_less (Sign.classes_of thy)
-        fun add_super sub super = class_less (sub, super) ? cons (sub, super)
-        fun add_supers sub = fold (add_super sub) supers
-      in fold add_supers subs [] end
-
-fun make_class_rel_clause (sub,super) =
-  ClassRelClause {name = sub ^ "_" ^ super,
-                  subclass = `make_type_class sub,
-                  superclass = `make_type_class super}
-
-fun make_class_rel_clauses thy subs supers =
-  map make_class_rel_clause (class_pairs thy subs supers);
-
-
-(** Isabelle arities **)
-
-fun arity_clause _ _ (_, []) = []
-  | arity_clause seen n (tcons, ("HOL.type",_)::ars) =  (*ignore*)
-      arity_clause seen n (tcons,ars)
-  | arity_clause seen n (tcons, (ar as (class,_)) :: ars) =
-      if member (op =) seen class then (*multiple arities for the same tycon, class pair*)
-          make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) ::
-          arity_clause seen (n+1) (tcons,ars)
-      else
-          make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class, ar) ::
-          arity_clause (class::seen) n (tcons,ars)
-
-fun multi_arity_clause [] = []
-  | multi_arity_clause ((tcons, ars) :: tc_arlists) =
-      arity_clause [] 1 (tcons, ars) @ multi_arity_clause tc_arlists
-
-(*Generate all pairs (tycon,class,sorts) such that tycon belongs to class in theory thy
-  provided its arguments have the corresponding sorts.*)
-fun type_class_pairs thy tycons classes =
-  let val alg = Sign.classes_of thy
-      fun domain_sorts tycon = Sorts.mg_domain alg tycon o single
-      fun add_class tycon class =
-        cons (class, domain_sorts tycon class)
-        handle Sorts.CLASS_ERROR _ => I
-      fun try_classes tycon = (tycon, fold (add_class tycon) classes [])
-  in  map try_classes tycons  end;
-
-(*Proving one (tycon, class) membership may require proving others, so iterate.*)
-fun iter_type_class_pairs _ _ [] = ([], [])
-  | iter_type_class_pairs thy tycons classes =
-      let val cpairs = type_class_pairs thy tycons classes
-          val newclasses = union_all (union_all (union_all (map (map #2 o #2) cpairs)))
-            |> subtract (op =) classes |> subtract (op =) HOLogic.typeS
-          val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
-      in (union (op =) classes' classes, union (op =) cpairs' cpairs) end;
-
-fun make_arity_clauses thy tycons classes =
-  let val (classes', cpairs) = iter_type_class_pairs thy tycons classes
-  in  (classes', multi_arity_clause cpairs)  end;
-
-datatype combtyp =
-  CombTVar of name |
-  CombTFree of name |
-  CombType of name * combtyp list
-
-datatype combterm =
-  CombConst of name * combtyp * combtyp list (* Const and Free *) |
-  CombVar of name * combtyp |
-  CombApp of combterm * combterm
-
-datatype fol_literal = FOLLiteral of bool * combterm
-
-(*********************************************************************)
-(* convert a clause with type Term.term to a clause with type clause *)
-(*********************************************************************)
-
-(*Result of a function type; no need to check that the argument type matches.*)
-fun result_type (CombType (_, [_, tp2])) = tp2
-  | result_type _ = raise Fail "non-function type"
-
-fun combtyp_of (CombConst (_, tp, _)) = tp
-  | combtyp_of (CombVar (_, tp)) = tp
-  | combtyp_of (CombApp (t1, _)) = result_type (combtyp_of t1)
-
-(*gets the head of a combinator application, along with the list of arguments*)
-fun strip_combterm_comb u =
-    let fun stripc (CombApp(t,u), ts) = stripc (t, u::ts)
-        |   stripc  x =  x
-    in stripc(u,[]) end
-
-fun type_of (Type (a, Ts)) =
-    let val (folTypes,ts) = types_of Ts in
-      (CombType (`make_fixed_type_const a, folTypes), ts)
-    end
-  | type_of (tp as TFree (a, _)) = (CombTFree (`make_fixed_type_var a), [tp])
-  | type_of (tp as TVar (x, _)) =
-    (CombTVar (make_schematic_type_var x, string_of_indexname x), [tp])
-and types_of Ts =
-    let val (folTyps, ts) = ListPair.unzip (map type_of Ts) in
-      (folTyps, union_all ts)
-    end
-
-(* same as above, but no gathering of sort information *)
-fun simp_type_of (Type (a, Ts)) =
-      CombType (`make_fixed_type_const a, map simp_type_of Ts)
-  | simp_type_of (TFree (a, _)) = CombTFree (`make_fixed_type_var a)
-  | simp_type_of (TVar (x, _)) =
-    CombTVar (make_schematic_type_var x, string_of_indexname x)
-
-(* Converts a term (with combinators) into a combterm. Also accummulates sort
-   infomation. *)
-fun combterm_from_term thy bs (P $ Q) =
-      let val (P', tsP) = combterm_from_term thy bs P
-          val (Q', tsQ) = combterm_from_term thy bs Q
-      in  (CombApp (P', Q'), union (op =) tsP tsQ)  end
-  | combterm_from_term thy _ (Const (c, T)) =
-      let
-        val (tp, ts) = type_of T
-        val tvar_list =
-          (if String.isPrefix skolem_theory_name c then
-             [] |> Term.add_tvarsT T |> map TVar
-           else
-             (c, T) |> Sign.const_typargs thy)
-          |> map simp_type_of
-        val c' = CombConst (`make_fixed_const c, tp, tvar_list)
-      in  (c',ts)  end
-  | combterm_from_term _ _ (Free (v, T)) =
-      let val (tp,ts) = type_of T
-          val v' = CombConst (`make_fixed_var v, tp, [])
-      in  (v',ts)  end
-  | combterm_from_term _ _ (Var (v, T)) =
-      let val (tp,ts) = type_of T
-          val v' = CombVar ((make_schematic_var v, string_of_indexname v), tp)
-      in  (v',ts)  end
-  | combterm_from_term _ bs (Bound j) =
-      let
-        val (s, T) = nth bs j
-        val (tp, ts) = type_of T
-        val v' = CombConst (`make_bound_var s, tp, [])
-      in (v', ts) end
-  | combterm_from_term _ _ (Abs _) = raise Fail "HOL clause: Abs"
-
-fun predicate_of thy ((@{const Not} $ P), pos) = predicate_of thy (P, not pos)
-  | predicate_of thy (t, pos) =
-    (combterm_from_term thy [] (Envir.eta_contract t), pos)
-
-fun literals_of_term1 args thy (@{const Trueprop} $ P) =
-    literals_of_term1 args thy P
-  | literals_of_term1 args thy (@{const HOL.disj} $ P $ Q) =
-    literals_of_term1 (literals_of_term1 args thy P) thy Q
-  | literals_of_term1 (lits, ts) thy P =
-    let val ((pred, ts'), pol) = predicate_of thy (P, true) in
-      (FOLLiteral (pol, pred) :: lits, union (op =) ts ts')
-    end
-val literals_of_term = literals_of_term1 ([], [])
-
-fun skolem_name i j num_T_args =
-  skolem_prefix ^ (space_implode "_" (map Int.toString [i, j, num_T_args])) ^
-  skolem_infix ^ "g"
-
-fun conceal_skolem_terms i skolems t =
-  if exists_Const (curry (op =) @{const_name skolem} o fst) t then
-    let
-      fun aux skolems
-              (t as (Const (@{const_name skolem}, Type (_, [_, T])) $ _)) =
-          let
-            val (skolems, s) =
-              if i = ~1 then
-                (skolems, @{const_name undefined})
-              else case AList.find (op aconv) skolems t of
-                s :: _ => (skolems, s)
-              | [] =>
-                let
-                  val s = skolem_theory_name ^ "." ^
-                          skolem_name i (length skolems)
-                                        (length (Term.add_tvarsT T []))
-                in ((s, t) :: skolems, s) end
-          in (skolems, Const (s, T)) end
-        | aux skolems (t1 $ t2) =
-          let
-            val (skolems, t1) = aux skolems t1
-            val (skolems, t2) = aux skolems t2
-          in (skolems, t1 $ t2) end
-        | aux skolems (Abs (s, T, t')) =
-          let val (skolems, t') = aux skolems t' in
-            (skolems, Abs (s, T, t'))
-          end
-        | aux skolems t = (skolems, t)
-    in aux skolems t end
-  else
-    (skolems, t)
-
-fun reveal_skolem_terms skolems =
-  map_aterms (fn t as Const (s, _) =>
-                 if String.isPrefix skolem_theory_name s then
-                   AList.lookup (op =) skolems s |> the
-                   |> map_types Type_Infer.paramify_vars
-                 else
-                   t
-               | t => t)
-
-
-(***************************************************************)
-(* Type Classes Present in the Axiom or Conjecture Clauses     *)
-(***************************************************************)
-
-fun set_insert (x, s) = Symtab.update (x, ()) s
-
-fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
-
-(*Remove this trivial type class*)
-fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset;
-
-fun tfree_classes_of_terms ts =
-  let val sorts_list = map (map #2 o OldTerm.term_tfrees) ts
-  in  Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list))  end;
-
-fun tvar_classes_of_terms ts =
-  let val sorts_list = map (map #2 o OldTerm.term_tvars) ts
-  in  Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list))  end;
-
-(*fold type constructors*)
-fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x))
-  | fold_type_consts _ _ x = x;
-
-(*Type constructors used to instantiate overloaded constants are the only ones needed.*)
-fun add_type_consts_in_term thy =
-  let
-    fun aux (Const x) =
-        fold (fold_type_consts set_insert) (Sign.const_typargs thy x)
-      | aux (Abs (_, _, u)) = aux u
-      | aux (Const (@{const_name skolem}, _) $ _) = I
-      | aux (t $ u) = aux t #> aux u
-      | aux _ = I
-  in aux end
-
-fun type_consts_of_terms thy ts =
-  Symtab.keys (fold (add_type_consts_in_term thy) ts Symtab.empty);
-
-end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Sledgehammer/metis_reconstruct.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -0,0 +1,514 @@
+(*  Title:      HOL/Tools/Sledgehammer/metis_reconstruct.ML
+    Author:     Kong W. Susanto, Cambridge University Computer Laboratory
+    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   Cambridge University 2007
+
+Proof reconstruction for Metis.
+*)
+
+signature METIS_RECONSTRUCT =
+sig
+  type mode = Metis_Translate.mode
+
+  val trace: bool Unsynchronized.ref
+  val lookth : (Metis_Thm.thm * 'a) list -> Metis_Thm.thm -> 'a
+  val replay_one_inference :
+    Proof.context -> mode -> (string * term) list
+    -> Metis_Thm.thm * Metis_Proof.inference -> (Metis_Thm.thm * thm) list
+    -> (Metis_Thm.thm * thm) list
+end;
+
+structure Metis_Reconstruct : METIS_RECONSTRUCT =
+struct
+
+open Metis_Translate
+
+val trace = Unsynchronized.ref false
+fun trace_msg msg = if !trace then tracing (msg ()) else ()
+
+datatype term_or_type = SomeTerm of term | SomeType of typ
+
+fun terms_of [] = []
+  | terms_of (SomeTerm t :: tts) = t :: terms_of tts
+  | terms_of (SomeType _ :: tts) = terms_of tts;
+
+fun types_of [] = []
+  | types_of (SomeTerm (Var ((a,idx), _)) :: tts) =
+      if String.isPrefix "_" a then
+          (*Variable generated by Metis, which might have been a type variable.*)
+          TVar (("'" ^ a, idx), HOLogic.typeS) :: types_of tts
+      else types_of tts
+  | types_of (SomeTerm _ :: tts) = types_of tts
+  | types_of (SomeType T :: tts) = T :: types_of tts;
+
+fun apply_list rator nargs rands =
+  let val trands = terms_of rands
+  in  if length trands = nargs then SomeTerm (list_comb(rator, trands))
+      else raise Fail
+        ("apply_list: wrong number of arguments: " ^ Syntax.string_of_term_global Pure.thy rator ^
+          " expected " ^ Int.toString nargs ^
+          " received " ^ commas (map (Syntax.string_of_term_global Pure.thy) trands))
+  end;
+
+fun infer_types ctxt =
+  Syntax.check_terms (ProofContext.set_mode ProofContext.mode_pattern ctxt);
+
+(*We use 1 rather than 0 because variable references in clauses may otherwise conflict
+  with variable constraints in the goal...at least, type inference often fails otherwise.
+  SEE ALSO axiom_inf below.*)
+fun mk_var (w, T) = Var ((w, 1), T)
+
+(*include the default sort, if available*)
+fun mk_tfree ctxt w =
+  let val ww = "'" ^ w
+  in  TFree(ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1)))  end;
+
+(*Remove the "apply" operator from an HO term*)
+fun strip_happ args (Metis_Term.Fn(".",[t,u])) = strip_happ (u::args) t
+  | strip_happ args x = (x, args);
+
+fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
+
+fun smart_invert_const "fequal" = @{const_name HOL.eq}
+  | smart_invert_const s = invert_const s
+
+fun hol_type_from_metis_term _ (Metis_Term.Var v) =
+     (case strip_prefix_and_unascii tvar_prefix v of
+          SOME w => make_tvar w
+        | NONE   => make_tvar v)
+  | hol_type_from_metis_term ctxt (Metis_Term.Fn(x, tys)) =
+     (case strip_prefix_and_unascii type_const_prefix x of
+          SOME tc => Type (smart_invert_const tc,
+                           map (hol_type_from_metis_term ctxt) tys)
+        | NONE    =>
+      case strip_prefix_and_unascii tfree_prefix x of
+          SOME tf => mk_tfree ctxt tf
+        | NONE    => raise Fail ("hol_type_from_metis_term: " ^ x));
+
+(*Maps metis terms to isabelle terms*)
+fun hol_term_from_metis_PT ctxt fol_tm =
+  let val thy = ProofContext.theory_of ctxt
+      val _ = trace_msg (fn () => "hol_term_from_metis_PT: " ^
+                                  Metis_Term.toString fol_tm)
+      fun tm_to_tt (Metis_Term.Var v) =
+             (case strip_prefix_and_unascii tvar_prefix v of
+                  SOME w => SomeType (make_tvar w)
+                | NONE =>
+              case strip_prefix_and_unascii schematic_var_prefix v of
+                  SOME w => SomeTerm (mk_var (w, HOLogic.typeT))
+                | NONE   => SomeTerm (mk_var (v, HOLogic.typeT)) )
+                    (*Var from Metis with a name like _nnn; possibly a type variable*)
+        | tm_to_tt (Metis_Term.Fn ("{}", [arg])) = tm_to_tt arg   (*hBOOL*)
+        | tm_to_tt (t as Metis_Term.Fn (".",_)) =
+            let val (rator,rands) = strip_happ [] t
+            in  case rator of
+                    Metis_Term.Fn(fname,ts) => applic_to_tt (fname, ts @ rands)
+                  | _ => case tm_to_tt rator of
+                             SomeTerm t => SomeTerm (list_comb(t, terms_of (map tm_to_tt rands)))
+                           | _ => raise Fail "tm_to_tt: HO application"
+            end
+        | tm_to_tt (Metis_Term.Fn (fname, args)) = applic_to_tt (fname,args)
+      and applic_to_tt ("=",ts) =
+            SomeTerm (list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), terms_of (map tm_to_tt ts)))
+        | applic_to_tt (a,ts) =
+            case strip_prefix_and_unascii const_prefix a of
+                SOME b =>
+                  let val c = smart_invert_const b
+                      val ntypes = num_type_args thy c
+                      val nterms = length ts - ntypes
+                      val tts = map tm_to_tt ts
+                      val tys = types_of (List.take(tts,ntypes))
+                  in if length tys = ntypes then
+                         apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes))
+                     else
+                       raise Fail ("Constant " ^ c ^ " expects " ^ Int.toString ntypes ^
+                                   " but gets " ^ Int.toString (length tys) ^
+                                   " type arguments\n" ^
+                                   cat_lines (map (Syntax.string_of_typ ctxt) tys) ^
+                                   " the terms are \n" ^
+                                   cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
+                     end
+              | NONE => (*Not a constant. Is it a type constructor?*)
+            case strip_prefix_and_unascii type_const_prefix a of
+                SOME b =>
+                SomeType (Type (smart_invert_const b, types_of (map tm_to_tt ts)))
+              | NONE => (*Maybe a TFree. Should then check that ts=[].*)
+            case strip_prefix_and_unascii tfree_prefix a of
+                SOME b => SomeType (mk_tfree ctxt b)
+              | NONE => (*a fixed variable? They are Skolem functions.*)
+            case strip_prefix_and_unascii fixed_var_prefix a of
+                SOME b =>
+                  let val opr = Free (b, HOLogic.typeT)
+                  in  apply_list opr (length ts) (map tm_to_tt ts)  end
+              | NONE => raise Fail ("unexpected metis function: " ^ a)
+  in
+    case tm_to_tt fol_tm of
+      SomeTerm t => t
+    | SomeType T => raise TYPE ("fol_tm_to_tt: Term expected", [T], [])
+  end
+
+(*Maps fully-typed metis terms to isabelle terms*)
+fun hol_term_from_metis_FT ctxt fol_tm =
+  let val _ = trace_msg (fn () => "hol_term_from_metis_FT: " ^
+                                  Metis_Term.toString fol_tm)
+      fun cvt (Metis_Term.Fn ("ti", [Metis_Term.Var v, _])) =
+             (case strip_prefix_and_unascii schematic_var_prefix v of
+                  SOME w =>  mk_var(w, dummyT)
+                | NONE   => mk_var(v, dummyT))
+        | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn ("=",[]), _])) =
+            Const (@{const_name HOL.eq}, HOLogic.typeT)
+        | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (x,[]), ty])) =
+           (case strip_prefix_and_unascii const_prefix x of
+                SOME c => Const (smart_invert_const c, dummyT)
+              | NONE => (*Not a constant. Is it a fixed variable??*)
+            case strip_prefix_and_unascii fixed_var_prefix x of
+                SOME v => Free (v, hol_type_from_metis_term ctxt ty)
+              | NONE => raise Fail ("hol_term_from_metis_FT bad constant: " ^ x))
+        | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (".",[tm1,tm2]), _])) =
+            cvt tm1 $ cvt tm2
+        | cvt (Metis_Term.Fn (".",[tm1,tm2])) = (*untyped application*)
+            cvt tm1 $ cvt tm2
+        | cvt (Metis_Term.Fn ("{}", [arg])) = cvt arg   (*hBOOL*)
+        | cvt (Metis_Term.Fn ("=", [tm1,tm2])) =
+            list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), map cvt [tm1,tm2])
+        | cvt (t as Metis_Term.Fn (x, [])) =
+           (case strip_prefix_and_unascii const_prefix x of
+                SOME c => Const (smart_invert_const c, dummyT)
+              | NONE => (*Not a constant. Is it a fixed variable??*)
+            case strip_prefix_and_unascii fixed_var_prefix x of
+                SOME v => Free (v, dummyT)
+              | NONE => (trace_msg (fn () => "hol_term_from_metis_FT bad const: " ^ x);
+                  hol_term_from_metis_PT ctxt t))
+        | cvt t = (trace_msg (fn () => "hol_term_from_metis_FT bad term: " ^ Metis_Term.toString t);
+            hol_term_from_metis_PT ctxt t)
+  in fol_tm |> cvt end
+
+fun hol_term_from_metis FT = hol_term_from_metis_FT
+  | hol_term_from_metis _ = hol_term_from_metis_PT
+
+fun hol_terms_from_fol ctxt mode skolems fol_tms =
+  let val ts = map (hol_term_from_metis mode ctxt) fol_tms
+      val _ = trace_msg (fn () => "  calling type inference:")
+      val _ = app (fn t => trace_msg (fn () => Syntax.string_of_term ctxt t)) ts
+      val ts' = ts |> map (reveal_skolem_terms skolems) |> infer_types ctxt
+      val _ = app (fn t => trace_msg
+                    (fn () => "  final term: " ^ Syntax.string_of_term ctxt t ^
+                              "  of type  " ^ Syntax.string_of_typ ctxt (type_of t)))
+                  ts'
+  in  ts'  end;
+
+(* ------------------------------------------------------------------------- *)
+(* FOL step Inference Rules                                                  *)
+(* ------------------------------------------------------------------------- *)
+
+(*for debugging only*)
+(*
+fun print_thpair (fth,th) =
+  (trace_msg (fn () => "=============================================");
+   trace_msg (fn () => "Metis: " ^ Metis_Thm.toString fth);
+   trace_msg (fn () => "Isabelle: " ^ Display.string_of_thm_without_context th));
+*)
+
+fun lookth thpairs (fth : Metis_Thm.thm) =
+  the (AList.lookup (uncurry Metis_Thm.equal) thpairs fth)
+  handle Option.Option =>
+         raise Fail ("Failed to find Metis theorem " ^ Metis_Thm.toString fth)
+
+fun cterm_incr_types thy idx = cterm_of thy o (map_types (Logic.incr_tvar idx));
+
+(* INFERENCE RULE: AXIOM *)
+
+fun axiom_inf thpairs th = Thm.incr_indexes 1 (lookth thpairs th);
+    (*This causes variables to have an index of 1 by default. SEE ALSO mk_var above.*)
+
+(* INFERENCE RULE: ASSUME *)
+
+val EXCLUDED_MIDDLE = @{lemma "P ==> ~ P ==> False" by (rule notE)}
+
+fun inst_excluded_middle thy i_atm =
+  let val th = EXCLUDED_MIDDLE
+      val [vx] = Term.add_vars (prop_of th) []
+      val substs = [(cterm_of thy (Var vx), cterm_of thy i_atm)]
+  in  cterm_instantiate substs th  end;
+
+fun assume_inf ctxt mode skolems atm =
+  inst_excluded_middle
+      (ProofContext.theory_of ctxt)
+      (singleton (hol_terms_from_fol ctxt mode skolems) (Metis_Term.Fn atm))
+
+(* INFERENCE RULE: INSTANTIATE (SUBST). Type instantiations are ignored. Trying
+   to reconstruct them admits new possibilities of errors, e.g. concerning
+   sorts. Instead we try to arrange that new TVars are distinct and that types
+   can be inferred from terms. *)
+
+fun inst_inf ctxt mode skolems thpairs fsubst th =
+  let val thy = ProofContext.theory_of ctxt
+      val i_th = lookth thpairs th
+      val i_th_vars = Term.add_vars (prop_of i_th) []
+      fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars)
+      fun subst_translation (x,y) =
+        let val v = find_var x
+            (* We call "reveal_skolem_terms" and "infer_types" below. *)
+            val t = hol_term_from_metis mode ctxt y
+        in  SOME (cterm_of thy (Var v), t)  end
+        handle Option.Option =>
+               (trace_msg (fn () => "\"find_var\" failed for " ^ x ^
+                                    " in " ^ Display.string_of_thm ctxt i_th);
+                NONE)
+             | TYPE _ =>
+               (trace_msg (fn () => "\"hol_term_from_metis\" failed for " ^ x ^
+                                    " in " ^ Display.string_of_thm ctxt i_th);
+                NONE)
+      fun remove_typeinst (a, t) =
+            case strip_prefix_and_unascii schematic_var_prefix a of
+                SOME b => SOME (b, t)
+              | NONE => case strip_prefix_and_unascii tvar_prefix a of
+                SOME _ => NONE          (*type instantiations are forbidden!*)
+              | NONE => SOME (a,t)    (*internal Metis var?*)
+      val _ = trace_msg (fn () => "  isa th: " ^ Display.string_of_thm ctxt i_th)
+      val substs = map_filter remove_typeinst (Metis_Subst.toList fsubst)
+      val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs)
+      val tms = rawtms |> map (reveal_skolem_terms skolems) |> infer_types ctxt
+      val ctm_of = cterm_incr_types thy (1 + Thm.maxidx_of i_th)
+      val substs' = ListPair.zip (vars, map ctm_of tms)
+      val _ = trace_msg (fn () =>
+        cat_lines ("subst_translations:" ::
+          (substs' |> map (fn (x, y) =>
+            Syntax.string_of_term ctxt (term_of x) ^ " |-> " ^
+            Syntax.string_of_term ctxt (term_of y)))));
+  in cterm_instantiate substs' i_th end
+  handle THM (msg, _, _) =>
+         error ("Cannot replay Metis proof in Isabelle:\n" ^ msg)
+
+(* INFERENCE RULE: RESOLVE *)
+
+(* Like RSN, but we rename apart only the type variables. Vars here typically
+   have an index of 1, and the use of RSN would increase this typically to 3.
+   Instantiations of those Vars could then fail. See comment on "mk_var". *)
+fun resolve_inc_tyvars thy tha i thb =
+  let
+    val tha = Drule.incr_type_indexes (1 + Thm.maxidx_of thb) tha
+    fun aux tha thb =
+      case Thm.bicompose false (false, tha, nprems_of tha) i thb
+           |> Seq.list_of |> distinct Thm.eq_thm of
+        [th] => th
+      | _ => raise THM ("resolve_inc_tyvars: unique result expected", i,
+                        [tha, thb])
+  in
+    aux tha thb
+    handle TERM z =>
+           (* The unifier, which is invoked from "Thm.bicompose", will sometimes
+              refuse to unify "?a::?'a" with "?a::?'b" or "?a::nat" and throw a
+              "TERM" exception (with "add_ffpair" as first argument). We then
+              perform unification of the types of variables by hand and try
+              again. We could do this the first time around but this error
+              occurs seldom and we don't want to break existing proofs in subtle
+              ways or slow them down needlessly. *)
+           case [] |> fold (Term.add_vars o prop_of) [tha, thb]
+                   |> AList.group (op =)
+                   |> maps (fn ((s, _), T :: Ts) =>
+                               map (fn T' => (Free (s, T), Free (s, T'))) Ts)
+                   |> rpair (Envir.empty ~1)
+                   |-> fold (Pattern.unify thy)
+                   |> Envir.type_env |> Vartab.dest
+                   |> map (fn (x, (S, T)) =>
+                              pairself (ctyp_of thy) (TVar (x, S), T)) of
+             [] => raise TERM z
+           | ps => aux (instantiate (ps, []) tha) (instantiate (ps, []) thb)
+  end
+
+fun mk_not (Const (@{const_name Not}, _) $ b) = b
+  | mk_not b = HOLogic.mk_not b
+
+(* Match untyped terms. *)
+fun untyped_aconv (Const (a, _)) (Const(b, _)) = (a = b)
+  | untyped_aconv (Free (a, _)) (Free (b, _)) = (a = b)
+  | untyped_aconv (Var ((a, _), _)) (Var ((b, _), _)) =
+    (a = b) (* The index is ignored, for some reason. *)
+  | untyped_aconv (Bound i) (Bound j) = (i = j)
+  | untyped_aconv (Abs (_, _, t)) (Abs (_, _, u)) = untyped_aconv t u
+  | untyped_aconv (t1 $ t2) (u1 $ u2) =
+    untyped_aconv t1 u1 andalso untyped_aconv t2 u2
+  | untyped_aconv _ _ = false
+
+(* Finding the relative location of an untyped term within a list of terms *)
+fun literal_index lit =
+  let
+    val lit = Envir.eta_contract lit
+    fun get _ [] = raise Empty
+      | get n (x :: xs) =
+        if untyped_aconv lit (Envir.eta_contract (HOLogic.dest_Trueprop x)) then
+          n
+        else
+          get (n+1) xs
+  in get 1 end
+
+fun resolve_inf ctxt mode skolems thpairs atm th1 th2 =
+  let
+    val thy = ProofContext.theory_of ctxt
+    val i_th1 = lookth thpairs th1 and i_th2 = lookth thpairs th2
+    val _ = trace_msg (fn () => "  isa th1 (pos): " ^ Display.string_of_thm ctxt i_th1)
+    val _ = trace_msg (fn () => "  isa th2 (neg): " ^ Display.string_of_thm ctxt i_th2)
+  in
+    (* Trivial cases where one operand is type info *)
+    if Thm.eq_thm (TrueI, i_th1) then
+      i_th2
+    else if Thm.eq_thm (TrueI, i_th2) then
+      i_th1
+    else
+      let
+        val i_atm = singleton (hol_terms_from_fol ctxt mode skolems)
+                              (Metis_Term.Fn atm)
+        val _ = trace_msg (fn () => "  atom: " ^ Syntax.string_of_term ctxt i_atm)
+        val prems_th1 = prems_of i_th1
+        val prems_th2 = prems_of i_th2
+        val index_th1 = literal_index (mk_not i_atm) prems_th1
+              handle Empty => raise Fail "Failed to find literal in th1"
+        val _ = trace_msg (fn () => "  index_th1: " ^ Int.toString index_th1)
+        val index_th2 = literal_index i_atm prems_th2
+              handle Empty => raise Fail "Failed to find literal in th2"
+        val _ = trace_msg (fn () => "  index_th2: " ^ Int.toString index_th2)
+    in
+      resolve_inc_tyvars thy (Meson.select_literal index_th1 i_th1) index_th2
+                         i_th2
+    end
+  end;
+
+(* INFERENCE RULE: REFL *)
+
+val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp}
+
+val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
+val refl_idx = 1 + Thm.maxidx_of REFL_THM;
+
+fun refl_inf ctxt mode skolems t =
+  let val thy = ProofContext.theory_of ctxt
+      val i_t = singleton (hol_terms_from_fol ctxt mode skolems) t
+      val _ = trace_msg (fn () => "  term: " ^ Syntax.string_of_term ctxt i_t)
+      val c_t = cterm_incr_types thy refl_idx i_t
+  in  cterm_instantiate [(refl_x, c_t)] REFL_THM  end;
+
+(* INFERENCE RULE: EQUALITY *)
+
+val subst_em = @{lemma "s = t ==> P s ==> ~ P t ==> False" by simp}
+val ssubst_em = @{lemma "s = t ==> P t ==> ~ P s ==> False" by simp}
+
+val metis_eq = Metis_Term.Fn ("=", []);
+
+fun get_ty_arg_size _ (Const (@{const_name HOL.eq}, _)) = 0  (*equality has no type arguments*)
+  | get_ty_arg_size thy (Const (c, _)) = (num_type_args thy c handle TYPE _ => 0)
+  | get_ty_arg_size _ _ = 0;
+
+fun equality_inf ctxt mode skolems (pos, atm) fp fr =
+  let val thy = ProofContext.theory_of ctxt
+      val m_tm = Metis_Term.Fn atm
+      val [i_atm,i_tm] = hol_terms_from_fol ctxt mode skolems [m_tm, fr]
+      val _ = trace_msg (fn () => "sign of the literal: " ^ Bool.toString pos)
+      fun replace_item_list lx 0 (_::ls) = lx::ls
+        | replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls
+      fun path_finder_FO tm [] = (tm, Bound 0)
+        | path_finder_FO tm (p::ps) =
+            let val (tm1,args) = strip_comb tm
+                val adjustment = get_ty_arg_size thy tm1
+                val p' = if adjustment > p then p else p-adjustment
+                val tm_p = List.nth(args,p')
+                  handle Subscript =>
+                         error ("Cannot replay Metis proof in Isabelle:\n" ^
+                                "equality_inf: " ^ Int.toString p ^ " adj " ^
+                                Int.toString adjustment ^ " term " ^
+                                Syntax.string_of_term ctxt tm)
+                val _ = trace_msg (fn () => "path_finder: " ^ Int.toString p ^
+                                      "  " ^ Syntax.string_of_term ctxt tm_p)
+                val (r,t) = path_finder_FO tm_p ps
+            in
+                (r, list_comb (tm1, replace_item_list t p' args))
+            end
+      fun path_finder_HO tm [] = (tm, Bound 0)
+        | path_finder_HO (t$u) (0::ps) = (fn(x,y) => (x, y$u)) (path_finder_HO t ps)
+        | path_finder_HO (t$u) (_::ps) = (fn(x,y) => (x, t$y)) (path_finder_HO u ps)
+        | path_finder_HO tm ps =
+          raise Fail ("Cannot replay Metis proof in Isabelle:\n" ^
+                      "equality_inf, path_finder_HO: path = " ^
+                      space_implode " " (map Int.toString ps) ^
+                      " isa-term: " ^  Syntax.string_of_term ctxt tm)
+      fun path_finder_FT tm [] _ = (tm, Bound 0)
+        | path_finder_FT tm (0::ps) (Metis_Term.Fn ("ti", [t1, _])) =
+            path_finder_FT tm ps t1
+        | path_finder_FT (t$u) (0::ps) (Metis_Term.Fn (".", [t1, _])) =
+            (fn(x,y) => (x, y$u)) (path_finder_FT t ps t1)
+        | path_finder_FT (t$u) (1::ps) (Metis_Term.Fn (".", [_, t2])) =
+            (fn(x,y) => (x, t$y)) (path_finder_FT u ps t2)
+        | path_finder_FT tm ps t =
+          raise Fail ("Cannot replay Metis proof in Isabelle:\n" ^
+                      "equality_inf, path_finder_FT: path = " ^
+                      space_implode " " (map Int.toString ps) ^
+                      " isa-term: " ^  Syntax.string_of_term ctxt tm ^
+                      " fol-term: " ^ Metis_Term.toString t)
+      fun path_finder FO tm ps _ = path_finder_FO tm ps
+        | path_finder HO (tm as Const(@{const_name HOL.eq},_) $ _ $ _) (p::ps) _ =
+             (*equality: not curried, as other predicates are*)
+             if p=0 then path_finder_HO tm (0::1::ps)  (*select first operand*)
+             else path_finder_HO tm (p::ps)        (*1 selects second operand*)
+        | path_finder HO tm (_ :: ps) (Metis_Term.Fn ("{}", [_])) =
+             path_finder_HO tm ps      (*if not equality, ignore head to skip hBOOL*)
+        | path_finder FT (tm as Const(@{const_name HOL.eq}, _) $ _ $ _) (p::ps)
+                            (Metis_Term.Fn ("=", [t1,t2])) =
+             (*equality: not curried, as other predicates are*)
+             if p=0 then path_finder_FT tm (0::1::ps)
+                          (Metis_Term.Fn (".", [Metis_Term.Fn (".", [metis_eq,t1]), t2]))
+                          (*select first operand*)
+             else path_finder_FT tm (p::ps)
+                   (Metis_Term.Fn (".", [metis_eq,t2]))
+                   (*1 selects second operand*)
+        | path_finder FT tm (_ :: ps) (Metis_Term.Fn ("{}", [t1])) = path_finder_FT tm ps t1
+             (*if not equality, ignore head to skip the hBOOL predicate*)
+        | path_finder FT tm ps t = path_finder_FT tm ps t  (*really an error case!*)
+      fun path_finder_lit ((nt as Const (@{const_name Not}, _)) $ tm_a) idx =
+            let val (tm, tm_rslt) = path_finder mode tm_a idx m_tm
+            in (tm, nt $ tm_rslt) end
+        | path_finder_lit tm_a idx = path_finder mode tm_a idx m_tm
+      val (tm_subst, body) = path_finder_lit i_atm fp
+      val tm_abs = Abs ("x", type_of tm_subst, body)
+      val _ = trace_msg (fn () => "abstraction: " ^ Syntax.string_of_term ctxt tm_abs)
+      val _ = trace_msg (fn () => "i_tm: " ^ Syntax.string_of_term ctxt i_tm)
+      val _ = trace_msg (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst)
+      val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst)  (*ill typed but gives right max*)
+      val subst' = Thm.incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
+      val _ = trace_msg (fn () => "subst' " ^ Display.string_of_thm ctxt subst')
+      val eq_terms = map (pairself (cterm_of thy))
+        (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
+  in  cterm_instantiate eq_terms subst'  end;
+
+val factor = Seq.hd o distinct_subgoals_tac;
+
+fun step ctxt mode skolems thpairs p =
+  case p of
+    (fol_th, Metis_Proof.Axiom _) => factor (axiom_inf thpairs fol_th)
+  | (_, Metis_Proof.Assume f_atm) => assume_inf ctxt mode skolems f_atm
+  | (_, Metis_Proof.Metis_Subst (f_subst, f_th1)) =>
+    factor (inst_inf ctxt mode skolems thpairs f_subst f_th1)
+  | (_, Metis_Proof.Resolve(f_atm, f_th1, f_th2)) =>
+    factor (resolve_inf ctxt mode skolems thpairs f_atm f_th1 f_th2)
+  | (_, Metis_Proof.Refl f_tm) => refl_inf ctxt mode skolems f_tm
+  | (_, Metis_Proof.Equality (f_lit, f_p, f_r)) =>
+    equality_inf ctxt mode skolems f_lit f_p f_r
+
+fun is_real_literal (_, (c, _)) = not (String.isPrefix class_prefix c)
+
+fun replay_one_inference ctxt mode skolems (fol_th, inf) thpairs =
+  let
+    val _ = trace_msg (fn () => "=============================================")
+    val _ = trace_msg (fn () => "METIS THM: " ^ Metis_Thm.toString fol_th)
+    val _ = trace_msg (fn () => "INFERENCE: " ^ Metis_Proof.inferenceToString inf)
+    val th = Meson.flexflex_first_order (step ctxt mode skolems
+                                              thpairs (fol_th, inf))
+    val _ = trace_msg (fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th)
+    val _ = trace_msg (fn () => "=============================================")
+    val n_metis_lits =
+      length (filter is_real_literal (Metis_LiteralSet.toList (Metis_Thm.clause fol_th)))
+    val _ = if nprems_of th = n_metis_lits then ()
+            else error "Cannot replay Metis proof in Isabelle: Out of sync."
+  in (fol_th, th) :: thpairs end
+
+end;
--- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -9,733 +9,34 @@
 
 signature METIS_TACTICS =
 sig
-  val trace: bool Unsynchronized.ref
-  val type_lits: bool Config.T
-  val metis_tac: Proof.context -> thm list -> int -> tactic
-  val metisF_tac: Proof.context -> thm list -> int -> tactic
-  val metisFT_tac: Proof.context -> thm list -> int -> tactic
-  val setup: theory -> theory
+  val trace : bool Unsynchronized.ref
+  val type_lits : bool Config.T
+  val metis_tac : Proof.context -> thm list -> int -> tactic
+  val metisF_tac : Proof.context -> thm list -> int -> tactic
+  val metisFT_tac : Proof.context -> thm list -> int -> tactic
+  val setup : theory -> theory
 end
 
 structure Metis_Tactics : METIS_TACTICS =
 struct
 
-open Metis_Clauses
+open Metis_Translate
+open Metis_Reconstruct
 
-val trace = Unsynchronized.ref false;
-fun trace_msg msg = if !trace then tracing (msg ()) else ();
+val trace = Unsynchronized.ref false
+fun trace_msg msg = if !trace then tracing (msg ()) else ()
 
 val (type_lits, type_lits_setup) = Attrib.config_bool "metis_type_lits" (K true);
 
-datatype mode = FO | HO | FT  (* first-order, higher-order, fully-typed *)
-
-(* ------------------------------------------------------------------------- *)
-(* Useful Theorems                                                           *)
-(* ------------------------------------------------------------------------- *)
-val EXCLUDED_MIDDLE = @{lemma "P ==> ~ P ==> False" by (rule notE)}
-val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp}
-val subst_em = @{lemma "s = t ==> P s ==> ~ P t ==> False" by simp}
-val ssubst_em = @{lemma "s = t ==> P t ==> ~ P s ==> False" by simp}
-
-(* ------------------------------------------------------------------------- *)
-(* Useful Functions                                                          *)
-(* ------------------------------------------------------------------------- *)
-
-(* Match untyped terms. *)
-fun untyped_aconv (Const (a, _)) (Const(b, _)) = (a = b)
-  | untyped_aconv (Free (a, _)) (Free (b, _)) = (a = b)
-  | untyped_aconv (Var ((a, _), _)) (Var ((b, _), _)) =
-    (a = b) (* The index is ignored, for some reason. *)
-  | untyped_aconv (Bound i) (Bound j) = (i = j)
-  | untyped_aconv (Abs (_, _, t)) (Abs (_, _, u)) = untyped_aconv t u
-  | untyped_aconv (t1 $ t2) (u1 $ u2) =
-    untyped_aconv t1 u1 andalso untyped_aconv t2 u2
-  | untyped_aconv _ _ = false
-
-(* Finding the relative location of an untyped term within a list of terms *)
-fun get_index lit =
-  let val lit = Envir.eta_contract lit
-      fun get _ [] = raise Empty
-        | get n (x::xs) = if untyped_aconv lit (Envir.eta_contract (HOLogic.dest_Trueprop x))
-                          then n  else get (n+1) xs
-  in get 1 end;
-
-(* ------------------------------------------------------------------------- *)
-(* HOL to FOL  (Isabelle to Metis)                                           *)
-(* ------------------------------------------------------------------------- *)
-
-fun fn_isa_to_met_sublevel "equal" = "=" (* FIXME: "c_fequal" *)
-  | fn_isa_to_met_sublevel x = x
-fun fn_isa_to_met_toplevel "equal" = "="
-  | fn_isa_to_met_toplevel x = x
-
-fun metis_lit b c args = (b, (c, args));
-
-fun metis_term_from_combtyp (CombTVar (s, _)) = Metis_Term.Var s
-  | metis_term_from_combtyp (CombTFree (s, _)) = Metis_Term.Fn (s, [])
-  | metis_term_from_combtyp (CombType ((s, _), tps)) =
-    Metis_Term.Fn (s, map metis_term_from_combtyp tps);
-
-(*These two functions insert type literals before the real literals. That is the
-  opposite order from TPTP linkup, but maybe OK.*)
-
-fun hol_term_to_fol_FO tm =
-  case strip_combterm_comb tm of
-      (CombConst ((c, _), _, tys), tms) =>
-        let val tyargs = map metis_term_from_combtyp tys
-            val args   = map hol_term_to_fol_FO tms
-        in Metis_Term.Fn (c, tyargs @ args) end
-    | (CombVar ((v, _), _), []) => Metis_Term.Var v
-    | _ => raise Fail "non-first-order combterm"
-
-fun hol_term_to_fol_HO (CombConst ((a, _), _, tylist)) =
-      Metis_Term.Fn (fn_isa_to_met_sublevel a, map metis_term_from_combtyp tylist)
-  | hol_term_to_fol_HO (CombVar ((s, _), _)) = Metis_Term.Var s
-  | hol_term_to_fol_HO (CombApp (tm1, tm2)) =
-       Metis_Term.Fn (".", map hol_term_to_fol_HO [tm1, tm2]);
-
-(*The fully-typed translation, to avoid type errors*)
-fun wrap_type (tm, ty) = Metis_Term.Fn("ti", [tm, metis_term_from_combtyp ty]);
-
-fun hol_term_to_fol_FT (CombVar ((s, _), ty)) = wrap_type (Metis_Term.Var s, ty)
-  | hol_term_to_fol_FT (CombConst((a, _), ty, _)) =
-      wrap_type (Metis_Term.Fn(fn_isa_to_met_sublevel a, []), ty)
-  | hol_term_to_fol_FT (tm as CombApp(tm1,tm2)) =
-       wrap_type (Metis_Term.Fn(".", map hol_term_to_fol_FT [tm1,tm2]),
-                  combtyp_of tm)
-
-fun hol_literal_to_fol FO (FOLLiteral (pos, tm)) =
-      let val (CombConst((p, _), _, tys), tms) = strip_combterm_comb tm
-          val tylits = if p = "equal" then [] else map metis_term_from_combtyp tys
-          val lits = map hol_term_to_fol_FO tms
-      in metis_lit pos (fn_isa_to_met_toplevel p) (tylits @ lits) end
-  | hol_literal_to_fol HO (FOLLiteral (pos, tm)) =
-     (case strip_combterm_comb tm of
-          (CombConst(("equal", _), _, _), tms) =>
-            metis_lit pos "=" (map hol_term_to_fol_HO tms)
-        | _ => metis_lit pos "{}" [hol_term_to_fol_HO tm])   (*hBOOL*)
-  | hol_literal_to_fol FT (FOLLiteral (pos, tm)) =
-     (case strip_combterm_comb tm of
-          (CombConst(("equal", _), _, _), tms) =>
-            metis_lit pos "=" (map hol_term_to_fol_FT tms)
-        | _ => metis_lit pos "{}" [hol_term_to_fol_FT tm])   (*hBOOL*);
-
-fun literals_of_hol_term thy mode t =
-      let val (lits, types_sorts) = literals_of_term thy t
-      in  (map (hol_literal_to_fol mode) lits, types_sorts) end;
-
-(*Sign should be "true" for conjecture type constraints, "false" for type lits in clauses.*)
-fun metis_of_type_literals pos (TyLitVar ((s, _), (s', _))) =
-    metis_lit pos s [Metis_Term.Var s']
-  | metis_of_type_literals pos (TyLitFree ((s, _), (s', _))) =
-    metis_lit pos s [Metis_Term.Fn (s',[])]
-
-fun default_sort _ (TVar _) = false
-  | default_sort ctxt (TFree (x, s)) = (s = the_default [] (Variable.def_sort ctxt (x, ~1)));
-
-fun metis_of_tfree tf =
-  Metis_Thm.axiom (Metis_LiteralSet.singleton (metis_of_type_literals true tf));
-
-fun hol_thm_to_fol is_conjecture ctxt mode j skolems th =
-  let
-    val thy = ProofContext.theory_of ctxt
-    val (skolems, (mlits, types_sorts)) =
-     th |> prop_of |> conceal_skolem_terms j skolems
-        ||> (HOLogic.dest_Trueprop #> literals_of_hol_term thy mode)
-  in
-      if is_conjecture then
-          (Metis_Thm.axiom (Metis_LiteralSet.fromList mlits),
-           type_literals_for_types types_sorts, skolems)
-      else
-        let val tylits = filter_out (default_sort ctxt) types_sorts
-                         |> type_literals_for_types
-            val mtylits = if Config.get ctxt type_lits
-                          then map (metis_of_type_literals false) tylits else []
-        in
-          (Metis_Thm.axiom (Metis_LiteralSet.fromList(mtylits @ mlits)), [],
-           skolems)
-        end
-  end;
-
-(* ARITY CLAUSE *)
-
-fun m_arity_cls (TConsLit ((c, _), (t, _), args)) =
-    metis_lit true c [Metis_Term.Fn(t, map (Metis_Term.Var o fst) args)]
-  | m_arity_cls (TVarLit ((c, _), (s, _))) =
-    metis_lit false c [Metis_Term.Var s]
-
-(*TrueI is returned as the Isabelle counterpart because there isn't any.*)
-fun arity_cls (ArityClause {conclLit, premLits, ...}) =
-  (TrueI,
-   Metis_Thm.axiom (Metis_LiteralSet.fromList (map m_arity_cls (conclLit :: premLits))));
-
-(* CLASSREL CLAUSE *)
-
-fun m_class_rel_cls (subclass, _) (superclass, _) =
-  [metis_lit false subclass [Metis_Term.Var "T"], metis_lit true superclass [Metis_Term.Var "T"]];
-
-fun class_rel_cls (ClassRelClause {subclass, superclass, ...}) =
-  (TrueI, Metis_Thm.axiom (Metis_LiteralSet.fromList (m_class_rel_cls subclass superclass)));
-
-(* ------------------------------------------------------------------------- *)
-(* FOL to HOL  (Metis to Isabelle)                                           *)
-(* ------------------------------------------------------------------------- *)
-
-datatype term_or_type = Term of Term.term | Type of Term.typ;
-
-fun terms_of [] = []
-  | terms_of (Term t :: tts) = t :: terms_of tts
-  | terms_of (Type _ :: tts) = terms_of tts;
-
-fun types_of [] = []
-  | types_of (Term (Term.Var ((a,idx), _)) :: tts) =
-      if String.isPrefix "_" a then
-          (*Variable generated by Metis, which might have been a type variable.*)
-          TVar (("'" ^ a, idx), HOLogic.typeS) :: types_of tts
-      else types_of tts
-  | types_of (Term _ :: tts) = types_of tts
-  | types_of (Type T :: tts) = T :: types_of tts;
-
-fun apply_list rator nargs rands =
-  let val trands = terms_of rands
-  in  if length trands = nargs then Term (list_comb(rator, trands))
-      else raise Fail
-        ("apply_list: wrong number of arguments: " ^ Syntax.string_of_term_global Pure.thy rator ^
-          " expected " ^ Int.toString nargs ^
-          " received " ^ commas (map (Syntax.string_of_term_global Pure.thy) trands))
-  end;
-
-fun infer_types ctxt =
-  Syntax.check_terms (ProofContext.set_mode ProofContext.mode_pattern ctxt);
-
-(*We use 1 rather than 0 because variable references in clauses may otherwise conflict
-  with variable constraints in the goal...at least, type inference often fails otherwise.
-  SEE ALSO axiom_inf below.*)
-fun mk_var (w,T) = Term.Var((w,1), T);
-
-(*include the default sort, if available*)
-fun mk_tfree ctxt w =
-  let val ww = "'" ^ w
-  in  TFree(ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1)))  end;
-
-(*Remove the "apply" operator from an HO term*)
-fun strip_happ args (Metis_Term.Fn(".",[t,u])) = strip_happ (u::args) t
-  | strip_happ args x = (x, args);
-
-fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
-
-fun smart_invert_const "fequal" = @{const_name HOL.eq}
-  | smart_invert_const s = invert_const s
-
-fun hol_type_from_metis_term _ (Metis_Term.Var v) =
-     (case strip_prefix_and_unascii tvar_prefix v of
-          SOME w => make_tvar w
-        | NONE   => make_tvar v)
-  | hol_type_from_metis_term ctxt (Metis_Term.Fn(x, tys)) =
-     (case strip_prefix_and_unascii type_const_prefix x of
-          SOME tc => Term.Type (smart_invert_const tc,
-                                map (hol_type_from_metis_term ctxt) tys)
-        | NONE    =>
-      case strip_prefix_and_unascii tfree_prefix x of
-          SOME tf => mk_tfree ctxt tf
-        | NONE    => raise Fail ("hol_type_from_metis_term: " ^ x));
-
-(*Maps metis terms to isabelle terms*)
-fun hol_term_from_metis_PT ctxt fol_tm =
-  let val thy = ProofContext.theory_of ctxt
-      val _ = trace_msg (fn () => "hol_term_from_metis_PT: " ^
-                                  Metis_Term.toString fol_tm)
-      fun tm_to_tt (Metis_Term.Var v) =
-             (case strip_prefix_and_unascii tvar_prefix v of
-                  SOME w => Type (make_tvar w)
-                | NONE =>
-              case strip_prefix_and_unascii schematic_var_prefix v of
-                  SOME w => Term (mk_var (w, HOLogic.typeT))
-                | NONE   => Term (mk_var (v, HOLogic.typeT)) )
-                    (*Var from Metis with a name like _nnn; possibly a type variable*)
-        | tm_to_tt (Metis_Term.Fn ("{}", [arg])) = tm_to_tt arg   (*hBOOL*)
-        | tm_to_tt (t as Metis_Term.Fn (".",_)) =
-            let val (rator,rands) = strip_happ [] t
-            in  case rator of
-                    Metis_Term.Fn(fname,ts) => applic_to_tt (fname, ts @ rands)
-                  | _ => case tm_to_tt rator of
-                             Term t => Term (list_comb(t, terms_of (map tm_to_tt rands)))
-                           | _ => raise Fail "tm_to_tt: HO application"
-            end
-        | tm_to_tt (Metis_Term.Fn (fname, args)) = applic_to_tt (fname,args)
-      and applic_to_tt ("=",ts) =
-            Term (list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), terms_of (map tm_to_tt ts)))
-        | applic_to_tt (a,ts) =
-            case strip_prefix_and_unascii const_prefix a of
-                SOME b =>
-                  let val c = smart_invert_const b
-                      val ntypes = num_type_args thy c
-                      val nterms = length ts - ntypes
-                      val tts = map tm_to_tt ts
-                      val tys = types_of (List.take(tts,ntypes))
-                  in if length tys = ntypes then
-                         apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes))
-                     else
-                       raise Fail ("Constant " ^ c ^ " expects " ^ Int.toString ntypes ^
-                                   " but gets " ^ Int.toString (length tys) ^
-                                   " type arguments\n" ^
-                                   cat_lines (map (Syntax.string_of_typ ctxt) tys) ^
-                                   " the terms are \n" ^
-                                   cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
-                     end
-              | NONE => (*Not a constant. Is it a type constructor?*)
-            case strip_prefix_and_unascii type_const_prefix a of
-                SOME b =>
-                  Type (Term.Type (smart_invert_const b, types_of (map tm_to_tt ts)))
-              | NONE => (*Maybe a TFree. Should then check that ts=[].*)
-            case strip_prefix_and_unascii tfree_prefix a of
-                SOME b => Type (mk_tfree ctxt b)
-              | NONE => (*a fixed variable? They are Skolem functions.*)
-            case strip_prefix_and_unascii fixed_var_prefix a of
-                SOME b =>
-                  let val opr = Term.Free(b, HOLogic.typeT)
-                  in  apply_list opr (length ts) (map tm_to_tt ts)  end
-              | NONE => raise Fail ("unexpected metis function: " ^ a)
-  in
-    case tm_to_tt fol_tm of
-      Term t => t
-    | _ => raise Fail "fol_tm_to_tt: Term expected"
-  end
+fun is_false t = t aconv (HOLogic.mk_Trueprop HOLogic.false_const);
 
-(*Maps fully-typed metis terms to isabelle terms*)
-fun hol_term_from_metis_FT ctxt fol_tm =
-  let val _ = trace_msg (fn () => "hol_term_from_metis_FT: " ^
-                                  Metis_Term.toString fol_tm)
-      fun cvt (Metis_Term.Fn ("ti", [Metis_Term.Var v, _])) =
-             (case strip_prefix_and_unascii schematic_var_prefix v of
-                  SOME w =>  mk_var(w, dummyT)
-                | NONE   => mk_var(v, dummyT))
-        | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn ("=",[]), _])) =
-            Const (@{const_name HOL.eq}, HOLogic.typeT)
-        | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (x,[]), ty])) =
-           (case strip_prefix_and_unascii const_prefix x of
-                SOME c => Const (smart_invert_const c, dummyT)
-              | NONE => (*Not a constant. Is it a fixed variable??*)
-            case strip_prefix_and_unascii fixed_var_prefix x of
-                SOME v => Free (v, hol_type_from_metis_term ctxt ty)
-              | NONE => raise Fail ("hol_term_from_metis_FT bad constant: " ^ x))
-        | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (".",[tm1,tm2]), _])) =
-            cvt tm1 $ cvt tm2
-        | cvt (Metis_Term.Fn (".",[tm1,tm2])) = (*untyped application*)
-            cvt tm1 $ cvt tm2
-        | cvt (Metis_Term.Fn ("{}", [arg])) = cvt arg   (*hBOOL*)
-        | cvt (Metis_Term.Fn ("=", [tm1,tm2])) =
-            list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), map cvt [tm1,tm2])
-        | cvt (t as Metis_Term.Fn (x, [])) =
-           (case strip_prefix_and_unascii const_prefix x of
-                SOME c => Const (smart_invert_const c, dummyT)
-              | NONE => (*Not a constant. Is it a fixed variable??*)
-            case strip_prefix_and_unascii fixed_var_prefix x of
-                SOME v => Free (v, dummyT)
-              | NONE => (trace_msg (fn () => "hol_term_from_metis_FT bad const: " ^ x);
-                  hol_term_from_metis_PT ctxt t))
-        | cvt t = (trace_msg (fn () => "hol_term_from_metis_FT bad term: " ^ Metis_Term.toString t);
-            hol_term_from_metis_PT ctxt t)
-  in fol_tm |> cvt end
-
-fun hol_term_from_metis FT = hol_term_from_metis_FT
-  | hol_term_from_metis _ = hol_term_from_metis_PT
-
-fun hol_terms_from_fol ctxt mode skolems fol_tms =
-  let val ts = map (hol_term_from_metis mode ctxt) fol_tms
-      val _ = trace_msg (fn () => "  calling type inference:")
-      val _ = app (fn t => trace_msg (fn () => Syntax.string_of_term ctxt t)) ts
-      val ts' = ts |> map (reveal_skolem_terms skolems) |> infer_types ctxt
-      val _ = app (fn t => trace_msg
-                    (fn () => "  final term: " ^ Syntax.string_of_term ctxt t ^
-                              "  of type  " ^ Syntax.string_of_typ ctxt (type_of t)))
-                  ts'
-  in  ts'  end;
-
-fun mk_not (Const (@{const_name Not}, _) $ b) = b
-  | mk_not b = HOLogic.mk_not b;
-
-val metis_eq = Metis_Term.Fn ("=", []);
-
-(* ------------------------------------------------------------------------- *)
-(* FOL step Inference Rules                                                  *)
-(* ------------------------------------------------------------------------- *)
-
-(*for debugging only*)
-(*
-fun print_thpair (fth,th) =
-  (trace_msg (fn () => "=============================================");
-   trace_msg (fn () => "Metis: " ^ Metis_Thm.toString fth);
-   trace_msg (fn () => "Isabelle: " ^ Display.string_of_thm_without_context th));
-*)
-
-fun lookth thpairs (fth : Metis_Thm.thm) =
-  the (AList.lookup (uncurry Metis_Thm.equal) thpairs fth)
-  handle Option =>
-         raise Fail ("Failed to find a Metis theorem " ^ Metis_Thm.toString fth);
-
-fun is_TrueI th = Thm.eq_thm(TrueI,th);
-
-fun cterm_incr_types thy idx = cterm_of thy o (map_types (Logic.incr_tvar idx));
-
-fun inst_excluded_middle thy i_atm =
-  let val th = EXCLUDED_MIDDLE
-      val [vx] = Term.add_vars (prop_of th) []
-      val substs = [(cterm_of thy (Var vx), cterm_of thy i_atm)]
-  in  cterm_instantiate substs th  end;
-
-(* INFERENCE RULE: AXIOM *)
-fun axiom_inf thpairs th = Thm.incr_indexes 1 (lookth thpairs th);
-    (*This causes variables to have an index of 1 by default. SEE ALSO mk_var above.*)
-
-(* INFERENCE RULE: ASSUME *)
-fun assume_inf ctxt mode skolems atm =
-  inst_excluded_middle
-      (ProofContext.theory_of ctxt)
-      (singleton (hol_terms_from_fol ctxt mode skolems) (Metis_Term.Fn atm))
-
-(* INFERENCE RULE: INSTANTIATE (Subst). Type instantiations are ignored. Trying to reconstruct
-   them admits new possibilities of errors, e.g. concerning sorts. Instead we try to arrange
-   that new TVars are distinct and that types can be inferred from terms.*)
-fun inst_inf ctxt mode skolems thpairs fsubst th =
-  let val thy = ProofContext.theory_of ctxt
-      val i_th   = lookth thpairs th
-      val i_th_vars = Term.add_vars (prop_of i_th) []
-      fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars)
-      fun subst_translation (x,y) =
-            let val v = find_var x
-                (* We call "reveal_skolem_terms" and "infer_types" below. *)
-                val t = hol_term_from_metis mode ctxt y
-            in  SOME (cterm_of thy (Var v), t)  end
-            handle Option =>
-                (trace_msg (fn() => "\"find_var\" failed for the variable " ^ x ^
-                                       " in " ^ Display.string_of_thm ctxt i_th);
-                 NONE)
-      fun remove_typeinst (a, t) =
-            case strip_prefix_and_unascii schematic_var_prefix a of
-                SOME b => SOME (b, t)
-              | NONE => case strip_prefix_and_unascii tvar_prefix a of
-                SOME _ => NONE          (*type instantiations are forbidden!*)
-              | NONE => SOME (a,t)    (*internal Metis var?*)
-      val _ = trace_msg (fn () => "  isa th: " ^ Display.string_of_thm ctxt i_th)
-      val substs = map_filter remove_typeinst (Metis_Subst.toList fsubst)
-      val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs)
-      val tms = rawtms |> map (reveal_skolem_terms skolems) |> infer_types ctxt
-      val ctm_of = cterm_incr_types thy (1 + Thm.maxidx_of i_th)
-      val substs' = ListPair.zip (vars, map ctm_of tms)
-      val _ = trace_msg (fn () =>
-        cat_lines ("subst_translations:" ::
-          (substs' |> map (fn (x, y) =>
-            Syntax.string_of_term ctxt (term_of x) ^ " |-> " ^
-            Syntax.string_of_term ctxt (term_of y)))));
-  in cterm_instantiate substs' i_th end
-  handle THM (msg, _, _) =>
-         error ("Cannot replay Metis proof in Isabelle:\n" ^ msg)
-
-(* INFERENCE RULE: RESOLVE *)
-
-(* Like RSN, but we rename apart only the type variables. Vars here typically
-   have an index of 1, and the use of RSN would increase this typically to 3.
-   Instantiations of those Vars could then fail. See comment on "mk_var". *)
-fun resolve_inc_tyvars thy tha i thb =
-  let
-    val tha = Drule.incr_type_indexes (1 + Thm.maxidx_of thb) tha
-    fun aux tha thb =
-      case Thm.bicompose false (false, tha, nprems_of tha) i thb
-           |> Seq.list_of |> distinct Thm.eq_thm of
-        [th] => th
-      | _ => raise THM ("resolve_inc_tyvars: unique result expected", i,
-                        [tha, thb])
-  in
-    aux tha thb
-    handle TERM z =>
-           (* The unifier, which is invoked from "Thm.bicompose", will sometimes
-              refuse to unify "?a::?'a" with "?a::?'b" or "?a::nat" and throw a
-              "TERM" exception (with "add_ffpair" as first argument). We then
-              perform unification of the types of variables by hand and try
-              again. We could do this the first time around but this error
-              occurs seldom and we don't want to break existing proofs in subtle
-              ways or slow them down needlessly. *)
-           case [] |> fold (Term.add_vars o prop_of) [tha, thb]
-                   |> AList.group (op =)
-                   |> maps (fn ((s, _), T :: Ts) =>
-                               map (fn T' => (Free (s, T), Free (s, T'))) Ts)
-                   |> rpair (Envir.empty ~1)
-                   |-> fold (Pattern.unify thy)
-                   |> Envir.type_env |> Vartab.dest
-                   |> map (fn (x, (S, T)) =>
-                              pairself (ctyp_of thy) (TVar (x, S), T)) of
-             [] => raise TERM z
-           | ps => aux (instantiate (ps, []) tha) (instantiate (ps, []) thb)
-  end
-
-fun resolve_inf ctxt mode skolems thpairs atm th1 th2 =
-  let
-    val thy = ProofContext.theory_of ctxt
-    val i_th1 = lookth thpairs th1 and i_th2 = lookth thpairs th2
-    val _ = trace_msg (fn () => "  isa th1 (pos): " ^ Display.string_of_thm ctxt i_th1)
-    val _ = trace_msg (fn () => "  isa th2 (neg): " ^ Display.string_of_thm ctxt i_th2)
-  in
-    if is_TrueI i_th1 then i_th2 (*Trivial cases where one operand is type info*)
-    else if is_TrueI i_th2 then i_th1
-    else
-      let
-        val i_atm = singleton (hol_terms_from_fol ctxt mode skolems)
-                              (Metis_Term.Fn atm)
-        val _ = trace_msg (fn () => "  atom: " ^ Syntax.string_of_term ctxt i_atm)
-        val prems_th1 = prems_of i_th1
-        val prems_th2 = prems_of i_th2
-        val index_th1 = get_index (mk_not i_atm) prems_th1
-              handle Empty => raise Fail "Failed to find literal in th1"
-        val _ = trace_msg (fn () => "  index_th1: " ^ Int.toString index_th1)
-        val index_th2 = get_index i_atm prems_th2
-              handle Empty => raise Fail "Failed to find literal in th2"
-        val _ = trace_msg (fn () => "  index_th2: " ^ Int.toString index_th2)
-    in
-      resolve_inc_tyvars thy (Meson.select_literal index_th1 i_th1) index_th2
-                         i_th2
-    end
-  end;
-
-(* INFERENCE RULE: REFL *)
-val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
-val refl_idx = 1 + Thm.maxidx_of REFL_THM;
-
-fun refl_inf ctxt mode skolems t =
-  let val thy = ProofContext.theory_of ctxt
-      val i_t = singleton (hol_terms_from_fol ctxt mode skolems) t
-      val _ = trace_msg (fn () => "  term: " ^ Syntax.string_of_term ctxt i_t)
-      val c_t = cterm_incr_types thy refl_idx i_t
-  in  cterm_instantiate [(refl_x, c_t)] REFL_THM  end;
-
-fun get_ty_arg_size _ (Const (@{const_name HOL.eq}, _)) = 0  (*equality has no type arguments*)
-  | get_ty_arg_size thy (Const (c, _)) = (num_type_args thy c handle TYPE _ => 0)
-  | get_ty_arg_size _ _ = 0;
-
-(* INFERENCE RULE: EQUALITY *)
-fun equality_inf ctxt mode skolems (pos, atm) fp fr =
-  let val thy = ProofContext.theory_of ctxt
-      val m_tm = Metis_Term.Fn atm
-      val [i_atm,i_tm] = hol_terms_from_fol ctxt mode skolems [m_tm, fr]
-      val _ = trace_msg (fn () => "sign of the literal: " ^ Bool.toString pos)
-      fun replace_item_list lx 0 (_::ls) = lx::ls
-        | replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls
-      fun path_finder_FO tm [] = (tm, Term.Bound 0)
-        | path_finder_FO tm (p::ps) =
-            let val (tm1,args) = strip_comb tm
-                val adjustment = get_ty_arg_size thy tm1
-                val p' = if adjustment > p then p else p-adjustment
-                val tm_p = List.nth(args,p')
-                  handle Subscript =>
-                         error ("Cannot replay Metis proof in Isabelle:\n" ^
-                                "equality_inf: " ^ Int.toString p ^ " adj " ^
-                                Int.toString adjustment ^ " term " ^
-                                Syntax.string_of_term ctxt tm)
-                val _ = trace_msg (fn () => "path_finder: " ^ Int.toString p ^
-                                      "  " ^ Syntax.string_of_term ctxt tm_p)
-                val (r,t) = path_finder_FO tm_p ps
-            in
-                (r, list_comb (tm1, replace_item_list t p' args))
-            end
-      fun path_finder_HO tm [] = (tm, Term.Bound 0)
-        | path_finder_HO (t$u) (0::ps) = (fn(x,y) => (x, y$u)) (path_finder_HO t ps)
-        | path_finder_HO (t$u) (_::ps) = (fn(x,y) => (x, t$y)) (path_finder_HO u ps)
-        | path_finder_HO tm ps =
-          raise Fail ("equality_inf, path_finder_HO: path = " ^
-                      space_implode " " (map Int.toString ps) ^
-                      " isa-term: " ^  Syntax.string_of_term ctxt tm)
-      fun path_finder_FT tm [] _ = (tm, Term.Bound 0)
-        | path_finder_FT tm (0::ps) (Metis_Term.Fn ("ti", [t1, _])) =
-            path_finder_FT tm ps t1
-        | path_finder_FT (t$u) (0::ps) (Metis_Term.Fn (".", [t1, _])) =
-            (fn(x,y) => (x, y$u)) (path_finder_FT t ps t1)
-        | path_finder_FT (t$u) (1::ps) (Metis_Term.Fn (".", [_, t2])) =
-            (fn(x,y) => (x, t$y)) (path_finder_FT u ps t2)
-        | path_finder_FT tm ps t =
-          raise Fail ("equality_inf, path_finder_FT: path = " ^
-                      space_implode " " (map Int.toString ps) ^
-                      " isa-term: " ^  Syntax.string_of_term ctxt tm ^
-                      " fol-term: " ^ Metis_Term.toString t)
-      fun path_finder FO tm ps _ = path_finder_FO tm ps
-        | path_finder HO (tm as Const(@{const_name HOL.eq},_) $ _ $ _) (p::ps) _ =
-             (*equality: not curried, as other predicates are*)
-             if p=0 then path_finder_HO tm (0::1::ps)  (*select first operand*)
-             else path_finder_HO tm (p::ps)        (*1 selects second operand*)
-        | path_finder HO tm (_ :: ps) (Metis_Term.Fn ("{}", [_])) =
-             path_finder_HO tm ps      (*if not equality, ignore head to skip hBOOL*)
-        | path_finder FT (tm as Const(@{const_name HOL.eq}, _) $ _ $ _) (p::ps)
-                            (Metis_Term.Fn ("=", [t1,t2])) =
-             (*equality: not curried, as other predicates are*)
-             if p=0 then path_finder_FT tm (0::1::ps)
-                          (Metis_Term.Fn (".", [Metis_Term.Fn (".", [metis_eq,t1]), t2]))
-                          (*select first operand*)
-             else path_finder_FT tm (p::ps)
-                   (Metis_Term.Fn (".", [metis_eq,t2]))
-                   (*1 selects second operand*)
-        | path_finder FT tm (_ :: ps) (Metis_Term.Fn ("{}", [t1])) = path_finder_FT tm ps t1
-             (*if not equality, ignore head to skip the hBOOL predicate*)
-        | path_finder FT tm ps t = path_finder_FT tm ps t  (*really an error case!*)
-      fun path_finder_lit ((nt as Const (@{const_name Not}, _)) $ tm_a) idx =
-            let val (tm, tm_rslt) = path_finder mode tm_a idx m_tm
-            in (tm, nt $ tm_rslt) end
-        | path_finder_lit tm_a idx = path_finder mode tm_a idx m_tm
-      val (tm_subst, body) = path_finder_lit i_atm fp
-      val tm_abs = Term.Abs("x", Term.type_of tm_subst, body)
-      val _ = trace_msg (fn () => "abstraction: " ^ Syntax.string_of_term ctxt tm_abs)
-      val _ = trace_msg (fn () => "i_tm: " ^ Syntax.string_of_term ctxt i_tm)
-      val _ = trace_msg (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst)
-      val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst)  (*ill typed but gives right max*)
-      val subst' = Thm.incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
-      val _ = trace_msg (fn () => "subst' " ^ Display.string_of_thm ctxt subst')
-      val eq_terms = map (pairself (cterm_of thy))
-        (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
-  in  cterm_instantiate eq_terms subst'  end;
-
-val factor = Seq.hd o distinct_subgoals_tac;
-
-fun step ctxt mode skolems thpairs p =
-  case p of
-    (fol_th, Metis_Proof.Axiom _) => factor (axiom_inf thpairs fol_th)
-  | (_, Metis_Proof.Assume f_atm) => assume_inf ctxt mode skolems f_atm
-  | (_, Metis_Proof.Metis_Subst (f_subst, f_th1)) =>
-    factor (inst_inf ctxt mode skolems thpairs f_subst f_th1)
-  | (_, Metis_Proof.Resolve(f_atm, f_th1, f_th2)) =>
-    factor (resolve_inf ctxt mode skolems thpairs f_atm f_th1 f_th2)
-  | (_, Metis_Proof.Refl f_tm) => refl_inf ctxt mode skolems f_tm
-  | (_, Metis_Proof.Equality (f_lit, f_p, f_r)) =>
-    equality_inf ctxt mode skolems f_lit f_p f_r
-
-fun real_literal (_, (c, _)) = not (String.isPrefix class_prefix c);
-
-fun translate_one ctxt mode skolems (fol_th, inf) thpairs =
-  let
-    val _ = trace_msg (fn () => "=============================================")
-    val _ = trace_msg (fn () => "METIS THM: " ^ Metis_Thm.toString fol_th)
-    val _ = trace_msg (fn () => "INFERENCE: " ^ Metis_Proof.inferenceToString inf)
-    val th = Meson.flexflex_first_order (step ctxt mode skolems
-                                              thpairs (fol_th, inf))
-    val _ = trace_msg (fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th)
-    val _ = trace_msg (fn () => "=============================================")
-    val n_metis_lits =
-      length (filter real_literal (Metis_LiteralSet.toList (Metis_Thm.clause fol_th)))
-    val _ = if nprems_of th = n_metis_lits then ()
-            else error "Cannot replay Metis proof in Isabelle."
-  in (fol_th, th) :: thpairs end
+fun have_common_thm ths1 ths2 =
+  exists (member Thm.eq_thm ths1) (map Meson.make_meta_clause ths2)
 
 (*Determining which axiom clauses are actually used*)
 fun used_axioms axioms (th, Metis_Proof.Axiom _) = SOME (lookth axioms th)
   | used_axioms _ _ = NONE;
 
-(* ------------------------------------------------------------------------- *)
-(* Translation of HO Clauses                                                 *)
-(* ------------------------------------------------------------------------- *)
-
-fun type_ext thy tms =
-  let val subs = tfree_classes_of_terms tms
-      val supers = tvar_classes_of_terms tms
-      and tycons = type_consts_of_terms thy tms
-      val (supers', arity_clauses) = make_arity_clauses thy tycons supers
-      val class_rel_clauses = make_class_rel_clauses thy subs supers'
-  in  map class_rel_cls class_rel_clauses @ map arity_cls arity_clauses
-  end;
-
-(* ------------------------------------------------------------------------- *)
-(* Logic maps manage the interface between HOL and first-order logic.        *)
-(* ------------------------------------------------------------------------- *)
-
-type logic_map =
-  {axioms: (Metis_Thm.thm * thm) list,
-   tfrees: type_literal list,
-   skolems: (string * term) list}
-
-fun const_in_metis c (pred, tm_list) =
-  let
-    fun in_mterm (Metis_Term.Var _) = false
-      | in_mterm (Metis_Term.Fn (".", tm_list)) = exists in_mterm tm_list
-      | in_mterm (Metis_Term.Fn (nm, tm_list)) = c=nm orelse exists in_mterm tm_list
-  in  c = pred orelse exists in_mterm tm_list  end;
-
-(*Extract TFree constraints from context to include as conjecture clauses*)
-fun init_tfrees ctxt =
-  let fun add ((a,i),s) Ts = if i = ~1 then TFree(a,s) :: Ts else Ts in
-    Vartab.fold add (#2 (Variable.constraints_of ctxt)) []
-    |> type_literals_for_types
-  end;
-
-(*transform isabelle type / arity clause to metis clause *)
-fun add_type_thm [] lmap = lmap
-  | add_type_thm ((ith, mth) :: cls) {axioms, tfrees, skolems} =
-      add_type_thm cls {axioms = (mth, ith) :: axioms, tfrees = tfrees,
-                        skolems = skolems}
-
-(*Insert non-logical axioms corresponding to all accumulated TFrees*)
-fun add_tfrees {axioms, tfrees, skolems} : logic_map =
-     {axioms = map (rpair TrueI o metis_of_tfree) (distinct (op =) tfrees) @
-               axioms,
-      tfrees = tfrees, skolems = skolems}
-
-fun string_of_mode FO = "FO"
-  | string_of_mode HO = "HO"
-  | string_of_mode FT = "FT"
-
-val helpers =
-  [("c_COMBI", (false, map (`I) @{thms COMBI_def})),
-   ("c_COMBK", (false, map (`I) @{thms COMBK_def})),
-   ("c_COMBB", (false, map (`I) @{thms COMBB_def})),
-   ("c_COMBC", (false, map (`I) @{thms COMBC_def})),
-   ("c_COMBS", (false, map (`I) @{thms COMBS_def})),
-   ("c_fequal", (false, map (rpair @{thm equal_imp_equal})
-                            @{thms fequal_imp_equal equal_imp_fequal})),
-   ("c_True", (true, map (`I) @{thms True_or_False})),
-   ("c_False", (true, map (`I) @{thms True_or_False})),
-   ("c_If", (true, map (`I) @{thms if_True if_False True_or_False}))]
-
-fun is_quasi_fol_clause thy =
-  Meson.is_fol_term thy o snd o conceal_skolem_terms ~1 [] o prop_of
-
-(* Function to generate metis clauses, including comb and type clauses *)
-fun build_map mode0 ctxt cls ths =
-  let val thy = ProofContext.theory_of ctxt
-      (*The modes FO and FT are sticky. HO can be downgraded to FO.*)
-      fun set_mode FO = FO
-        | set_mode HO =
-          if forall (is_quasi_fol_clause thy) (cls @ ths) then FO else HO
-        | set_mode FT = FT
-      val mode = set_mode mode0
-      (*transform isabelle clause to metis clause *)
-      fun add_thm is_conjecture (metis_ith, isa_ith) {axioms, tfrees, skolems}
-                  : logic_map =
-        let
-          val (mth, tfree_lits, skolems) =
-            hol_thm_to_fol is_conjecture ctxt mode (length axioms) skolems
-                           metis_ith
-        in
-           {axioms = (mth, Meson.make_meta_clause isa_ith) :: axioms,
-            tfrees = union (op =) tfree_lits tfrees, skolems = skolems}
-        end;
-      val lmap = {axioms = [], tfrees = init_tfrees ctxt, skolems = []}
-                 |> fold (add_thm true o `I) cls
-                 |> add_tfrees
-                 |> fold (add_thm false o `I) ths
-      val clause_lists = map (Metis_Thm.clause o #1) (#axioms lmap)
-      fun is_used c =
-        exists (Metis_LiteralSet.exists (const_in_metis c o #2)) clause_lists
-      val lmap =
-        if mode = FO then
-          lmap
-        else
-          let
-            val helper_ths =
-              helpers |> filter (is_used o fst)
-                      |> maps (fn (c, (needs_full_types, thms)) =>
-                                  if not (is_used c) orelse
-                                     needs_full_types andalso mode <> FT then
-                                    []
-                                  else
-                                    thms)
-          in lmap |> fold (add_thm false) helper_ths end
-  in (mode, add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap) end
-
 val clause_params =
   {ordering = Metis_KnuthBendixOrder.default,
    orderLiterals = Metis_Clause.UnsignedLiteralOrder,
@@ -749,20 +50,12 @@
    variablesWeight = 0.0,
    literalsWeight = 0.0,
    models = []}
-val refute_params = {active = active_params, waiting = waiting_params}
-
-fun refute cls =
-  Metis_Resolution.new refute_params {axioms = cls, conjecture = []}
-  |> Metis_Resolution.loop
-
-fun is_false t = t aconv (HOLogic.mk_Trueprop HOLogic.false_const);
-
-fun common_thm ths1 ths2 = exists (member Thm.eq_thm ths1) (map Meson.make_meta_clause ths2);
-
+val resolution_params = {active = active_params, waiting = waiting_params}
 
 (* Main function to start Metis proof and reconstruction *)
 fun FOL_SOLVE mode ctxt cls ths0 =
   let val thy = ProofContext.theory_of ctxt
+      val type_lits = Config.get ctxt type_lits
       val th_cls_pairs =
         map (fn th => (Thm.get_name_hint th, Clausifier.cnf_axiom thy th)) ths0
       val ths = maps #2 th_cls_pairs
@@ -770,7 +63,8 @@
       val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) cls
       val _ = trace_msg (fn () => "THEOREM CLAUSES")
       val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) ths
-      val (mode, {axioms, tfrees, skolems}) = build_map mode ctxt cls ths
+      val (mode, {axioms, tfrees, skolems}) =
+        build_logic_map mode ctxt type_lits cls ths
       val _ = if null tfrees then ()
               else (trace_msg (fn () => "TFREE CLAUSES");
                     app (fn TyLitFree ((s, _), (s', _)) =>
@@ -784,21 +78,22 @@
       case filter (is_false o prop_of) cls of
           false_th::_ => [false_th RS @{thm FalseE}]
         | [] =>
-      case refute thms of
+      case Metis_Resolution.new resolution_params {axioms = thms, conjecture = []}
+           |> Metis_Resolution.loop of
           Metis_Resolution.Contradiction mth =>
             let val _ = trace_msg (fn () => "METIS RECONSTRUCTION START: " ^
                           Metis_Thm.toString mth)
                 val ctxt' = fold Variable.declare_constraints (map prop_of cls) ctxt
                              (*add constraints arising from converting goal to clause form*)
                 val proof = Metis_Proof.proof mth
-                val result = fold (translate_one ctxt' mode skolems) proof axioms
+                val result = fold (replay_one_inference ctxt' mode skolems) proof axioms
                 and used = map_filter (used_axioms axioms) proof
                 val _ = trace_msg (fn () => "METIS COMPLETED...clauses actually used:")
                 val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) used
                 val unused = th_cls_pairs |> map_filter (fn (name, cls) =>
-                  if common_thm used cls then NONE else SOME name)
+                  if have_common_thm used cls then NONE else SOME name)
             in
-                if not (null cls) andalso not (common_thm used cls) then
+                if not (null cls) andalso not (have_common_thm used cls) then
                   warning "Metis: The assumptions are inconsistent."
                 else
                   ();
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Sledgehammer/metis_translate.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -0,0 +1,738 @@
+(*  Title:      HOL/Tools/Sledgehammer/metis_translate.ML
+    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
+    Author:     Kong W. Susanto, Cambridge University Computer Laboratory
+    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Translation of HOL to FOL for Metis.
+*)
+
+signature METIS_TRANSLATE =
+sig
+  type name = string * string
+  datatype type_literal =
+    TyLitVar of name * name |
+    TyLitFree of name * name
+  datatype arLit =
+    TConsLit of name * name * name list |
+    TVarLit of name * name
+  datatype arity_clause =
+    ArityClause of {name: string, conclLit: arLit, premLits: arLit list}
+  datatype class_rel_clause =
+    ClassRelClause of {name: string, subclass: name, superclass: name}
+  datatype combtyp =
+    CombTVar of name |
+    CombTFree of name |
+    CombType of name * combtyp list
+  datatype combterm =
+    CombConst of name * combtyp * combtyp list (* Const and Free *) |
+    CombVar of name * combtyp |
+    CombApp of combterm * combterm
+  datatype fol_literal = FOLLiteral of bool * combterm
+
+  datatype mode = FO | HO | FT
+  type logic_map =
+    {axioms: (Metis_Thm.thm * thm) list,
+     tfrees: type_literal list,
+     skolems: (string * term) list}
+
+  val type_wrapper_name : string
+  val bound_var_prefix : string
+  val schematic_var_prefix: string
+  val fixed_var_prefix: string
+  val tvar_prefix: string
+  val tfree_prefix: string
+  val const_prefix: string
+  val type_const_prefix: string
+  val class_prefix: string
+  val invert_const: string -> string
+  val ascii_of: string -> string
+  val unascii_of: string -> string
+  val strip_prefix_and_unascii: string -> string -> string option
+  val make_bound_var : string -> string
+  val make_schematic_var : string * int -> string
+  val make_fixed_var : string -> string
+  val make_schematic_type_var : string * int -> string
+  val make_fixed_type_var : string -> string
+  val make_fixed_const : string -> string
+  val make_fixed_type_const : string -> string
+  val make_type_class : string -> string
+  val num_type_args: theory -> string -> int
+  val type_literals_for_types : typ list -> type_literal list
+  val make_class_rel_clauses :
+    theory -> class list -> class list -> class_rel_clause list
+  val make_arity_clauses :
+    theory -> string list -> class list -> class list * arity_clause list
+  val combtyp_of : combterm -> combtyp
+  val strip_combterm_comb : combterm -> combterm * combterm list
+  val combterm_from_term :
+    theory -> (string * typ) list -> term -> combterm * typ list
+  val reveal_skolem_terms : (string * term) list -> term -> term
+  val tfree_classes_of_terms : term list -> string list
+  val tvar_classes_of_terms : term list -> string list
+  val type_consts_of_terms : theory -> term list -> string list
+  val string_of_mode : mode -> string
+  val build_logic_map :
+    mode -> Proof.context -> bool -> thm list -> thm list -> mode * logic_map
+end
+
+structure Metis_Translate : METIS_TRANSLATE =
+struct
+
+val type_wrapper_name = "ti"
+
+val bound_var_prefix = "B_"
+val schematic_var_prefix = "V_"
+val fixed_var_prefix = "v_"
+
+val tvar_prefix = "T_";
+val tfree_prefix = "t_";
+
+val const_prefix = "c_";
+val type_const_prefix = "tc_";
+val class_prefix = "class_";
+
+fun union_all xss = fold (union (op =)) xss []
+
+(* Readable names for the more common symbolic functions. Do not mess with the
+   last nine entries of the table unless you know what you are doing. *)
+val const_trans_table =
+  Symtab.make [(@{type_name Product_Type.prod}, "prod"),
+               (@{type_name Sum_Type.sum}, "sum"),
+               (@{const_name HOL.eq}, "equal"),
+               (@{const_name HOL.conj}, "and"),
+               (@{const_name HOL.disj}, "or"),
+               (@{const_name HOL.implies}, "implies"),
+               (@{const_name Set.member}, "member"),
+               (@{const_name fequal}, "fequal"),
+               (@{const_name COMBI}, "COMBI"),
+               (@{const_name COMBK}, "COMBK"),
+               (@{const_name COMBB}, "COMBB"),
+               (@{const_name COMBC}, "COMBC"),
+               (@{const_name COMBS}, "COMBS"),
+               (@{const_name True}, "True"),
+               (@{const_name False}, "False"),
+               (@{const_name If}, "If")]
+
+(* Invert the table of translations between Isabelle and ATPs. *)
+val const_trans_table_inv =
+  Symtab.update ("fequal", @{const_name HOL.eq})
+                (Symtab.make (map swap (Symtab.dest const_trans_table)))
+
+val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
+
+(*Escaping of special characters.
+  Alphanumeric characters are left unchanged.
+  The character _ goes to __
+  Characters in the range ASCII space to / go to _A to _P, respectively.
+  Other characters go to _nnn where nnn is the decimal ASCII code.*)
+val A_minus_space = Char.ord #"A" - Char.ord #" ";
+
+fun stringN_of_int 0 _ = ""
+  | stringN_of_int k n = stringN_of_int (k-1) (n div 10) ^ Int.toString (n mod 10);
+
+fun ascii_of_c c =
+  if Char.isAlphaNum c then String.str c
+  else if c = #"_" then "__"
+  else if #" " <= c andalso c <= #"/"
+       then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
+  else ("_" ^ stringN_of_int 3 (Char.ord c))  (*fixed width, in case more digits follow*)
+
+val ascii_of = String.translate ascii_of_c;
+
+(** Remove ASCII armouring from names in proof files **)
+
+(*We don't raise error exceptions because this code can run inside the watcher.
+  Also, the errors are "impossible" (hah!)*)
+fun unascii_aux rcs [] = String.implode(rev rcs)
+  | unascii_aux rcs [#"_"] = unascii_aux (#"_"::rcs) []  (*ERROR*)
+      (*Three types of _ escapes: __, _A to _P, _nnn*)
+  | unascii_aux rcs (#"_" :: #"_" :: cs) = unascii_aux (#"_"::rcs) cs
+  | unascii_aux rcs (#"_" :: c :: cs) =
+      if #"A" <= c andalso c<= #"P"  (*translation of #" " to #"/"*)
+      then unascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
+      else
+        let val digits = List.take (c::cs, 3) handle Subscript => []
+        in
+            case Int.fromString (String.implode digits) of
+                NONE => unascii_aux (c:: #"_"::rcs) cs  (*ERROR*)
+              | SOME n => unascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
+        end
+  | unascii_aux rcs (c::cs) = unascii_aux (c::rcs) cs
+val unascii_of = unascii_aux [] o String.explode
+
+(* If string s has the prefix s1, return the result of deleting it,
+   un-ASCII'd. *)
+fun strip_prefix_and_unascii s1 s =
+  if String.isPrefix s1 s then
+    SOME (unascii_of (String.extract (s, size s1, NONE)))
+  else
+    NONE
+
+(*Remove the initial ' character from a type variable, if it is present*)
+fun trim_type_var s =
+  if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
+  else error ("trim_type: Malformed type variable encountered: " ^ s);
+
+fun ascii_of_indexname (v,0) = ascii_of v
+  | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
+
+fun make_bound_var x = bound_var_prefix ^ ascii_of x
+fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v
+fun make_fixed_var x = fixed_var_prefix ^ ascii_of x
+
+fun make_schematic_type_var (x,i) =
+      tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
+fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
+
+fun lookup_const c =
+  case Symtab.lookup const_trans_table c of
+    SOME c' => c'
+  | NONE => ascii_of c
+
+(* HOL.eq MUST BE "equal" because it's built into ATPs. *)
+fun make_fixed_const @{const_name HOL.eq} = "equal"
+  | make_fixed_const c = const_prefix ^ lookup_const c
+
+fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
+
+fun make_type_class clas = class_prefix ^ ascii_of clas;
+
+val skolem_prefix = "Sledgehammer" ^ Long_Name.separator ^ "Sko"
+
+(* The number of type arguments of a constant, zero if it's monomorphic. For
+   (instances of) Skolem pseudoconstants, this information is encoded in the
+   constant name. *)
+fun num_type_args thy s =
+  if String.isPrefix skolem_prefix s then
+    s |> space_explode Long_Name.separator |> List.last |> Int.fromString |> the
+  else
+    (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
+
+(**** Definitions and functions for FOL clauses for TPTP format output ****)
+
+type name = string * string
+
+(**** Isabelle FOL clauses ****)
+
+(* The first component is the type class; the second is a TVar or TFree. *)
+datatype type_literal =
+  TyLitVar of name * name |
+  TyLitFree of name * name
+
+(*Make literals for sorted type variables*)
+fun sorts_on_typs_aux (_, [])   = []
+  | sorts_on_typs_aux ((x,i),  s::ss) =
+      let val sorts = sorts_on_typs_aux ((x,i), ss)
+      in
+          if s = "HOL.type" then sorts
+          else if i = ~1 then TyLitFree (`make_type_class s, `make_fixed_type_var x) :: sorts
+          else TyLitVar (`make_type_class s, (make_schematic_type_var (x,i), x)) :: sorts
+      end;
+
+fun sorts_on_typs (TFree (a,s)) = sorts_on_typs_aux ((a,~1),s)
+  | sorts_on_typs (TVar (v,s))  = sorts_on_typs_aux (v,s);
+
+(*Given a list of sorted type variables, return a list of type literals.*)
+fun type_literals_for_types Ts =
+  fold (union (op =)) (map sorts_on_typs Ts) []
+
+(** make axiom and conjecture clauses. **)
+
+(**** Isabelle arities ****)
+
+datatype arLit =
+  TConsLit of name * name * name list |
+  TVarLit of name * name
+
+datatype arity_clause =
+  ArityClause of {name: string, conclLit: arLit, premLits: arLit list}
+
+
+fun gen_TVars 0 = []
+  | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
+
+fun pack_sort(_,[])  = []
+  | pack_sort(tvar, "HOL.type"::srt) = pack_sort (tvar, srt)   (*IGNORE sort "type"*)
+  | pack_sort(tvar, cls::srt) =
+    (`make_type_class cls, (tvar, tvar)) :: pack_sort (tvar, srt)
+
+(*Arity of type constructor tcon :: (arg1,...,argN)res*)
+fun make_axiom_arity_clause (tcons, name, (cls,args)) =
+  let
+    val tvars = gen_TVars (length args)
+    val tvars_srts = ListPair.zip (tvars, args)
+  in
+    ArityClause {name = name,
+                 conclLit = TConsLit (`make_type_class cls,
+                                      `make_fixed_type_const tcons,
+                                      tvars ~~ tvars),
+                 premLits = map TVarLit (union_all (map pack_sort tvars_srts))}
+  end
+
+
+(**** Isabelle class relations ****)
+
+datatype class_rel_clause =
+  ClassRelClause of {name: string, subclass: name, superclass: name}
+
+(*Generate all pairs (sub,super) such that sub is a proper subclass of super in theory thy.*)
+fun class_pairs _ [] _ = []
+  | class_pairs thy subs supers =
+      let
+        val class_less = Sorts.class_less (Sign.classes_of thy)
+        fun add_super sub super = class_less (sub, super) ? cons (sub, super)
+        fun add_supers sub = fold (add_super sub) supers
+      in fold add_supers subs [] end
+
+fun make_class_rel_clause (sub,super) =
+  ClassRelClause {name = sub ^ "_" ^ super,
+                  subclass = `make_type_class sub,
+                  superclass = `make_type_class super}
+
+fun make_class_rel_clauses thy subs supers =
+  map make_class_rel_clause (class_pairs thy subs supers);
+
+
+(** Isabelle arities **)
+
+fun arity_clause _ _ (_, []) = []
+  | arity_clause seen n (tcons, ("HOL.type",_)::ars) =  (*ignore*)
+      arity_clause seen n (tcons,ars)
+  | arity_clause seen n (tcons, (ar as (class,_)) :: ars) =
+      if member (op =) seen class then (*multiple arities for the same tycon, class pair*)
+          make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) ::
+          arity_clause seen (n+1) (tcons,ars)
+      else
+          make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class, ar) ::
+          arity_clause (class::seen) n (tcons,ars)
+
+fun multi_arity_clause [] = []
+  | multi_arity_clause ((tcons, ars) :: tc_arlists) =
+      arity_clause [] 1 (tcons, ars) @ multi_arity_clause tc_arlists
+
+(*Generate all pairs (tycon,class,sorts) such that tycon belongs to class in theory thy
+  provided its arguments have the corresponding sorts.*)
+fun type_class_pairs thy tycons classes =
+  let val alg = Sign.classes_of thy
+      fun domain_sorts tycon = Sorts.mg_domain alg tycon o single
+      fun add_class tycon class =
+        cons (class, domain_sorts tycon class)
+        handle Sorts.CLASS_ERROR _ => I
+      fun try_classes tycon = (tycon, fold (add_class tycon) classes [])
+  in  map try_classes tycons  end;
+
+(*Proving one (tycon, class) membership may require proving others, so iterate.*)
+fun iter_type_class_pairs _ _ [] = ([], [])
+  | iter_type_class_pairs thy tycons classes =
+      let val cpairs = type_class_pairs thy tycons classes
+          val newclasses = union_all (union_all (union_all (map (map #2 o #2) cpairs)))
+            |> subtract (op =) classes |> subtract (op =) HOLogic.typeS
+          val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
+      in (union (op =) classes' classes, union (op =) cpairs' cpairs) end;
+
+fun make_arity_clauses thy tycons classes =
+  let val (classes', cpairs) = iter_type_class_pairs thy tycons classes
+  in  (classes', multi_arity_clause cpairs)  end;
+
+datatype combtyp =
+  CombTVar of name |
+  CombTFree of name |
+  CombType of name * combtyp list
+
+datatype combterm =
+  CombConst of name * combtyp * combtyp list (* Const and Free *) |
+  CombVar of name * combtyp |
+  CombApp of combterm * combterm
+
+datatype fol_literal = FOLLiteral of bool * combterm
+
+(*********************************************************************)
+(* convert a clause with type Term.term to a clause with type clause *)
+(*********************************************************************)
+
+(*Result of a function type; no need to check that the argument type matches.*)
+fun result_type (CombType (_, [_, tp2])) = tp2
+  | result_type _ = raise Fail "non-function type"
+
+fun combtyp_of (CombConst (_, tp, _)) = tp
+  | combtyp_of (CombVar (_, tp)) = tp
+  | combtyp_of (CombApp (t1, _)) = result_type (combtyp_of t1)
+
+(*gets the head of a combinator application, along with the list of arguments*)
+fun strip_combterm_comb u =
+    let fun stripc (CombApp(t,u), ts) = stripc (t, u::ts)
+        |   stripc  x =  x
+    in stripc(u,[]) end
+
+fun combtype_of (Type (a, Ts)) =
+    let val (folTypes, ts) = combtypes_of Ts in
+      (CombType (`make_fixed_type_const a, folTypes), ts)
+    end
+  | combtype_of (tp as TFree (a, _)) = (CombTFree (`make_fixed_type_var a), [tp])
+  | combtype_of (tp as TVar (x, _)) =
+    (CombTVar (make_schematic_type_var x, string_of_indexname x), [tp])
+and combtypes_of Ts =
+  let val (folTyps, ts) = ListPair.unzip (map combtype_of Ts) in
+    (folTyps, union_all ts)
+  end
+
+(* same as above, but no gathering of sort information *)
+fun simple_combtype_of (Type (a, Ts)) =
+    CombType (`make_fixed_type_const a, map simple_combtype_of Ts)
+  | simple_combtype_of (TFree (a, _)) = CombTFree (`make_fixed_type_var a)
+  | simple_combtype_of (TVar (x, _)) =
+    CombTVar (make_schematic_type_var x, string_of_indexname x)
+
+(* Converts a term (with combinators) into a combterm. Also accummulates sort
+   infomation. *)
+fun combterm_from_term thy bs (P $ Q) =
+      let val (P', tsP) = combterm_from_term thy bs P
+          val (Q', tsQ) = combterm_from_term thy bs Q
+      in  (CombApp (P', Q'), union (op =) tsP tsQ)  end
+  | combterm_from_term thy _ (Const (c, T)) =
+      let
+        val (tp, ts) = combtype_of T
+        val tvar_list =
+          (if String.isPrefix skolem_prefix c then
+             [] |> Term.add_tvarsT T |> map TVar
+           else
+             (c, T) |> Sign.const_typargs thy)
+          |> map simple_combtype_of
+        val c' = CombConst (`make_fixed_const c, tp, tvar_list)
+      in  (c',ts)  end
+  | combterm_from_term _ _ (Free (v, T)) =
+      let val (tp, ts) = combtype_of T
+          val v' = CombConst (`make_fixed_var v, tp, [])
+      in  (v',ts)  end
+  | combterm_from_term _ _ (Var (v, T)) =
+      let val (tp,ts) = combtype_of T
+          val v' = CombVar ((make_schematic_var v, string_of_indexname v), tp)
+      in  (v',ts)  end
+  | combterm_from_term _ bs (Bound j) =
+      let
+        val (s, T) = nth bs j
+        val (tp, ts) = combtype_of T
+        val v' = CombConst (`make_bound_var s, tp, [])
+      in (v', ts) end
+  | combterm_from_term _ _ (Abs _) = raise Fail "HOL clause: Abs"
+
+fun predicate_of thy ((@{const Not} $ P), pos) = predicate_of thy (P, not pos)
+  | predicate_of thy (t, pos) =
+    (combterm_from_term thy [] (Envir.eta_contract t), pos)
+
+fun literals_of_term1 args thy (@{const Trueprop} $ P) =
+    literals_of_term1 args thy P
+  | literals_of_term1 args thy (@{const HOL.disj} $ P $ Q) =
+    literals_of_term1 (literals_of_term1 args thy P) thy Q
+  | literals_of_term1 (lits, ts) thy P =
+    let val ((pred, ts'), pol) = predicate_of thy (P, true) in
+      (FOLLiteral (pol, pred) :: lits, union (op =) ts ts')
+    end
+val literals_of_term = literals_of_term1 ([], [])
+
+fun skolem_name i j num_T_args =
+  skolem_prefix ^ Long_Name.separator ^
+  (space_implode Long_Name.separator (map Int.toString [i, j, num_T_args]))
+
+fun conceal_skolem_terms i skolems t =
+  if exists_Const (curry (op =) @{const_name skolem} o fst) t then
+    let
+      fun aux skolems
+              (t as (Const (@{const_name skolem}, Type (_, [_, T])) $ _)) =
+          let
+            val (skolems, s) =
+              if i = ~1 then
+                (skolems, @{const_name undefined})
+              else case AList.find (op aconv) skolems t of
+                s :: _ => (skolems, s)
+              | [] =>
+                let
+                  val s = skolem_name i (length skolems)
+                                      (length (Term.add_tvarsT T []))
+                in ((s, t) :: skolems, s) end
+          in (skolems, Const (s, T)) end
+        | aux skolems (t1 $ t2) =
+          let
+            val (skolems, t1) = aux skolems t1
+            val (skolems, t2) = aux skolems t2
+          in (skolems, t1 $ t2) end
+        | aux skolems (Abs (s, T, t')) =
+          let val (skolems, t') = aux skolems t' in
+            (skolems, Abs (s, T, t'))
+          end
+        | aux skolems t = (skolems, t)
+    in aux skolems t end
+  else
+    (skolems, t)
+
+fun reveal_skolem_terms skolems =
+  map_aterms (fn t as Const (s, _) =>
+                 if String.isPrefix skolem_prefix s then
+                   AList.lookup (op =) skolems s |> the
+                   |> map_types Type_Infer.paramify_vars
+                 else
+                   t
+               | t => t)
+
+
+(***************************************************************)
+(* Type Classes Present in the Axiom or Conjecture Clauses     *)
+(***************************************************************)
+
+fun set_insert (x, s) = Symtab.update (x, ()) s
+
+fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
+
+(*Remove this trivial type class*)
+fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset;
+
+fun tfree_classes_of_terms ts =
+  let val sorts_list = map (map #2 o OldTerm.term_tfrees) ts
+  in  Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list))  end;
+
+fun tvar_classes_of_terms ts =
+  let val sorts_list = map (map #2 o OldTerm.term_tvars) ts
+  in  Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list))  end;
+
+(*fold type constructors*)
+fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x))
+  | fold_type_consts _ _ x = x;
+
+(*Type constructors used to instantiate overloaded constants are the only ones needed.*)
+fun add_type_consts_in_term thy =
+  let
+    fun aux (Const x) =
+        fold (fold_type_consts set_insert) (Sign.const_typargs thy x)
+      | aux (Abs (_, _, u)) = aux u
+      | aux (Const (@{const_name skolem}, _) $ _) = I
+      | aux (t $ u) = aux t #> aux u
+      | aux _ = I
+  in aux end
+
+fun type_consts_of_terms thy ts =
+  Symtab.keys (fold (add_type_consts_in_term thy) ts Symtab.empty);
+
+(* ------------------------------------------------------------------------- *)
+(* HOL to FOL  (Isabelle to Metis)                                           *)
+(* ------------------------------------------------------------------------- *)
+
+datatype mode = FO | HO | FT  (* first-order, higher-order, fully-typed *)
+
+fun string_of_mode FO = "FO"
+  | string_of_mode HO = "HO"
+  | string_of_mode FT = "FT"
+
+fun fn_isa_to_met_sublevel "equal" = "=" (* FIXME: "c_fequal" *)
+  | fn_isa_to_met_sublevel x = x
+fun fn_isa_to_met_toplevel "equal" = "="
+  | fn_isa_to_met_toplevel x = x
+
+fun metis_lit b c args = (b, (c, args));
+
+fun metis_term_from_combtyp (CombTVar (s, _)) = Metis_Term.Var s
+  | metis_term_from_combtyp (CombTFree (s, _)) = Metis_Term.Fn (s, [])
+  | metis_term_from_combtyp (CombType ((s, _), tps)) =
+    Metis_Term.Fn (s, map metis_term_from_combtyp tps);
+
+(*These two functions insert type literals before the real literals. That is the
+  opposite order from TPTP linkup, but maybe OK.*)
+
+fun hol_term_to_fol_FO tm =
+  case strip_combterm_comb tm of
+      (CombConst ((c, _), _, tys), tms) =>
+        let val tyargs = map metis_term_from_combtyp tys
+            val args   = map hol_term_to_fol_FO tms
+        in Metis_Term.Fn (c, tyargs @ args) end
+    | (CombVar ((v, _), _), []) => Metis_Term.Var v
+    | _ => raise Fail "non-first-order combterm"
+
+fun hol_term_to_fol_HO (CombConst ((a, _), _, tylist)) =
+      Metis_Term.Fn (fn_isa_to_met_sublevel a, map metis_term_from_combtyp tylist)
+  | hol_term_to_fol_HO (CombVar ((s, _), _)) = Metis_Term.Var s
+  | hol_term_to_fol_HO (CombApp (tm1, tm2)) =
+       Metis_Term.Fn (".", map hol_term_to_fol_HO [tm1, tm2]);
+
+(*The fully-typed translation, to avoid type errors*)
+fun wrap_type (tm, ty) = Metis_Term.Fn("ti", [tm, metis_term_from_combtyp ty]);
+
+fun hol_term_to_fol_FT (CombVar ((s, _), ty)) = wrap_type (Metis_Term.Var s, ty)
+  | hol_term_to_fol_FT (CombConst((a, _), ty, _)) =
+      wrap_type (Metis_Term.Fn(fn_isa_to_met_sublevel a, []), ty)
+  | hol_term_to_fol_FT (tm as CombApp(tm1,tm2)) =
+       wrap_type (Metis_Term.Fn(".", map hol_term_to_fol_FT [tm1,tm2]),
+                  combtyp_of tm)
+
+fun hol_literal_to_fol FO (FOLLiteral (pos, tm)) =
+      let val (CombConst((p, _), _, tys), tms) = strip_combterm_comb tm
+          val tylits = if p = "equal" then [] else map metis_term_from_combtyp tys
+          val lits = map hol_term_to_fol_FO tms
+      in metis_lit pos (fn_isa_to_met_toplevel p) (tylits @ lits) end
+  | hol_literal_to_fol HO (FOLLiteral (pos, tm)) =
+     (case strip_combterm_comb tm of
+          (CombConst(("equal", _), _, _), tms) =>
+            metis_lit pos "=" (map hol_term_to_fol_HO tms)
+        | _ => metis_lit pos "{}" [hol_term_to_fol_HO tm])   (*hBOOL*)
+  | hol_literal_to_fol FT (FOLLiteral (pos, tm)) =
+     (case strip_combterm_comb tm of
+          (CombConst(("equal", _), _, _), tms) =>
+            metis_lit pos "=" (map hol_term_to_fol_FT tms)
+        | _ => metis_lit pos "{}" [hol_term_to_fol_FT tm])   (*hBOOL*);
+
+fun literals_of_hol_term thy mode t =
+      let val (lits, types_sorts) = literals_of_term thy t
+      in  (map (hol_literal_to_fol mode) lits, types_sorts) end;
+
+(*Sign should be "true" for conjecture type constraints, "false" for type lits in clauses.*)
+fun metis_of_type_literals pos (TyLitVar ((s, _), (s', _))) =
+    metis_lit pos s [Metis_Term.Var s']
+  | metis_of_type_literals pos (TyLitFree ((s, _), (s', _))) =
+    metis_lit pos s [Metis_Term.Fn (s',[])]
+
+fun default_sort _ (TVar _) = false
+  | default_sort ctxt (TFree (x, s)) = (s = the_default [] (Variable.def_sort ctxt (x, ~1)));
+
+fun metis_of_tfree tf =
+  Metis_Thm.axiom (Metis_LiteralSet.singleton (metis_of_type_literals true tf));
+
+fun hol_thm_to_fol is_conjecture ctxt type_lits mode j skolems th =
+  let
+    val thy = ProofContext.theory_of ctxt
+    val (skolems, (mlits, types_sorts)) =
+     th |> prop_of |> conceal_skolem_terms j skolems
+        ||> (HOLogic.dest_Trueprop #> literals_of_hol_term thy mode)
+  in
+    if is_conjecture then
+      (Metis_Thm.axiom (Metis_LiteralSet.fromList mlits),
+       type_literals_for_types types_sorts, skolems)
+    else
+      let
+        val tylits = filter_out (default_sort ctxt) types_sorts
+                     |> type_literals_for_types
+        val mtylits =
+          if type_lits then map (metis_of_type_literals false) tylits else []
+      in
+        (Metis_Thm.axiom (Metis_LiteralSet.fromList(mtylits @ mlits)), [],
+         skolems)
+      end
+  end;
+
+val helpers =
+  [("c_COMBI", (false, map (`I) @{thms COMBI_def})),
+   ("c_COMBK", (false, map (`I) @{thms COMBK_def})),
+   ("c_COMBB", (false, map (`I) @{thms COMBB_def})),
+   ("c_COMBC", (false, map (`I) @{thms COMBC_def})),
+   ("c_COMBS", (false, map (`I) @{thms COMBS_def})),
+   ("c_fequal", (false, map (rpair @{thm equal_imp_equal})
+                            @{thms fequal_imp_equal equal_imp_fequal})),
+   ("c_True", (true, map (`I) @{thms True_or_False})),
+   ("c_False", (true, map (`I) @{thms True_or_False})),
+   ("c_If", (true, map (`I) @{thms if_True if_False True_or_False}))]
+
+(* ------------------------------------------------------------------------- *)
+(* Logic maps manage the interface between HOL and first-order logic.        *)
+(* ------------------------------------------------------------------------- *)
+
+type logic_map =
+  {axioms: (Metis_Thm.thm * thm) list,
+   tfrees: type_literal list,
+   skolems: (string * term) list}
+
+fun is_quasi_fol_clause thy =
+  Meson.is_fol_term thy o snd o conceal_skolem_terms ~1 [] o prop_of
+
+(*Extract TFree constraints from context to include as conjecture clauses*)
+fun init_tfrees ctxt =
+  let fun add ((a,i),s) Ts = if i = ~1 then TFree(a,s) :: Ts else Ts in
+    Vartab.fold add (#2 (Variable.constraints_of ctxt)) []
+    |> type_literals_for_types
+  end;
+
+(*Insert non-logical axioms corresponding to all accumulated TFrees*)
+fun add_tfrees {axioms, tfrees, skolems} : logic_map =
+     {axioms = map (rpair TrueI o metis_of_tfree) (distinct (op =) tfrees) @
+               axioms,
+      tfrees = tfrees, skolems = skolems}
+
+(*transform isabelle type / arity clause to metis clause *)
+fun add_type_thm [] lmap = lmap
+  | add_type_thm ((ith, mth) :: cls) {axioms, tfrees, skolems} =
+      add_type_thm cls {axioms = (mth, ith) :: axioms, tfrees = tfrees,
+                        skolems = skolems}
+
+fun const_in_metis c (pred, tm_list) =
+  let
+    fun in_mterm (Metis_Term.Var _) = false
+      | in_mterm (Metis_Term.Fn (".", tm_list)) = exists in_mterm tm_list
+      | in_mterm (Metis_Term.Fn (nm, tm_list)) = c=nm orelse exists in_mterm tm_list
+  in  c = pred orelse exists in_mterm tm_list  end;
+
+(* ARITY CLAUSE *)
+fun m_arity_cls (TConsLit ((c, _), (t, _), args)) =
+    metis_lit true c [Metis_Term.Fn(t, map (Metis_Term.Var o fst) args)]
+  | m_arity_cls (TVarLit ((c, _), (s, _))) =
+    metis_lit false c [Metis_Term.Var s]
+(*TrueI is returned as the Isabelle counterpart because there isn't any.*)
+fun arity_cls (ArityClause {conclLit, premLits, ...}) =
+  (TrueI,
+   Metis_Thm.axiom (Metis_LiteralSet.fromList (map m_arity_cls (conclLit :: premLits))));
+
+(* CLASSREL CLAUSE *)
+fun m_class_rel_cls (subclass, _) (superclass, _) =
+  [metis_lit false subclass [Metis_Term.Var "T"], metis_lit true superclass [Metis_Term.Var "T"]];
+fun class_rel_cls (ClassRelClause {subclass, superclass, ...}) =
+  (TrueI, Metis_Thm.axiom (Metis_LiteralSet.fromList (m_class_rel_cls subclass superclass)));
+
+fun type_ext thy tms =
+  let val subs = tfree_classes_of_terms tms
+      val supers = tvar_classes_of_terms tms
+      and tycons = type_consts_of_terms thy tms
+      val (supers', arity_clauses) = make_arity_clauses thy tycons supers
+      val class_rel_clauses = make_class_rel_clauses thy subs supers'
+  in  map class_rel_cls class_rel_clauses @ map arity_cls arity_clauses
+  end;
+
+(* Function to generate metis clauses, including comb and type clauses *)
+fun build_logic_map mode0 ctxt type_lits cls ths =
+  let val thy = ProofContext.theory_of ctxt
+      (*The modes FO and FT are sticky. HO can be downgraded to FO.*)
+      fun set_mode FO = FO
+        | set_mode HO =
+          if forall (is_quasi_fol_clause thy) (cls @ ths) then FO else HO
+        | set_mode FT = FT
+      val mode = set_mode mode0
+      (*transform isabelle clause to metis clause *)
+      fun add_thm is_conjecture (metis_ith, isa_ith) {axioms, tfrees, skolems}
+                  : logic_map =
+        let
+          val (mth, tfree_lits, skolems) =
+            hol_thm_to_fol is_conjecture ctxt type_lits mode (length axioms)
+                           skolems metis_ith
+        in
+           {axioms = (mth, Meson.make_meta_clause isa_ith) :: axioms,
+            tfrees = union (op =) tfree_lits tfrees, skolems = skolems}
+        end;
+      val lmap = {axioms = [], tfrees = init_tfrees ctxt, skolems = []}
+                 |> fold (add_thm true o `I) cls
+                 |> add_tfrees
+                 |> fold (add_thm false o `I) ths
+      val clause_lists = map (Metis_Thm.clause o #1) (#axioms lmap)
+      fun is_used c =
+        exists (Metis_LiteralSet.exists (const_in_metis c o #2)) clause_lists
+      val lmap =
+        if mode = FO then
+          lmap
+        else
+          let
+            val helper_ths =
+              helpers |> filter (is_used o fst)
+                      |> maps (fn (c, (needs_full_types, thms)) =>
+                                  if not (is_used c) orelse
+                                     needs_full_types andalso mode <> FT then
+                                    []
+                                  else
+                                    thms)
+          in lmap |> fold (add_thm false) helper_ths end
+  in (mode, add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap) end
+
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -13,6 +13,7 @@
   type relevance_override = Sledgehammer_Filter.relevance_override
   type fol_formula = Sledgehammer_Translate.fol_formula
   type minimize_command = Sledgehammer_Reconstruct.minimize_command
+
   type params =
     {blocking: bool,
      debug: bool,
@@ -27,12 +28,14 @@
      isar_shrink_factor: int,
      timeout: Time.time,
      expect: string}
+
   type problem =
     {state: Proof.state,
      goal: thm,
      subgoal: int,
      axioms: (term * ((string * locality) * fol_formula) option) list,
      only: bool}
+
   type prover_result =
     {outcome: failure option,
      message: string,
@@ -43,6 +46,7 @@
      tstplike_proof: string,
      axiom_names: (string * locality) list vector,
      conjecture_shape: int list list}
+
   type prover = params -> minimize_command -> problem -> prover_result
 
   val dest_dir : string Config.T
@@ -62,8 +66,9 @@
 struct
 
 open ATP_Problem
+open ATP_Proof
 open ATP_Systems
-open Metis_Clauses
+open Metis_Translate
 open Sledgehammer_Util
 open Sledgehammer_Filter
 open Sledgehammer_Translate
@@ -135,102 +140,11 @@
   |> Exn.release
   |> tap (after path)
 
-fun extract_delimited (begin_delim, end_delim) output =
-  output |> first_field begin_delim |> the |> snd
-         |> first_field end_delim |> the |> fst
-         |> first_field "\n" |> the |> snd
-  handle Option.Option => ""
-
-val tstp_important_message_delims =
-  ("% SZS start RequiredInformation", "% SZS end RequiredInformation")
-
-fun extract_important_message output =
-  case extract_delimited tstp_important_message_delims output of
-    "" => ""
-  | s => s |> space_explode "\n" |> filter_out (curry (op =) "")
-           |> map (perhaps (try (unprefix "%")))
-           |> map (perhaps (try (unprefix " ")))
-           |> space_implode "\n " |> quote
-
-(* Splits by the first possible of a list of delimiters. *)
-fun extract_tstplike_proof delims output =
-  case pairself (find_first (fn s => String.isSubstring s output))
-                (ListPair.unzip delims) of
-    (SOME begin_delim, SOME end_delim) =>
-    extract_delimited (begin_delim, end_delim) output
-  | _ => ""
-
-fun extract_tstplike_proof_and_outcome complete res_code proof_delims
-                                       known_failures output =
-  case known_failure_in_output output known_failures of
-    NONE => (case extract_tstplike_proof proof_delims output of
-             "" => ("", SOME (if res_code = 0 andalso output = "" then
-                                Interrupted
-                              else
-                                 UnknownError))
-           | tstplike_proof => if res_code = 0 then (tstplike_proof, NONE)
-                               else ("", SOME UnknownError))
-  | SOME failure =>
-    ("", SOME (if failure = IncompleteUnprovable andalso complete then
-                 Unprovable
-               else
-                 failure))
+(* generic TPTP-based provers *)
 
-fun extract_clause_sequence output =
-  let
-    val tokens_of = String.tokens (not o Char.isAlphaNum)
-    fun extract_num ("clause" :: (ss as _ :: _)) =
-    Int.fromString (List.last ss)
-      | extract_num _ = NONE
-  in output |> split_lines |> map_filter (extract_num o tokens_of) end
-
-val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation"
-
-val parse_clause_formula_pair =
-  $$ "(" |-- scan_integer --| $$ ","
-  -- (Symbol.scan_id ::: Scan.repeat ($$ "," |-- Symbol.scan_id)) --| $$ ")"
-  --| Scan.option ($$ ",")
-val parse_clause_formula_relation =
-  Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
-  |-- Scan.repeat parse_clause_formula_pair
-val extract_clause_formula_relation =
-  Substring.full #> Substring.position set_ClauseFormulaRelationN
-  #> snd #> Substring.position "." #> fst #> Substring.string
-  #> explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation
-  #> fst
-
-(* TODO: move to "Sledgehammer_Reconstruct" *)
-fun repair_conjecture_shape_and_theorem_names output conjecture_shape
-                                              axiom_names =
-  if String.isSubstring set_ClauseFormulaRelationN output then
-    (* This is a hack required for keeping track of axioms after they have been
-       clausified by SPASS's Flotter tool. The "ATP/scripts/spass" script is
-       also part of this hack. *)
-    let
-      val j0 = hd (hd conjecture_shape)
-      val seq = extract_clause_sequence output
-      val name_map = extract_clause_formula_relation output
-      fun renumber_conjecture j =
-        conjecture_prefix ^ string_of_int (j - j0)
-        |> AList.find (fn (s, ss) => member (op =) ss s) name_map
-        |> map (fn s => find_index (curry (op =) s) seq + 1)
-      fun names_for_number j =
-        j |> AList.lookup (op =) name_map |> these
-          |> map_filter (try (unprefix axiom_prefix)) |> map unascii_of
-          |> map (fn name =>
-                     (name, name |> find_first_in_list_vector axiom_names
-                                 |> the)
-                     handle Option.Option =>
-                            error ("No such fact: " ^ quote name ^ "."))
-    in
-      (conjecture_shape |> map (maps renumber_conjecture),
-       seq |> map names_for_number |> Vector.fromList)
-    end
-  else
-    (conjecture_shape, axiom_names)
-
-
-(* generic TPTP-based provers *)
+(* Important messages are important but not so important that users want to see
+   them each time. *)
+val important_message_keep_factor = 0.1
 
 fun prover_fun auto atp_name
         {exec, required_execs, arguments, has_incomplete_mode, proof_delims,
@@ -342,9 +256,13 @@
          (output, msecs, tstplike_proof, outcome)) =
       with_path cleanup export run_on problem_path_name
     val (conjecture_shape, axiom_names) =
-      repair_conjecture_shape_and_theorem_names output conjecture_shape
-                                                axiom_names
-    val important_message = extract_important_message output
+      repair_conjecture_shape_and_axiom_names output conjecture_shape
+                                              axiom_names
+    val important_message =
+      if random () <= important_message_keep_factor then
+        extract_important_message output
+      else
+        ""
     val banner = if auto then "Sledgehammer found a proof"
                  else "Try this command"
     val (message, used_thm_names) =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -20,7 +20,7 @@
 structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE =
 struct
 
-open ATP_Systems
+open ATP_Proof
 open Sledgehammer_Util
 open Sledgehammer_Filter
 open Sledgehammer_Translate
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -3,7 +3,7 @@
     Author:     Claire Quigley, Cambridge University Computer Laboratory
     Author:     Jasmin Blanchette, TU Muenchen
 
-Transfer of proofs from external provers.
+Proof reconstruction for Sledgehammer.
 *)
 
 signature SLEDGEHAMMER_RECONSTRUCT =
@@ -17,6 +17,9 @@
     string Symtab.table * bool * int * Proof.context * int list list
   type text_result = string * (string * locality) list
 
+  val repair_conjecture_shape_and_axiom_names :
+    string -> int list list -> (string * locality) list vector
+    -> int list list * (string * locality) list vector
   val metis_proof_text : metis_params -> text_result
   val isar_proof_text : isar_params -> metis_params -> text_result
   val proof_text : bool -> isar_params -> metis_params -> text_result
@@ -27,7 +30,7 @@
 
 open ATP_Problem
 open ATP_Proof
-open Metis_Clauses
+open Metis_Translate
 open Sledgehammer_Util
 open Sledgehammer_Filter
 open Sledgehammer_Translate
@@ -40,6 +43,69 @@
   string Symtab.table * bool * int * Proof.context * int list list
 type text_result = string * (string * locality) list
 
+fun is_head_digit s = Char.isDigit (String.sub (s, 0))
+val scan_integer = Scan.many1 is_head_digit >> (the o Int.fromString o implode)
+
+fun find_first_in_list_vector vec key =
+  Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
+                 | (_, value) => value) NONE vec
+
+
+(** SPASS's Flotter hack **)
+
+(* This is a hack required for keeping track of axioms after they have been
+   clausified by SPASS's Flotter tool. The "ATP/scripts/spass" script is also
+   part of this hack. *)
+
+val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation"
+
+fun extract_clause_sequence output =
+  let
+    val tokens_of = String.tokens (not o Char.isAlphaNum)
+    fun extract_num ("clause" :: (ss as _ :: _)) =
+    Int.fromString (List.last ss)
+      | extract_num _ = NONE
+  in output |> split_lines |> map_filter (extract_num o tokens_of) end
+
+val parse_clause_formula_pair =
+  $$ "(" |-- scan_integer --| $$ ","
+  -- (Symbol.scan_id ::: Scan.repeat ($$ "," |-- Symbol.scan_id)) --| $$ ")"
+  --| Scan.option ($$ ",")
+val parse_clause_formula_relation =
+  Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
+  |-- Scan.repeat parse_clause_formula_pair
+val extract_clause_formula_relation =
+  Substring.full #> Substring.position set_ClauseFormulaRelationN
+  #> snd #> Substring.position "." #> fst #> Substring.string
+  #> explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation
+  #> fst
+
+fun repair_conjecture_shape_and_axiom_names output conjecture_shape
+                                            axiom_names =
+  if String.isSubstring set_ClauseFormulaRelationN output then
+    let
+      val j0 = hd (hd conjecture_shape)
+      val seq = extract_clause_sequence output
+      val name_map = extract_clause_formula_relation output
+      fun renumber_conjecture j =
+        conjecture_prefix ^ string_of_int (j - j0)
+        |> AList.find (fn (s, ss) => member (op =) ss s) name_map
+        |> map (fn s => find_index (curry (op =) s) seq + 1)
+      fun names_for_number j =
+        j |> AList.lookup (op =) name_map |> these
+          |> map_filter (try (unprefix axiom_prefix)) |> map unascii_of
+          |> map (fn name =>
+                     (name, name |> find_first_in_list_vector axiom_names
+                                 |> the)
+                     handle Option.Option =>
+                            error ("No such fact: " ^ quote name ^ "."))
+    in
+      (conjecture_shape |> map (maps renumber_conjecture),
+       seq |> map names_for_number |> Vector.fromList)
+    end
+  else
+    (conjecture_shape, axiom_names)
+
 
 (** Soft-core proof reconstruction: Metis one-liner **)
 
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -3,7 +3,7 @@
     Author:     Makarius
     Author:     Jasmin Blanchette, TU Muenchen
 
-Translation of HOL to FOL.
+Translation of HOL to FOL for Sledgehammer.
 *)
 
 signature SLEDGEHAMMER_TRANSLATE =
@@ -30,7 +30,7 @@
 struct
 
 open ATP_Problem
-open Metis_Clauses
+open Metis_Translate
 open Sledgehammer_Util
 
 val axiom_prefix = "ax_"
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -6,13 +6,11 @@
 
 signature SLEDGEHAMMER_UTIL =
 sig
-  val find_first_in_list_vector : (''a * 'b) list vector -> ''a -> 'b option
   val plural_s : int -> string
   val serial_commas : string -> string list -> string list
   val simplify_spaces : string -> string
   val parse_bool_option : bool -> string -> string -> bool option
   val parse_time_option : string -> string -> Time.time option
-  val scan_integer : string list -> int * string list
   val nat_subscript : int -> string
   val unyxml : string -> string
   val maybe_quote : string -> string
@@ -28,10 +26,6 @@
 structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
 struct
 
-fun find_first_in_list_vector vec key =
-  Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
-                 | (_, value) => value) NONE vec
-
 fun plural_s n = if n = 1 then "" else "s"
 
 fun serial_commas _ [] = ["??"]
@@ -72,9 +66,6 @@
         SOME (Time.fromMilliseconds msecs)
     end
 
-fun is_head_digit s = Char.isDigit (String.sub (s, 0))
-val scan_integer = Scan.many1 is_head_digit >> (the o Int.fromString o implode)
-
 val subscript = implode o map (prefix "\<^isub>") o explode
 fun nat_subscript n =
   n |> string_of_int |> print_mode_active Symbol.xsymbolsN ? subscript
--- a/src/Tools/Metis/README	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/Tools/Metis/README	Fri Sep 17 16:38:11 2010 +0200
@@ -8,7 +8,7 @@
  1. The files "Makefile" and "script/mlpp" and the directory "src/"
     were initially copied from Joe Hurd's official Metis package. The
     package that was used when these notes where written was Metis 2.3
-    (15 Sept. 2010).
+    (16 Sept. 2010).
 
  2. The license in each source file will probably not be something we
     can use in Isabelle. The "fix_metis_license" script can be run to
@@ -20,11 +20,11 @@
         as part of Isabelle, with the Metis code covered under the
         Isabelle BSD license.
 
- 3. Some modifications have to be done manually to the source files.
-    The ultimate way to track them down is to use Mercurial. The
-    command
+ 3. Some modifications might have to be done manually to the source
+    files. The ultimate way to track them down is to use Mercurial.
+    The command
 
-        hg diff -rbeabb8443ee4: src
+        hg diff -rcffceed8e7fa: src
 
     should do the trick. You might need to specify a different
     revision number if somebody updated the Metis sources without
@@ -55,4 +55,4 @@
 Good luck!
 
     Jasmin Blanchette
-    16 Sept. 2010
+    17 Sept. 2010
--- a/src/Tools/Metis/metis.ML	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/Tools/Metis/metis.ML	Fri Sep 17 16:38:11 2010 +0200
@@ -967,49 +967,32 @@
 end;
 
 local
-  fun calcPrimes ps n i =
-      if List.exists (fn p => divides p i) ps then calcPrimes ps n (i + 1)
+  fun calcPrimes mode ps i n =
+      if n = 0 then ps
+      else if List.exists (fn p => divides p i) ps then
+        let
+          val i = i + 1
+          and n = if mode then n else n - 1
+        in
+          calcPrimes mode ps i n
+        end
       else
         let
           val ps = ps @ [i]
+          and i = i + 1
           and n = n - 1
         in
-          if n = 0 then ps else calcPrimes ps n (i + 1)
+          calcPrimes mode ps i n
         end;
-
-  val primesList = Unsynchronized.ref [2];
-in
-  fun primes n = Metis_Portable.critical (fn () =>
-      let
-        val Unsynchronized.ref ps = primesList
-
-        val k = n - length ps
-      in
-        if k <= 0 then List.take (ps,n)
-        else
-          let
-            val ps = calcPrimes ps k (List.last ps + 1)
-
-            val () = primesList := ps
-          in
-            ps
-          end
-      end) ();
-end;
-
-fun primesUpTo n = Metis_Portable.critical (fn () =>
-    let
-      fun f k =
-          let
-            val l = primes k
-
-            val p = List.last l
-          in
-            if p < n then f (2 * k) else takeWhile (fn j => j <= n) l
-          end
-    in
-      f 8
-    end) ();
+in
+  fun primes n =
+      if n <= 0 then []
+      else calcPrimes true [2] 3 (n - 1);
+
+  fun primesUpTo n =
+      if n < 2 then []
+      else calcPrimes false [2] 3 (n - 2);
+end;
 
 (* ------------------------------------------------------------------------- *)
 (* Strings.                                                                  *)
@@ -1228,23 +1211,30 @@
 
 local
   val generator = Unsynchronized.ref 0
-in
-  fun newInt () = Metis_Portable.critical (fn () =>
+
+  fun newIntThunk () =
       let
         val n = !generator
+
         val () = generator := n + 1
       in
         n
-      end) ();
-
-  fun newInts 0 = []
-    | newInts k = Metis_Portable.critical (fn () =>
+      end;
+
+  fun newIntsThunk k () =
       let
         val n = !generator
+
         val () = generator := n + k
       in
         interval n k
-      end) ();
+      end;
+in
+  fun newInt () = Metis_Portable.critical newIntThunk ();
+
+  fun newInts k =
+      if k <= 0 then []
+      else Metis_Portable.critical (newIntsThunk k) ();
 end;
 
 fun withRef (r,new) f x =
@@ -14197,24 +14187,23 @@
 (* Basic conjunctive normal form.                                            *)
 (* ------------------------------------------------------------------------- *)
 
-val newSkolemFunction =
-    let
-      val counter : int Metis_StringMap.map Unsynchronized.ref = Unsynchronized.ref (Metis_StringMap.new ())
-
-      fun new n () =
-          let
-            val Unsynchronized.ref m = counter
-            val s = Metis_Name.toString n
-            val i = Option.getOpt (Metis_StringMap.peek m s, 0)
-            val () = counter := Metis_StringMap.insert m (s, i + 1)
-            val i = if i = 0 then "" else "_" ^ Int.toString i
-            val s = skolemPrefix ^ "_" ^ s ^ i
-          in
-            Metis_Name.fromString s
-          end
-    in
-      fn n => Metis_Portable.critical (new n) ()
-    end;
+local
+  val counter : int Metis_StringMap.map Unsynchronized.ref = Unsynchronized.ref (Metis_StringMap.new ());
+
+  fun new n () =
+      let
+        val Unsynchronized.ref m = counter
+        val s = Metis_Name.toString n
+        val i = Option.getOpt (Metis_StringMap.peek m s, 0)
+        val () = counter := Metis_StringMap.insert m (s, i + 1)
+        val i = if i = 0 then "" else "_" ^ Int.toString i
+        val s = skolemPrefix ^ "_" ^ s ^ i
+      in
+        Metis_Name.fromString s
+      end;
+in
+  fun newSkolemFunction n = Metis_Portable.critical (new n) ();
+end;
 
 fun skolemize fv bv fm =
     let
@@ -14753,18 +14742,19 @@
 (* Definitions.                                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val newDefinitionRelation =
-    let
-      val counter : int Unsynchronized.ref = Unsynchronized.ref 0
-    in
-      fn () => Metis_Portable.critical (fn () =>
-         let
-           val Unsynchronized.ref i = counter
-           val () = counter := i + 1
-         in
-           definitionPrefix ^ "_" ^ Int.toString i
-         end) ()
-    end;
+local
+  val counter : int Unsynchronized.ref = Unsynchronized.ref 0;
+
+  fun new () =
+      let
+        val Unsynchronized.ref i = counter
+        val () = counter := i + 1
+      in
+        definitionPrefix ^ "_" ^ Int.toString i
+      end;
+in
+  fun newDefinitionRelation () = Metis_Portable.critical new ();
+end;
 
 fun newDefinition def =
     let
--- a/src/Tools/Metis/src/Normalize.sml	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/Tools/Metis/src/Normalize.sml	Fri Sep 17 16:38:11 2010 +0200
@@ -685,24 +685,23 @@
 (* Basic conjunctive normal form.                                            *)
 (* ------------------------------------------------------------------------- *)
 
-val newSkolemFunction =
-    let
-      val counter : int StringMap.map ref = ref (StringMap.new ())
+local
+  val counter : int StringMap.map ref = ref (StringMap.new ());
 
-      fun new n () =
-          let
-            val ref m = counter
-            val s = Name.toString n
-            val i = Option.getOpt (StringMap.peek m s, 0)
-            val () = counter := StringMap.insert m (s, i + 1)
-            val i = if i = 0 then "" else "_" ^ Int.toString i
-            val s = skolemPrefix ^ "_" ^ s ^ i
-          in
-            Name.fromString s
-          end
-    in
-      fn n => Portable.critical (new n) ()
-    end;
+  fun new n () =
+      let
+        val ref m = counter
+        val s = Name.toString n
+        val i = Option.getOpt (StringMap.peek m s, 0)
+        val () = counter := StringMap.insert m (s, i + 1)
+        val i = if i = 0 then "" else "_" ^ Int.toString i
+        val s = skolemPrefix ^ "_" ^ s ^ i
+      in
+        Name.fromString s
+      end;
+in
+  fun newSkolemFunction n = Portable.critical (new n) ();
+end;
 
 fun skolemize fv bv fm =
     let
@@ -1241,18 +1240,19 @@
 (* Definitions.                                                              *)
 (* ------------------------------------------------------------------------- *)
 
-val newDefinitionRelation =
-    let
-      val counter : int ref = ref 0
-    in
-      fn () => Portable.critical (fn () =>
-         let
-           val ref i = counter
-           val () = counter := i + 1
-         in
-           definitionPrefix ^ "_" ^ Int.toString i
-         end) ()
-    end;
+local
+  val counter : int ref = ref 0;
+
+  fun new () =
+      let
+        val ref i = counter
+        val () = counter := i + 1
+      in
+        definitionPrefix ^ "_" ^ Int.toString i
+      end;
+in
+  fun newDefinitionRelation () = Portable.critical new ();
+end;
 
 fun newDefinition def =
     let
--- a/src/Tools/Metis/src/Useful.sml	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/Tools/Metis/src/Useful.sml	Fri Sep 17 16:38:11 2010 +0200
@@ -462,50 +462,33 @@
 end;
 
 local
-  fun calcPrimes ps n i =
-      if List.exists (fn p => divides p i) ps then calcPrimes ps n (i + 1)
+  fun calcPrimes mode ps i n =
+      if n = 0 then ps
+      else if List.exists (fn p => divides p i) ps then
+        let
+          val i = i + 1
+          and n = if mode then n else n - 1
+        in
+          calcPrimes mode ps i n
+        end
       else
         let
           val ps = ps @ [i]
+          and i = i + 1
           and n = n - 1
         in
-          if n = 0 then ps else calcPrimes ps n (i + 1)
+          calcPrimes mode ps i n
         end;
-
-  val primesList = ref [2];
 in
-  fun primes n = Portable.critical (fn () =>
-      let
-        val ref ps = primesList
-
-        val k = n - length ps
-      in
-        if k <= 0 then List.take (ps,n)
-        else
-          let
-            val ps = calcPrimes ps k (List.last ps + 1)
+  fun primes n =
+      if n <= 0 then []
+      else calcPrimes true [2] 3 (n - 1);
 
-            val () = primesList := ps
-          in
-            ps
-          end
-      end) ();
+  fun primesUpTo n =
+      if n < 2 then []
+      else calcPrimes false [2] 3 (n - 2);
 end;
 
-fun primesUpTo n = Portable.critical (fn () =>
-    let
-      fun f k =
-          let
-            val l = primes k
-
-            val p = List.last l
-          in
-            if p < n then f (2 * k) else takeWhile (fn j => j <= n) l
-          end
-    in
-      f 8
-    end) ();
-
 (* ------------------------------------------------------------------------- *)
 (* Strings.                                                                  *)
 (* ------------------------------------------------------------------------- *)
@@ -723,23 +706,30 @@
 
 local
   val generator = ref 0
-in
-  fun newInt () = Portable.critical (fn () =>
+
+  fun newIntThunk () =
       let
         val n = !generator
+
         val () = generator := n + 1
       in
         n
-      end) ();
+      end;
 
-  fun newInts 0 = []
-    | newInts k = Portable.critical (fn () =>
+  fun newIntsThunk k () =
       let
         val n = !generator
+
         val () = generator := n + k
       in
         interval n k
-      end) ();
+      end;
+in
+  fun newInt () = Portable.critical newIntThunk ();
+
+  fun newInts k =
+      if k <= 0 then []
+      else Portable.critical (newIntsThunk k) ();
 end;
 
 fun withRef (r,new) f x =
--- a/src/Tools/Metis/src/metis.sml	Fri Sep 17 16:15:45 2010 +0200
+++ b/src/Tools/Metis/src/metis.sml	Fri Sep 17 16:38:11 2010 +0200
@@ -13,7 +13,7 @@
 
 val VERSION = "2.3";
 
-val versionString = PROGRAM^" "^VERSION^" (release 20100915)"^"\n";
+val versionString = PROGRAM^" "^VERSION^" (release 20100916)"^"\n";
 
 (* ------------------------------------------------------------------------- *)
 (* Program options.                                                          *)