removed Sledgehammer's support for the DFG syntax;
authorblanchet
Tue, 22 Jun 2010 14:28:22 +0200
changeset 37498 b426cbdb5a23
parent 37497 71fdbffe3275
child 37499 5ff37037fbec
removed Sledgehammer's support for the DFG syntax; this removes 350 buggy lines from Sledgehammer. SPASS 3.5 and above support the TPTP syntax.
doc-src/Sledgehammer/sledgehammer.tex
src/HOL/Tools/ATP_Manager/atp_systems.ML
src/HOL/Tools/Sledgehammer/meson_tactic.ML
src/HOL/Tools/Sledgehammer/metis_tactics.ML
src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimizer.ML
src/HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML
src/HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML
src/HOL/Tools/Sledgehammer/sledgehammer_hol_clause.ML
src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML
src/HOL/Tools/Sledgehammer/sledgehammer_util.ML
--- a/doc-src/Sledgehammer/sledgehammer.tex	Tue Jun 22 13:17:59 2010 +0200
+++ b/doc-src/Sledgehammer/sledgehammer.tex	Tue Jun 22 14:28:22 2010 +0200
@@ -341,11 +341,6 @@
 download page. Sledgehammer requires version 3.5 or above. See
 \S\ref{installation} for details.
 
-\item[$\bullet$] \textbf{\textit{spass\_dfg}:} Same as the above, except that
-Sledgehammer communicates with SPASS using the native DFG syntax rather than the
-TPTP syntax. Sledgehammer requires version 3.0 or above. This ATP is provided
-for compatibility reasons.
-
 \item[$\bullet$] \textbf{\textit{vampire}:} Vampire is an ATP developed by
 Andrei Voronkov and his colleagues \cite{riazanov-voronkov-2002}. To use
 Vampire, set the environment variable \texttt{VAMPIRE\_HOME} to the directory
--- a/src/HOL/Tools/ATP_Manager/atp_systems.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/ATP_Manager/atp_systems.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -110,7 +110,7 @@
 
 fun shape_of_clauses _ [] = []
   | shape_of_clauses j ([] :: clauses) = [] :: shape_of_clauses j clauses
-  | shape_of_clauses j ((lit :: lits) :: clauses) =
+  | shape_of_clauses j ((_ :: lits) :: clauses) =
     let val shape = shape_of_clauses (j + 1) (lits :: clauses) in
       (j :: hd shape) :: tl shape
     end
@@ -123,7 +123,7 @@
          : problem) =
   let
     (* get clauses and prepare them for writing *)
-    val (ctxt, (chained_ths, th)) = goal;
+    val (ctxt, (_, th)) = goal;
     val thy = ProofContext.theory_of ctxt;
     val goal_clss = #1 (neg_conjecture_clauses ctxt th subgoal)
     val goal_cls = List.concat goal_clss
@@ -136,7 +136,7 @@
         NONE => the_filtered_clauses
       | SOME axcls => axcls);
     val (internal_thm_names, clauses) =
-      prepare goal_cls chained_ths the_axiom_clauses the_filtered_clauses thy;
+      prepare goal_cls the_axiom_clauses the_filtered_clauses thy
 
     (* path to unique problem file *)
     val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
@@ -243,16 +243,15 @@
                 max_axiom_clauses, prefers_theory_relevant})
         (params as {debug, overlord, full_types, respect_no_atp,
                     relevance_threshold, relevance_convergence, theory_relevant,
-                    defs_relevant, isar_proof, ...})
+                    defs_relevant, ...})
         minimize_command timeout =
   generic_prover overlord
       (relevant_facts full_types respect_no_atp relevance_threshold
                       relevance_convergence defs_relevant max_axiom_clauses
                       (the_default prefers_theory_relevant theory_relevant))
-      (prepare_clauses false full_types)
-      (write_tptp_file (debug andalso overlord)) home_var
-      executable (arguments timeout) proof_delims known_failures name params
-      minimize_command
+      (prepare_clauses full_types) (write_tptp_file (debug andalso overlord))
+      home_var executable (arguments timeout) proof_delims known_failures name
+      params minimize_command
 
 fun tptp_prover name p = (name, generic_tptp_prover (name, p));
 
@@ -313,31 +312,14 @@
 
 (* SPASS *)
 
-fun generic_dfg_prover
-        (name, {home_var, executable, arguments, proof_delims, known_failures,
-                max_axiom_clauses, prefers_theory_relevant})
-        (params as {overlord, full_types, respect_no_atp, relevance_threshold,
-                    relevance_convergence, theory_relevant, defs_relevant, ...})
-        minimize_command timeout =
-  generic_prover overlord
-      (relevant_facts full_types respect_no_atp relevance_threshold
-                      relevance_convergence defs_relevant max_axiom_clauses
-                      (the_default prefers_theory_relevant theory_relevant))
-      (prepare_clauses true full_types) write_dfg_file home_var executable
-      (arguments timeout) proof_delims known_failures name params
-      minimize_command
-
-fun dfg_prover name p = (name, generic_dfg_prover (name, p))
-
 (* The "-VarWeight=3" option helps the higher-order problems, probably by
    counteracting the presence of "hAPP". *)
-fun generic_spass_config dfg : prover_config =
+val spass_config : prover_config =
   {home_var = "SPASS_HOME",
    executable = "SPASS",
    arguments = fn timeout =>
-     "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \
-     \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_generous_secs timeout)
-     |> not dfg ? prefix "-TPTP ",
+     "-TPTP -Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \
+     \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_generous_secs timeout),
    proof_delims = [("Here is a proof", "Formulae used in the proof")],
    known_failures =
      [(IncompleteUnprovable, "SPASS beiseite: Completion found"),
@@ -349,10 +331,6 @@
       (OldSpass, "Unrecognized option TPTP")],
    max_axiom_clauses = 40,
    prefers_theory_relevant = true}
-val spass_dfg_config = generic_spass_config true
-val spass_dfg = dfg_prover "spass_dfg" spass_dfg_config
-
-val spass_config = generic_spass_config false
 val spass = tptp_prover "spass" spass_config
 
 (* remote prover invocation via SystemOnTPTP *)
@@ -414,8 +392,7 @@
   space_implode " " [maybe_remote e e_config, maybe_remote spass spass_config,
                      remotify (fst vampire)]
 
-val provers =
-  [spass, spass_dfg, vampire, e, remote_vampire, remote_spass, remote_e]
+val provers = [spass, vampire, e, remote_vampire, remote_spass, remote_e]
 val prover_setup = fold add_prover provers
 
 val setup =
--- a/src/HOL/Tools/Sledgehammer/meson_tactic.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/meson_tactic.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -17,11 +17,11 @@
 open Sledgehammer_Fact_Preprocessor
 
 (*Expand all new definitions of abstraction or Skolem functions in a proof state.*)
-fun is_absko (Const (@{const_name "=="}, _) $ Free (a, _) $ u) =
+fun is_absko (Const (@{const_name "=="}, _) $ Free (a, _) $ _) =
     String.isPrefix skolem_prefix a
   | is_absko _ = false;
 
-fun is_okdef xs (Const (@{const_name "=="}, _) $ t $ u) =   (*Definition of Free, not in certain terms*)
+fun is_okdef xs (Const (@{const_name "=="}, _) $ t $ _) =   (*Definition of Free, not in certain terms*)
       is_Free t andalso not (member (op aconv) xs t)
   | is_okdef _ _ = false
 
--- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -58,7 +58,7 @@
 (* 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 n [] = raise Empty
+      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;
@@ -238,7 +238,7 @@
         | NONE    => raise Fail ("fol_type_to_isa: " ^ x));
 
 fun reveal_skolem_somes skolem_somes =
-  map_aterms (fn t as Const (s, T) =>
+  map_aterms (fn t as Const (s, _) =>
                  if String.isPrefix skolem_theory_name s then
                    AList.lookup (op =) skolem_somes s
                    |> the |> map_types Type_Infer.paramify_vars
@@ -348,8 +348,7 @@
   | hol_term_from_fol _ = hol_term_from_fol_PT
 
 fun hol_terms_from_fol ctxt mode skolem_somes fol_tms =
-  let val thy = ProofContext.theory_of ctxt
-      val ts = map (hol_term_from_fol mode ctxt) fol_tms
+  let val ts = map (hol_term_from_fol 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_somes skolem_somes) |> infer_types ctxt
@@ -694,11 +693,10 @@
             tfrees = union (op =) tfree_lits tfrees,
             skolem_somes = skolem_somes}
         end;
-      val lmap as {skolem_somes, ...} =
-        {axioms = [], tfrees = init_tfrees ctxt, skolem_somes = []}
-        |> fold (add_thm true) cls
-        |> add_tfrees
-        |> fold (add_thm false) ths
+      val lmap = {axioms = [], tfrees = init_tfrees ctxt, skolem_somes = []}
+                 |> fold (add_thm true) cls
+                 |> add_tfrees
+                 |> fold (add_thm false) 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
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -27,8 +27,7 @@
     -> Proof.context * (thm list * 'a) -> thm list
     -> (thm * (string * int)) list
   val prepare_clauses :
-    bool -> bool -> thm list -> thm list
-    -> (thm * (axiom_name * hol_clause_id)) list
+    bool -> thm list -> (thm * (axiom_name * hol_clause_id)) list
     -> (thm * (axiom_name * hol_clause_id)) list -> theory
     -> axiom_name vector
        * (hol_clause list * hol_clause list * hol_clause list *
@@ -265,7 +264,7 @@
   end
 
 fun relevant_clauses ctxt relevance_convergence defs_relevant max_new
-                     (relevance_override as {add, del, only}) ctab =
+                     ({add, del, ...} : relevance_override) ctab =
   let
     val thy = ProofContext.theory_of ctxt
     val add_thms = cnf_for_facts ctxt add
@@ -398,7 +397,7 @@
 fun multi_name a th (n, pairs) =
   (n + 1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs);
 
-fun add_names (a, []) pairs = pairs
+fun add_names (_, []) pairs = pairs
   | add_names (a, [th]) pairs = (a, th) :: pairs
   | add_names (a, ths) pairs = #2 (fold (multi_name a) ths (1, pairs))
 
@@ -478,7 +477,7 @@
 (*Ensures that no higher-order theorems "leak out"*)
 fun restrict_to_logic thy true cls =
     filter (is_quasi_fol_term thy o prop_of o fst) cls
-  | restrict_to_logic thy false cls = cls;
+  | restrict_to_logic _ false cls = cls
 
 (**** Predicates to detect unwanted clauses (prolific or likely to cause
       unsoundness) ****)
@@ -566,7 +565,7 @@
 
 (* prepare for passing to writer,
    create additional clauses based on the information from extra_cls *)
-fun prepare_clauses dfg full_types goal_cls chained_ths axcls extra_cls thy =
+fun prepare_clauses full_types goal_cls axcls extra_cls thy =
   let
     val is_FO = is_fol_goal thy goal_cls
     val ccls = subtract_cls extra_cls goal_cls
@@ -577,12 +576,12 @@
     and supers = tvar_classes_of_terms axtms
     and tycons = type_consts_of_terms thy (ccltms @ axtms)
     (*TFrees in conjecture clauses; TVars in axiom clauses*)
-    val conjectures = make_conjecture_clauses dfg thy ccls
-    val (_, extra_clauses) = ListPair.unzip (make_axiom_clauses dfg thy extra_cls)
-    val (clnames, axiom_clauses) = ListPair.unzip (make_axiom_clauses dfg thy axcls)
+    val conjectures = make_conjecture_clauses thy ccls
+    val (_, extra_clauses) = ListPair.unzip (make_axiom_clauses thy extra_cls)
+    val (clnames, axiom_clauses) = ListPair.unzip (make_axiom_clauses thy axcls)
     val helper_clauses =
-      get_helper_clauses dfg thy is_FO full_types conjectures extra_cls
-    val (supers', arity_clauses) = make_arity_clauses_dfg dfg thy tycons supers
+      get_helper_clauses thy is_FO full_types conjectures extra_cls
+    val (supers', arity_clauses) = make_arity_clauses thy tycons supers
     val classrel_clauses = make_classrel_clauses thy subs supers'
   in
     (Vector.fromList clnames,
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimizer.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimizer.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -47,8 +47,8 @@
 fun string_for_outcome NONE = "Success."
   | string_for_outcome (SOME failure) = string_for_failure failure
 
-fun sledgehammer_test_theorems (params as {full_types, ...} : params) prover
-        timeout subgoal state filtered_clauses name_thms_pairs =
+fun sledgehammer_test_theorems (params : params) prover timeout subgoal state
+                               filtered_clauses name_thms_pairs =
   let
     val num_theorems = length name_thms_pairs
     val _ = priority ("Testing " ^ string_of_int num_theorems ^
@@ -88,7 +88,7 @@
         (result as {outcome = NONE, ...}) => SOME result
       | _ => NONE
 
-    val {context = ctxt, facts, goal} = Proof.goal state;
+    val {context = ctxt, goal, ...} = Proof.goal state;
   in
     (* try prove first to check result and get used theorems *)
     (case test_thms_fun NONE name_thms_pairs of
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -21,9 +21,8 @@
   val cnf_rules_pairs:
     theory -> (string * thm) list -> (thm * (string * int)) list
   val saturate_skolem_cache: theory -> theory option
-  val use_skolem_cache: bool Unsynchronized.ref
+  val auto_saturate_skolem_cache: bool Unsynchronized.ref
     (* for emergency use where the Skolem cache causes problems *)
-  val strip_subgoal : thm -> int -> (string * typ) list * term list * term
   val neg_clausify: thm -> thm list
   val neg_conjecture_clauses:
     Proof.context -> thm -> int -> thm list list * (string * typ) list
@@ -133,7 +132,7 @@
       | dec_sko (@{const "op &"} $ p $ q) thx = dec_sko q (dec_sko p thx)
       | dec_sko (@{const "op |"} $ p $ q) thx = dec_sko q (dec_sko p thx)
       | dec_sko (@{const Trueprop} $ p) thx = dec_sko p thx
-      | dec_sko t thx = thx
+      | dec_sko _ thx = thx
   in dec_sko (prop_of th) ([], thy) end
 
 fun mk_skolem_id t =
@@ -168,7 +167,7 @@
       | dec_sko (@{const "op &"} $ p $ q) defs = dec_sko q (dec_sko p defs)
       | dec_sko (@{const "op |"} $ p $ q) defs = dec_sko q (dec_sko p defs)
       | dec_sko (@{const Trueprop} $ p) defs = dec_sko p defs
-      | dec_sko t defs = defs (*Do nothing otherwise*)
+      | dec_sko _ defs = defs
   in  dec_sko (prop_of th) []  end;
 
 
@@ -520,22 +519,11 @@
 
 end;
 
-val use_skolem_cache = Unsynchronized.ref true
-
-fun clause_cache_endtheory thy =
-  if !use_skolem_cache then saturate_skolem_cache thy else NONE
-
+val auto_saturate_skolem_cache = Unsynchronized.ref true
 
-(* The cache can be kept smaller by inspecting the prop of each thm. Can ignore
-   all that are quasi-lambda-free, but then the individual theory caches become
-   much bigger. *)
+fun conditionally_saturate_skolem_cache thy =
+  if !auto_saturate_skolem_cache then saturate_skolem_cache thy else NONE
 
-fun strip_subgoal goal i =
-  let
-    val (t, frees) = Logic.goal_params (prop_of goal) i
-    val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees)
-    val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
-  in (rev (map dest_Free frees), hyp_ts, concl_t) end
 
 (*** Converting a subgoal into negated conjecture clauses. ***)
 
@@ -574,7 +562,7 @@
 (** setup **)
 
 val setup =
-  perhaps saturate_skolem_cache
-  #> Theory.at_end clause_cache_endtheory
+  perhaps conditionally_saturate_skolem_cache
+  #> Theory.at_end conditionally_saturate_skolem_cache
 
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -28,8 +28,8 @@
   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 : bool -> string -> string
-  val make_fixed_type_const : bool -> string -> string
+  val make_fixed_const : string -> string
+  val make_fixed_type_const : string -> string
   val make_type_class : string -> string
   type name = string * string
   type name_pool = string Symtab.table * string Symtab.table
@@ -49,7 +49,6 @@
     TyLitFree of string * name
   exception CLAUSE of string * term
   val type_literals_for_types : typ list -> type_literal list
-  val get_tvar_strs: typ list -> string list
   datatype arLit =
       TConsLit of class * string * string list
     | TVarLit of class * string
@@ -58,28 +57,7 @@
   datatype classrel_clause = ClassrelClause of
    {axiom_name: axiom_name, subclass: class, superclass: class}
   val make_classrel_clauses: theory -> class list -> class list -> classrel_clause list
-  val make_arity_clauses_dfg: bool -> theory -> string list -> class list -> class list * arity_clause list
   val make_arity_clauses: theory -> string list -> class list -> class list * arity_clause list
-  val add_type_sort_preds: typ -> int Symtab.table -> int Symtab.table
-  val add_classrel_clause_preds :
-    classrel_clause -> int Symtab.table -> int Symtab.table
-  val class_of_arityLit: arLit -> class
-  val add_arity_clause_preds: arity_clause -> int Symtab.table -> int Symtab.table
-  val add_fol_type_funcs: fol_type -> int Symtab.table -> int Symtab.table
-  val add_arity_clause_funcs:
-    arity_clause -> int Symtab.table -> int Symtab.table
-  val init_functab: int Symtab.table
-  val dfg_sign: bool -> string -> string
-  val dfg_of_type_literal: bool -> type_literal -> string
-  val gen_dfg_cls: int * string * kind * string list * string list * string list -> string
-  val string_of_preds: (string * Int.int) list -> string
-  val string_of_funcs: (string * int) list -> string
-  val string_of_symbols: string -> string -> string
-  val string_of_start: string -> string
-  val string_of_descrip : string -> string
-  val dfg_tfree_clause : string -> string
-  val dfg_classrel_clause: classrel_clause -> string
-  val dfg_arity_clause: arity_clause -> string
   val tptp_sign: bool -> string -> string
   val tptp_of_type_literal :
     bool -> type_literal -> name_pool option -> string * name_pool option
@@ -196,31 +174,21 @@
       tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
 fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
 
-val max_dfg_symbol_length = 63
-
-(* HACK because SPASS 3.0 truncates identifiers to 63 characters. *)
-fun controlled_length dfg s =
-  if dfg andalso size s > max_dfg_symbol_length then
-    String.extract (s, 0, SOME (max_dfg_symbol_length div 2 - 1)) ^ "__" ^
-    String.extract (s, size s - max_dfg_symbol_length div 2 + 1, NONE)
-  else
-    s
+fun lookup_const c =
+  case Symtab.lookup const_trans_table c of
+    SOME c' => c'
+  | NONE => ascii_of c
 
-fun lookup_const dfg c =
-    case Symtab.lookup const_trans_table c of
-        SOME c' => c'
-      | NONE => controlled_length dfg (ascii_of c);
-
-fun lookup_type_const dfg c =
-    case Symtab.lookup type_const_trans_table c of
-        SOME c' => c'
-      | NONE => controlled_length dfg (ascii_of c);
+fun lookup_type_const c =
+  case Symtab.lookup type_const_trans_table c of
+    SOME c' => c'
+  | NONE => ascii_of c
 
 (* "op =" MUST BE "equal" because it's built into ATPs. *)
-fun make_fixed_const _ (@{const_name "op ="}) = "equal"
-  | make_fixed_const dfg c = const_prefix ^ lookup_const dfg c;
+fun make_fixed_const @{const_name "op ="} = "equal"
+  | make_fixed_const c = const_prefix ^ lookup_const c
 
-fun make_fixed_type_const dfg c = tconst_prefix ^ lookup_type_const dfg c;
+fun make_fixed_type_const c = tconst_prefix ^ lookup_type_const c
 
 fun make_type_class clas = class_prefix ^ ascii_of clas;
 
@@ -287,8 +255,7 @@
                             the_pool
               |> apsnd SOME
 
-(**** Definitions and functions for FOL clauses, for conversion to TPTP or DFG
-      format ****)
+(**** Definitions and functions for FOL clauses for TPTP format output ****)
 
 datatype kind = Axiom | Conjecture;
 
@@ -329,30 +296,12 @@
 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);
 
-fun pred_of_sort (TyLitVar (s, _)) = (s, 1)
-  | pred_of_sort (TyLitFree (s, _)) = (s, 1)
-
 (*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) []
 
-(*The correct treatment of TFrees like 'a in lemmas (axiom clauses) is not clear.
-  *  Ignoring them leads to unsound proofs, since we do nothing to ensure that 'a
-    in a lemma has the same sort as 'a in the conjecture.
-  * Deleting such clauses will lead to problems with locales in other use of local results
-    where 'a is fixed. Probably we should delete clauses unless the sorts agree.
-  * Currently we include a class constraint in the clause, exactly as with TVars.
-*)
-
 (** make axiom and conjecture clauses. **)
 
-fun get_tvar_strs [] = []
-  | get_tvar_strs ((TVar (indx,s))::Ts) =
-      insert (op =) (make_schematic_type_var indx) (get_tvar_strs Ts)
-  | get_tvar_strs((TFree _)::Ts) = get_tvar_strs Ts
-
-
-
 (**** Isabelle arities ****)
 
 datatype arLit = TConsLit of class * string * string list
@@ -372,13 +321,13 @@
   | pack_sort(tvar, cls::srt) =  (cls, tvar) :: pack_sort(tvar, srt);
 
 (*Arity of type constructor tcon :: (arg1,...,argN)res*)
-fun make_axiom_arity_clause dfg (tcons, axiom_name, (cls,args)) =
+fun make_axiom_arity_clause (tcons, axiom_name, (cls,args)) =
    let val tvars = gen_TVars (length args)
        val tvars_srts = ListPair.zip (tvars,args)
    in
-      ArityClause {axiom_name = axiom_name, 
-                   conclLit = TConsLit (cls, make_fixed_type_const dfg tcons, tvars),
-                   premLits = map TVarLit (union_all(map pack_sort tvars_srts))}
+     ArityClause {axiom_name = axiom_name, 
+                  conclLit = TConsLit (cls, make_fixed_type_const tcons, tvars),
+                  premLits = map TVarLit (union_all(map pack_sort tvars_srts))}
    end;
 
 
@@ -390,7 +339,7 @@
                             superclass: class};
 
 (*Generate all pairs (sub,super) such that sub is a proper subclass of super in theory thy.*)
-fun class_pairs thy [] supers = []
+fun class_pairs _ [] _ = []
   | class_pairs thy subs supers =
       let
         val class_less = Sorts.class_less (Sign.classes_of thy)
@@ -409,20 +358,20 @@
 
 (** Isabelle arities **)
 
-fun arity_clause dfg _ _ (tcons, []) = []
-  | arity_clause dfg seen n (tcons, ("HOL.type",_)::ars) =  (*ignore*)
-      arity_clause dfg seen n (tcons,ars)
-  | arity_clause dfg seen n (tcons, (ar as (class,_)) :: ars) =
+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 dfg (tcons, lookup_type_const dfg tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) ::
-          arity_clause dfg seen (n+1) (tcons,ars)
+          make_axiom_arity_clause (tcons, lookup_type_const tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) ::
+          arity_clause seen (n+1) (tcons,ars)
       else
-          make_axiom_arity_clause dfg (tcons, lookup_type_const dfg tcons ^ "_" ^ class, ar) ::
-          arity_clause dfg (class::seen) n (tcons,ars)
+          make_axiom_arity_clause (tcons, lookup_type_const tcons ^ "_" ^ class, ar) ::
+          arity_clause (class::seen) n (tcons,ars)
 
-fun multi_arity_clause dfg [] = []
-  | multi_arity_clause dfg ((tcons, ars) :: tc_arlists) =
-      arity_clause dfg [] 1 (tcons, ars) @ multi_arity_clause dfg tc_arlists
+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.*)
@@ -436,7 +385,7 @@
   in  map try_classes tycons  end;
 
 (*Proving one (tycon, class) membership may require proving others, so iterate.*)
-fun iter_type_class_pairs thy tycons [] = ([], [])
+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)))
@@ -444,130 +393,16 @@
           val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
       in (union (op =) classes' classes, union (op =) cpairs' cpairs) end;
 
-fun make_arity_clauses_dfg dfg thy tycons classes =
+fun make_arity_clauses thy tycons classes =
   let val (classes', cpairs) = iter_type_class_pairs thy tycons classes
-  in  (classes', multi_arity_clause dfg cpairs)  end;
-val make_arity_clauses = make_arity_clauses_dfg false;
-
-(**** Find occurrences of predicates in clauses ****)
-
-(*FIXME: multiple-arity checking doesn't work, as update_new is the wrong
-  function (it flags repeated declarations of a function, even with the same arity)*)
-
-fun update_many keypairs = fold Symtab.update keypairs
-
-val add_type_sort_preds = update_many o map pred_of_sort o sorts_on_typs
-
-fun add_classrel_clause_preds (ClassrelClause {subclass, superclass, ...}) =
-  Symtab.update (subclass, 1) #> Symtab.update (superclass, 1)
-
-fun class_of_arityLit (TConsLit (tclass, _, _)) = tclass
-  | class_of_arityLit (TVarLit (tclass, _)) = tclass;
-
-fun add_arity_clause_preds (ArityClause {conclLit, premLits, ...}) =
-  let
-    val classes = map (make_type_class o class_of_arityLit)
-                      (conclLit :: premLits)
-  in fold (Symtab.update o rpair 1) classes end;
-
-(*** Find occurrences of functions in clauses ***)
-
-fun add_fol_type_funcs (TyVar _) = I
-  | add_fol_type_funcs (TyFree (s, _)) = Symtab.update (s, 0)
-  | add_fol_type_funcs (TyConstr ((s, _), tys)) =
-    Symtab.update (s, length tys) #> fold add_fol_type_funcs tys
-
-(*TFrees are recorded as constants*)
-fun add_type_sort_funcs (TVar _, funcs) = funcs
-  | add_type_sort_funcs (TFree (a, _), funcs) =
-      Symtab.update (make_fixed_type_var a, 0) funcs
-
-fun add_arity_clause_funcs (ArityClause {conclLit,...}) funcs =
-  let val TConsLit (_, tcons, tvars) = conclLit
-  in  Symtab.update (tcons, length tvars) funcs  end;
-
-(*This type can be overlooked because it is built-in...*)
-val init_functab = Symtab.update ("tc_itself", 1) Symtab.empty;
-
-
-(**** String-oriented operations ****)
-
-fun string_of_clausename (cls_id,ax_name) =
-    clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id;
-
-fun string_of_type_clsname (cls_id,ax_name,idx) =
-    string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);
-
-
-(**** Producing DFG files ****)
-
-(*Attach sign in DFG syntax: false means negate.*)
-fun dfg_sign true s = s
-  | dfg_sign false s = "not(" ^ s ^ ")"
-
-fun dfg_of_type_literal pos (TyLitVar (s, (s', _))) =
-    dfg_sign pos (s ^ "(" ^ s' ^ ")")
-  | dfg_of_type_literal pos (TyLitFree (s, (s', _))) =
-    dfg_sign pos (s ^ "(" ^ s' ^ ")");
-
-(*Enclose the clause body by quantifiers, if necessary*)
-fun dfg_forall [] body = body
-  | dfg_forall vars body = "forall([" ^ commas vars ^ "],\n" ^ body ^ ")"
-
-fun gen_dfg_cls (cls_id, ax_name, Axiom, lits, tylits, vars) =
-      "clause( %(axiom)\n" ^
-      dfg_forall vars ("or(" ^ commas (tylits@lits) ^ ")") ^ ",\n" ^
-      string_of_clausename (cls_id,ax_name) ^  ").\n\n"
-  | gen_dfg_cls (cls_id, ax_name, Conjecture, lits, _, vars) =
-      "clause( %(negated_conjecture)\n" ^
-      dfg_forall vars ("or(" ^ commas lits ^ ")") ^ ",\n" ^
-      string_of_clausename (cls_id,ax_name) ^  ").\n\n";
-
-fun string_of_arity (name, arity) =  "(" ^ name ^ ", " ^ Int.toString arity ^ ")"
-
-fun string_of_preds [] = ""
-  | string_of_preds preds = "predicates[" ^ commas(map string_of_arity preds) ^ "].\n";
-
-fun string_of_funcs [] = ""
-  | string_of_funcs funcs = "functions[" ^ commas(map string_of_arity funcs) ^ "].\n" ;
-
-fun string_of_symbols predstr funcstr =
-  "list_of_symbols.\n" ^ predstr  ^ funcstr  ^ "end_of_list.\n\n";
-
-fun string_of_start name = "begin_problem(" ^ name ^ ").\n\n";
-
-fun string_of_descrip name =
-  "list_of_descriptions.\nname({*" ^ name ^
-  "*}).\nauthor({*Isabelle*}).\nstatus(unknown).\ndescription({*auto-generated*}).\nend_of_list.\n\n"
-
-fun dfg_tfree_clause tfree_lit =
-  "clause( %(negated_conjecture)\n" ^ "or( " ^ tfree_lit ^ "),\n" ^ "tfree_tcs" ^ ").\n\n"
-
-fun dfg_of_arLit (TConsLit (c,t,args)) =
-      dfg_sign true (make_type_class c ^ "(" ^ t ^ paren_pack args ^ ")")
-  | dfg_of_arLit (TVarLit (c,str)) =
-      dfg_sign false (make_type_class c ^ "(" ^ str ^ ")")
-
-fun dfg_classrelLits sub sup =  "not(" ^ sub ^ "(T)), " ^ sup ^ "(T)";
-
-fun dfg_classrel_clause (ClassrelClause {axiom_name,subclass,superclass,...}) =
-  "clause(forall([T],\nor( " ^ dfg_classrelLits subclass superclass ^ ")),\n" ^
-  axiom_name ^ ").\n\n";
-
-fun string_of_ar axiom_name = arclause_prefix ^ ascii_of axiom_name;
-
-fun dfg_arity_clause (ArityClause{axiom_name,conclLit,premLits,...}) =
-  let val TConsLit (_,_,tvars) = conclLit
-      val lits = map dfg_of_arLit (conclLit :: premLits)
-  in
-      "clause( %(axiom)\n" ^
-      dfg_forall tvars ("or( " ^ commas lits ^ ")") ^ ",\n" ^
-      string_of_ar axiom_name ^ ").\n\n"
-  end;
+  in  (classes', multi_arity_clause cpairs)  end;
 
 
 (**** Produce TPTP files ****)
 
+fun string_of_clausename (cls_id, ax_name) =
+    clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id
+
 fun tptp_sign true s = s
   | tptp_sign false s = "~ " ^ s
 
@@ -594,8 +429,8 @@
   | tptp_of_arLit (TVarLit (c,str)) =
       tptp_sign false (make_type_class c ^ "(" ^ str ^ ")")
 
-fun tptp_arity_clause (ArityClause{axiom_name,conclLit,premLits,...}) =
-  tptp_cnf (string_of_ar axiom_name) "axiom"
+fun tptp_arity_clause (ArityClause {axiom_name, conclLit, premLits, ...}) =
+  tptp_cnf (arclause_prefix ^ ascii_of axiom_name) "axiom"
            (tptp_clause (map tptp_of_arLit (conclLit :: premLits)))
 
 fun tptp_classrelLits sub sup =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_hol_clause.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_hol_clause.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -32,19 +32,17 @@
   val conceal_skolem_somes :
     int -> (string * term) list -> term -> (string * term) list * term
   exception TRIVIAL
-  val make_conjecture_clauses : bool -> theory -> thm list -> hol_clause list
-  val make_axiom_clauses : bool -> theory ->
-       (thm * (axiom_name * hol_clause_id)) list -> (axiom_name * hol_clause) list
+  val make_conjecture_clauses : theory -> thm list -> hol_clause list
+  val make_axiom_clauses :
+    theory -> (thm * (axiom_name * hol_clause_id)) list
+    -> (axiom_name * hol_clause) list
   val type_wrapper_name : string
   val get_helper_clauses :
-    bool -> theory -> bool -> bool -> hol_clause list
-      -> (thm * (axiom_name * hol_clause_id)) list -> hol_clause list
+    theory -> bool -> bool -> hol_clause list
+    -> (thm * (axiom_name * hol_clause_id)) list -> hol_clause list
   val write_tptp_file : bool -> bool -> bool -> Path.T ->
     hol_clause list * hol_clause list * hol_clause list * hol_clause list *
     classrel_clause list * arity_clause list -> name_pool option * int
-  val write_dfg_file : bool -> bool -> Path.T ->
-    hol_clause list * hol_clause list * hol_clause list * hol_clause list *
-    classrel_clause list * arity_clause list -> name_pool option * int
 end
 
 structure Sledgehammer_HOL_Clause : SLEDGEHAMMER_HOL_CLAUSE =
@@ -98,65 +96,66 @@
 
 fun isTaut (HOLClause {literals,...}) = exists isTrue literals;
 
-fun type_of dfg (Type (a, Ts)) =
-    let val (folTypes,ts) = types_of dfg Ts in
-      (TyConstr (`(make_fixed_type_const dfg) a, folTypes), ts)
+fun type_of (Type (a, Ts)) =
+    let val (folTypes,ts) = types_of Ts in
+      (TyConstr (`make_fixed_type_const a, folTypes), ts)
     end
-  | type_of _ (tp as TFree (a, _)) = (TyFree (`make_fixed_type_var a), [tp])
-  | type_of _ (tp as TVar (x, _)) =
+  | type_of (tp as TFree (a, _)) = (TyFree (`make_fixed_type_var a), [tp])
+  | type_of (tp as TVar (x, _)) =
     (TyVar (make_schematic_type_var x, string_of_indexname x), [tp])
-and types_of dfg Ts =
-      let val (folTyps,ts) = ListPair.unzip (map (type_of dfg) Ts)
-      in  (folTyps, union_all ts)  end;
+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 dfg (Type (a, Ts)) =
-      TyConstr (`(make_fixed_type_const dfg) a, map (simp_type_of dfg) Ts)
-  | simp_type_of _ (TFree (a, _)) = TyFree (`make_fixed_type_var a)
-  | simp_type_of _ (TVar (x, _)) =
-    TyVar (make_schematic_type_var x, string_of_indexname x);
+fun simp_type_of (Type (a, Ts)) =
+      TyConstr (`make_fixed_type_const a, map simp_type_of Ts)
+  | simp_type_of (TFree (a, _)) = TyFree (`make_fixed_type_var a)
+  | simp_type_of (TVar (x, _)) =
+    TyVar (make_schematic_type_var x, string_of_indexname x)
 
 
 (* convert a Term.term (with combinators) into a combterm, also accummulate sort info *)
-fun combterm_of dfg thy (Const (c, T)) =
+fun combterm_of thy (Const (c, T)) =
       let
-        val (tp, ts) = type_of dfg T
+        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 dfg)
-        val c' = CombConst (`(make_fixed_const dfg) c, tp, tvar_list)
+          |> map simp_type_of
+        val c' = CombConst (`make_fixed_const c, tp, tvar_list)
       in  (c',ts)  end
-  | combterm_of dfg _ (Free(v, T)) =
-      let val (tp,ts) = type_of dfg T
+  | combterm_of _ (Free(v, T)) =
+      let val (tp,ts) = type_of T
           val v' = CombConst (`make_fixed_var v, tp, [])
       in  (v',ts)  end
-  | combterm_of dfg _ (Var(v, T)) =
-      let val (tp,ts) = type_of dfg T
+  | combterm_of _ (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_of dfg thy (P $ Q) =
-      let val (P',tsP) = combterm_of dfg thy P
-          val (Q',tsQ) = combterm_of dfg thy Q
-      in  (CombApp(P',Q'), union (op =) tsP tsQ)  end
-  | combterm_of _ _ (t as Abs _) = raise CLAUSE ("HOL clause", t);
-
-fun predicate_of dfg thy ((@{const Not} $ P), polarity) = predicate_of dfg thy (P, not polarity)
-  | predicate_of dfg thy (t,polarity) = (combterm_of dfg thy (Envir.eta_contract t), polarity);
+  | combterm_of thy (P $ Q) =
+      let val (P', tsP) = combterm_of thy P
+          val (Q', tsQ) = combterm_of thy Q
+      in  (CombApp (P', Q'), union (op =) tsP tsQ)  end
+  | combterm_of _ (t as Abs _) = raise CLAUSE ("HOL clause", t)
 
-fun literals_of_term1 dfg thy args (@{const Trueprop} $ P) = literals_of_term1 dfg thy args P
-  | literals_of_term1 dfg thy args (@{const "op |"} $ P $ Q) =
-      literals_of_term1 dfg thy (literals_of_term1 dfg thy args P) Q
-  | literals_of_term1 dfg thy (lits,ts) P =
-      let val ((pred,ts'),pol) = predicate_of dfg thy (P,true)
-      in
-          (Literal(pol,pred)::lits, union (op =) ts ts')
-      end;
+fun predicate_of thy ((@{const Not} $ P), polarity) =
+    predicate_of thy (P, not polarity)
+  | predicate_of thy (t, polarity) =
+    (combterm_of thy (Envir.eta_contract t), polarity)
 
-fun literals_of_term_dfg dfg thy P = literals_of_term1 dfg thy ([],[]) P;
-val literals_of_term = literals_of_term_dfg false;
+fun literals_of_term1 args thy (@{const Trueprop} $ P) =
+    literals_of_term1 args thy P
+  | literals_of_term1 args thy (@{const "op |"} $ 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
+      (Literal (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])) ^
@@ -199,11 +198,11 @@
 exception TRIVIAL;
 
 (* making axiom and conjecture clauses *)
-fun make_clause dfg thy (clause_id, axiom_name, kind, th) skolem_somes =
+fun make_clause thy (clause_id, axiom_name, kind, th) skolem_somes =
   let
     val (skolem_somes, t) =
       th |> prop_of |> conceal_skolem_somes clause_id skolem_somes
-    val (lits, ctypes_sorts) = literals_of_term_dfg dfg thy t
+    val (lits, ctypes_sorts) = literals_of_term thy t
   in
     if forall isFalse lits then
       raise TRIVIAL
@@ -213,29 +212,26 @@
                   kind = kind, literals = lits, ctypes_sorts = ctypes_sorts})
   end
 
-fun add_axiom_clause dfg thy (th, (name, id)) (skolem_somes, clss) =
+fun add_axiom_clause thy (th, (name, id)) (skolem_somes, clss) =
   let
-    val (skolem_somes, cls) =
-      make_clause dfg thy (id, name, Axiom, th) skolem_somes
+    val (skolem_somes, cls) = make_clause thy (id, name, Axiom, th) skolem_somes
   in (skolem_somes, clss |> not (isTaut cls) ? cons (name, cls)) end
 
-fun make_axiom_clauses dfg thy clauses =
-  ([], []) |> fold (add_axiom_clause dfg thy) clauses |> snd
+fun make_axiom_clauses thy clauses =
+  ([], []) |> fold (add_axiom_clause thy) clauses |> snd
 
-fun make_conjecture_clauses dfg thy =
+fun make_conjecture_clauses thy =
   let
     fun aux _ _ [] = []
       | aux n skolem_somes (th :: ths) =
         let
           val (skolem_somes, cls) =
-            make_clause dfg thy (n, "conjecture", Conjecture, th) skolem_somes
+            make_clause thy (n, "conjecture", Conjecture, th) skolem_somes
         in cls :: aux (n + 1) skolem_somes ths end
   in aux 0 [] end
 
 (**********************************************************************)
-(* convert clause into ATP specific formats:                          *)
-(* TPTP used by Vampire and E                                         *)
-(* DFG used by SPASS                                                  *)
+(* convert clause into TPTP format                                    *)
 (**********************************************************************)
 
 (*Result of a function type; no need to check that the argument type matches.*)
@@ -315,17 +311,11 @@
   string_of_combterm params c #>> prefix "hBOOL" o paren_pack o single
 
 fun string_of_predicate (params as (_, explicit_apply, _, cnh)) t =
-  case t of
-    (CombApp (CombApp (CombConst (("equal", _), _, _), t1), t2)) =>
-    (* DFG only: new TPTP prefers infix equality *)
-    pool_map (string_of_combterm params) [t1, t2]
-    #>> prefix "equal" o paren_pack
-  | _ =>
-    case #1 (strip_combterm_comb t) of
-      CombConst ((s, _), _, _) =>
-      (if needs_hBOOL explicit_apply cnh s then boolify else string_of_combterm)
-          params t
-    | _ => boolify params t
+  case #1 (strip_combterm_comb t) of
+    CombConst ((s, _), _, _) =>
+    (if needs_hBOOL explicit_apply cnh s then boolify else string_of_combterm)
+        params t
+  | _ => boolify params t
 
 
 (*** TPTP format ***)
@@ -353,7 +343,7 @@
 
 fun tptp_clause params (cls as HOLClause {axiom_name, clause_id, kind, ...})
                 pool =
-let
+  let
     val ((lits, tylits), pool) =
       tptp_type_literals params (kind = Conjecture) cls pool
   in
@@ -361,94 +351,6 @@
   end
 
 
-(*** DFG format ***)
-
-fun dfg_literal params (Literal (pos, pred)) =
-  string_of_predicate params pred #>> dfg_sign pos
-
-fun dfg_type_literals params pos (HOLClause {literals, ctypes_sorts, ...}) =
-  pool_map (dfg_literal params) literals
-  #>> rpair (map (dfg_of_type_literal pos)
-                 (type_literals_for_types ctypes_sorts))
-
-fun get_uvars (CombConst _) vars pool = (vars, pool)
-  | get_uvars (CombVar (name, _)) vars pool =
-    nice_name name pool |>> (fn var => var :: vars)
-  | get_uvars (CombApp (P, Q)) vars pool =
-    let val (vars, pool) = get_uvars P vars pool in get_uvars Q vars pool end
-
-fun get_uvars_l (Literal (_, c)) = get_uvars c [];
-
-fun dfg_vars (HOLClause {literals, ...}) =
-  pool_map get_uvars_l literals #>> union_all
-
-fun dfg_clause params (cls as HOLClause {axiom_name, clause_id, kind,
-                                         ctypes_sorts, ...}) pool =
-  let
-    val ((lits, tylits), pool) =
-      dfg_type_literals params (kind = Conjecture) cls pool
-    val (vars, pool) = dfg_vars cls pool
-    val tvars = get_tvar_strs ctypes_sorts
-  in
-    ((gen_dfg_cls (clause_id, axiom_name, kind, lits, tylits, tvars @ vars),
-      tylits), pool)
-  end
-
-
-(** For DFG format: accumulate function and predicate declarations **)
-
-fun add_types tvars = fold add_fol_type_funcs tvars
-
-fun add_decls (full_types, explicit_apply, cma, cnh)
-              (CombConst ((c, _), ctp, tvars)) (funcs, preds) =
-      (if c = "equal" then
-         (add_types tvars funcs, preds)
-       else
-         let val arity = min_arity_of cma c
-             val ntys = if not full_types then length tvars else 0
-             val addit = Symtab.update(c, arity + ntys)
-         in
-             if needs_hBOOL explicit_apply cnh c then
-               (add_types tvars (addit funcs), preds)
-             else
-               (add_types tvars funcs, addit preds)
-         end) |>> full_types ? add_fol_type_funcs ctp
-  | add_decls _ (CombVar (_, ctp)) (funcs, preds) =
-      (add_fol_type_funcs ctp funcs, preds)
-  | add_decls params (CombApp (P, Q)) decls =
-      decls |> add_decls params P |> add_decls params Q
-
-fun add_literal_decls params (Literal (_, c)) = add_decls params c
-
-fun add_clause_decls params (HOLClause {literals, ...}) decls =
-    fold (add_literal_decls params) literals decls
-    handle Symtab.DUP a => error ("function " ^ a ^ " has multiple arities")
-
-fun decls_of_clauses (params as (full_types, explicit_apply, _, _)) clauses
-                     arity_clauses =
-  let
-    val functab =
-      init_functab
-      |> fold Symtab.update [(type_wrapper_name, 2), ("hAPP", 2),
-                             ("tc_bool", 0)]
-      val predtab = Symtab.empty |> Symtab.update ("hBOOL", 1)
-      val (functab, predtab) =
-        (functab, predtab) |> fold (add_clause_decls params) clauses
-                           |>> fold add_arity_clause_funcs arity_clauses
-  in (Symtab.dest functab, Symtab.dest predtab) end
-
-fun add_clause_preds (HOLClause {ctypes_sorts, ...}) preds =
-  fold add_type_sort_preds ctypes_sorts preds
-  handle Symtab.DUP a => error ("predicate " ^ a ^ " has multiple arities")
-
-(*Higher-order clauses have only the predicates hBOOL and type classes.*)
-fun preds_of_clauses clauses clsrel_clauses arity_clauses =
-  Symtab.empty
-  |> fold add_clause_preds clauses
-  |> fold add_arity_clause_preds arity_clauses
-  |> fold add_classrel_clause_preds clsrel_clauses
-  |> Symtab.dest
-
 (**********************************************************************)
 (* write clauses to files                                             *)
 (**********************************************************************)
@@ -480,9 +382,9 @@
   Symtab.make (maps (maps (map (rpair 0) o fst))
                     [optional_helpers, optional_typed_helpers])
 
-fun get_helper_clauses dfg thy is_FO full_types conjectures axcls =
+fun get_helper_clauses thy is_FO full_types conjectures axcls =
   let
-    val axclauses = map snd (make_axiom_clauses dfg thy axcls)
+    val axclauses = map snd (make_axiom_clauses thy axcls)
     val ct = fold (fold count_clause) [conjectures, axclauses] init_counters
     fun is_needed c = the (Symtab.lookup ct c) > 0
     val cnfs =
@@ -492,7 +394,7 @@
                    if exists is_needed ss then cnf_helper_thms thy raw ths
                    else []))
       @ (if is_FO then [] else cnf_helper_thms thy false mandatory_helpers)
-  in map snd (make_axiom_clauses dfg thy cnfs) end
+  in map snd (make_axiom_clauses thy cnfs) end
 
 (*Find the minimal arity of each function mentioned in the term. Also, note which uses
   are not at top level, to see if hBOOL is needed.*)
@@ -521,7 +423,7 @@
                            (const_min_arity, const_needs_hBOOL) =
   fold count_constants_lit literals (const_min_arity, const_needs_hBOOL);
 
-fun display_arity explicit_apply const_needs_hBOOL (c,n) =
+fun display_arity explicit_apply const_needs_hBOOL (c, n) =
   trace_msg (fn () => "Constant: " ^ c ^
                 " arity:\t" ^ Int.toString n ^
                 (if needs_hBOOL explicit_apply const_needs_hBOOL c then
@@ -541,10 +443,6 @@
      in (const_min_arity, const_needs_hBOOL) end
   else (Symtab.empty, Symtab.empty);
 
-fun header () =
-  "% This file was generated by Isabelle (most likely Sledgehammer)\n" ^
-  "% " ^ timestamp () ^ "\n"
-
 (* TPTP format *)
 
 fun write_tptp_file readable_names full_types explicit_apply file clauses =
@@ -572,7 +470,8 @@
       + length helper_clss
     val _ =
       File.write_list file
-          (header () ::
+          ("% This file was generated by Isabelle (most likely Sledgehammer)\n\
+           \% " ^ timestamp () ^ "\n" ::
            section "Relevant fact" ax_clss @
            section "Class relationship" classrel_clss @
            section "Arity declaration" arity_clss @
@@ -581,53 +480,4 @@
            section "Type variable" tfree_clss)
   in (pool, conjecture_offset) end
 
-(* DFG format *)
-
-fun dfg_tfree_predicate s = (first_field "(" s |> the |> fst, 1)
-
-fun write_dfg_file full_types explicit_apply file clauses =
-  let
-    (* Some of the helper functions below are not name-pool-aware. However,
-       there's no point in adding name pool support if the DFG format is on its
-       way out anyway. *)
-    val pool = NONE
-    val (conjectures, axclauses, _, helper_clauses,
-      classrel_clauses, arity_clauses) = clauses
-    val (cma, cnh) = count_constants explicit_apply clauses
-    val params = (full_types, explicit_apply, cma, cnh)
-    val ((conjecture_clss, tfree_litss), pool) =
-      pool_map (dfg_clause params) conjectures pool |>> ListPair.unzip
-    val tfree_lits = union_all tfree_litss
-    val problem_name = Path.implode (Path.base file)
-    val (axstrs, pool) = pool_map (apfst fst oo dfg_clause params) axclauses pool
-    val tfree_clss = map dfg_tfree_clause tfree_lits
-    val tfree_preds = map dfg_tfree_predicate tfree_lits
-    val (helper_clauses_strs, pool) =
-      pool_map (apfst fst oo dfg_clause params) helper_clauses pool
-    val (funcs, cl_preds) =
-      decls_of_clauses params (helper_clauses @ conjectures @ axclauses) arity_clauses
-    val ty_preds = preds_of_clauses axclauses classrel_clauses arity_clauses
-    val preds = tfree_preds @ cl_preds @ ty_preds
-    val conjecture_offset =
-      length axclauses + length classrel_clauses + length arity_clauses
-      + length helper_clauses
-    val _ =
-      File.write_list file
-          (header () ::
-           string_of_start problem_name ::
-           string_of_descrip problem_name ::
-           string_of_symbols (string_of_funcs funcs)
-                             (string_of_preds preds) ::
-           "list_of_clauses(axioms, cnf).\n" ::
-           axstrs @
-           map dfg_classrel_clause classrel_clauses @
-           map dfg_arity_clause arity_clauses @
-           helper_clauses_strs @
-           ["end_of_list.\n\nlist_of_clauses(conjectures, cnf).\n"] @
-           conjecture_clss @
-           tfree_clss @
-           ["end_of_list.\n\n",
-            "end_problem.\n"])
-  in (pool, conjecture_offset) end
-
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -58,7 +58,6 @@
   {add = [], del = ns, only = false}
 fun only_relevance_override ns : relevance_override =
   {add = ns, del = [], only = true}
-val no_relevance_override = add_to_relevance_override []
 fun merge_relevance_override_pairwise (r1 : relevance_override)
                                       (r2 : relevance_override) =
   {add = #add r1 @ #add r2, del = #del r1 @ #del r2,
@@ -166,7 +165,7 @@
 val infinity_time_in_secs = 60 * 60 * 24 * 365
 val the_timeout = the_default (Time.fromSeconds infinity_time_in_secs)
 
-fun extract_params thy default_params override_params =
+fun extract_params default_params override_params =
   let
     val override_params = map unalias_raw_param override_params
     val raw_params = rev override_params @ rev default_params
@@ -216,7 +215,7 @@
      timeout = timeout, minimize_timeout = minimize_timeout}
   end
 
-fun get_params thy = extract_params thy (default_raw_params thy)
+fun get_params thy = extract_params (default_raw_params thy)
 fun default_params thy = get_params thy o map (apsnd single)
 
 val subgoal_count = Logic.count_prems o prop_of o #goal o Proof.goal
@@ -264,7 +263,6 @@
 val running_atpsN = "running_atps"
 val kill_atpsN = "kill_atps"
 val refresh_tptpN = "refresh_tptp"
-val helpN = "help"
 
 fun minimizize_raw_param_name "timeout" = "minimize_timeout"
   | minimizize_raw_param_name name = name
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -368,7 +368,7 @@
   | add_type_constraint _ = I
 
 fun is_positive_literal (@{const Not} $ _) = false
-  | is_positive_literal t = true
+  | is_positive_literal _ = true
 
 fun negate_term thy (Const (@{const_name All}, T) $ Abs (s, T', t')) =
     Const (@{const_name Ex}, T) $ Abs (s, T', negate_term thy t')
@@ -551,11 +551,11 @@
 (* Vampire is keen on producing these. *)
 fun is_trivial_formula (@{const Not} $ (Const (@{const_name "op ="}, _)
                                         $ t1 $ t2)) = (t1 aconv t2)
-  | is_trivial_formula t = false
+  | is_trivial_formula _ = false
 
-fun add_desired_line _ _ _ _ _ (line as Definition (num, _, _)) (j, lines) =
+fun add_desired_line _ _ _ _ (line as Definition (num, _, _)) (j, lines) =
     (j, line :: map (replace_deps_in_line (num, [])) lines)
-  | add_desired_line ctxt isar_shrink_factor conjecture_shape thm_names frees
+  | add_desired_line isar_shrink_factor conjecture_shape thm_names frees
                      (Inference (num, t, deps)) (j, lines) =
     (j + 1,
      if is_axiom_clause_number thm_names num orelse
@@ -667,7 +667,7 @@
 
 fun forall_vars t = fold_rev forall_of (map Var (Term.add_vars t [])) t
 
-fun step_for_line _ _ (Definition (num, t1, t2)) = Let (t1, t2)
+fun step_for_line _ _ (Definition (_, t1, t2)) = Let (t1, t2)
   | step_for_line _ _ (Inference (num, t, [])) = Assume ((raw_prefix, num), t)
   | step_for_line thm_names j (Inference (num, t, deps)) =
     Have (if j = 1 then [Show] else [], (raw_prefix, num),
@@ -683,7 +683,7 @@
       |> decode_lines ctxt full_types tfrees
       |> rpair [] |-> fold_rev (add_line conjecture_shape thm_names)
       |> rpair [] |-> fold_rev add_nontrivial_line
-      |> rpair (0, []) |-> fold_rev (add_desired_line ctxt isar_shrink_factor
+      |> rpair (0, []) |-> fold_rev (add_desired_line isar_shrink_factor
                                                conjecture_shape thm_names frees)
       |> snd
   in
@@ -722,7 +722,7 @@
         else if forall (curry (op =) (hd proof1) o hd) (tl proofs) then
           aux (hd proof1 :: proof_tail) (map tl proofs)
         else case hd proof1 of
-          Have ([], l, t, by) =>
+          Have ([], l, t, _) => (* FIXME: should we really ignore the "by"? *)
           if forall (fn Have ([], l', t', _) :: _ => (l, t) = (l', t')
                       | _ => false) (tl proofs) andalso
              not (exists (member (op =) (maps new_labels_of proofs))
@@ -755,7 +755,7 @@
         first_pass (proof, contra) |>> cons step
       | first_pass ((step as Let _) :: proof, contra) =
         first_pass (proof, contra) |>> cons step
-      | first_pass ((step as Assume (l as (_, num), t)) :: proof, contra) =
+      | first_pass ((step as Assume (l as (_, num), _)) :: proof, contra) =
         if member (op =) concl_ls l then
           first_pass (proof, contra ||> l = hd concl_ls ? cons step)
         else
@@ -992,6 +992,13 @@
         do_indent 0 ^ (if n <> 1 then "next" else "qed") ^ "\n"
   in do_proof end
 
+fun strip_subgoal goal i =
+  let
+    val (t, frees) = Logic.goal_params (prop_of goal) i
+    val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees)
+    val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
+  in (rev (map dest_Free frees), hyp_ts, concl_t) end
+
 fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
                     (other_params as (full_types, _, atp_proof, thm_names, goal,
                                       i)) =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Tue Jun 22 13:17:59 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Tue Jun 22 14:28:22 2010 +0200
@@ -6,7 +6,6 @@
 
 signature SLEDGEHAMMER_UTIL =
 sig
-  val pairf : ('a -> 'b) -> ('a -> 'c) -> 'a -> 'b * 'c
   val plural_s : int -> string
   val serial_commas : string -> string list -> string list
   val replace_all : string -> string -> string -> string
@@ -24,8 +23,6 @@
 structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
 struct
 
-fun pairf f g x = (f x, g x)
-
 fun plural_s n = if n = 1 then "" else "s"
 
 fun serial_commas _ [] = ["??"]