shorten a few file names
authorblanchet
Tue, 31 Aug 2010 23:46:23 +0200
changeset 38986 e34c1b09bb5e
parent 38985 162bbbea4e4d
child 38987 96fae8916d8b
shorten a few file names
src/HOL/IsaMakefile
src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML
src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML
src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML
src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML
src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
--- a/src/HOL/IsaMakefile	Tue Aug 31 23:43:23 2010 +0200
+++ b/src/HOL/IsaMakefile	Tue Aug 31 23:46:23 2010 +0200
@@ -319,10 +319,10 @@
   Tools/Sledgehammer/metis_clauses.ML \
   Tools/Sledgehammer/metis_tactics.ML \
   Tools/Sledgehammer/sledgehammer.ML \
-  Tools/Sledgehammer/sledgehammer_fact_filter.ML \
-  Tools/Sledgehammer/sledgehammer_fact_minimize.ML \
+  Tools/Sledgehammer/sledgehammer_filter.ML \
+  Tools/Sledgehammer/sledgehammer_minimize.ML \
   Tools/Sledgehammer/sledgehammer_isar.ML \
-  Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML \
+  Tools/Sledgehammer/sledgehammer_reconstruct.ML \
   Tools/Sledgehammer/sledgehammer_translate.ML \
   Tools/Sledgehammer/sledgehammer_util.ML \
   Tools/SMT/cvc3_solver.ML \
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML	Tue Aug 31 23:43:23 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,800 +0,0 @@
-(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
-    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
-    Author:     Jasmin Blanchette, TU Muenchen
-*)
-
-signature SLEDGEHAMMER_FACT_FILTER =
-sig
-  datatype locality = General | Intro | Elim | Simp | Local | Chained
-
-  type relevance_override =
-    {add: Facts.ref list,
-     del: Facts.ref list,
-     only: bool}
-
-  val trace : bool Unsynchronized.ref
-  val worse_irrel_freq : real Unsynchronized.ref
-  val higher_order_irrel_weight : real Unsynchronized.ref
-  val abs_rel_weight : real Unsynchronized.ref
-  val abs_irrel_weight : real Unsynchronized.ref
-  val skolem_irrel_weight : real Unsynchronized.ref
-  val intro_bonus : real Unsynchronized.ref
-  val elim_bonus : real Unsynchronized.ref
-  val simp_bonus : real Unsynchronized.ref
-  val local_bonus : real Unsynchronized.ref
-  val chained_bonus : real Unsynchronized.ref
-  val max_imperfect : real Unsynchronized.ref
-  val max_imperfect_exp : real Unsynchronized.ref
-  val threshold_divisor : real Unsynchronized.ref
-  val ridiculous_threshold : real Unsynchronized.ref
-  val name_thm_pairs_from_ref :
-    Proof.context -> unit Symtab.table -> thm list -> Facts.ref
-    -> ((string * locality) * thm) list
-  val relevant_facts :
-    Proof.context -> bool -> real * real -> int -> bool -> relevance_override
-    -> thm list -> term list -> term -> ((string * locality) * thm) list
-end;
-
-structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
-struct
-
-open Sledgehammer_Util
-
-val trace = Unsynchronized.ref false
-fun trace_msg msg = if !trace then tracing (msg ()) else ()
-
-(* experimental feature *)
-val term_patterns = false
-
-val respect_no_atp = true
-
-datatype locality = General | Intro | Elim | Simp | Local | Chained
-
-type relevance_override =
-  {add: Facts.ref list,
-   del: Facts.ref list,
-   only: bool}
-
-val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
-
-fun repair_name reserved multi j name =
-  (name |> Symtab.defined reserved name ? quote) ^
-  (if multi then "(" ^ Int.toString j ^ ")" else "")
-
-fun name_thm_pairs_from_ref ctxt reserved chained_ths xref =
-  let
-    val ths = ProofContext.get_fact ctxt xref
-    val name = Facts.string_of_ref xref
-    val multi = length ths > 1
-  in
-    (ths, (1, []))
-    |-> fold (fn th => fn (j, rest) =>
-                 (j + 1, ((repair_name reserved multi j name,
-                          if member Thm.eq_thm chained_ths th then Chained
-                          else General), th) :: rest))
-    |> snd
-  end
-
-(***************************************************************)
-(* Relevance Filtering                                         *)
-(***************************************************************)
-
-(*** constants with types ***)
-
-fun order_of_type (Type (@{type_name fun}, [T1, @{typ bool}])) =
-    order_of_type T1 (* cheat: pretend sets are first-order *)
-  | order_of_type (Type (@{type_name fun}, [T1, T2])) =
-    Int.max (order_of_type T1 + 1, order_of_type T2)
-  | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0
-  | order_of_type _ = 0
-
-(* An abstraction of Isabelle types and first-order terms *)
-datatype pattern = PVar | PApp of string * pattern list
-datatype ptype = PType of int * pattern list
-
-fun string_for_pattern PVar = "_"
-  | string_for_pattern (PApp (s, ps)) =
-    if null ps then s else s ^ string_for_patterns ps
-and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")"
-fun string_for_ptype (PType (_, ps)) = string_for_patterns ps
-
-(*Is the second type an instance of the first one?*)
-fun match_pattern (PVar, _) = true
-  | match_pattern (PApp _, PVar) = false
-  | match_pattern (PApp (s, ps), PApp (t, qs)) =
-    s = t andalso match_patterns (ps, qs)
-and match_patterns (_, []) = true
-  | match_patterns ([], _) = false
-  | match_patterns (p :: ps, q :: qs) =
-    match_pattern (p, q) andalso match_patterns (ps, qs)
-fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs)
-
-(* Is there a unifiable constant? *)
-fun pconst_mem f consts (s, ps) =
-  exists (curry (match_ptype o f) ps)
-         (map snd (filter (curry (op =) s o fst) consts))
-fun pconst_hyper_mem f const_tab (s, ps) =
-  exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s))
-
-fun pattern_for_type (Type (s, Ts)) = PApp (s, map pattern_for_type Ts)
-  | pattern_for_type (TFree (s, _)) = PApp (s, [])
-  | pattern_for_type (TVar _) = PVar
-
-fun pterm thy t =
-  case strip_comb t of
-    (Const x, ts) => PApp (pconst thy true x ts)
-  | (Free x, ts) => PApp (pconst thy false x ts)
-  | (Var x, []) => PVar
-  | _ => PApp ("?", [])  (* equivalence class of higher-order constructs *)
-(* Pairs a constant with the list of its type instantiations. *)
-and ptype thy const x ts =
-  (if const then map pattern_for_type (these (try (Sign.const_typargs thy) x))
-   else []) @
-  (if term_patterns then map (pterm thy) ts else [])
-and pconst thy const (s, T) ts = (s, ptype thy const (s, T) ts)
-and rich_ptype thy const (s, T) ts =
-  PType (order_of_type T, ptype thy const (s, T) ts)
-and rich_pconst thy const (s, T) ts = (s, rich_ptype thy const (s, T) ts)
-
-fun string_for_hyper_pconst (s, ps) =
-  s ^ "{" ^ commas (map string_for_ptype ps) ^ "}"
-
-val abs_name = "Sledgehammer.abs"
-val skolem_prefix = "Sledgehammer.sko"
-
-(* These are typically simplified away by "Meson.presimplify". Equality is
-   handled specially via "fequal". *)
-val boring_consts =
-  [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let},
-   @{const_name HOL.eq}]
-
-(* Add a pconstant to the table, but a [] entry means a standard
-   connective, which we ignore.*)
-fun add_pconst_to_table also_skolem (c, p) =
-  if member (op =) boring_consts c orelse
-     (not also_skolem andalso String.isPrefix skolem_prefix c) then
-    I
-  else
-    Symtab.map_default (c, [p]) (insert (op =) p)
-
-fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
-
-fun pconsts_in_terms thy also_skolems pos ts =
-  let
-    val flip = Option.map not
-    (* We include free variables, as well as constants, to handle locales. For
-       each quantifiers that must necessarily be skolemized by the ATP, we
-       introduce a fresh constant to simulate the effect of Skolemization. *)
-    fun do_const const (s, T) ts =
-      add_pconst_to_table also_skolems (rich_pconst thy const (s, T) ts)
-      #> fold do_term ts
-    and do_term t =
-      case strip_comb t of
-        (Const x, ts) => do_const true x ts
-      | (Free x, ts) => do_const false x ts
-      | (Abs (_, T, t'), ts) =>
-        (null ts
-         ? add_pconst_to_table true (abs_name, PType (order_of_type T + 1, [])))
-        #> fold do_term (t' :: ts)
-      | (_, ts) => fold do_term ts
-    fun do_quantifier will_surely_be_skolemized abs_T body_t =
-      do_formula pos body_t
-      #> (if also_skolems andalso will_surely_be_skolemized then
-            add_pconst_to_table true
-                         (gensym skolem_prefix, PType (order_of_type abs_T, []))
-          else
-            I)
-    and do_term_or_formula T =
-      if is_formula_type T then do_formula NONE else do_term
-    and do_formula pos t =
-      case t of
-        Const (@{const_name all}, _) $ Abs (_, T, t') =>
-        do_quantifier (pos = SOME false) T t'
-      | @{const "==>"} $ t1 $ t2 =>
-        do_formula (flip pos) t1 #> do_formula pos t2
-      | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
-        fold (do_term_or_formula T) [t1, t2]
-      | @{const Trueprop} $ t1 => do_formula pos t1
-      | @{const Not} $ t1 => do_formula (flip pos) t1
-      | Const (@{const_name All}, _) $ Abs (_, T, t') =>
-        do_quantifier (pos = SOME false) T t'
-      | Const (@{const_name Ex}, _) $ Abs (_, T, t') =>
-        do_quantifier (pos = SOME true) T t'
-      | @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
-      | @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
-      | @{const HOL.implies} $ t1 $ t2 =>
-        do_formula (flip pos) t1 #> do_formula pos t2
-      | Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 =>
-        fold (do_term_or_formula T) [t1, t2]
-      | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
-        $ t1 $ t2 $ t3 =>
-        do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3]
-      | Const (@{const_name Ex1}, _) $ Abs (_, T, t') =>
-        do_quantifier (is_some pos) T t'
-      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') =>
-        do_quantifier (pos = SOME false) T
-                      (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, t'))
-      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') =>
-        do_quantifier (pos = SOME true) T
-                      (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, t'))
-      | (t0 as Const (_, @{typ bool})) $ t1 =>
-        do_term t0 #> do_formula pos t1  (* theory constant *)
-      | _ => do_term t
-  in Symtab.empty |> fold (do_formula pos) ts end
-
-(*Inserts a dummy "constant" referring to the theory name, so that relevance
-  takes the given theory into account.*)
-fun theory_const_prop_of theory_relevant th =
-  if theory_relevant then
-    let
-      val name = Context.theory_name (theory_of_thm th)
-      val t = Const (name ^ ". 1", @{typ bool})
-    in t $ prop_of th end
-  else
-    prop_of th
-
-(**** Constant / Type Frequencies ****)
-
-(* A two-dimensional symbol table counts frequencies of constants. It's keyed
-   first by constant name and second by its list of type instantiations. For the
-   latter, we need a linear ordering on "pattern list". *)
-
-fun pattern_ord p =
-  case p of
-    (PVar, PVar) => EQUAL
-  | (PVar, PApp _) => LESS
-  | (PApp _, PVar) => GREATER
-  | (PApp q1, PApp q2) =>
-    prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
-fun ptype_ord (PType p, PType q) =
-  prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q)
-
-structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
-
-fun count_axiom_consts theory_relevant thy =
-  let
-    fun do_const const (s, T) ts =
-      (* Two-dimensional table update. Constant maps to types maps to count. *)
-      PType_Tab.map_default (rich_ptype thy const (s, T) ts, 0) (Integer.add 1)
-      |> Symtab.map_default (s, PType_Tab.empty)
-      #> fold do_term ts
-    and do_term t =
-      case strip_comb t of
-        (Const x, ts) => do_const true x ts
-      | (Free x, ts) => do_const false x ts
-      | (Abs (_, _, t'), ts) => fold do_term (t' :: ts)
-      | (_, ts) => fold do_term ts
-  in do_term o theory_const_prop_of theory_relevant o snd end
-
-
-(**** Actual Filtering Code ****)
-
-fun pow_int x 0 = 1.0
-  | pow_int x 1 = x
-  | pow_int x n = if n > 0 then x * pow_int x (n - 1) else pow_int x (n + 1) / x
-
-(*The frequency of a constant is the sum of those of all instances of its type.*)
-fun pconst_freq match const_tab (c, ps) =
-  PType_Tab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m)
-                 (the (Symtab.lookup const_tab c)) 0
-
-
-(* A surprising number of theorems contain only a few significant constants.
-   These include all induction rules, and other general theorems. *)
-
-(* "log" seems best in practice. A constant function of one ignores the constant
-   frequencies. Rare constants give more points if they are relevant than less
-   rare ones. *)
-fun rel_weight_for order freq = 1.0 + 2.0 / Math.ln (Real.fromInt freq + 1.0)
-
-(* FUDGE *)
-val worse_irrel_freq = Unsynchronized.ref 100.0
-val higher_order_irrel_weight = Unsynchronized.ref 1.05
-
-(* Irrelevant constants are treated differently. We associate lower penalties to
-   very rare constants and very common ones -- the former because they can't
-   lead to the inclusion of too many new facts, and the latter because they are
-   so common as to be of little interest. *)
-fun irrel_weight_for order freq =
-  let val (k, x) = !worse_irrel_freq |> `Real.ceil in
-    (if freq < k then Math.ln (Real.fromInt (freq + 1)) / Math.ln x
-     else rel_weight_for order freq / rel_weight_for order k)
-    * pow_int (!higher_order_irrel_weight) (order - 1)
-  end
-
-(* FUDGE *)
-val abs_rel_weight = Unsynchronized.ref 0.5
-val abs_irrel_weight = Unsynchronized.ref 2.0
-val skolem_irrel_weight = Unsynchronized.ref 0.75
-
-(* Computes a constant's weight, as determined by its frequency. *)
-fun generic_pconst_weight abs_weight skolem_weight weight_for f const_tab
-                          (c as (s, PType (m, _))) =
-  if s = abs_name then abs_weight
-  else if String.isPrefix skolem_prefix s then skolem_weight
-  else weight_for m (pconst_freq (match_ptype o f) const_tab c)
-
-fun rel_pconst_weight const_tab =
-  generic_pconst_weight (!abs_rel_weight) 0.0 rel_weight_for I const_tab
-fun irrel_pconst_weight const_tab =
-  generic_pconst_weight (!abs_irrel_weight) (!skolem_irrel_weight)
-                        irrel_weight_for swap const_tab
-
-(* FUDGE *)
-val intro_bonus = Unsynchronized.ref 0.15
-val elim_bonus = Unsynchronized.ref 0.15
-val simp_bonus = Unsynchronized.ref 0.15
-val local_bonus = Unsynchronized.ref 0.55
-val chained_bonus = Unsynchronized.ref 1.5
-
-fun locality_bonus General = 0.0
-  | locality_bonus Intro = !intro_bonus
-  | locality_bonus Elim = !elim_bonus
-  | locality_bonus Simp = !simp_bonus
-  | locality_bonus Local = !local_bonus
-  | locality_bonus Chained = !chained_bonus
-
-fun axiom_weight loc const_tab relevant_consts axiom_consts =
-  case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
-                    ||> filter_out (pconst_hyper_mem swap relevant_consts) of
-    ([], _) => 0.0
-  | (rel, irrel) =>
-    let
-      val irrel = irrel |> filter_out (pconst_mem swap rel)
-      val rel_weight =
-        0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
-      val irrel_weight =
-        ~ (locality_bonus loc)
-        |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
-      val res = rel_weight / (rel_weight + irrel_weight)
-    in if Real.isFinite res then res else 0.0 end
-
-(* FIXME: experiment
-fun debug_axiom_weight loc const_tab relevant_consts axiom_consts =
-  case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
-                    ||> filter_out (pconst_hyper_mem swap relevant_consts) of
-    ([], _) => 0.0
-  | (rel, irrel) =>
-    let
-      val irrel = irrel |> filter_out (pconst_mem swap rel)
-      val rels_weight =
-        0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
-      val irrels_weight =
-        ~ (locality_bonus loc)
-        |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
-val _ = tracing (PolyML.makestring ("REL: ", map (`(rel_pconst_weight const_tab)) rel))
-val _ = tracing (PolyML.makestring ("IRREL: ", map (`(irrel_pconst_weight const_tab)) irrel))
-      val res = rels_weight / (rels_weight + irrels_weight)
-    in if Real.isFinite res then res else 0.0 end
-*)
-
-fun pconsts_in_axiom thy t =
-  Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
-              (pconsts_in_terms thy true (SOME true) [t]) []
-fun pair_consts_axiom theory_relevant thy axiom =
-  case axiom |> snd |> theory_const_prop_of theory_relevant
-             |> pconsts_in_axiom thy of
-    [] => NONE
-  | consts => SOME ((axiom, consts), NONE)
-
-type annotated_thm =
-  (((unit -> string) * locality) * thm) * (string * ptype) list
-
-(* FUDGE *)
-val max_imperfect = Unsynchronized.ref 11.5
-val max_imperfect_exp = Unsynchronized.ref 1.0
-
-fun take_most_relevant max_relevant remaining_max
-                       (candidates : (annotated_thm * real) list) =
-  let
-    val max_imperfect =
-      Real.ceil (Math.pow (!max_imperfect,
-                    Math.pow (Real.fromInt remaining_max
-                              / Real.fromInt max_relevant, !max_imperfect_exp)))
-    val (perfect, imperfect) =
-      candidates |> sort (Real.compare o swap o pairself snd)
-                 |> take_prefix (fn (_, w) => w > 0.99999)
-    val ((accepts, more_rejects), rejects) =
-      chop max_imperfect imperfect |>> append perfect |>> chop remaining_max
-  in
-    trace_msg (fn () =>
-        "Actually passed (" ^ Int.toString (length accepts) ^ " of " ^
-        Int.toString (length candidates) ^ "): " ^
-        (accepts |> map (fn ((((name, _), _), _), weight) =>
-                            name () ^ " [" ^ Real.toString weight ^ "]")
-                 |> commas));
-    (accepts, more_rejects @ rejects)
-  end
-
-fun if_empty_replace_with_locality thy axioms loc tab =
-  if Symtab.is_empty tab then
-    pconsts_in_terms thy false (SOME false)
-        (map_filter (fn ((_, loc'), th) =>
-                        if loc' = loc then SOME (prop_of th) else NONE) axioms)
-  else
-    tab
-
-(* FUDGE *)
-val threshold_divisor = Unsynchronized.ref 2.0
-val ridiculous_threshold = Unsynchronized.ref 0.1
-
-fun relevance_filter ctxt threshold0 decay max_relevant theory_relevant
-                     ({add, del, ...} : relevance_override) axioms goal_ts =
-  let
-    val thy = ProofContext.theory_of ctxt
-    val const_tab =
-      fold (count_axiom_consts theory_relevant thy) axioms Symtab.empty
-    val goal_const_tab =
-      pconsts_in_terms thy false (SOME false) goal_ts
-      |> fold (if_empty_replace_with_locality thy axioms) [Chained, Local]
-    val add_thms = maps (ProofContext.get_fact ctxt) add
-    val del_thms = maps (ProofContext.get_fact ctxt) del
-    fun iter j remaining_max threshold rel_const_tab hopeless hopeful =
-      let
-        fun game_over rejects =
-          (* Add "add:" facts. *)
-          if null add_thms then
-            []
-          else
-            map_filter (fn ((p as (_, th), _), _) =>
-                           if member Thm.eq_thm add_thms th then SOME p
-                           else NONE) rejects
-        fun relevant [] rejects [] =
-            (* Nothing has been added this iteration. *)
-            if j = 0 andalso threshold >= !ridiculous_threshold then
-              (* First iteration? Try again. *)
-              iter 0 max_relevant (threshold / !threshold_divisor) rel_const_tab
-                   hopeless hopeful
-            else
-              game_over (rejects @ hopeless)
-          | relevant candidates rejects [] =
-            let
-              val (accepts, more_rejects) =
-                take_most_relevant max_relevant remaining_max candidates
-              val rel_const_tab' =
-                rel_const_tab
-                |> fold (add_pconst_to_table false) (maps (snd o fst) accepts)
-              fun is_dirty (c, _) =
-                Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
-              val (hopeful_rejects, hopeless_rejects) =
-                 (rejects @ hopeless, ([], []))
-                 |-> fold (fn (ax as (_, consts), old_weight) =>
-                              if exists is_dirty consts then
-                                apfst (cons (ax, NONE))
-                              else
-                                apsnd (cons (ax, old_weight)))
-                 |>> append (more_rejects
-                             |> map (fn (ax as (_, consts), old_weight) =>
-                                        (ax, if exists is_dirty consts then NONE
-                                             else SOME old_weight)))
-              val threshold =
-                1.0 - (1.0 - threshold)
-                      * Math.pow (decay, Real.fromInt (length accepts))
-              val remaining_max = remaining_max - length accepts
-            in
-              trace_msg (fn () => "New or updated constants: " ^
-                  commas (rel_const_tab' |> Symtab.dest
-                          |> subtract (op =) (rel_const_tab |> Symtab.dest)
-                          |> map string_for_hyper_pconst));
-              map (fst o fst) accepts @
-              (if remaining_max = 0 then
-                 game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects)
-               else
-                 iter (j + 1) remaining_max threshold rel_const_tab'
-                      hopeless_rejects hopeful_rejects)
-            end
-          | relevant candidates rejects
-                     (((ax as (((_, loc), th), axiom_consts)), cached_weight)
-                      :: hopeful) =
-            let
-              val weight =
-                case cached_weight of
-                  SOME w => w
-                | NONE => axiom_weight loc const_tab rel_const_tab axiom_consts
-(* FIXME: experiment
-val name = fst (fst (fst ax)) ()
-val _ = if String.isSubstring "positive_minus" name orelse String.isSubstring "not_exp_le_zero" name then
-tracing ("*** " ^ name ^ PolyML.makestring (debug_axiom_weight loc const_tab rel_const_tab axiom_consts))
-else
-()
-*)
-            in
-              if weight >= threshold then
-                relevant ((ax, weight) :: candidates) rejects hopeful
-              else
-                relevant candidates ((ax, weight) :: rejects) hopeful
-            end
-        in
-          trace_msg (fn () =>
-              "ITERATION " ^ string_of_int j ^ ": current threshold: " ^
-              Real.toString threshold ^ ", constants: " ^
-              commas (rel_const_tab |> Symtab.dest
-                      |> filter (curry (op <>) [] o snd)
-                      |> map string_for_hyper_pconst));
-          relevant [] [] hopeful
-        end
-  in
-    axioms |> filter_out (member Thm.eq_thm del_thms o snd)
-           |> map_filter (pair_consts_axiom theory_relevant thy)
-           |> iter 0 max_relevant threshold0 goal_const_tab []
-           |> tap (fn res => trace_msg (fn () =>
-                                "Total relevant: " ^ Int.toString (length res)))
-  end
-
-
-(***************************************************************)
-(* Retrieving and filtering lemmas                             *)
-(***************************************************************)
-
-(*** retrieve lemmas and filter them ***)
-
-(*Reject theorems with names like "List.filter.filter_list_def" or
-  "Accessible_Part.acc.defs", as these are definitions arising from packages.*)
-fun is_package_def a =
-  let val names = Long_Name.explode a
-  in
-     length names > 2 andalso
-     not (hd names = "local") andalso
-     String.isSuffix "_def" a  orelse  String.isSuffix "_defs" a
-  end;
-
-fun mk_fact_table f xs =
-  fold (Termtab.update o `(prop_of o f)) xs Termtab.empty
-fun uniquify xs = Termtab.fold (cons o snd) (mk_fact_table snd xs) []
-
-(* FIXME: put other record thms here, or declare as "no_atp" *)
-val multi_base_blacklist =
-  ["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits",
-   "split_asm", "cases", "ext_cases", "eq.simps", "eq.refl", "nchotomy",
-   "case_cong", "weak_case_cong"]
-  |> map (prefix ".")
-
-val max_lambda_nesting = 3
-
-fun term_has_too_many_lambdas max (t1 $ t2) =
-    exists (term_has_too_many_lambdas max) [t1, t2]
-  | term_has_too_many_lambdas max (Abs (_, _, t)) =
-    max = 0 orelse term_has_too_many_lambdas (max - 1) t
-  | term_has_too_many_lambdas _ _ = false
-
-(* Don't count nested lambdas at the level of formulas, since they are
-   quantifiers. *)
-fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) =
-    formula_has_too_many_lambdas (T :: Ts) t
-  | formula_has_too_many_lambdas Ts t =
-    if is_formula_type (fastype_of1 (Ts, t)) then
-      exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t))
-    else
-      term_has_too_many_lambdas max_lambda_nesting t
-
-(* The max apply depth of any "metis" call in "Metis_Examples" (on 2007-10-31)
-   was 11. *)
-val max_apply_depth = 15
-
-fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
-  | apply_depth (Abs (_, _, t)) = apply_depth t
-  | apply_depth _ = 0
-
-fun is_formula_too_complex t =
-  apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t
-
-val exists_sledgehammer_const =
-  exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
-
-(* FIXME: make more reliable *)
-val exists_low_level_class_const =
-  exists_Const (fn (s, _) =>
-     String.isSubstring (Long_Name.separator ^ "class" ^ Long_Name.separator) s)
-
-fun is_metastrange_theorem th =
-  case head_of (concl_of th) of
-      Const (a, _) => (a <> @{const_name Trueprop} andalso
-                       a <> @{const_name "=="})
-    | _ => false
-
-fun is_that_fact th =
-  String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th)
-  andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
-                           | _ => false) (prop_of th)
-
-val type_has_top_sort =
-  exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
-
-(**** Predicates to detect unwanted facts (prolific or likely to cause
-      unsoundness) ****)
-
-(* Too general means, positive equality literal with a variable X as one
-   operand, when X does not occur properly in the other operand. This rules out
-   clearly inconsistent facts such as X = a | X = b, though it by no means
-   guarantees soundness. *)
-
-(* Unwanted equalities are those between a (bound or schematic) variable that
-   does not properly occur in the second operand. *)
-val is_exhaustive_finite =
-  let
-    fun is_bad_equal (Var z) t =
-        not (exists_subterm (fn Var z' => z = z' | _ => false) t)
-      | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
-      | is_bad_equal _ _ = false
-    fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
-    fun do_formula pos t =
-      case (pos, t) of
-        (_, @{const Trueprop} $ t1) => do_formula pos t1
-      | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
-        do_formula pos t'
-      | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
-        do_formula pos t'
-      | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
-        do_formula pos t'
-      | (_, @{const "==>"} $ t1 $ t2) =>
-        do_formula (not pos) t1 andalso
-        (t2 = @{prop False} orelse do_formula pos t2)
-      | (_, @{const HOL.implies} $ t1 $ t2) =>
-        do_formula (not pos) t1 andalso
-        (t2 = @{const False} orelse do_formula pos t2)
-      | (_, @{const Not} $ t1) => do_formula (not pos) t1
-      | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
-      | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
-      | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
-      | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
-      | _ => false
-  in do_formula true end
-
-fun has_bound_or_var_of_type tycons =
-  exists_subterm (fn Var (_, Type (s, _)) => member (op =) tycons s
-                   | Abs (_, Type (s, _), _) => member (op =) tycons s
-                   | _ => false)
-
-(* Facts are forbidden to contain variables of these types. The typical reason
-   is that they lead to unsoundness. Note that "unit" satisfies numerous
-   equations like "?x = ()". The resulting clauses will have no type constraint,
-   yielding false proofs. Even "bool" leads to many unsound proofs, though only
-   for higher-order problems. *)
-val dangerous_types = [@{type_name unit}, @{type_name bool}, @{type_name prop}];
-
-(* Facts containing variables of type "unit" or "bool" or of the form
-   "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
-   are omitted. *)
-fun is_dangerous_term full_types t =
-  not full_types andalso
-  let val t = transform_elim_term t in
-    has_bound_or_var_of_type dangerous_types t orelse
-    is_exhaustive_finite t
-  end
-
-fun is_theorem_bad_for_atps full_types thm =
-  let val t = prop_of thm in
-    is_formula_too_complex t orelse exists_type type_has_top_sort t orelse
-    is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse
-    exists_low_level_class_const t orelse is_metastrange_theorem thm orelse
-    is_that_fact thm
-  end
-
-fun clasimpset_rules_of ctxt =
-  let
-    val {safeIs, safeEs, hazIs, hazEs, ...} = ctxt |> claset_of |> rep_cs
-    val intros = safeIs @ hazIs
-    val elims = map Classical.classical_rule (safeEs @ hazEs)
-    val simps = ctxt |> simpset_of |> dest_ss |> #simps |> map snd
-  in (mk_fact_table I intros, mk_fact_table I elims, mk_fact_table I simps) end
-
-fun all_name_thms_pairs ctxt reserved full_types add_thms chained_ths =
-  let
-    val thy = ProofContext.theory_of ctxt
-    val global_facts = PureThy.facts_of thy
-    val local_facts = ProofContext.facts_of ctxt
-    val named_locals = local_facts |> Facts.dest_static []
-    val is_chained = member Thm.eq_thm chained_ths
-    val (intros, elims, simps) =
-      if exists (curry (op <) 0.0) [!intro_bonus, !elim_bonus, !simp_bonus] then
-        clasimpset_rules_of ctxt
-      else
-        (Termtab.empty, Termtab.empty, Termtab.empty)
-    (* Unnamed nonchained formulas with schematic variables are omitted, because
-       they are rejected by the backticks (`...`) parser for some reason. *)
-    fun is_good_unnamed_local th =
-      not (Thm.has_name_hint th) andalso
-      (not (exists_subterm is_Var (prop_of th)) orelse (is_chained th)) andalso
-      forall (fn (_, ths) => not (member Thm.eq_thm ths th)) named_locals
-    val unnamed_locals =
-      union Thm.eq_thm (Facts.props local_facts) chained_ths
-      |> filter is_good_unnamed_local |> map (pair "" o single)
-    val full_space =
-      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
-    fun add_facts global foldx facts =
-      foldx (fn (name0, ths) =>
-        if name0 <> "" andalso
-           forall (not o member Thm.eq_thm add_thms) ths andalso
-           (Facts.is_concealed facts name0 orelse
-            (respect_no_atp andalso is_package_def name0) orelse
-            exists (fn s => String.isSuffix s name0) multi_base_blacklist orelse
-            String.isSuffix "_def_raw" (* FIXME: crude hack *) name0) then
-          I
-        else
-          let
-            val multi = length ths > 1
-            fun backquotify th =
-              "`" ^ Print_Mode.setmp [Print_Mode.input]
-                                 (Syntax.string_of_term ctxt) (prop_of th) ^ "`"
-              |> String.translate (fn c => if Char.isPrint c then str c else "")
-              |> simplify_spaces
-            fun check_thms a =
-              case try (ProofContext.get_thms ctxt) a of
-                NONE => false
-              | SOME ths' => Thm.eq_thms (ths, ths')
-          in
-            pair 1
-            #> fold (fn th => fn (j, rest) =>
-                 (j + 1,
-                  if is_theorem_bad_for_atps full_types th andalso
-                     not (member Thm.eq_thm add_thms th) then
-                    rest
-                  else
-                    (((fn () =>
-                          if name0 = "" then
-                            th |> backquotify
-                          else
-                            let
-                              val name1 = Facts.extern facts name0
-                              val name2 = Name_Space.extern full_space name0
-                            in
-                              case find_first check_thms [name1, name2, name0] of
-                                SOME name => repair_name reserved multi j name
-                              | NONE => ""
-                            end),
-                      let val t = prop_of th in
-                        if is_chained th then Chained
-                        else if not global then Local
-                        else if Termtab.defined intros t then Intro
-                        else if Termtab.defined elims t then Elim
-                        else if Termtab.defined simps t then Simp
-                        else General
-                      end),
-                      (multi, th)) :: rest)) ths
-            #> snd
-          end)
-  in
-    [] |> add_facts false fold local_facts (unnamed_locals @ named_locals)
-       |> add_facts true Facts.fold_static global_facts global_facts
-  end
-
-(* The single-name theorems go after the multiple-name ones, so that single
-   names are preferred when both are available. *)
-fun name_thm_pairs ctxt respect_no_atp =
-  List.partition (fst o snd) #> op @ #> map (apsnd snd)
-  #> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd)
-
-(***************************************************************)
-(* ATP invocation methods setup                                *)
-(***************************************************************)
-
-fun relevant_facts ctxt full_types (threshold0, threshold1) max_relevant
-                   theory_relevant (relevance_override as {add, del, only})
-                   chained_ths hyp_ts concl_t =
-  let
-    val decay = Math.pow ((1.0 - threshold1) / (1.0 - threshold0),
-                          1.0 / Real.fromInt (max_relevant + 1))
-    val add_thms = maps (ProofContext.get_fact ctxt) add
-    val reserved = reserved_isar_keyword_table ()
-    val axioms =
-      (if only then
-         maps (map (fn ((name, loc), th) => ((K name, loc), (true, th)))
-               o name_thm_pairs_from_ref ctxt reserved chained_ths) add
-       else
-         all_name_thms_pairs ctxt reserved full_types add_thms chained_ths)
-      |> name_thm_pairs ctxt (respect_no_atp andalso not only)
-      |> uniquify
-  in
-    trace_msg (fn () => "Considering " ^ Int.toString (length axioms) ^
-                        " theorems");
-    (if threshold0 > 1.0 orelse threshold0 > threshold1 then
-       []
-     else if threshold0 < 0.0 then
-       axioms
-     else
-       relevance_filter ctxt threshold0 decay max_relevant theory_relevant
-                        relevance_override axioms (concl_t :: hyp_ts))
-    |> map (apfst (apfst (fn f => f ())))
-  end
-
-end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML	Tue Aug 31 23:43:23 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,162 +0,0 @@
-(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML
-    Author:     Philipp Meyer, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-
-Minimization of theorem list for Metis using automatic theorem provers.
-*)
-
-signature SLEDGEHAMMER_FACT_MINIMIZE =
-sig
-  type locality = Sledgehammer_Fact_Filter.locality
-  type params = Sledgehammer.params
-
-  val minimize_theorems :
-    params -> int -> int -> Proof.state -> ((string * locality) * thm list) list
-    -> ((string * locality) * thm list) list option * string
-  val run_minimize : params -> int -> Facts.ref list -> Proof.state -> unit
-end;
-
-structure Sledgehammer_Fact_Minimize : SLEDGEHAMMER_FACT_MINIMIZE =
-struct
-
-open ATP_Systems
-open Sledgehammer_Util
-open Sledgehammer_Fact_Filter
-open Sledgehammer_Proof_Reconstruct
-open Sledgehammer
-
-(* wrapper for calling external prover *)
-
-fun string_for_failure Unprovable = "Unprovable."
-  | string_for_failure TimedOut = "Timed out."
-  | string_for_failure _ = "Unknown error."
-
-fun n_theorems names =
-  let val n = length names in
-    string_of_int n ^ " theorem" ^ plural_s n ^
-    (if n > 0 then
-       ": " ^ (names |> map fst
-                     |> sort_distinct string_ord |> space_implode " ")
-     else
-       "")
-  end
-
-fun test_theorems ({debug, verbose, overlord, atps, full_types, isar_proof,
-                    isar_shrink_factor, ...} : params)
-                  (prover : prover) explicit_apply timeout subgoal state
-                  axioms =
-  let
-    val _ =
-      priority ("Testing " ^ n_theorems (map fst axioms) ^ "...")
-    val params =
-      {blocking = true, debug = debug, verbose = verbose, overlord = overlord,
-       atps = atps, full_types = full_types, explicit_apply = explicit_apply,
-       relevance_thresholds = (1.01, 1.01), max_relevant = NONE,
-       theory_relevant = NONE, isar_proof = isar_proof,
-       isar_shrink_factor = isar_shrink_factor, timeout = timeout, expect = ""}
-    val axioms = maps (fn (n, ths) => map (pair n) ths) axioms
-    val {context = ctxt, facts, goal} = Proof.goal state
-    val problem =
-     {subgoal = subgoal, goal = (ctxt, (facts, goal)),
-      relevance_override = {add = [], del = [], only = false},
-      axioms = SOME axioms}
-    val result as {outcome, used_thm_names, ...} = prover params (K "") problem
-  in
-    priority (case outcome of
-                NONE =>
-                if length used_thm_names = length axioms then
-                  "Found proof."
-                else
-                  "Found proof with " ^ n_theorems used_thm_names ^ "."
-              | SOME failure => string_for_failure failure);
-    result
-  end
-
-(* minimalization of thms *)
-
-fun filter_used_facts used = filter (member (op =) used o fst)
-
-fun sublinear_minimize _ [] p = p
-  | sublinear_minimize test (x :: xs) (seen, result) =
-    case test (xs @ seen) of
-      result as {outcome = NONE, proof, used_thm_names, ...} : prover_result =>
-      sublinear_minimize test (filter_used_facts used_thm_names xs)
-                         (filter_used_facts used_thm_names seen, result)
-    | _ => sublinear_minimize test xs (x :: seen, result)
-
-(* Give the ATP some slack. The ATP gets further slack because the Sledgehammer
-   preprocessing time is included in the estimate below but isn't part of the
-   timeout. *)
-val fudge_msecs = 1000
-
-fun minimize_theorems {atps = [], ...} _ _ _ _ = error "No ATP is set."
-  | minimize_theorems (params as {debug, atps = atp :: _, full_types,
-                                  isar_proof, isar_shrink_factor, timeout, ...})
-                      i n state axioms =
-  let
-    val thy = Proof.theory_of state
-    val prover = get_prover_fun thy atp
-    val msecs = Time.toMilliseconds timeout
-    val _ = priority ("Sledgehammer minimize: ATP " ^ quote atp ^ ".")
-    val {context = ctxt, goal, ...} = Proof.goal state
-    val (_, hyp_ts, concl_t) = strip_subgoal goal i
-    val explicit_apply =
-      not (forall (Meson.is_fol_term thy)
-                  (concl_t :: hyp_ts @ maps (map prop_of o snd) axioms))
-    fun do_test timeout =
-      test_theorems params prover explicit_apply timeout i state
-    val timer = Timer.startRealTimer ()
-  in
-    (case do_test timeout axioms of
-       result as {outcome = NONE, pool, used_thm_names,
-                  conjecture_shape, ...} =>
-       let
-         val time = Timer.checkRealTimer timer
-         val new_timeout =
-           Int.min (msecs, Time.toMilliseconds time + fudge_msecs)
-           |> Time.fromMilliseconds
-         val (min_thms, {proof, axiom_names, ...}) =
-           sublinear_minimize (do_test new_timeout)
-               (filter_used_facts used_thm_names axioms) ([], result)
-         val n = length min_thms
-         val _ = priority (cat_lines
-           ["Minimized: " ^ string_of_int n ^ " theorem" ^ plural_s n] ^
-            (case length (filter (curry (op =) Chained o snd o fst) min_thms) of
-               0 => ""
-             | n => " (including " ^ Int.toString n ^ " chained)") ^ ".")
-       in
-         (SOME min_thms,
-          proof_text isar_proof
-              (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
-              (full_types, K "", proof, axiom_names, goal, i) |> fst)
-       end
-     | {outcome = SOME TimedOut, ...} =>
-       (NONE, "Timeout: You can increase the time limit using the \"timeout\" \
-              \option (e.g., \"timeout = " ^
-              string_of_int (10 + msecs div 1000) ^ " s\").")
-     | {outcome = SOME UnknownError, ...} =>
-       (* Failure sometimes mean timeout, unfortunately. *)
-       (NONE, "Failure: No proof was found with the current time limit. You \
-              \can increase the time limit using the \"timeout\" \
-              \option (e.g., \"timeout = " ^
-              string_of_int (10 + msecs div 1000) ^ " s\").")
-     | {message, ...} => (NONE, "ATP error: " ^ message))
-    handle ERROR msg => (NONE, "Error: " ^ msg)
-  end
-
-fun run_minimize params i refs state =
-  let
-    val ctxt = Proof.context_of state
-    val reserved = reserved_isar_keyword_table ()
-    val chained_ths = #facts (Proof.goal state)
-    val axioms =
-      maps (map (apsnd single)
-            o name_thm_pairs_from_ref ctxt reserved chained_ths) refs
-  in
-    case subgoal_count state of
-      0 => priority "No subgoal!"
-    | n =>
-      (kill_atps (); priority (#2 (minimize_theorems params i n state axioms)))
-  end
-
-end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML	Tue Aug 31 23:46:23 2010 +0200
@@ -0,0 +1,800 @@
+(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
+    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
+    Author:     Jasmin Blanchette, TU Muenchen
+*)
+
+signature SLEDGEHAMMER_FACT_FILTER =
+sig
+  datatype locality = General | Intro | Elim | Simp | Local | Chained
+
+  type relevance_override =
+    {add: Facts.ref list,
+     del: Facts.ref list,
+     only: bool}
+
+  val trace : bool Unsynchronized.ref
+  val worse_irrel_freq : real Unsynchronized.ref
+  val higher_order_irrel_weight : real Unsynchronized.ref
+  val abs_rel_weight : real Unsynchronized.ref
+  val abs_irrel_weight : real Unsynchronized.ref
+  val skolem_irrel_weight : real Unsynchronized.ref
+  val intro_bonus : real Unsynchronized.ref
+  val elim_bonus : real Unsynchronized.ref
+  val simp_bonus : real Unsynchronized.ref
+  val local_bonus : real Unsynchronized.ref
+  val chained_bonus : real Unsynchronized.ref
+  val max_imperfect : real Unsynchronized.ref
+  val max_imperfect_exp : real Unsynchronized.ref
+  val threshold_divisor : real Unsynchronized.ref
+  val ridiculous_threshold : real Unsynchronized.ref
+  val name_thm_pairs_from_ref :
+    Proof.context -> unit Symtab.table -> thm list -> Facts.ref
+    -> ((string * locality) * thm) list
+  val relevant_facts :
+    Proof.context -> bool -> real * real -> int -> bool -> relevance_override
+    -> thm list -> term list -> term -> ((string * locality) * thm) list
+end;
+
+structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
+struct
+
+open Sledgehammer_Util
+
+val trace = Unsynchronized.ref false
+fun trace_msg msg = if !trace then tracing (msg ()) else ()
+
+(* experimental feature *)
+val term_patterns = false
+
+val respect_no_atp = true
+
+datatype locality = General | Intro | Elim | Simp | Local | Chained
+
+type relevance_override =
+  {add: Facts.ref list,
+   del: Facts.ref list,
+   only: bool}
+
+val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
+
+fun repair_name reserved multi j name =
+  (name |> Symtab.defined reserved name ? quote) ^
+  (if multi then "(" ^ Int.toString j ^ ")" else "")
+
+fun name_thm_pairs_from_ref ctxt reserved chained_ths xref =
+  let
+    val ths = ProofContext.get_fact ctxt xref
+    val name = Facts.string_of_ref xref
+    val multi = length ths > 1
+  in
+    (ths, (1, []))
+    |-> fold (fn th => fn (j, rest) =>
+                 (j + 1, ((repair_name reserved multi j name,
+                          if member Thm.eq_thm chained_ths th then Chained
+                          else General), th) :: rest))
+    |> snd
+  end
+
+(***************************************************************)
+(* Relevance Filtering                                         *)
+(***************************************************************)
+
+(*** constants with types ***)
+
+fun order_of_type (Type (@{type_name fun}, [T1, @{typ bool}])) =
+    order_of_type T1 (* cheat: pretend sets are first-order *)
+  | order_of_type (Type (@{type_name fun}, [T1, T2])) =
+    Int.max (order_of_type T1 + 1, order_of_type T2)
+  | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0
+  | order_of_type _ = 0
+
+(* An abstraction of Isabelle types and first-order terms *)
+datatype pattern = PVar | PApp of string * pattern list
+datatype ptype = PType of int * pattern list
+
+fun string_for_pattern PVar = "_"
+  | string_for_pattern (PApp (s, ps)) =
+    if null ps then s else s ^ string_for_patterns ps
+and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")"
+fun string_for_ptype (PType (_, ps)) = string_for_patterns ps
+
+(*Is the second type an instance of the first one?*)
+fun match_pattern (PVar, _) = true
+  | match_pattern (PApp _, PVar) = false
+  | match_pattern (PApp (s, ps), PApp (t, qs)) =
+    s = t andalso match_patterns (ps, qs)
+and match_patterns (_, []) = true
+  | match_patterns ([], _) = false
+  | match_patterns (p :: ps, q :: qs) =
+    match_pattern (p, q) andalso match_patterns (ps, qs)
+fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs)
+
+(* Is there a unifiable constant? *)
+fun pconst_mem f consts (s, ps) =
+  exists (curry (match_ptype o f) ps)
+         (map snd (filter (curry (op =) s o fst) consts))
+fun pconst_hyper_mem f const_tab (s, ps) =
+  exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s))
+
+fun pattern_for_type (Type (s, Ts)) = PApp (s, map pattern_for_type Ts)
+  | pattern_for_type (TFree (s, _)) = PApp (s, [])
+  | pattern_for_type (TVar _) = PVar
+
+fun pterm thy t =
+  case strip_comb t of
+    (Const x, ts) => PApp (pconst thy true x ts)
+  | (Free x, ts) => PApp (pconst thy false x ts)
+  | (Var x, []) => PVar
+  | _ => PApp ("?", [])  (* equivalence class of higher-order constructs *)
+(* Pairs a constant with the list of its type instantiations. *)
+and ptype thy const x ts =
+  (if const then map pattern_for_type (these (try (Sign.const_typargs thy) x))
+   else []) @
+  (if term_patterns then map (pterm thy) ts else [])
+and pconst thy const (s, T) ts = (s, ptype thy const (s, T) ts)
+and rich_ptype thy const (s, T) ts =
+  PType (order_of_type T, ptype thy const (s, T) ts)
+and rich_pconst thy const (s, T) ts = (s, rich_ptype thy const (s, T) ts)
+
+fun string_for_hyper_pconst (s, ps) =
+  s ^ "{" ^ commas (map string_for_ptype ps) ^ "}"
+
+val abs_name = "Sledgehammer.abs"
+val skolem_prefix = "Sledgehammer.sko"
+
+(* These are typically simplified away by "Meson.presimplify". Equality is
+   handled specially via "fequal". *)
+val boring_consts =
+  [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let},
+   @{const_name HOL.eq}]
+
+(* Add a pconstant to the table, but a [] entry means a standard
+   connective, which we ignore.*)
+fun add_pconst_to_table also_skolem (c, p) =
+  if member (op =) boring_consts c orelse
+     (not also_skolem andalso String.isPrefix skolem_prefix c) then
+    I
+  else
+    Symtab.map_default (c, [p]) (insert (op =) p)
+
+fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
+
+fun pconsts_in_terms thy also_skolems pos ts =
+  let
+    val flip = Option.map not
+    (* We include free variables, as well as constants, to handle locales. For
+       each quantifiers that must necessarily be skolemized by the ATP, we
+       introduce a fresh constant to simulate the effect of Skolemization. *)
+    fun do_const const (s, T) ts =
+      add_pconst_to_table also_skolems (rich_pconst thy const (s, T) ts)
+      #> fold do_term ts
+    and do_term t =
+      case strip_comb t of
+        (Const x, ts) => do_const true x ts
+      | (Free x, ts) => do_const false x ts
+      | (Abs (_, T, t'), ts) =>
+        (null ts
+         ? add_pconst_to_table true (abs_name, PType (order_of_type T + 1, [])))
+        #> fold do_term (t' :: ts)
+      | (_, ts) => fold do_term ts
+    fun do_quantifier will_surely_be_skolemized abs_T body_t =
+      do_formula pos body_t
+      #> (if also_skolems andalso will_surely_be_skolemized then
+            add_pconst_to_table true
+                         (gensym skolem_prefix, PType (order_of_type abs_T, []))
+          else
+            I)
+    and do_term_or_formula T =
+      if is_formula_type T then do_formula NONE else do_term
+    and do_formula pos t =
+      case t of
+        Const (@{const_name all}, _) $ Abs (_, T, t') =>
+        do_quantifier (pos = SOME false) T t'
+      | @{const "==>"} $ t1 $ t2 =>
+        do_formula (flip pos) t1 #> do_formula pos t2
+      | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
+        fold (do_term_or_formula T) [t1, t2]
+      | @{const Trueprop} $ t1 => do_formula pos t1
+      | @{const Not} $ t1 => do_formula (flip pos) t1
+      | Const (@{const_name All}, _) $ Abs (_, T, t') =>
+        do_quantifier (pos = SOME false) T t'
+      | Const (@{const_name Ex}, _) $ Abs (_, T, t') =>
+        do_quantifier (pos = SOME true) T t'
+      | @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
+      | @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
+      | @{const HOL.implies} $ t1 $ t2 =>
+        do_formula (flip pos) t1 #> do_formula pos t2
+      | Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 =>
+        fold (do_term_or_formula T) [t1, t2]
+      | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
+        $ t1 $ t2 $ t3 =>
+        do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3]
+      | Const (@{const_name Ex1}, _) $ Abs (_, T, t') =>
+        do_quantifier (is_some pos) T t'
+      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') =>
+        do_quantifier (pos = SOME false) T
+                      (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, t'))
+      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') =>
+        do_quantifier (pos = SOME true) T
+                      (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, t'))
+      | (t0 as Const (_, @{typ bool})) $ t1 =>
+        do_term t0 #> do_formula pos t1  (* theory constant *)
+      | _ => do_term t
+  in Symtab.empty |> fold (do_formula pos) ts end
+
+(*Inserts a dummy "constant" referring to the theory name, so that relevance
+  takes the given theory into account.*)
+fun theory_const_prop_of theory_relevant th =
+  if theory_relevant then
+    let
+      val name = Context.theory_name (theory_of_thm th)
+      val t = Const (name ^ ". 1", @{typ bool})
+    in t $ prop_of th end
+  else
+    prop_of th
+
+(**** Constant / Type Frequencies ****)
+
+(* A two-dimensional symbol table counts frequencies of constants. It's keyed
+   first by constant name and second by its list of type instantiations. For the
+   latter, we need a linear ordering on "pattern list". *)
+
+fun pattern_ord p =
+  case p of
+    (PVar, PVar) => EQUAL
+  | (PVar, PApp _) => LESS
+  | (PApp _, PVar) => GREATER
+  | (PApp q1, PApp q2) =>
+    prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
+fun ptype_ord (PType p, PType q) =
+  prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q)
+
+structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
+
+fun count_axiom_consts theory_relevant thy =
+  let
+    fun do_const const (s, T) ts =
+      (* Two-dimensional table update. Constant maps to types maps to count. *)
+      PType_Tab.map_default (rich_ptype thy const (s, T) ts, 0) (Integer.add 1)
+      |> Symtab.map_default (s, PType_Tab.empty)
+      #> fold do_term ts
+    and do_term t =
+      case strip_comb t of
+        (Const x, ts) => do_const true x ts
+      | (Free x, ts) => do_const false x ts
+      | (Abs (_, _, t'), ts) => fold do_term (t' :: ts)
+      | (_, ts) => fold do_term ts
+  in do_term o theory_const_prop_of theory_relevant o snd end
+
+
+(**** Actual Filtering Code ****)
+
+fun pow_int x 0 = 1.0
+  | pow_int x 1 = x
+  | pow_int x n = if n > 0 then x * pow_int x (n - 1) else pow_int x (n + 1) / x
+
+(*The frequency of a constant is the sum of those of all instances of its type.*)
+fun pconst_freq match const_tab (c, ps) =
+  PType_Tab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m)
+                 (the (Symtab.lookup const_tab c)) 0
+
+
+(* A surprising number of theorems contain only a few significant constants.
+   These include all induction rules, and other general theorems. *)
+
+(* "log" seems best in practice. A constant function of one ignores the constant
+   frequencies. Rare constants give more points if they are relevant than less
+   rare ones. *)
+fun rel_weight_for order freq = 1.0 + 2.0 / Math.ln (Real.fromInt freq + 1.0)
+
+(* FUDGE *)
+val worse_irrel_freq = Unsynchronized.ref 100.0
+val higher_order_irrel_weight = Unsynchronized.ref 1.05
+
+(* Irrelevant constants are treated differently. We associate lower penalties to
+   very rare constants and very common ones -- the former because they can't
+   lead to the inclusion of too many new facts, and the latter because they are
+   so common as to be of little interest. *)
+fun irrel_weight_for order freq =
+  let val (k, x) = !worse_irrel_freq |> `Real.ceil in
+    (if freq < k then Math.ln (Real.fromInt (freq + 1)) / Math.ln x
+     else rel_weight_for order freq / rel_weight_for order k)
+    * pow_int (!higher_order_irrel_weight) (order - 1)
+  end
+
+(* FUDGE *)
+val abs_rel_weight = Unsynchronized.ref 0.5
+val abs_irrel_weight = Unsynchronized.ref 2.0
+val skolem_irrel_weight = Unsynchronized.ref 0.75
+
+(* Computes a constant's weight, as determined by its frequency. *)
+fun generic_pconst_weight abs_weight skolem_weight weight_for f const_tab
+                          (c as (s, PType (m, _))) =
+  if s = abs_name then abs_weight
+  else if String.isPrefix skolem_prefix s then skolem_weight
+  else weight_for m (pconst_freq (match_ptype o f) const_tab c)
+
+fun rel_pconst_weight const_tab =
+  generic_pconst_weight (!abs_rel_weight) 0.0 rel_weight_for I const_tab
+fun irrel_pconst_weight const_tab =
+  generic_pconst_weight (!abs_irrel_weight) (!skolem_irrel_weight)
+                        irrel_weight_for swap const_tab
+
+(* FUDGE *)
+val intro_bonus = Unsynchronized.ref 0.15
+val elim_bonus = Unsynchronized.ref 0.15
+val simp_bonus = Unsynchronized.ref 0.15
+val local_bonus = Unsynchronized.ref 0.55
+val chained_bonus = Unsynchronized.ref 1.5
+
+fun locality_bonus General = 0.0
+  | locality_bonus Intro = !intro_bonus
+  | locality_bonus Elim = !elim_bonus
+  | locality_bonus Simp = !simp_bonus
+  | locality_bonus Local = !local_bonus
+  | locality_bonus Chained = !chained_bonus
+
+fun axiom_weight loc const_tab relevant_consts axiom_consts =
+  case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
+                    ||> filter_out (pconst_hyper_mem swap relevant_consts) of
+    ([], _) => 0.0
+  | (rel, irrel) =>
+    let
+      val irrel = irrel |> filter_out (pconst_mem swap rel)
+      val rel_weight =
+        0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
+      val irrel_weight =
+        ~ (locality_bonus loc)
+        |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
+      val res = rel_weight / (rel_weight + irrel_weight)
+    in if Real.isFinite res then res else 0.0 end
+
+(* FIXME: experiment
+fun debug_axiom_weight loc const_tab relevant_consts axiom_consts =
+  case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
+                    ||> filter_out (pconst_hyper_mem swap relevant_consts) of
+    ([], _) => 0.0
+  | (rel, irrel) =>
+    let
+      val irrel = irrel |> filter_out (pconst_mem swap rel)
+      val rels_weight =
+        0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
+      val irrels_weight =
+        ~ (locality_bonus loc)
+        |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
+val _ = tracing (PolyML.makestring ("REL: ", map (`(rel_pconst_weight const_tab)) rel))
+val _ = tracing (PolyML.makestring ("IRREL: ", map (`(irrel_pconst_weight const_tab)) irrel))
+      val res = rels_weight / (rels_weight + irrels_weight)
+    in if Real.isFinite res then res else 0.0 end
+*)
+
+fun pconsts_in_axiom thy t =
+  Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
+              (pconsts_in_terms thy true (SOME true) [t]) []
+fun pair_consts_axiom theory_relevant thy axiom =
+  case axiom |> snd |> theory_const_prop_of theory_relevant
+             |> pconsts_in_axiom thy of
+    [] => NONE
+  | consts => SOME ((axiom, consts), NONE)
+
+type annotated_thm =
+  (((unit -> string) * locality) * thm) * (string * ptype) list
+
+(* FUDGE *)
+val max_imperfect = Unsynchronized.ref 11.5
+val max_imperfect_exp = Unsynchronized.ref 1.0
+
+fun take_most_relevant max_relevant remaining_max
+                       (candidates : (annotated_thm * real) list) =
+  let
+    val max_imperfect =
+      Real.ceil (Math.pow (!max_imperfect,
+                    Math.pow (Real.fromInt remaining_max
+                              / Real.fromInt max_relevant, !max_imperfect_exp)))
+    val (perfect, imperfect) =
+      candidates |> sort (Real.compare o swap o pairself snd)
+                 |> take_prefix (fn (_, w) => w > 0.99999)
+    val ((accepts, more_rejects), rejects) =
+      chop max_imperfect imperfect |>> append perfect |>> chop remaining_max
+  in
+    trace_msg (fn () =>
+        "Actually passed (" ^ Int.toString (length accepts) ^ " of " ^
+        Int.toString (length candidates) ^ "): " ^
+        (accepts |> map (fn ((((name, _), _), _), weight) =>
+                            name () ^ " [" ^ Real.toString weight ^ "]")
+                 |> commas));
+    (accepts, more_rejects @ rejects)
+  end
+
+fun if_empty_replace_with_locality thy axioms loc tab =
+  if Symtab.is_empty tab then
+    pconsts_in_terms thy false (SOME false)
+        (map_filter (fn ((_, loc'), th) =>
+                        if loc' = loc then SOME (prop_of th) else NONE) axioms)
+  else
+    tab
+
+(* FUDGE *)
+val threshold_divisor = Unsynchronized.ref 2.0
+val ridiculous_threshold = Unsynchronized.ref 0.1
+
+fun relevance_filter ctxt threshold0 decay max_relevant theory_relevant
+                     ({add, del, ...} : relevance_override) axioms goal_ts =
+  let
+    val thy = ProofContext.theory_of ctxt
+    val const_tab =
+      fold (count_axiom_consts theory_relevant thy) axioms Symtab.empty
+    val goal_const_tab =
+      pconsts_in_terms thy false (SOME false) goal_ts
+      |> fold (if_empty_replace_with_locality thy axioms) [Chained, Local]
+    val add_thms = maps (ProofContext.get_fact ctxt) add
+    val del_thms = maps (ProofContext.get_fact ctxt) del
+    fun iter j remaining_max threshold rel_const_tab hopeless hopeful =
+      let
+        fun game_over rejects =
+          (* Add "add:" facts. *)
+          if null add_thms then
+            []
+          else
+            map_filter (fn ((p as (_, th), _), _) =>
+                           if member Thm.eq_thm add_thms th then SOME p
+                           else NONE) rejects
+        fun relevant [] rejects [] =
+            (* Nothing has been added this iteration. *)
+            if j = 0 andalso threshold >= !ridiculous_threshold then
+              (* First iteration? Try again. *)
+              iter 0 max_relevant (threshold / !threshold_divisor) rel_const_tab
+                   hopeless hopeful
+            else
+              game_over (rejects @ hopeless)
+          | relevant candidates rejects [] =
+            let
+              val (accepts, more_rejects) =
+                take_most_relevant max_relevant remaining_max candidates
+              val rel_const_tab' =
+                rel_const_tab
+                |> fold (add_pconst_to_table false) (maps (snd o fst) accepts)
+              fun is_dirty (c, _) =
+                Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
+              val (hopeful_rejects, hopeless_rejects) =
+                 (rejects @ hopeless, ([], []))
+                 |-> fold (fn (ax as (_, consts), old_weight) =>
+                              if exists is_dirty consts then
+                                apfst (cons (ax, NONE))
+                              else
+                                apsnd (cons (ax, old_weight)))
+                 |>> append (more_rejects
+                             |> map (fn (ax as (_, consts), old_weight) =>
+                                        (ax, if exists is_dirty consts then NONE
+                                             else SOME old_weight)))
+              val threshold =
+                1.0 - (1.0 - threshold)
+                      * Math.pow (decay, Real.fromInt (length accepts))
+              val remaining_max = remaining_max - length accepts
+            in
+              trace_msg (fn () => "New or updated constants: " ^
+                  commas (rel_const_tab' |> Symtab.dest
+                          |> subtract (op =) (rel_const_tab |> Symtab.dest)
+                          |> map string_for_hyper_pconst));
+              map (fst o fst) accepts @
+              (if remaining_max = 0 then
+                 game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects)
+               else
+                 iter (j + 1) remaining_max threshold rel_const_tab'
+                      hopeless_rejects hopeful_rejects)
+            end
+          | relevant candidates rejects
+                     (((ax as (((_, loc), th), axiom_consts)), cached_weight)
+                      :: hopeful) =
+            let
+              val weight =
+                case cached_weight of
+                  SOME w => w
+                | NONE => axiom_weight loc const_tab rel_const_tab axiom_consts
+(* FIXME: experiment
+val name = fst (fst (fst ax)) ()
+val _ = if String.isSubstring "positive_minus" name orelse String.isSubstring "not_exp_le_zero" name then
+tracing ("*** " ^ name ^ PolyML.makestring (debug_axiom_weight loc const_tab rel_const_tab axiom_consts))
+else
+()
+*)
+            in
+              if weight >= threshold then
+                relevant ((ax, weight) :: candidates) rejects hopeful
+              else
+                relevant candidates ((ax, weight) :: rejects) hopeful
+            end
+        in
+          trace_msg (fn () =>
+              "ITERATION " ^ string_of_int j ^ ": current threshold: " ^
+              Real.toString threshold ^ ", constants: " ^
+              commas (rel_const_tab |> Symtab.dest
+                      |> filter (curry (op <>) [] o snd)
+                      |> map string_for_hyper_pconst));
+          relevant [] [] hopeful
+        end
+  in
+    axioms |> filter_out (member Thm.eq_thm del_thms o snd)
+           |> map_filter (pair_consts_axiom theory_relevant thy)
+           |> iter 0 max_relevant threshold0 goal_const_tab []
+           |> tap (fn res => trace_msg (fn () =>
+                                "Total relevant: " ^ Int.toString (length res)))
+  end
+
+
+(***************************************************************)
+(* Retrieving and filtering lemmas                             *)
+(***************************************************************)
+
+(*** retrieve lemmas and filter them ***)
+
+(*Reject theorems with names like "List.filter.filter_list_def" or
+  "Accessible_Part.acc.defs", as these are definitions arising from packages.*)
+fun is_package_def a =
+  let val names = Long_Name.explode a
+  in
+     length names > 2 andalso
+     not (hd names = "local") andalso
+     String.isSuffix "_def" a  orelse  String.isSuffix "_defs" a
+  end;
+
+fun mk_fact_table f xs =
+  fold (Termtab.update o `(prop_of o f)) xs Termtab.empty
+fun uniquify xs = Termtab.fold (cons o snd) (mk_fact_table snd xs) []
+
+(* FIXME: put other record thms here, or declare as "no_atp" *)
+val multi_base_blacklist =
+  ["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits",
+   "split_asm", "cases", "ext_cases", "eq.simps", "eq.refl", "nchotomy",
+   "case_cong", "weak_case_cong"]
+  |> map (prefix ".")
+
+val max_lambda_nesting = 3
+
+fun term_has_too_many_lambdas max (t1 $ t2) =
+    exists (term_has_too_many_lambdas max) [t1, t2]
+  | term_has_too_many_lambdas max (Abs (_, _, t)) =
+    max = 0 orelse term_has_too_many_lambdas (max - 1) t
+  | term_has_too_many_lambdas _ _ = false
+
+(* Don't count nested lambdas at the level of formulas, since they are
+   quantifiers. *)
+fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) =
+    formula_has_too_many_lambdas (T :: Ts) t
+  | formula_has_too_many_lambdas Ts t =
+    if is_formula_type (fastype_of1 (Ts, t)) then
+      exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t))
+    else
+      term_has_too_many_lambdas max_lambda_nesting t
+
+(* The max apply depth of any "metis" call in "Metis_Examples" (on 2007-10-31)
+   was 11. *)
+val max_apply_depth = 15
+
+fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
+  | apply_depth (Abs (_, _, t)) = apply_depth t
+  | apply_depth _ = 0
+
+fun is_formula_too_complex t =
+  apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t
+
+val exists_sledgehammer_const =
+  exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
+
+(* FIXME: make more reliable *)
+val exists_low_level_class_const =
+  exists_Const (fn (s, _) =>
+     String.isSubstring (Long_Name.separator ^ "class" ^ Long_Name.separator) s)
+
+fun is_metastrange_theorem th =
+  case head_of (concl_of th) of
+      Const (a, _) => (a <> @{const_name Trueprop} andalso
+                       a <> @{const_name "=="})
+    | _ => false
+
+fun is_that_fact th =
+  String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th)
+  andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
+                           | _ => false) (prop_of th)
+
+val type_has_top_sort =
+  exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
+
+(**** Predicates to detect unwanted facts (prolific or likely to cause
+      unsoundness) ****)
+
+(* Too general means, positive equality literal with a variable X as one
+   operand, when X does not occur properly in the other operand. This rules out
+   clearly inconsistent facts such as X = a | X = b, though it by no means
+   guarantees soundness. *)
+
+(* Unwanted equalities are those between a (bound or schematic) variable that
+   does not properly occur in the second operand. *)
+val is_exhaustive_finite =
+  let
+    fun is_bad_equal (Var z) t =
+        not (exists_subterm (fn Var z' => z = z' | _ => false) t)
+      | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
+      | is_bad_equal _ _ = false
+    fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
+    fun do_formula pos t =
+      case (pos, t) of
+        (_, @{const Trueprop} $ t1) => do_formula pos t1
+      | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
+        do_formula pos t'
+      | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
+        do_formula pos t'
+      | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
+        do_formula pos t'
+      | (_, @{const "==>"} $ t1 $ t2) =>
+        do_formula (not pos) t1 andalso
+        (t2 = @{prop False} orelse do_formula pos t2)
+      | (_, @{const HOL.implies} $ t1 $ t2) =>
+        do_formula (not pos) t1 andalso
+        (t2 = @{const False} orelse do_formula pos t2)
+      | (_, @{const Not} $ t1) => do_formula (not pos) t1
+      | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
+      | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
+      | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
+      | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
+      | _ => false
+  in do_formula true end
+
+fun has_bound_or_var_of_type tycons =
+  exists_subterm (fn Var (_, Type (s, _)) => member (op =) tycons s
+                   | Abs (_, Type (s, _), _) => member (op =) tycons s
+                   | _ => false)
+
+(* Facts are forbidden to contain variables of these types. The typical reason
+   is that they lead to unsoundness. Note that "unit" satisfies numerous
+   equations like "?x = ()". The resulting clauses will have no type constraint,
+   yielding false proofs. Even "bool" leads to many unsound proofs, though only
+   for higher-order problems. *)
+val dangerous_types = [@{type_name unit}, @{type_name bool}, @{type_name prop}];
+
+(* Facts containing variables of type "unit" or "bool" or of the form
+   "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
+   are omitted. *)
+fun is_dangerous_term full_types t =
+  not full_types andalso
+  let val t = transform_elim_term t in
+    has_bound_or_var_of_type dangerous_types t orelse
+    is_exhaustive_finite t
+  end
+
+fun is_theorem_bad_for_atps full_types thm =
+  let val t = prop_of thm in
+    is_formula_too_complex t orelse exists_type type_has_top_sort t orelse
+    is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse
+    exists_low_level_class_const t orelse is_metastrange_theorem thm orelse
+    is_that_fact thm
+  end
+
+fun clasimpset_rules_of ctxt =
+  let
+    val {safeIs, safeEs, hazIs, hazEs, ...} = ctxt |> claset_of |> rep_cs
+    val intros = safeIs @ hazIs
+    val elims = map Classical.classical_rule (safeEs @ hazEs)
+    val simps = ctxt |> simpset_of |> dest_ss |> #simps |> map snd
+  in (mk_fact_table I intros, mk_fact_table I elims, mk_fact_table I simps) end
+
+fun all_name_thms_pairs ctxt reserved full_types add_thms chained_ths =
+  let
+    val thy = ProofContext.theory_of ctxt
+    val global_facts = PureThy.facts_of thy
+    val local_facts = ProofContext.facts_of ctxt
+    val named_locals = local_facts |> Facts.dest_static []
+    val is_chained = member Thm.eq_thm chained_ths
+    val (intros, elims, simps) =
+      if exists (curry (op <) 0.0) [!intro_bonus, !elim_bonus, !simp_bonus] then
+        clasimpset_rules_of ctxt
+      else
+        (Termtab.empty, Termtab.empty, Termtab.empty)
+    (* Unnamed nonchained formulas with schematic variables are omitted, because
+       they are rejected by the backticks (`...`) parser for some reason. *)
+    fun is_good_unnamed_local th =
+      not (Thm.has_name_hint th) andalso
+      (not (exists_subterm is_Var (prop_of th)) orelse (is_chained th)) andalso
+      forall (fn (_, ths) => not (member Thm.eq_thm ths th)) named_locals
+    val unnamed_locals =
+      union Thm.eq_thm (Facts.props local_facts) chained_ths
+      |> filter is_good_unnamed_local |> map (pair "" o single)
+    val full_space =
+      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
+    fun add_facts global foldx facts =
+      foldx (fn (name0, ths) =>
+        if name0 <> "" andalso
+           forall (not o member Thm.eq_thm add_thms) ths andalso
+           (Facts.is_concealed facts name0 orelse
+            (respect_no_atp andalso is_package_def name0) orelse
+            exists (fn s => String.isSuffix s name0) multi_base_blacklist orelse
+            String.isSuffix "_def_raw" (* FIXME: crude hack *) name0) then
+          I
+        else
+          let
+            val multi = length ths > 1
+            fun backquotify th =
+              "`" ^ Print_Mode.setmp [Print_Mode.input]
+                                 (Syntax.string_of_term ctxt) (prop_of th) ^ "`"
+              |> String.translate (fn c => if Char.isPrint c then str c else "")
+              |> simplify_spaces
+            fun check_thms a =
+              case try (ProofContext.get_thms ctxt) a of
+                NONE => false
+              | SOME ths' => Thm.eq_thms (ths, ths')
+          in
+            pair 1
+            #> fold (fn th => fn (j, rest) =>
+                 (j + 1,
+                  if is_theorem_bad_for_atps full_types th andalso
+                     not (member Thm.eq_thm add_thms th) then
+                    rest
+                  else
+                    (((fn () =>
+                          if name0 = "" then
+                            th |> backquotify
+                          else
+                            let
+                              val name1 = Facts.extern facts name0
+                              val name2 = Name_Space.extern full_space name0
+                            in
+                              case find_first check_thms [name1, name2, name0] of
+                                SOME name => repair_name reserved multi j name
+                              | NONE => ""
+                            end),
+                      let val t = prop_of th in
+                        if is_chained th then Chained
+                        else if not global then Local
+                        else if Termtab.defined intros t then Intro
+                        else if Termtab.defined elims t then Elim
+                        else if Termtab.defined simps t then Simp
+                        else General
+                      end),
+                      (multi, th)) :: rest)) ths
+            #> snd
+          end)
+  in
+    [] |> add_facts false fold local_facts (unnamed_locals @ named_locals)
+       |> add_facts true Facts.fold_static global_facts global_facts
+  end
+
+(* The single-name theorems go after the multiple-name ones, so that single
+   names are preferred when both are available. *)
+fun name_thm_pairs ctxt respect_no_atp =
+  List.partition (fst o snd) #> op @ #> map (apsnd snd)
+  #> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd)
+
+(***************************************************************)
+(* ATP invocation methods setup                                *)
+(***************************************************************)
+
+fun relevant_facts ctxt full_types (threshold0, threshold1) max_relevant
+                   theory_relevant (relevance_override as {add, del, only})
+                   chained_ths hyp_ts concl_t =
+  let
+    val decay = Math.pow ((1.0 - threshold1) / (1.0 - threshold0),
+                          1.0 / Real.fromInt (max_relevant + 1))
+    val add_thms = maps (ProofContext.get_fact ctxt) add
+    val reserved = reserved_isar_keyword_table ()
+    val axioms =
+      (if only then
+         maps (map (fn ((name, loc), th) => ((K name, loc), (true, th)))
+               o name_thm_pairs_from_ref ctxt reserved chained_ths) add
+       else
+         all_name_thms_pairs ctxt reserved full_types add_thms chained_ths)
+      |> name_thm_pairs ctxt (respect_no_atp andalso not only)
+      |> uniquify
+  in
+    trace_msg (fn () => "Considering " ^ Int.toString (length axioms) ^
+                        " theorems");
+    (if threshold0 > 1.0 orelse threshold0 > threshold1 then
+       []
+     else if threshold0 < 0.0 then
+       axioms
+     else
+       relevance_filter ctxt threshold0 decay max_relevant theory_relevant
+                        relevance_override axioms (concl_t :: hyp_ts))
+    |> map (apfst (apfst (fn f => f ())))
+  end
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Tue Aug 31 23:46:23 2010 +0200
@@ -0,0 +1,162 @@
+(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML
+    Author:     Philipp Meyer, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Minimization of theorem list for Metis using automatic theorem provers.
+*)
+
+signature SLEDGEHAMMER_FACT_MINIMIZE =
+sig
+  type locality = Sledgehammer_Fact_Filter.locality
+  type params = Sledgehammer.params
+
+  val minimize_theorems :
+    params -> int -> int -> Proof.state -> ((string * locality) * thm list) list
+    -> ((string * locality) * thm list) list option * string
+  val run_minimize : params -> int -> Facts.ref list -> Proof.state -> unit
+end;
+
+structure Sledgehammer_Fact_Minimize : SLEDGEHAMMER_FACT_MINIMIZE =
+struct
+
+open ATP_Systems
+open Sledgehammer_Util
+open Sledgehammer_Fact_Filter
+open Sledgehammer_Proof_Reconstruct
+open Sledgehammer
+
+(* wrapper for calling external prover *)
+
+fun string_for_failure Unprovable = "Unprovable."
+  | string_for_failure TimedOut = "Timed out."
+  | string_for_failure _ = "Unknown error."
+
+fun n_theorems names =
+  let val n = length names in
+    string_of_int n ^ " theorem" ^ plural_s n ^
+    (if n > 0 then
+       ": " ^ (names |> map fst
+                     |> sort_distinct string_ord |> space_implode " ")
+     else
+       "")
+  end
+
+fun test_theorems ({debug, verbose, overlord, atps, full_types, isar_proof,
+                    isar_shrink_factor, ...} : params)
+                  (prover : prover) explicit_apply timeout subgoal state
+                  axioms =
+  let
+    val _ =
+      priority ("Testing " ^ n_theorems (map fst axioms) ^ "...")
+    val params =
+      {blocking = true, debug = debug, verbose = verbose, overlord = overlord,
+       atps = atps, full_types = full_types, explicit_apply = explicit_apply,
+       relevance_thresholds = (1.01, 1.01), max_relevant = NONE,
+       theory_relevant = NONE, isar_proof = isar_proof,
+       isar_shrink_factor = isar_shrink_factor, timeout = timeout, expect = ""}
+    val axioms = maps (fn (n, ths) => map (pair n) ths) axioms
+    val {context = ctxt, facts, goal} = Proof.goal state
+    val problem =
+     {subgoal = subgoal, goal = (ctxt, (facts, goal)),
+      relevance_override = {add = [], del = [], only = false},
+      axioms = SOME axioms}
+    val result as {outcome, used_thm_names, ...} = prover params (K "") problem
+  in
+    priority (case outcome of
+                NONE =>
+                if length used_thm_names = length axioms then
+                  "Found proof."
+                else
+                  "Found proof with " ^ n_theorems used_thm_names ^ "."
+              | SOME failure => string_for_failure failure);
+    result
+  end
+
+(* minimalization of thms *)
+
+fun filter_used_facts used = filter (member (op =) used o fst)
+
+fun sublinear_minimize _ [] p = p
+  | sublinear_minimize test (x :: xs) (seen, result) =
+    case test (xs @ seen) of
+      result as {outcome = NONE, proof, used_thm_names, ...} : prover_result =>
+      sublinear_minimize test (filter_used_facts used_thm_names xs)
+                         (filter_used_facts used_thm_names seen, result)
+    | _ => sublinear_minimize test xs (x :: seen, result)
+
+(* Give the ATP some slack. The ATP gets further slack because the Sledgehammer
+   preprocessing time is included in the estimate below but isn't part of the
+   timeout. *)
+val fudge_msecs = 1000
+
+fun minimize_theorems {atps = [], ...} _ _ _ _ = error "No ATP is set."
+  | minimize_theorems (params as {debug, atps = atp :: _, full_types,
+                                  isar_proof, isar_shrink_factor, timeout, ...})
+                      i n state axioms =
+  let
+    val thy = Proof.theory_of state
+    val prover = get_prover_fun thy atp
+    val msecs = Time.toMilliseconds timeout
+    val _ = priority ("Sledgehammer minimize: ATP " ^ quote atp ^ ".")
+    val {context = ctxt, goal, ...} = Proof.goal state
+    val (_, hyp_ts, concl_t) = strip_subgoal goal i
+    val explicit_apply =
+      not (forall (Meson.is_fol_term thy)
+                  (concl_t :: hyp_ts @ maps (map prop_of o snd) axioms))
+    fun do_test timeout =
+      test_theorems params prover explicit_apply timeout i state
+    val timer = Timer.startRealTimer ()
+  in
+    (case do_test timeout axioms of
+       result as {outcome = NONE, pool, used_thm_names,
+                  conjecture_shape, ...} =>
+       let
+         val time = Timer.checkRealTimer timer
+         val new_timeout =
+           Int.min (msecs, Time.toMilliseconds time + fudge_msecs)
+           |> Time.fromMilliseconds
+         val (min_thms, {proof, axiom_names, ...}) =
+           sublinear_minimize (do_test new_timeout)
+               (filter_used_facts used_thm_names axioms) ([], result)
+         val n = length min_thms
+         val _ = priority (cat_lines
+           ["Minimized: " ^ string_of_int n ^ " theorem" ^ plural_s n] ^
+            (case length (filter (curry (op =) Chained o snd o fst) min_thms) of
+               0 => ""
+             | n => " (including " ^ Int.toString n ^ " chained)") ^ ".")
+       in
+         (SOME min_thms,
+          proof_text isar_proof
+              (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
+              (full_types, K "", proof, axiom_names, goal, i) |> fst)
+       end
+     | {outcome = SOME TimedOut, ...} =>
+       (NONE, "Timeout: You can increase the time limit using the \"timeout\" \
+              \option (e.g., \"timeout = " ^
+              string_of_int (10 + msecs div 1000) ^ " s\").")
+     | {outcome = SOME UnknownError, ...} =>
+       (* Failure sometimes mean timeout, unfortunately. *)
+       (NONE, "Failure: No proof was found with the current time limit. You \
+              \can increase the time limit using the \"timeout\" \
+              \option (e.g., \"timeout = " ^
+              string_of_int (10 + msecs div 1000) ^ " s\").")
+     | {message, ...} => (NONE, "ATP error: " ^ message))
+    handle ERROR msg => (NONE, "Error: " ^ msg)
+  end
+
+fun run_minimize params i refs state =
+  let
+    val ctxt = Proof.context_of state
+    val reserved = reserved_isar_keyword_table ()
+    val chained_ths = #facts (Proof.goal state)
+    val axioms =
+      maps (map (apsnd single)
+            o name_thm_pairs_from_ref ctxt reserved chained_ths) refs
+  in
+    case subgoal_count state of
+      0 => priority "No subgoal!"
+    | n =>
+      (kill_atps (); priority (#2 (minimize_theorems params i n state axioms)))
+  end
+
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML	Tue Aug 31 23:43:23 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1038 +0,0 @@
-(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML
-    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
-    Author:     Claire Quigley, Cambridge University Computer Laboratory
-    Author:     Jasmin Blanchette, TU Muenchen
-
-Transfer of proofs from external provers.
-*)
-
-signature SLEDGEHAMMER_PROOF_RECONSTRUCT =
-sig
-  type locality = Sledgehammer_Fact_Filter.locality
-  type minimize_command = string list -> string
-  type metis_params =
-    bool * minimize_command * string * (string * locality) list vector * thm
-    * int
-  type isar_params =
-    string Symtab.table * bool * int * Proof.context * int list list
-  type text_result = string * (string * locality) list
-
-  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
-end;
-
-structure Sledgehammer_Proof_Reconstruct : SLEDGEHAMMER_PROOF_RECONSTRUCT =
-struct
-
-open ATP_Problem
-open Metis_Clauses
-open Sledgehammer_Util
-open Sledgehammer_Fact_Filter
-open Sledgehammer_Translate
-
-type minimize_command = string list -> string
-type metis_params =
-  bool * minimize_command * string * (string * locality) list vector * thm * int
-type isar_params =
-  string Symtab.table * bool * int * Proof.context * int list list
-type text_result = string * (string * locality) list
-
-(* Simple simplifications to ensure that sort annotations don't leave a trail of
-   spurious "True"s. *)
-fun s_not @{const False} = @{const True}
-  | s_not @{const True} = @{const False}
-  | s_not (@{const Not} $ t) = t
-  | s_not t = @{const Not} $ t
-fun s_conj (@{const True}, t2) = t2
-  | s_conj (t1, @{const True}) = t1
-  | s_conj p = HOLogic.mk_conj p
-fun s_disj (@{const False}, t2) = t2
-  | s_disj (t1, @{const False}) = t1
-  | s_disj p = HOLogic.mk_disj p
-fun s_imp (@{const True}, t2) = t2
-  | s_imp (t1, @{const False}) = s_not t1
-  | s_imp p = HOLogic.mk_imp p
-fun s_iff (@{const True}, t2) = t2
-  | s_iff (t1, @{const True}) = t1
-  | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
-
-fun mk_anot (AConn (ANot, [phi])) = phi
-  | mk_anot phi = AConn (ANot, [phi])
-fun mk_aconn c (phi1, phi2) = AConn (c, [phi1, phi2])
-
-fun index_in_shape x = find_index (exists (curry (op =) x))
-fun is_axiom_number axiom_names num =
-  num > 0 andalso num <= Vector.length axiom_names andalso
-  not (null (Vector.sub (axiom_names, num - 1)))
-fun is_conjecture_number conjecture_shape num =
-  index_in_shape num conjecture_shape >= 0
-
-fun negate_term (Const (@{const_name All}, T) $ Abs (s, T', t')) =
-    Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t')
-  | negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
-    Const (@{const_name All}, T) $ Abs (s, T', negate_term t')
-  | negate_term (@{const HOL.implies} $ t1 $ t2) =
-    @{const HOL.conj} $ t1 $ negate_term t2
-  | negate_term (@{const HOL.conj} $ t1 $ t2) =
-    @{const HOL.disj} $ negate_term t1 $ negate_term t2
-  | negate_term (@{const HOL.disj} $ t1 $ t2) =
-    @{const HOL.conj} $ negate_term t1 $ negate_term t2
-  | negate_term (@{const Not} $ t) = t
-  | negate_term t = @{const Not} $ t
-
-datatype ('a, 'b, 'c, 'd, 'e) raw_step =
-  Definition of 'a * 'b * 'c |
-  Inference of 'a * 'd * 'e list
-
-fun raw_step_number (Definition (num, _, _)) = num
-  | raw_step_number (Inference (num, _, _)) = num
-
-(**** PARSING OF TSTP FORMAT ****)
-
-(*Strings enclosed in single quotes, e.g. filenames*)
-val scan_quoted = $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode;
-
-val scan_dollar_name =
-  Scan.repeat ($$ "$") -- Symbol.scan_id >> (fn (ss, s) => implode ss ^ s)
-
-fun repair_name _ "$true" = "c_True"
-  | repair_name _ "$false" = "c_False"
-  | repair_name _ "$$e" = "c_equal" (* seen in Vampire proofs *)
-  | repair_name _ "equal" = "c_equal" (* needed by SPASS? *)
-  | repair_name pool s =
-    case Symtab.lookup pool s of
-      SOME s' => s'
-    | NONE =>
-      if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then
-        "c_equal" (* seen in Vampire proofs *)
-      else
-        s
-(* Generalized first-order terms, which include file names, numbers, etc. *)
-val parse_potential_integer =
-  (scan_dollar_name || scan_quoted) >> K NONE
-  || scan_integer >> SOME
-fun parse_annotation x =
-  ((parse_potential_integer ::: Scan.repeat ($$ " " |-- parse_potential_integer)
-    >> map_filter I) -- Scan.optional parse_annotation []
-     >> uncurry (union (op =))
-   || $$ "(" |-- parse_annotations --| $$ ")"
-   || $$ "[" |-- parse_annotations --| $$ "]") x
-and parse_annotations x =
-  (Scan.optional (parse_annotation
-                  ::: Scan.repeat ($$ "," |-- parse_annotation)) []
-   >> (fn numss => fold (union (op =)) numss [])) x
-
-(* Vampire proof lines sometimes contain needless information such as "(0:3)",
-   which can be hard to disambiguate from function application in an LL(1)
-   parser. As a workaround, we extend the TPTP term syntax with such detritus
-   and ignore it. *)
-fun parse_vampire_detritus x =
-  (scan_integer |-- $$ ":" --| scan_integer >> K []) x
-
-fun parse_term pool x =
-  ((scan_dollar_name >> repair_name pool)
-    -- Scan.optional ($$ "(" |-- (parse_vampire_detritus || parse_terms pool)
-                      --| $$ ")") []
-    --| Scan.optional ($$ "(" |-- parse_vampire_detritus --| $$ ")") []
-   >> ATerm) x
-and parse_terms pool x =
-  (parse_term pool ::: Scan.repeat ($$ "," |-- parse_term pool)) x
-
-fun parse_atom pool =
-  parse_term pool -- Scan.option (Scan.option ($$ "!") --| $$ "="
-                                  -- parse_term pool)
-  >> (fn (u1, NONE) => AAtom u1
-       | (u1, SOME (NONE, u2)) => AAtom (ATerm ("c_equal", [u1, u2]))
-       | (u1, SOME (SOME _, u2)) =>
-         mk_anot (AAtom (ATerm ("c_equal", [u1, u2]))))
-
-fun fo_term_head (ATerm (s, _)) = s
-
-(* TPTP formulas are fully parenthesized, so we don't need to worry about
-   operator precedence. *)
-fun parse_formula pool x =
-  (($$ "(" |-- parse_formula pool --| $$ ")"
-    || ($$ "!" >> K AForall || $$ "?" >> K AExists)
-       --| $$ "[" -- parse_terms pool --| $$ "]" --| $$ ":"
-       -- parse_formula pool
-       >> (fn ((q, ts), phi) => AQuant (q, map fo_term_head ts, phi))
-    || $$ "~" |-- parse_formula pool >> mk_anot
-    || parse_atom pool)
-   -- Scan.option ((Scan.this_string "=>" >> K AImplies
-                    || Scan.this_string "<=>" >> K AIff
-                    || Scan.this_string "<~>" >> K ANotIff
-                    || Scan.this_string "<=" >> K AIf
-                    || $$ "|" >> K AOr || $$ "&" >> K AAnd)
-                   -- parse_formula pool)
-   >> (fn (phi1, NONE) => phi1
-        | (phi1, SOME (c, phi2)) => mk_aconn c (phi1, phi2))) x
-
-val parse_tstp_extra_arguments =
-  Scan.optional ($$ "," |-- parse_annotation
-                 --| Scan.option ($$ "," |-- parse_annotations)) []
-
-(* Syntax: (fof|cnf)\(<num>, <formula_role>, <formula> <extra_arguments>\).
-   The <num> could be an identifier, but we assume integers. *)
- fun parse_tstp_line pool =
-   ((Scan.this_string "fof" || Scan.this_string "cnf") -- $$ "(")
-     |-- scan_integer --| $$ "," -- Symbol.scan_id --| $$ ","
-     -- parse_formula pool -- parse_tstp_extra_arguments --| $$ ")" --| $$ "."
-    >> (fn (((num, role), phi), deps) =>
-           case role of
-             "definition" =>
-             (case phi of
-                AConn (AIff, [phi1 as AAtom _, phi2]) =>
-                Definition (num, phi1, phi2)
-              | AAtom (ATerm ("c_equal", _)) =>
-                Inference (num, phi, deps) (* Vampire's equality proxy axiom *)
-              | _ => raise Fail "malformed definition")
-           | _ => Inference (num, phi, deps))
-
-(**** PARSING OF VAMPIRE OUTPUT ****)
-
-(* Syntax: <num>. <formula> <annotation> *)
-fun parse_vampire_line pool =
-  scan_integer --| $$ "." -- parse_formula pool -- parse_annotation
-  >> (fn ((num, phi), deps) => Inference (num, phi, deps))
-
-(**** PARSING OF SPASS OUTPUT ****)
-
-(* SPASS returns clause references of the form "x.y". We ignore "y", whose role
-   is not clear anyway. *)
-val parse_dot_name = scan_integer --| $$ "." --| scan_integer
-
-val parse_spass_annotations =
-  Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name
-                                         --| Scan.option ($$ ","))) []
-
-(* It is not clear why some literals are followed by sequences of stars and/or
-   pluses. We ignore them. *)
-fun parse_decorated_atom pool =
-  parse_atom pool --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")
-
-fun mk_horn ([], []) = AAtom (ATerm ("c_False", []))
-  | mk_horn ([], pos_lits) = foldr1 (mk_aconn AOr) pos_lits
-  | mk_horn (neg_lits, []) = mk_anot (foldr1 (mk_aconn AAnd) neg_lits)
-  | mk_horn (neg_lits, pos_lits) =
-    mk_aconn AImplies (foldr1 (mk_aconn AAnd) neg_lits,
-                       foldr1 (mk_aconn AOr) pos_lits)
-
-fun parse_horn_clause pool =
-  Scan.repeat (parse_decorated_atom pool) --| $$ "|" --| $$ "|"
-    -- Scan.repeat (parse_decorated_atom pool) --| $$ "-" --| $$ ">"
-    -- Scan.repeat (parse_decorated_atom pool)
-  >> (mk_horn o apfst (op @))
-
-(* Syntax: <num>[0:<inference><annotations>]
-   <atoms> || <atoms> -> <atoms>. *)
-fun parse_spass_line pool =
-  scan_integer --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id
-    -- parse_spass_annotations --| $$ "]" -- parse_horn_clause pool --| $$ "."
-  >> (fn ((num, deps), u) => Inference (num, u, deps))
-
-fun parse_line pool =
-  parse_tstp_line pool || parse_vampire_line pool || parse_spass_line pool
-fun parse_lines pool = Scan.repeat1 (parse_line pool)
-fun parse_proof pool =
-  fst o Scan.finite Symbol.stopper
-            (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output")
-                            (parse_lines pool)))
-  o explode o strip_spaces_except_between_ident_chars
-
-(**** INTERPRETATION OF TSTP SYNTAX TREES ****)
-
-exception FO_TERM of string fo_term list
-exception FORMULA of (string, string fo_term) formula list
-exception SAME of unit
-
-(* Type variables are given the basic sort "HOL.type". Some will later be
-   constrained by information from type literals, or by type inference. *)
-fun type_from_fo_term tfrees (u as ATerm (a, us)) =
-  let val Ts = map (type_from_fo_term tfrees) us in
-    case strip_prefix_and_unascii type_const_prefix a of
-      SOME b => Type (invert_const b, Ts)
-    | NONE =>
-      if not (null us) then
-        raise FO_TERM [u]  (* only "tconst"s have type arguments *)
-      else case strip_prefix_and_unascii tfree_prefix a of
-        SOME b =>
-        let val s = "'" ^ b in
-          TFree (s, AList.lookup (op =) tfrees s |> the_default HOLogic.typeS)
-        end
-      | NONE =>
-        case strip_prefix_and_unascii tvar_prefix a of
-          SOME b => TVar (("'" ^ b, 0), HOLogic.typeS)
-        | NONE =>
-          (* Variable from the ATP, say "X1" *)
-          Type_Infer.param 0 (a, HOLogic.typeS)
-  end
-
-(* Type class literal applied to a type. Returns triple of polarity, class,
-   type. *)
-fun type_constraint_from_term pos tfrees (u as ATerm (a, us)) =
-  case (strip_prefix_and_unascii class_prefix a,
-        map (type_from_fo_term tfrees) us) of
-    (SOME b, [T]) => (pos, b, T)
-  | _ => raise FO_TERM [u]
-
-(** Accumulate type constraints in a formula: negative type literals **)
-fun add_var (key, z)  = Vartab.map_default (key, []) (cons z)
-fun add_type_constraint (false, cl, TFree (a ,_)) = add_var ((a, ~1), cl)
-  | add_type_constraint (false, cl, TVar (ix, _)) = add_var (ix, cl)
-  | add_type_constraint _ = I
-
-fun repair_atp_variable_name f s =
-  let
-    fun subscript_name s n = s ^ nat_subscript n
-    val s = String.map f s
-  in
-    case space_explode "_" s of
-      [_] => (case take_suffix Char.isDigit (String.explode s) of
-                (cs1 as _ :: _, cs2 as _ :: _) =>
-                subscript_name (String.implode cs1)
-                               (the (Int.fromString (String.implode cs2)))
-              | (_, _) => s)
-    | [s1, s2] => (case Int.fromString s2 of
-                     SOME n => subscript_name s1 n
-                   | NONE => s)
-    | _ => s
-  end
-
-(* First-order translation. No types are known for variables. "HOLogic.typeT"
-   should allow them to be inferred. *)
-fun raw_term_from_pred thy full_types tfrees =
-  let
-    fun aux opt_T extra_us u =
-      case u of
-        ATerm ("hBOOL", [u1]) => aux (SOME @{typ bool}) [] u1
-      | ATerm ("hAPP", [u1, u2]) => aux opt_T (u2 :: extra_us) u1
-      | ATerm (a, us) =>
-        if a = type_wrapper_name then
-          case us of
-            [typ_u, term_u] =>
-            aux (SOME (type_from_fo_term tfrees typ_u)) extra_us term_u
-          | _ => raise FO_TERM us
-        else case strip_prefix_and_unascii const_prefix a of
-          SOME "equal" =>
-          list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT),
-                     map (aux NONE []) us)
-        | SOME b =>
-          let
-            val c = invert_const b
-            val num_type_args = num_type_args thy c
-            val (type_us, term_us) =
-              chop (if full_types then 0 else num_type_args) us
-            (* Extra args from "hAPP" come after any arguments given directly to
-               the constant. *)
-            val term_ts = map (aux NONE []) term_us
-            val extra_ts = map (aux NONE []) extra_us
-            val t =
-              Const (c, if full_types then
-                          case opt_T of
-                            SOME T => map fastype_of term_ts ---> T
-                          | NONE =>
-                            if num_type_args = 0 then
-                              Sign.const_instance thy (c, [])
-                            else
-                              raise Fail ("no type information for " ^ quote c)
-                        else
-                          Sign.const_instance thy (c,
-                              map (type_from_fo_term tfrees) type_us))
-          in list_comb (t, term_ts @ extra_ts) end
-        | NONE => (* a free or schematic variable *)
-          let
-            val ts = map (aux NONE []) (us @ extra_us)
-            val T = map fastype_of ts ---> HOLogic.typeT
-            val t =
-              case strip_prefix_and_unascii fixed_var_prefix a of
-                SOME b => Free (b, T)
-              | NONE =>
-                case strip_prefix_and_unascii schematic_var_prefix a of
-                  SOME b => Var ((b, 0), T)
-                | NONE =>
-                  if is_tptp_variable a then
-                    Var ((repair_atp_variable_name Char.toLower a, 0), T)
-                  else
-                    (* Skolem constants? *)
-                    Var ((repair_atp_variable_name Char.toUpper a, 0), T)
-          in list_comb (t, ts) end
-  in aux (SOME HOLogic.boolT) [] end
-
-fun term_from_pred thy full_types tfrees pos (u as ATerm (s, _)) =
-  if String.isPrefix class_prefix s then
-    add_type_constraint (type_constraint_from_term pos tfrees u)
-    #> pair @{const True}
-  else
-    pair (raw_term_from_pred thy full_types tfrees u)
-
-val combinator_table =
-  [(@{const_name COMBI}, @{thm COMBI_def_raw}),
-   (@{const_name COMBK}, @{thm COMBK_def_raw}),
-   (@{const_name COMBB}, @{thm COMBB_def_raw}),
-   (@{const_name COMBC}, @{thm COMBC_def_raw}),
-   (@{const_name COMBS}, @{thm COMBS_def_raw})]
-
-fun uncombine_term (t1 $ t2) = betapply (pairself uncombine_term (t1, t2))
-  | uncombine_term (Abs (s, T, t')) = Abs (s, T, uncombine_term t')
-  | uncombine_term (t as Const (x as (s, _))) =
-    (case AList.lookup (op =) combinator_table s of
-       SOME thm => thm |> prop_of |> specialize_type @{theory} x |> Logic.dest_equals |> snd
-     | NONE => t)
-  | uncombine_term t = t
-
-(* Update schematic type variables with detected sort constraints. It's not
-   totally clear when this code is necessary. *)
-fun repair_tvar_sorts (t, tvar_tab) =
-  let
-    fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)
-      | do_type (TVar (xi, s)) =
-        TVar (xi, the_default s (Vartab.lookup tvar_tab xi))
-      | do_type (TFree z) = TFree z
-    fun do_term (Const (a, T)) = Const (a, do_type T)
-      | do_term (Free (a, T)) = Free (a, do_type T)
-      | do_term (Var (xi, T)) = Var (xi, do_type T)
-      | do_term (t as Bound _) = t
-      | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t)
-      | do_term (t1 $ t2) = do_term t1 $ do_term t2
-  in t |> not (Vartab.is_empty tvar_tab) ? do_term end
-
-fun quantify_over_free quant_s free_s body_t =
-  case Term.add_frees body_t [] |> filter (curry (op =) free_s o fst) of
-    [] => body_t
-  | frees as (_, free_T) :: _ =>
-    Abs (free_s, free_T, fold (curry abstract_over) (map Free frees) body_t)
-
-(* Interpret an ATP formula as a HOL term, extracting sort constraints as they
-   appear in the formula. *)
-fun prop_from_formula thy full_types tfrees phi =
-  let
-    fun do_formula pos phi =
-      case phi of
-        AQuant (_, [], phi) => do_formula pos phi
-      | AQuant (q, x :: xs, phi') =>
-        do_formula pos (AQuant (q, xs, phi'))
-        #>> quantify_over_free (case q of
-                                  AForall => @{const_name All}
-                                | AExists => @{const_name Ex})
-                               (repair_atp_variable_name Char.toLower x)
-      | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
-      | AConn (c, [phi1, phi2]) =>
-        do_formula (pos |> c = AImplies ? not) phi1
-        ##>> do_formula pos phi2
-        #>> (case c of
-               AAnd => s_conj
-             | AOr => s_disj
-             | AImplies => s_imp
-             | AIf => s_imp o swap
-             | AIff => s_iff
-             | ANotIff => s_not o s_iff)
-      | AAtom tm => term_from_pred thy full_types tfrees pos tm
-      | _ => raise FORMULA [phi]
-  in repair_tvar_sorts (do_formula true phi Vartab.empty) end
-
-fun check_formula ctxt =
-  Type_Infer.constrain HOLogic.boolT
-  #> Syntax.check_term (ProofContext.set_mode ProofContext.mode_schematic ctxt)
-
-
-(**** Translation of TSTP files to Isar Proofs ****)
-
-fun unvarify_term (Var ((s, 0), T)) = Free (s, T)
-  | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
-
-fun decode_line full_types tfrees (Definition (num, phi1, phi2)) ctxt =
-    let
-      val thy = ProofContext.theory_of ctxt
-      val t1 = prop_from_formula thy full_types tfrees phi1
-      val vars = snd (strip_comb t1)
-      val frees = map unvarify_term vars
-      val unvarify_args = subst_atomic (vars ~~ frees)
-      val t2 = prop_from_formula thy full_types tfrees phi2
-      val (t1, t2) =
-        HOLogic.eq_const HOLogic.typeT $ t1 $ t2
-        |> unvarify_args |> uncombine_term |> check_formula ctxt
-        |> HOLogic.dest_eq
-    in
-      (Definition (num, t1, t2),
-       fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt)
-    end
-  | decode_line full_types tfrees (Inference (num, u, deps)) ctxt =
-    let
-      val thy = ProofContext.theory_of ctxt
-      val t = u |> prop_from_formula thy full_types tfrees
-                |> uncombine_term |> check_formula ctxt
-    in
-      (Inference (num, t, deps),
-       fold Variable.declare_term (OldTerm.term_frees t) ctxt)
-    end
-fun decode_lines ctxt full_types tfrees lines =
-  fst (fold_map (decode_line full_types tfrees) lines ctxt)
-
-fun is_same_inference _ (Definition _) = false
-  | is_same_inference t (Inference (_, t', _)) = t aconv t'
-
-(* No "real" literals means only type information (tfree_tcs, clsrel, or
-   clsarity). *)
-val is_only_type_information = curry (op aconv) HOLogic.true_const
-
-fun replace_one_dep (old, new) dep = if dep = old then new else [dep]
-fun replace_deps_in_line _ (line as Definition _) = line
-  | replace_deps_in_line p (Inference (num, t, deps)) =
-    Inference (num, t, fold (union (op =) o replace_one_dep p) deps [])
-
-(* Discard axioms; consolidate adjacent lines that prove the same formula, since
-   they differ only in type information.*)
-fun add_line _ _ (line as Definition _) lines = line :: lines
-  | add_line conjecture_shape axiom_names (Inference (num, t, [])) lines =
-    (* No dependencies: axiom, conjecture, or (for Vampire) internal axioms or
-       definitions. *)
-    if is_axiom_number axiom_names num then
-      (* Axioms are not proof lines. *)
-      if is_only_type_information t then
-        map (replace_deps_in_line (num, [])) lines
-      (* Is there a repetition? If so, replace later line by earlier one. *)
-      else case take_prefix (not o is_same_inference t) lines of
-        (_, []) => lines (*no repetition of proof line*)
-      | (pre, Inference (num', _, _) :: post) =>
-        pre @ map (replace_deps_in_line (num', [num])) post
-    else if is_conjecture_number conjecture_shape num then
-      Inference (num, negate_term t, []) :: lines
-    else
-      map (replace_deps_in_line (num, [])) lines
-  | add_line _ _ (Inference (num, t, deps)) lines =
-    (* Type information will be deleted later; skip repetition test. *)
-    if is_only_type_information t then
-      Inference (num, t, deps) :: lines
-    (* Is there a repetition? If so, replace later line by earlier one. *)
-    else case take_prefix (not o is_same_inference t) lines of
-      (* FIXME: Doesn't this code risk conflating proofs involving different
-         types? *)
-       (_, []) => Inference (num, t, deps) :: lines
-     | (pre, Inference (num', t', _) :: post) =>
-       Inference (num, t', deps) ::
-       pre @ map (replace_deps_in_line (num', [num])) post
-
-(* Recursively delete empty lines (type information) from the proof. *)
-fun add_nontrivial_line (Inference (num, t, [])) lines =
-    if is_only_type_information t then delete_dep num lines
-    else Inference (num, t, []) :: lines
-  | add_nontrivial_line line lines = line :: lines
-and delete_dep num lines =
-  fold_rev add_nontrivial_line (map (replace_deps_in_line (num, [])) lines) []
-
-(* ATPs sometimes reuse free variable names in the strangest ways. Removing
-   offending lines often does the trick. *)
-fun is_bad_free frees (Free x) = not (member (op =) frees x)
-  | is_bad_free _ _ = false
-
-(* Vampire is keen on producing these. *)
-fun is_trivial_formula (@{const Not} $ (Const (@{const_name HOL.eq}, _)
-                                        $ t1 $ t2)) = (t1 aconv t2)
-  | is_trivial_formula _ = false
-
-fun add_desired_line _ _ _ _ (line as Definition (num, _, _)) (j, lines) =
-    (j, line :: map (replace_deps_in_line (num, [])) lines)
-  | add_desired_line isar_shrink_factor conjecture_shape axiom_names frees
-                     (Inference (num, t, deps)) (j, lines) =
-    (j + 1,
-     if is_axiom_number axiom_names num orelse
-        is_conjecture_number conjecture_shape num orelse
-        (not (is_only_type_information t) andalso
-         null (Term.add_tvars t []) andalso
-         not (exists_subterm (is_bad_free frees) t) andalso
-         not (is_trivial_formula t) andalso
-         (null lines orelse (* last line must be kept *)
-          (length deps >= 2 andalso j mod isar_shrink_factor = 0))) then
-       Inference (num, t, deps) :: lines  (* keep line *)
-     else
-       map (replace_deps_in_line (num, deps)) lines)  (* drop line *)
-
-(** EXTRACTING LEMMAS **)
-
-(* Like "split_line", but ignores "\n" that follow a comma (as in SNARK's
-   output). *)
-val split_proof_lines =
-  let
-    fun aux [] [] = []
-      | aux line [] = [implode (rev line)]
-      | aux line ("," :: "\n" :: rest) = aux ("," :: line) rest
-      | aux line ("\n" :: rest) = aux line [] @ aux [] rest
-      | aux line (s :: rest) = aux (s :: line) rest
-  in aux [] o explode end
-
-(* A list consisting of the first number in each line is returned. For TSTP,
-   interesting lines have the form "fof(108, axiom, ...)", where the number
-   (108) is extracted. For SPASS, lines have the form "108[0:Inp] ...", where
-   the first number (108) is extracted. For Vampire, we look for
-   "108. ... [input]". *)
-fun used_facts_in_atp_proof axiom_names atp_proof =
-  let
-    fun axiom_names_at_index num =
-      let val j = Int.fromString num |> the_default ~1 in
-        if is_axiom_number axiom_names j then Vector.sub (axiom_names, j - 1)
-        else []
-      end
-    val tokens_of =
-      String.tokens (fn c => not (Char.isAlphaNum c) andalso c <> #"_")
-    fun do_line (tag :: num :: "axiom" :: (rest as _ :: _)) =
-        if tag = "cnf" orelse tag = "fof" then
-          (case strip_prefix_and_unascii axiom_prefix (List.last rest) of
-             SOME name =>
-             if member (op =) rest "file" then
-               ([(name, name |> find_first_in_list_vector axiom_names |> the)]
-                handle Option.Option =>
-                       error ("No such fact: " ^ quote name ^ "."))
-             else
-               axiom_names_at_index num
-           | NONE => axiom_names_at_index num)
-        else
-          []
-      | do_line (num :: "0" :: "Inp" :: _) = axiom_names_at_index num
-      | do_line (num :: rest) =
-        (case List.last rest of "input" => axiom_names_at_index num | _ => [])
-      | do_line _ = []
-  in atp_proof |> split_proof_lines |> maps (do_line o tokens_of) end
-
-val indent_size = 2
-val no_label = ("", ~1)
-
-val raw_prefix = "X"
-val assum_prefix = "A"
-val fact_prefix = "F"
-
-fun string_for_label (s, num) = s ^ string_of_int num
-
-fun metis_using [] = ""
-  | metis_using ls =
-    "using " ^ space_implode " " (map string_for_label ls) ^ " "
-fun metis_apply _ 1 = "by "
-  | metis_apply 1 _ = "apply "
-  | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply "
-fun metis_name full_types = if full_types then "metisFT" else "metis"
-fun metis_call full_types [] = metis_name full_types
-  | metis_call full_types ss =
-    "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")"
-fun metis_command full_types i n (ls, ss) =
-  metis_using ls ^ metis_apply i n ^ metis_call full_types ss
-fun metis_line full_types i n ss =
-  "Try this command: " ^
-  Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "."
-fun minimize_line _ [] = ""
-  | minimize_line minimize_command ss =
-    case minimize_command ss of
-      "" => ""
-    | command =>
-      "\nTo minimize the number of lemmas, try this: " ^
-      Markup.markup Markup.sendback command ^ "."
-
-fun used_facts axiom_names =
-  used_facts_in_atp_proof axiom_names
-  #> List.partition (curry (op =) Chained o snd)
-  #> pairself (sort_distinct (string_ord o pairself fst))
-
-fun metis_proof_text (full_types, minimize_command, atp_proof, axiom_names,
-                      goal, i) =
-  let
-    val (chained_lemmas, other_lemmas) = used_facts axiom_names atp_proof
-    val n = Logic.count_prems (prop_of goal)
-  in
-    (metis_line full_types i n (map fst other_lemmas) ^
-     minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
-     other_lemmas @ chained_lemmas)
-  end
-
-(** Isar proof construction and manipulation **)
-
-fun merge_fact_sets (ls1, ss1) (ls2, ss2) =
-  (union (op =) ls1 ls2, union (op =) ss1 ss2)
-
-type label = string * int
-type facts = label list * string list
-
-datatype qualifier = Show | Then | Moreover | Ultimately
-
-datatype step =
-  Fix of (string * typ) list |
-  Let of term * term |
-  Assume of label * term |
-  Have of qualifier list * label * term * byline
-and byline =
-  ByMetis of facts |
-  CaseSplit of step list list * facts
-
-fun smart_case_split [] facts = ByMetis facts
-  | smart_case_split proofs facts = CaseSplit (proofs, facts)
-
-fun add_fact_from_dep axiom_names num =
-  if is_axiom_number axiom_names num then
-    apsnd (union (op =) (map fst (Vector.sub (axiom_names, num - 1))))
-  else
-    apfst (insert (op =) (raw_prefix, num))
-
-fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t
-fun forall_vars t = fold_rev forall_of (map Var (Term.add_vars t [])) t
-
-fun step_for_line _ _ (Definition (_, t1, t2)) = Let (t1, t2)
-  | step_for_line _ _ (Inference (num, t, [])) = Assume ((raw_prefix, num), t)
-  | step_for_line axiom_names j (Inference (num, t, deps)) =
-    Have (if j = 1 then [Show] else [], (raw_prefix, num),
-          forall_vars t,
-          ByMetis (fold (add_fact_from_dep axiom_names) deps ([], [])))
-
-fun proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
-                         atp_proof conjecture_shape axiom_names params frees =
-  let
-    val lines =
-      atp_proof ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
-      |> parse_proof pool
-      |> sort (int_ord o pairself raw_step_number)
-      |> decode_lines ctxt full_types tfrees
-      |> rpair [] |-> fold_rev (add_line conjecture_shape axiom_names)
-      |> rpair [] |-> fold_rev add_nontrivial_line
-      |> rpair (0, []) |-> fold_rev (add_desired_line isar_shrink_factor
-                                             conjecture_shape axiom_names frees)
-      |> snd
-  in
-    (if null params then [] else [Fix params]) @
-    map2 (step_for_line axiom_names) (length lines downto 1) lines
-  end
-
-(* When redirecting proofs, we keep information about the labels seen so far in
-   the "backpatches" data structure. The first component indicates which facts
-   should be associated with forthcoming proof steps. The second component is a
-   pair ("assum_ls", "drop_ls"), where "assum_ls" are the labels that should
-   become assumptions and "drop_ls" are the labels that should be dropped in a
-   case split. *)
-type backpatches = (label * facts) list * (label list * label list)
-
-fun used_labels_of_step (Have (_, _, _, by)) =
-    (case by of
-       ByMetis (ls, _) => ls
-     | CaseSplit (proofs, (ls, _)) =>
-       fold (union (op =) o used_labels_of) proofs ls)
-  | used_labels_of_step _ = []
-and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof []
-
-fun new_labels_of_step (Fix _) = []
-  | new_labels_of_step (Let _) = []
-  | new_labels_of_step (Assume (l, _)) = [l]
-  | new_labels_of_step (Have (_, l, _, _)) = [l]
-val new_labels_of = maps new_labels_of_step
-
-val join_proofs =
-  let
-    fun aux _ [] = NONE
-      | aux proof_tail (proofs as (proof1 :: _)) =
-        if exists null proofs then
-          NONE
-        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, _) => (* 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))
-                         (used_labels_of proof_tail)) then
-            SOME (l, t, map rev proofs, proof_tail)
-          else
-            NONE
-        | _ => NONE
-  in aux [] o map rev end
-
-fun case_split_qualifiers proofs =
-  case length proofs of
-    0 => []
-  | 1 => [Then]
-  | _ => [Ultimately]
-
-fun redirect_proof conjecture_shape hyp_ts concl_t proof =
-  let
-    (* The first pass outputs those steps that are independent of the negated
-       conjecture. The second pass flips the proof by contradiction to obtain a
-       direct proof, introducing case splits when an inference depends on
-       several facts that depend on the negated conjecture. *)
-    fun find_hyp num =
-      nth hyp_ts (index_in_shape num conjecture_shape)
-      handle Subscript =>
-             raise Fail ("Cannot find hypothesis " ^ Int.toString num)
-     val concl_ls = map (pair raw_prefix) (List.last conjecture_shape)
-     val canonicalize_labels =
-       map (fn l => if member (op =) concl_ls l then hd concl_ls else l)
-       #> distinct (op =)
-     fun first_pass ([], contra) = ([], contra)
-       | first_pass ((step as Fix _) :: proof, contra) =
-         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), _)) :: proof, contra) =
-         if member (op =) concl_ls l then
-           first_pass (proof, contra ||> l = hd concl_ls ? cons step)
-         else
-           first_pass (proof, contra) |>> cons (Assume (l, find_hyp num))
-       | first_pass (Have (qs, l, t, ByMetis (ls, ss)) :: proof, contra) =
-         let
-           val ls = canonicalize_labels ls
-           val step = Have (qs, l, t, ByMetis (ls, ss))
-         in
-           if exists (member (op =) (fst contra)) ls then
-             first_pass (proof, contra |>> cons l ||> cons step)
-           else
-             first_pass (proof, contra) |>> cons step
-         end
-       | first_pass _ = raise Fail "malformed proof"
-    val (proof_top, (contra_ls, contra_proof)) =
-      first_pass (proof, (concl_ls, []))
-    val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst
-    fun backpatch_labels patches ls =
-      fold merge_fact_sets (map (backpatch_label patches) ls) ([], [])
-    fun second_pass end_qs ([], assums, patches) =
-        ([Have (end_qs, no_label, concl_t,
-                ByMetis (backpatch_labels patches (map snd assums)))], patches)
-      | second_pass end_qs (Assume (l, t) :: proof, assums, patches) =
-        second_pass end_qs (proof, (t, l) :: assums, patches)
-      | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums,
-                            patches) =
-        if member (op =) (snd (snd patches)) l andalso
-           not (member (op =) (fst (snd patches)) l) andalso
-           not (AList.defined (op =) (fst patches) l) then
-          second_pass end_qs (proof, assums, patches ||> apsnd (append ls))
-        else
-          (case List.partition (member (op =) contra_ls) ls of
-             ([contra_l], co_ls) =>
-             if member (op =) qs Show then
-               second_pass end_qs (proof, assums,
-                                   patches |>> cons (contra_l, (co_ls, ss)))
-             else
-               second_pass end_qs
-                           (proof, assums,
-                            patches |>> cons (contra_l, (l :: co_ls, ss)))
-               |>> cons (if member (op =) (fst (snd patches)) l then
-                           Assume (l, negate_term t)
-                         else
-                           Have (qs, l, negate_term t,
-                                 ByMetis (backpatch_label patches l)))
-           | (contra_ls as _ :: _, co_ls) =>
-             let
-               val proofs =
-                 map_filter
-                     (fn l =>
-                         if member (op =) concl_ls l then
-                           NONE
-                         else
-                           let
-                             val drop_ls = filter (curry (op <>) l) contra_ls
-                           in
-                             second_pass []
-                                 (proof, assums,
-                                  patches ||> apfst (insert (op =) l)
-                                          ||> apsnd (union (op =) drop_ls))
-                             |> fst |> SOME
-                           end) contra_ls
-               val (assumes, facts) =
-                 if member (op =) (fst (snd patches)) l then
-                   ([Assume (l, negate_term t)], (l :: co_ls, ss))
-                 else
-                   ([], (co_ls, ss))
-             in
-               (case join_proofs proofs of
-                  SOME (l, t, proofs, proof_tail) =>
-                  Have (case_split_qualifiers proofs @
-                        (if null proof_tail then end_qs else []), l, t,
-                        smart_case_split proofs facts) :: proof_tail
-                | NONE =>
-                  [Have (case_split_qualifiers proofs @ end_qs, no_label,
-                         concl_t, smart_case_split proofs facts)],
-                patches)
-               |>> append assumes
-             end
-           | _ => raise Fail "malformed proof")
-       | second_pass _ _ = raise Fail "malformed proof"
-    val proof_bottom =
-      second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst
-  in proof_top @ proof_bottom end
-
-(* FIXME: Still needed? Probably not. *)
-val kill_duplicate_assumptions_in_proof =
-  let
-    fun relabel_facts subst =
-      apfst (map (fn l => AList.lookup (op =) subst l |> the_default l))
-    fun do_step (step as Assume (l, t)) (proof, subst, assums) =
-        (case AList.lookup (op aconv) assums t of
-           SOME l' => (proof, (l, l') :: subst, assums)
-         | NONE => (step :: proof, subst, (t, l) :: assums))
-      | do_step (Have (qs, l, t, by)) (proof, subst, assums) =
-        (Have (qs, l, t,
-               case by of
-                 ByMetis facts => ByMetis (relabel_facts subst facts)
-               | CaseSplit (proofs, facts) =>
-                 CaseSplit (map do_proof proofs, relabel_facts subst facts)) ::
-         proof, subst, assums)
-      | do_step step (proof, subst, assums) = (step :: proof, subst, assums)
-    and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev
-  in do_proof end
-
-val then_chain_proof =
-  let
-    fun aux _ [] = []
-      | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof
-      | aux l' (Have (qs, l, t, by) :: proof) =
-        (case by of
-           ByMetis (ls, ss) =>
-           Have (if member (op =) ls l' then
-                   (Then :: qs, l, t,
-                    ByMetis (filter_out (curry (op =) l') ls, ss))
-                 else
-                   (qs, l, t, ByMetis (ls, ss)))
-         | CaseSplit (proofs, facts) =>
-           Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) ::
-        aux l proof
-      | aux _ (step :: proof) = step :: aux no_label proof
-  in aux no_label end
-
-fun kill_useless_labels_in_proof proof =
-  let
-    val used_ls = used_labels_of proof
-    fun do_label l = if member (op =) used_ls l then l else no_label
-    fun do_step (Assume (l, t)) = Assume (do_label l, t)
-      | do_step (Have (qs, l, t, by)) =
-        Have (qs, do_label l, t,
-              case by of
-                CaseSplit (proofs, facts) =>
-                CaseSplit (map (map do_step) proofs, facts)
-              | _ => by)
-      | do_step step = step
-  in map do_step proof end
-
-fun prefix_for_depth n = replicate_string (n + 1)
-
-val relabel_proof =
-  let
-    fun aux _ _ _ [] = []
-      | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) =
-        if l = no_label then
-          Assume (l, t) :: aux subst depth (next_assum, next_fact) proof
-        else
-          let val l' = (prefix_for_depth depth assum_prefix, next_assum) in
-            Assume (l', t) ::
-            aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof
-          end
-      | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) =
-        let
-          val (l', subst, next_fact) =
-            if l = no_label then
-              (l, subst, next_fact)
-            else
-              let
-                val l' = (prefix_for_depth depth fact_prefix, next_fact)
-              in (l', (l, l') :: subst, next_fact + 1) end
-          val relabel_facts =
-            apfst (map (fn l =>
-                           case AList.lookup (op =) subst l of
-                             SOME l' => l'
-                           | NONE => raise Fail ("unknown label " ^
-                                                 quote (string_for_label l))))
-          val by =
-            case by of
-              ByMetis facts => ByMetis (relabel_facts facts)
-            | CaseSplit (proofs, facts) =>
-              CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs,
-                         relabel_facts facts)
-        in
-          Have (qs, l', t, by) ::
-          aux subst depth (next_assum, next_fact) proof
-        end
-      | aux subst depth nextp (step :: proof) =
-        step :: aux subst depth nextp proof
-  in aux [] 0 (1, 1) end
-
-fun string_for_proof ctxt full_types i n =
-  let
-    fun fix_print_mode f x =
-      setmp_CRITICAL show_no_free_types true
-          (setmp_CRITICAL show_types true
-               (Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
-                                         (print_mode_value ())) f)) x
-    fun do_indent ind = replicate_string (ind * indent_size) " "
-    fun do_free (s, T) =
-      maybe_quote s ^ " :: " ^
-      maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T)
-    fun do_label l = if l = no_label then "" else string_for_label l ^ ": "
-    fun do_have qs =
-      (if member (op =) qs Moreover then "moreover " else "") ^
-      (if member (op =) qs Ultimately then "ultimately " else "") ^
-      (if member (op =) qs Then then
-         if member (op =) qs Show then "thus" else "hence"
-       else
-         if member (op =) qs Show then "show" else "have")
-    val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt)
-    fun do_facts (ls, ss) =
-      metis_command full_types 1 1
-                    (ls |> sort_distinct (prod_ord string_ord int_ord),
-                     ss |> sort_distinct string_ord)
-    and do_step ind (Fix xs) =
-        do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n"
-      | do_step ind (Let (t1, t2)) =
-        do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n"
-      | do_step ind (Assume (l, t)) =
-        do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n"
-      | do_step ind (Have (qs, l, t, ByMetis facts)) =
-        do_indent ind ^ do_have qs ^ " " ^
-        do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n"
-      | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) =
-        space_implode (do_indent ind ^ "moreover\n")
-                      (map (do_block ind) proofs) ^
-        do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^
-        do_facts facts ^ "\n"
-    and do_steps prefix suffix ind steps =
-      let val s = implode (map (do_step ind) steps) in
-        replicate_string (ind * indent_size - size prefix) " " ^ prefix ^
-        String.extract (s, ind * indent_size,
-                        SOME (size s - ind * indent_size - 1)) ^
-        suffix ^ "\n"
-      end
-    and do_block ind proof = do_steps "{ " " }" (ind + 1) proof
-    (* One-step proofs are pointless; better use the Metis one-liner
-       directly. *)
-    and do_proof [Have (_, _, _, ByMetis _)] = ""
-      | do_proof proof =
-        (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
-        do_indent 0 ^ "proof -\n" ^
-        do_steps "" "" 1 proof ^
-        do_indent 0 ^ (if n <> 1 then "next" else "qed")
-  in do_proof end
-
-fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
-                    (other_params as (full_types, _, atp_proof, axiom_names,
-                                      goal, i)) =
-  let
-    val (params, hyp_ts, concl_t) = strip_subgoal goal i
-    val frees = fold Term.add_frees (concl_t :: hyp_ts) []
-    val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) []
-    val n = Logic.count_prems (prop_of goal)
-    val (one_line_proof, lemma_names) = metis_proof_text other_params
-    fun isar_proof_for () =
-      case proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
-                                atp_proof conjecture_shape axiom_names params
-                                frees
-           |> redirect_proof conjecture_shape hyp_ts concl_t
-           |> kill_duplicate_assumptions_in_proof
-           |> then_chain_proof
-           |> kill_useless_labels_in_proof
-           |> relabel_proof
-           |> string_for_proof ctxt full_types i n of
-        "" => "\nNo structured proof available."
-      | proof => "\n\nStructured proof:\n" ^ Markup.markup Markup.sendback proof
-    val isar_proof =
-      if debug then
-        isar_proof_for ()
-      else
-        try isar_proof_for ()
-        |> the_default "\nWarning: The Isar proof construction failed."
-  in (one_line_proof ^ isar_proof, lemma_names) end
-
-fun proof_text isar_proof isar_params other_params =
-  (if isar_proof then isar_proof_text isar_params else metis_proof_text)
-      other_params
-
-end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Tue Aug 31 23:46:23 2010 +0200
@@ -0,0 +1,1038 @@
+(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML
+    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
+    Author:     Claire Quigley, Cambridge University Computer Laboratory
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Transfer of proofs from external provers.
+*)
+
+signature SLEDGEHAMMER_PROOF_RECONSTRUCT =
+sig
+  type locality = Sledgehammer_Fact_Filter.locality
+  type minimize_command = string list -> string
+  type metis_params =
+    bool * minimize_command * string * (string * locality) list vector * thm
+    * int
+  type isar_params =
+    string Symtab.table * bool * int * Proof.context * int list list
+  type text_result = string * (string * locality) list
+
+  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
+end;
+
+structure Sledgehammer_Proof_Reconstruct : SLEDGEHAMMER_PROOF_RECONSTRUCT =
+struct
+
+open ATP_Problem
+open Metis_Clauses
+open Sledgehammer_Util
+open Sledgehammer_Fact_Filter
+open Sledgehammer_Translate
+
+type minimize_command = string list -> string
+type metis_params =
+  bool * minimize_command * string * (string * locality) list vector * thm * int
+type isar_params =
+  string Symtab.table * bool * int * Proof.context * int list list
+type text_result = string * (string * locality) list
+
+(* Simple simplifications to ensure that sort annotations don't leave a trail of
+   spurious "True"s. *)
+fun s_not @{const False} = @{const True}
+  | s_not @{const True} = @{const False}
+  | s_not (@{const Not} $ t) = t
+  | s_not t = @{const Not} $ t
+fun s_conj (@{const True}, t2) = t2
+  | s_conj (t1, @{const True}) = t1
+  | s_conj p = HOLogic.mk_conj p
+fun s_disj (@{const False}, t2) = t2
+  | s_disj (t1, @{const False}) = t1
+  | s_disj p = HOLogic.mk_disj p
+fun s_imp (@{const True}, t2) = t2
+  | s_imp (t1, @{const False}) = s_not t1
+  | s_imp p = HOLogic.mk_imp p
+fun s_iff (@{const True}, t2) = t2
+  | s_iff (t1, @{const True}) = t1
+  | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
+
+fun mk_anot (AConn (ANot, [phi])) = phi
+  | mk_anot phi = AConn (ANot, [phi])
+fun mk_aconn c (phi1, phi2) = AConn (c, [phi1, phi2])
+
+fun index_in_shape x = find_index (exists (curry (op =) x))
+fun is_axiom_number axiom_names num =
+  num > 0 andalso num <= Vector.length axiom_names andalso
+  not (null (Vector.sub (axiom_names, num - 1)))
+fun is_conjecture_number conjecture_shape num =
+  index_in_shape num conjecture_shape >= 0
+
+fun negate_term (Const (@{const_name All}, T) $ Abs (s, T', t')) =
+    Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t')
+  | negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
+    Const (@{const_name All}, T) $ Abs (s, T', negate_term t')
+  | negate_term (@{const HOL.implies} $ t1 $ t2) =
+    @{const HOL.conj} $ t1 $ negate_term t2
+  | negate_term (@{const HOL.conj} $ t1 $ t2) =
+    @{const HOL.disj} $ negate_term t1 $ negate_term t2
+  | negate_term (@{const HOL.disj} $ t1 $ t2) =
+    @{const HOL.conj} $ negate_term t1 $ negate_term t2
+  | negate_term (@{const Not} $ t) = t
+  | negate_term t = @{const Not} $ t
+
+datatype ('a, 'b, 'c, 'd, 'e) raw_step =
+  Definition of 'a * 'b * 'c |
+  Inference of 'a * 'd * 'e list
+
+fun raw_step_number (Definition (num, _, _)) = num
+  | raw_step_number (Inference (num, _, _)) = num
+
+(**** PARSING OF TSTP FORMAT ****)
+
+(*Strings enclosed in single quotes, e.g. filenames*)
+val scan_quoted = $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode;
+
+val scan_dollar_name =
+  Scan.repeat ($$ "$") -- Symbol.scan_id >> (fn (ss, s) => implode ss ^ s)
+
+fun repair_name _ "$true" = "c_True"
+  | repair_name _ "$false" = "c_False"
+  | repair_name _ "$$e" = "c_equal" (* seen in Vampire proofs *)
+  | repair_name _ "equal" = "c_equal" (* needed by SPASS? *)
+  | repair_name pool s =
+    case Symtab.lookup pool s of
+      SOME s' => s'
+    | NONE =>
+      if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then
+        "c_equal" (* seen in Vampire proofs *)
+      else
+        s
+(* Generalized first-order terms, which include file names, numbers, etc. *)
+val parse_potential_integer =
+  (scan_dollar_name || scan_quoted) >> K NONE
+  || scan_integer >> SOME
+fun parse_annotation x =
+  ((parse_potential_integer ::: Scan.repeat ($$ " " |-- parse_potential_integer)
+    >> map_filter I) -- Scan.optional parse_annotation []
+     >> uncurry (union (op =))
+   || $$ "(" |-- parse_annotations --| $$ ")"
+   || $$ "[" |-- parse_annotations --| $$ "]") x
+and parse_annotations x =
+  (Scan.optional (parse_annotation
+                  ::: Scan.repeat ($$ "," |-- parse_annotation)) []
+   >> (fn numss => fold (union (op =)) numss [])) x
+
+(* Vampire proof lines sometimes contain needless information such as "(0:3)",
+   which can be hard to disambiguate from function application in an LL(1)
+   parser. As a workaround, we extend the TPTP term syntax with such detritus
+   and ignore it. *)
+fun parse_vampire_detritus x =
+  (scan_integer |-- $$ ":" --| scan_integer >> K []) x
+
+fun parse_term pool x =
+  ((scan_dollar_name >> repair_name pool)
+    -- Scan.optional ($$ "(" |-- (parse_vampire_detritus || parse_terms pool)
+                      --| $$ ")") []
+    --| Scan.optional ($$ "(" |-- parse_vampire_detritus --| $$ ")") []
+   >> ATerm) x
+and parse_terms pool x =
+  (parse_term pool ::: Scan.repeat ($$ "," |-- parse_term pool)) x
+
+fun parse_atom pool =
+  parse_term pool -- Scan.option (Scan.option ($$ "!") --| $$ "="
+                                  -- parse_term pool)
+  >> (fn (u1, NONE) => AAtom u1
+       | (u1, SOME (NONE, u2)) => AAtom (ATerm ("c_equal", [u1, u2]))
+       | (u1, SOME (SOME _, u2)) =>
+         mk_anot (AAtom (ATerm ("c_equal", [u1, u2]))))
+
+fun fo_term_head (ATerm (s, _)) = s
+
+(* TPTP formulas are fully parenthesized, so we don't need to worry about
+   operator precedence. *)
+fun parse_formula pool x =
+  (($$ "(" |-- parse_formula pool --| $$ ")"
+    || ($$ "!" >> K AForall || $$ "?" >> K AExists)
+       --| $$ "[" -- parse_terms pool --| $$ "]" --| $$ ":"
+       -- parse_formula pool
+       >> (fn ((q, ts), phi) => AQuant (q, map fo_term_head ts, phi))
+    || $$ "~" |-- parse_formula pool >> mk_anot
+    || parse_atom pool)
+   -- Scan.option ((Scan.this_string "=>" >> K AImplies
+                    || Scan.this_string "<=>" >> K AIff
+                    || Scan.this_string "<~>" >> K ANotIff
+                    || Scan.this_string "<=" >> K AIf
+                    || $$ "|" >> K AOr || $$ "&" >> K AAnd)
+                   -- parse_formula pool)
+   >> (fn (phi1, NONE) => phi1
+        | (phi1, SOME (c, phi2)) => mk_aconn c (phi1, phi2))) x
+
+val parse_tstp_extra_arguments =
+  Scan.optional ($$ "," |-- parse_annotation
+                 --| Scan.option ($$ "," |-- parse_annotations)) []
+
+(* Syntax: (fof|cnf)\(<num>, <formula_role>, <formula> <extra_arguments>\).
+   The <num> could be an identifier, but we assume integers. *)
+ fun parse_tstp_line pool =
+   ((Scan.this_string "fof" || Scan.this_string "cnf") -- $$ "(")
+     |-- scan_integer --| $$ "," -- Symbol.scan_id --| $$ ","
+     -- parse_formula pool -- parse_tstp_extra_arguments --| $$ ")" --| $$ "."
+    >> (fn (((num, role), phi), deps) =>
+           case role of
+             "definition" =>
+             (case phi of
+                AConn (AIff, [phi1 as AAtom _, phi2]) =>
+                Definition (num, phi1, phi2)
+              | AAtom (ATerm ("c_equal", _)) =>
+                Inference (num, phi, deps) (* Vampire's equality proxy axiom *)
+              | _ => raise Fail "malformed definition")
+           | _ => Inference (num, phi, deps))
+
+(**** PARSING OF VAMPIRE OUTPUT ****)
+
+(* Syntax: <num>. <formula> <annotation> *)
+fun parse_vampire_line pool =
+  scan_integer --| $$ "." -- parse_formula pool -- parse_annotation
+  >> (fn ((num, phi), deps) => Inference (num, phi, deps))
+
+(**** PARSING OF SPASS OUTPUT ****)
+
+(* SPASS returns clause references of the form "x.y". We ignore "y", whose role
+   is not clear anyway. *)
+val parse_dot_name = scan_integer --| $$ "." --| scan_integer
+
+val parse_spass_annotations =
+  Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name
+                                         --| Scan.option ($$ ","))) []
+
+(* It is not clear why some literals are followed by sequences of stars and/or
+   pluses. We ignore them. *)
+fun parse_decorated_atom pool =
+  parse_atom pool --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")
+
+fun mk_horn ([], []) = AAtom (ATerm ("c_False", []))
+  | mk_horn ([], pos_lits) = foldr1 (mk_aconn AOr) pos_lits
+  | mk_horn (neg_lits, []) = mk_anot (foldr1 (mk_aconn AAnd) neg_lits)
+  | mk_horn (neg_lits, pos_lits) =
+    mk_aconn AImplies (foldr1 (mk_aconn AAnd) neg_lits,
+                       foldr1 (mk_aconn AOr) pos_lits)
+
+fun parse_horn_clause pool =
+  Scan.repeat (parse_decorated_atom pool) --| $$ "|" --| $$ "|"
+    -- Scan.repeat (parse_decorated_atom pool) --| $$ "-" --| $$ ">"
+    -- Scan.repeat (parse_decorated_atom pool)
+  >> (mk_horn o apfst (op @))
+
+(* Syntax: <num>[0:<inference><annotations>]
+   <atoms> || <atoms> -> <atoms>. *)
+fun parse_spass_line pool =
+  scan_integer --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id
+    -- parse_spass_annotations --| $$ "]" -- parse_horn_clause pool --| $$ "."
+  >> (fn ((num, deps), u) => Inference (num, u, deps))
+
+fun parse_line pool =
+  parse_tstp_line pool || parse_vampire_line pool || parse_spass_line pool
+fun parse_lines pool = Scan.repeat1 (parse_line pool)
+fun parse_proof pool =
+  fst o Scan.finite Symbol.stopper
+            (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output")
+                            (parse_lines pool)))
+  o explode o strip_spaces_except_between_ident_chars
+
+(**** INTERPRETATION OF TSTP SYNTAX TREES ****)
+
+exception FO_TERM of string fo_term list
+exception FORMULA of (string, string fo_term) formula list
+exception SAME of unit
+
+(* Type variables are given the basic sort "HOL.type". Some will later be
+   constrained by information from type literals, or by type inference. *)
+fun type_from_fo_term tfrees (u as ATerm (a, us)) =
+  let val Ts = map (type_from_fo_term tfrees) us in
+    case strip_prefix_and_unascii type_const_prefix a of
+      SOME b => Type (invert_const b, Ts)
+    | NONE =>
+      if not (null us) then
+        raise FO_TERM [u]  (* only "tconst"s have type arguments *)
+      else case strip_prefix_and_unascii tfree_prefix a of
+        SOME b =>
+        let val s = "'" ^ b in
+          TFree (s, AList.lookup (op =) tfrees s |> the_default HOLogic.typeS)
+        end
+      | NONE =>
+        case strip_prefix_and_unascii tvar_prefix a of
+          SOME b => TVar (("'" ^ b, 0), HOLogic.typeS)
+        | NONE =>
+          (* Variable from the ATP, say "X1" *)
+          Type_Infer.param 0 (a, HOLogic.typeS)
+  end
+
+(* Type class literal applied to a type. Returns triple of polarity, class,
+   type. *)
+fun type_constraint_from_term pos tfrees (u as ATerm (a, us)) =
+  case (strip_prefix_and_unascii class_prefix a,
+        map (type_from_fo_term tfrees) us) of
+    (SOME b, [T]) => (pos, b, T)
+  | _ => raise FO_TERM [u]
+
+(** Accumulate type constraints in a formula: negative type literals **)
+fun add_var (key, z)  = Vartab.map_default (key, []) (cons z)
+fun add_type_constraint (false, cl, TFree (a ,_)) = add_var ((a, ~1), cl)
+  | add_type_constraint (false, cl, TVar (ix, _)) = add_var (ix, cl)
+  | add_type_constraint _ = I
+
+fun repair_atp_variable_name f s =
+  let
+    fun subscript_name s n = s ^ nat_subscript n
+    val s = String.map f s
+  in
+    case space_explode "_" s of
+      [_] => (case take_suffix Char.isDigit (String.explode s) of
+                (cs1 as _ :: _, cs2 as _ :: _) =>
+                subscript_name (String.implode cs1)
+                               (the (Int.fromString (String.implode cs2)))
+              | (_, _) => s)
+    | [s1, s2] => (case Int.fromString s2 of
+                     SOME n => subscript_name s1 n
+                   | NONE => s)
+    | _ => s
+  end
+
+(* First-order translation. No types are known for variables. "HOLogic.typeT"
+   should allow them to be inferred. *)
+fun raw_term_from_pred thy full_types tfrees =
+  let
+    fun aux opt_T extra_us u =
+      case u of
+        ATerm ("hBOOL", [u1]) => aux (SOME @{typ bool}) [] u1
+      | ATerm ("hAPP", [u1, u2]) => aux opt_T (u2 :: extra_us) u1
+      | ATerm (a, us) =>
+        if a = type_wrapper_name then
+          case us of
+            [typ_u, term_u] =>
+            aux (SOME (type_from_fo_term tfrees typ_u)) extra_us term_u
+          | _ => raise FO_TERM us
+        else case strip_prefix_and_unascii const_prefix a of
+          SOME "equal" =>
+          list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT),
+                     map (aux NONE []) us)
+        | SOME b =>
+          let
+            val c = invert_const b
+            val num_type_args = num_type_args thy c
+            val (type_us, term_us) =
+              chop (if full_types then 0 else num_type_args) us
+            (* Extra args from "hAPP" come after any arguments given directly to
+               the constant. *)
+            val term_ts = map (aux NONE []) term_us
+            val extra_ts = map (aux NONE []) extra_us
+            val t =
+              Const (c, if full_types then
+                          case opt_T of
+                            SOME T => map fastype_of term_ts ---> T
+                          | NONE =>
+                            if num_type_args = 0 then
+                              Sign.const_instance thy (c, [])
+                            else
+                              raise Fail ("no type information for " ^ quote c)
+                        else
+                          Sign.const_instance thy (c,
+                              map (type_from_fo_term tfrees) type_us))
+          in list_comb (t, term_ts @ extra_ts) end
+        | NONE => (* a free or schematic variable *)
+          let
+            val ts = map (aux NONE []) (us @ extra_us)
+            val T = map fastype_of ts ---> HOLogic.typeT
+            val t =
+              case strip_prefix_and_unascii fixed_var_prefix a of
+                SOME b => Free (b, T)
+              | NONE =>
+                case strip_prefix_and_unascii schematic_var_prefix a of
+                  SOME b => Var ((b, 0), T)
+                | NONE =>
+                  if is_tptp_variable a then
+                    Var ((repair_atp_variable_name Char.toLower a, 0), T)
+                  else
+                    (* Skolem constants? *)
+                    Var ((repair_atp_variable_name Char.toUpper a, 0), T)
+          in list_comb (t, ts) end
+  in aux (SOME HOLogic.boolT) [] end
+
+fun term_from_pred thy full_types tfrees pos (u as ATerm (s, _)) =
+  if String.isPrefix class_prefix s then
+    add_type_constraint (type_constraint_from_term pos tfrees u)
+    #> pair @{const True}
+  else
+    pair (raw_term_from_pred thy full_types tfrees u)
+
+val combinator_table =
+  [(@{const_name COMBI}, @{thm COMBI_def_raw}),
+   (@{const_name COMBK}, @{thm COMBK_def_raw}),
+   (@{const_name COMBB}, @{thm COMBB_def_raw}),
+   (@{const_name COMBC}, @{thm COMBC_def_raw}),
+   (@{const_name COMBS}, @{thm COMBS_def_raw})]
+
+fun uncombine_term (t1 $ t2) = betapply (pairself uncombine_term (t1, t2))
+  | uncombine_term (Abs (s, T, t')) = Abs (s, T, uncombine_term t')
+  | uncombine_term (t as Const (x as (s, _))) =
+    (case AList.lookup (op =) combinator_table s of
+       SOME thm => thm |> prop_of |> specialize_type @{theory} x |> Logic.dest_equals |> snd
+     | NONE => t)
+  | uncombine_term t = t
+
+(* Update schematic type variables with detected sort constraints. It's not
+   totally clear when this code is necessary. *)
+fun repair_tvar_sorts (t, tvar_tab) =
+  let
+    fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)
+      | do_type (TVar (xi, s)) =
+        TVar (xi, the_default s (Vartab.lookup tvar_tab xi))
+      | do_type (TFree z) = TFree z
+    fun do_term (Const (a, T)) = Const (a, do_type T)
+      | do_term (Free (a, T)) = Free (a, do_type T)
+      | do_term (Var (xi, T)) = Var (xi, do_type T)
+      | do_term (t as Bound _) = t
+      | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t)
+      | do_term (t1 $ t2) = do_term t1 $ do_term t2
+  in t |> not (Vartab.is_empty tvar_tab) ? do_term end
+
+fun quantify_over_free quant_s free_s body_t =
+  case Term.add_frees body_t [] |> filter (curry (op =) free_s o fst) of
+    [] => body_t
+  | frees as (_, free_T) :: _ =>
+    Abs (free_s, free_T, fold (curry abstract_over) (map Free frees) body_t)
+
+(* Interpret an ATP formula as a HOL term, extracting sort constraints as they
+   appear in the formula. *)
+fun prop_from_formula thy full_types tfrees phi =
+  let
+    fun do_formula pos phi =
+      case phi of
+        AQuant (_, [], phi) => do_formula pos phi
+      | AQuant (q, x :: xs, phi') =>
+        do_formula pos (AQuant (q, xs, phi'))
+        #>> quantify_over_free (case q of
+                                  AForall => @{const_name All}
+                                | AExists => @{const_name Ex})
+                               (repair_atp_variable_name Char.toLower x)
+      | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
+      | AConn (c, [phi1, phi2]) =>
+        do_formula (pos |> c = AImplies ? not) phi1
+        ##>> do_formula pos phi2
+        #>> (case c of
+               AAnd => s_conj
+             | AOr => s_disj
+             | AImplies => s_imp
+             | AIf => s_imp o swap
+             | AIff => s_iff
+             | ANotIff => s_not o s_iff)
+      | AAtom tm => term_from_pred thy full_types tfrees pos tm
+      | _ => raise FORMULA [phi]
+  in repair_tvar_sorts (do_formula true phi Vartab.empty) end
+
+fun check_formula ctxt =
+  Type_Infer.constrain HOLogic.boolT
+  #> Syntax.check_term (ProofContext.set_mode ProofContext.mode_schematic ctxt)
+
+
+(**** Translation of TSTP files to Isar Proofs ****)
+
+fun unvarify_term (Var ((s, 0), T)) = Free (s, T)
+  | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
+
+fun decode_line full_types tfrees (Definition (num, phi1, phi2)) ctxt =
+    let
+      val thy = ProofContext.theory_of ctxt
+      val t1 = prop_from_formula thy full_types tfrees phi1
+      val vars = snd (strip_comb t1)
+      val frees = map unvarify_term vars
+      val unvarify_args = subst_atomic (vars ~~ frees)
+      val t2 = prop_from_formula thy full_types tfrees phi2
+      val (t1, t2) =
+        HOLogic.eq_const HOLogic.typeT $ t1 $ t2
+        |> unvarify_args |> uncombine_term |> check_formula ctxt
+        |> HOLogic.dest_eq
+    in
+      (Definition (num, t1, t2),
+       fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt)
+    end
+  | decode_line full_types tfrees (Inference (num, u, deps)) ctxt =
+    let
+      val thy = ProofContext.theory_of ctxt
+      val t = u |> prop_from_formula thy full_types tfrees
+                |> uncombine_term |> check_formula ctxt
+    in
+      (Inference (num, t, deps),
+       fold Variable.declare_term (OldTerm.term_frees t) ctxt)
+    end
+fun decode_lines ctxt full_types tfrees lines =
+  fst (fold_map (decode_line full_types tfrees) lines ctxt)
+
+fun is_same_inference _ (Definition _) = false
+  | is_same_inference t (Inference (_, t', _)) = t aconv t'
+
+(* No "real" literals means only type information (tfree_tcs, clsrel, or
+   clsarity). *)
+val is_only_type_information = curry (op aconv) HOLogic.true_const
+
+fun replace_one_dep (old, new) dep = if dep = old then new else [dep]
+fun replace_deps_in_line _ (line as Definition _) = line
+  | replace_deps_in_line p (Inference (num, t, deps)) =
+    Inference (num, t, fold (union (op =) o replace_one_dep p) deps [])
+
+(* Discard axioms; consolidate adjacent lines that prove the same formula, since
+   they differ only in type information.*)
+fun add_line _ _ (line as Definition _) lines = line :: lines
+  | add_line conjecture_shape axiom_names (Inference (num, t, [])) lines =
+    (* No dependencies: axiom, conjecture, or (for Vampire) internal axioms or
+       definitions. *)
+    if is_axiom_number axiom_names num then
+      (* Axioms are not proof lines. *)
+      if is_only_type_information t then
+        map (replace_deps_in_line (num, [])) lines
+      (* Is there a repetition? If so, replace later line by earlier one. *)
+      else case take_prefix (not o is_same_inference t) lines of
+        (_, []) => lines (*no repetition of proof line*)
+      | (pre, Inference (num', _, _) :: post) =>
+        pre @ map (replace_deps_in_line (num', [num])) post
+    else if is_conjecture_number conjecture_shape num then
+      Inference (num, negate_term t, []) :: lines
+    else
+      map (replace_deps_in_line (num, [])) lines
+  | add_line _ _ (Inference (num, t, deps)) lines =
+    (* Type information will be deleted later; skip repetition test. *)
+    if is_only_type_information t then
+      Inference (num, t, deps) :: lines
+    (* Is there a repetition? If so, replace later line by earlier one. *)
+    else case take_prefix (not o is_same_inference t) lines of
+      (* FIXME: Doesn't this code risk conflating proofs involving different
+         types? *)
+       (_, []) => Inference (num, t, deps) :: lines
+     | (pre, Inference (num', t', _) :: post) =>
+       Inference (num, t', deps) ::
+       pre @ map (replace_deps_in_line (num', [num])) post
+
+(* Recursively delete empty lines (type information) from the proof. *)
+fun add_nontrivial_line (Inference (num, t, [])) lines =
+    if is_only_type_information t then delete_dep num lines
+    else Inference (num, t, []) :: lines
+  | add_nontrivial_line line lines = line :: lines
+and delete_dep num lines =
+  fold_rev add_nontrivial_line (map (replace_deps_in_line (num, [])) lines) []
+
+(* ATPs sometimes reuse free variable names in the strangest ways. Removing
+   offending lines often does the trick. *)
+fun is_bad_free frees (Free x) = not (member (op =) frees x)
+  | is_bad_free _ _ = false
+
+(* Vampire is keen on producing these. *)
+fun is_trivial_formula (@{const Not} $ (Const (@{const_name HOL.eq}, _)
+                                        $ t1 $ t2)) = (t1 aconv t2)
+  | is_trivial_formula _ = false
+
+fun add_desired_line _ _ _ _ (line as Definition (num, _, _)) (j, lines) =
+    (j, line :: map (replace_deps_in_line (num, [])) lines)
+  | add_desired_line isar_shrink_factor conjecture_shape axiom_names frees
+                     (Inference (num, t, deps)) (j, lines) =
+    (j + 1,
+     if is_axiom_number axiom_names num orelse
+        is_conjecture_number conjecture_shape num orelse
+        (not (is_only_type_information t) andalso
+         null (Term.add_tvars t []) andalso
+         not (exists_subterm (is_bad_free frees) t) andalso
+         not (is_trivial_formula t) andalso
+         (null lines orelse (* last line must be kept *)
+          (length deps >= 2 andalso j mod isar_shrink_factor = 0))) then
+       Inference (num, t, deps) :: lines  (* keep line *)
+     else
+       map (replace_deps_in_line (num, deps)) lines)  (* drop line *)
+
+(** EXTRACTING LEMMAS **)
+
+(* Like "split_line", but ignores "\n" that follow a comma (as in SNARK's
+   output). *)
+val split_proof_lines =
+  let
+    fun aux [] [] = []
+      | aux line [] = [implode (rev line)]
+      | aux line ("," :: "\n" :: rest) = aux ("," :: line) rest
+      | aux line ("\n" :: rest) = aux line [] @ aux [] rest
+      | aux line (s :: rest) = aux (s :: line) rest
+  in aux [] o explode end
+
+(* A list consisting of the first number in each line is returned. For TSTP,
+   interesting lines have the form "fof(108, axiom, ...)", where the number
+   (108) is extracted. For SPASS, lines have the form "108[0:Inp] ...", where
+   the first number (108) is extracted. For Vampire, we look for
+   "108. ... [input]". *)
+fun used_facts_in_atp_proof axiom_names atp_proof =
+  let
+    fun axiom_names_at_index num =
+      let val j = Int.fromString num |> the_default ~1 in
+        if is_axiom_number axiom_names j then Vector.sub (axiom_names, j - 1)
+        else []
+      end
+    val tokens_of =
+      String.tokens (fn c => not (Char.isAlphaNum c) andalso c <> #"_")
+    fun do_line (tag :: num :: "axiom" :: (rest as _ :: _)) =
+        if tag = "cnf" orelse tag = "fof" then
+          (case strip_prefix_and_unascii axiom_prefix (List.last rest) of
+             SOME name =>
+             if member (op =) rest "file" then
+               ([(name, name |> find_first_in_list_vector axiom_names |> the)]
+                handle Option.Option =>
+                       error ("No such fact: " ^ quote name ^ "."))
+             else
+               axiom_names_at_index num
+           | NONE => axiom_names_at_index num)
+        else
+          []
+      | do_line (num :: "0" :: "Inp" :: _) = axiom_names_at_index num
+      | do_line (num :: rest) =
+        (case List.last rest of "input" => axiom_names_at_index num | _ => [])
+      | do_line _ = []
+  in atp_proof |> split_proof_lines |> maps (do_line o tokens_of) end
+
+val indent_size = 2
+val no_label = ("", ~1)
+
+val raw_prefix = "X"
+val assum_prefix = "A"
+val fact_prefix = "F"
+
+fun string_for_label (s, num) = s ^ string_of_int num
+
+fun metis_using [] = ""
+  | metis_using ls =
+    "using " ^ space_implode " " (map string_for_label ls) ^ " "
+fun metis_apply _ 1 = "by "
+  | metis_apply 1 _ = "apply "
+  | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply "
+fun metis_name full_types = if full_types then "metisFT" else "metis"
+fun metis_call full_types [] = metis_name full_types
+  | metis_call full_types ss =
+    "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")"
+fun metis_command full_types i n (ls, ss) =
+  metis_using ls ^ metis_apply i n ^ metis_call full_types ss
+fun metis_line full_types i n ss =
+  "Try this command: " ^
+  Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "."
+fun minimize_line _ [] = ""
+  | minimize_line minimize_command ss =
+    case minimize_command ss of
+      "" => ""
+    | command =>
+      "\nTo minimize the number of lemmas, try this: " ^
+      Markup.markup Markup.sendback command ^ "."
+
+fun used_facts axiom_names =
+  used_facts_in_atp_proof axiom_names
+  #> List.partition (curry (op =) Chained o snd)
+  #> pairself (sort_distinct (string_ord o pairself fst))
+
+fun metis_proof_text (full_types, minimize_command, atp_proof, axiom_names,
+                      goal, i) =
+  let
+    val (chained_lemmas, other_lemmas) = used_facts axiom_names atp_proof
+    val n = Logic.count_prems (prop_of goal)
+  in
+    (metis_line full_types i n (map fst other_lemmas) ^
+     minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
+     other_lemmas @ chained_lemmas)
+  end
+
+(** Isar proof construction and manipulation **)
+
+fun merge_fact_sets (ls1, ss1) (ls2, ss2) =
+  (union (op =) ls1 ls2, union (op =) ss1 ss2)
+
+type label = string * int
+type facts = label list * string list
+
+datatype qualifier = Show | Then | Moreover | Ultimately
+
+datatype step =
+  Fix of (string * typ) list |
+  Let of term * term |
+  Assume of label * term |
+  Have of qualifier list * label * term * byline
+and byline =
+  ByMetis of facts |
+  CaseSplit of step list list * facts
+
+fun smart_case_split [] facts = ByMetis facts
+  | smart_case_split proofs facts = CaseSplit (proofs, facts)
+
+fun add_fact_from_dep axiom_names num =
+  if is_axiom_number axiom_names num then
+    apsnd (union (op =) (map fst (Vector.sub (axiom_names, num - 1))))
+  else
+    apfst (insert (op =) (raw_prefix, num))
+
+fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t
+fun forall_vars t = fold_rev forall_of (map Var (Term.add_vars t [])) t
+
+fun step_for_line _ _ (Definition (_, t1, t2)) = Let (t1, t2)
+  | step_for_line _ _ (Inference (num, t, [])) = Assume ((raw_prefix, num), t)
+  | step_for_line axiom_names j (Inference (num, t, deps)) =
+    Have (if j = 1 then [Show] else [], (raw_prefix, num),
+          forall_vars t,
+          ByMetis (fold (add_fact_from_dep axiom_names) deps ([], [])))
+
+fun proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
+                         atp_proof conjecture_shape axiom_names params frees =
+  let
+    val lines =
+      atp_proof ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
+      |> parse_proof pool
+      |> sort (int_ord o pairself raw_step_number)
+      |> decode_lines ctxt full_types tfrees
+      |> rpair [] |-> fold_rev (add_line conjecture_shape axiom_names)
+      |> rpair [] |-> fold_rev add_nontrivial_line
+      |> rpair (0, []) |-> fold_rev (add_desired_line isar_shrink_factor
+                                             conjecture_shape axiom_names frees)
+      |> snd
+  in
+    (if null params then [] else [Fix params]) @
+    map2 (step_for_line axiom_names) (length lines downto 1) lines
+  end
+
+(* When redirecting proofs, we keep information about the labels seen so far in
+   the "backpatches" data structure. The first component indicates which facts
+   should be associated with forthcoming proof steps. The second component is a
+   pair ("assum_ls", "drop_ls"), where "assum_ls" are the labels that should
+   become assumptions and "drop_ls" are the labels that should be dropped in a
+   case split. *)
+type backpatches = (label * facts) list * (label list * label list)
+
+fun used_labels_of_step (Have (_, _, _, by)) =
+    (case by of
+       ByMetis (ls, _) => ls
+     | CaseSplit (proofs, (ls, _)) =>
+       fold (union (op =) o used_labels_of) proofs ls)
+  | used_labels_of_step _ = []
+and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof []
+
+fun new_labels_of_step (Fix _) = []
+  | new_labels_of_step (Let _) = []
+  | new_labels_of_step (Assume (l, _)) = [l]
+  | new_labels_of_step (Have (_, l, _, _)) = [l]
+val new_labels_of = maps new_labels_of_step
+
+val join_proofs =
+  let
+    fun aux _ [] = NONE
+      | aux proof_tail (proofs as (proof1 :: _)) =
+        if exists null proofs then
+          NONE
+        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, _) => (* 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))
+                         (used_labels_of proof_tail)) then
+            SOME (l, t, map rev proofs, proof_tail)
+          else
+            NONE
+        | _ => NONE
+  in aux [] o map rev end
+
+fun case_split_qualifiers proofs =
+  case length proofs of
+    0 => []
+  | 1 => [Then]
+  | _ => [Ultimately]
+
+fun redirect_proof conjecture_shape hyp_ts concl_t proof =
+  let
+    (* The first pass outputs those steps that are independent of the negated
+       conjecture. The second pass flips the proof by contradiction to obtain a
+       direct proof, introducing case splits when an inference depends on
+       several facts that depend on the negated conjecture. *)
+    fun find_hyp num =
+      nth hyp_ts (index_in_shape num conjecture_shape)
+      handle Subscript =>
+             raise Fail ("Cannot find hypothesis " ^ Int.toString num)
+     val concl_ls = map (pair raw_prefix) (List.last conjecture_shape)
+     val canonicalize_labels =
+       map (fn l => if member (op =) concl_ls l then hd concl_ls else l)
+       #> distinct (op =)
+     fun first_pass ([], contra) = ([], contra)
+       | first_pass ((step as Fix _) :: proof, contra) =
+         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), _)) :: proof, contra) =
+         if member (op =) concl_ls l then
+           first_pass (proof, contra ||> l = hd concl_ls ? cons step)
+         else
+           first_pass (proof, contra) |>> cons (Assume (l, find_hyp num))
+       | first_pass (Have (qs, l, t, ByMetis (ls, ss)) :: proof, contra) =
+         let
+           val ls = canonicalize_labels ls
+           val step = Have (qs, l, t, ByMetis (ls, ss))
+         in
+           if exists (member (op =) (fst contra)) ls then
+             first_pass (proof, contra |>> cons l ||> cons step)
+           else
+             first_pass (proof, contra) |>> cons step
+         end
+       | first_pass _ = raise Fail "malformed proof"
+    val (proof_top, (contra_ls, contra_proof)) =
+      first_pass (proof, (concl_ls, []))
+    val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst
+    fun backpatch_labels patches ls =
+      fold merge_fact_sets (map (backpatch_label patches) ls) ([], [])
+    fun second_pass end_qs ([], assums, patches) =
+        ([Have (end_qs, no_label, concl_t,
+                ByMetis (backpatch_labels patches (map snd assums)))], patches)
+      | second_pass end_qs (Assume (l, t) :: proof, assums, patches) =
+        second_pass end_qs (proof, (t, l) :: assums, patches)
+      | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums,
+                            patches) =
+        if member (op =) (snd (snd patches)) l andalso
+           not (member (op =) (fst (snd patches)) l) andalso
+           not (AList.defined (op =) (fst patches) l) then
+          second_pass end_qs (proof, assums, patches ||> apsnd (append ls))
+        else
+          (case List.partition (member (op =) contra_ls) ls of
+             ([contra_l], co_ls) =>
+             if member (op =) qs Show then
+               second_pass end_qs (proof, assums,
+                                   patches |>> cons (contra_l, (co_ls, ss)))
+             else
+               second_pass end_qs
+                           (proof, assums,
+                            patches |>> cons (contra_l, (l :: co_ls, ss)))
+               |>> cons (if member (op =) (fst (snd patches)) l then
+                           Assume (l, negate_term t)
+                         else
+                           Have (qs, l, negate_term t,
+                                 ByMetis (backpatch_label patches l)))
+           | (contra_ls as _ :: _, co_ls) =>
+             let
+               val proofs =
+                 map_filter
+                     (fn l =>
+                         if member (op =) concl_ls l then
+                           NONE
+                         else
+                           let
+                             val drop_ls = filter (curry (op <>) l) contra_ls
+                           in
+                             second_pass []
+                                 (proof, assums,
+                                  patches ||> apfst (insert (op =) l)
+                                          ||> apsnd (union (op =) drop_ls))
+                             |> fst |> SOME
+                           end) contra_ls
+               val (assumes, facts) =
+                 if member (op =) (fst (snd patches)) l then
+                   ([Assume (l, negate_term t)], (l :: co_ls, ss))
+                 else
+                   ([], (co_ls, ss))
+             in
+               (case join_proofs proofs of
+                  SOME (l, t, proofs, proof_tail) =>
+                  Have (case_split_qualifiers proofs @
+                        (if null proof_tail then end_qs else []), l, t,
+                        smart_case_split proofs facts) :: proof_tail
+                | NONE =>
+                  [Have (case_split_qualifiers proofs @ end_qs, no_label,
+                         concl_t, smart_case_split proofs facts)],
+                patches)
+               |>> append assumes
+             end
+           | _ => raise Fail "malformed proof")
+       | second_pass _ _ = raise Fail "malformed proof"
+    val proof_bottom =
+      second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst
+  in proof_top @ proof_bottom end
+
+(* FIXME: Still needed? Probably not. *)
+val kill_duplicate_assumptions_in_proof =
+  let
+    fun relabel_facts subst =
+      apfst (map (fn l => AList.lookup (op =) subst l |> the_default l))
+    fun do_step (step as Assume (l, t)) (proof, subst, assums) =
+        (case AList.lookup (op aconv) assums t of
+           SOME l' => (proof, (l, l') :: subst, assums)
+         | NONE => (step :: proof, subst, (t, l) :: assums))
+      | do_step (Have (qs, l, t, by)) (proof, subst, assums) =
+        (Have (qs, l, t,
+               case by of
+                 ByMetis facts => ByMetis (relabel_facts subst facts)
+               | CaseSplit (proofs, facts) =>
+                 CaseSplit (map do_proof proofs, relabel_facts subst facts)) ::
+         proof, subst, assums)
+      | do_step step (proof, subst, assums) = (step :: proof, subst, assums)
+    and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev
+  in do_proof end
+
+val then_chain_proof =
+  let
+    fun aux _ [] = []
+      | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof
+      | aux l' (Have (qs, l, t, by) :: proof) =
+        (case by of
+           ByMetis (ls, ss) =>
+           Have (if member (op =) ls l' then
+                   (Then :: qs, l, t,
+                    ByMetis (filter_out (curry (op =) l') ls, ss))
+                 else
+                   (qs, l, t, ByMetis (ls, ss)))
+         | CaseSplit (proofs, facts) =>
+           Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) ::
+        aux l proof
+      | aux _ (step :: proof) = step :: aux no_label proof
+  in aux no_label end
+
+fun kill_useless_labels_in_proof proof =
+  let
+    val used_ls = used_labels_of proof
+    fun do_label l = if member (op =) used_ls l then l else no_label
+    fun do_step (Assume (l, t)) = Assume (do_label l, t)
+      | do_step (Have (qs, l, t, by)) =
+        Have (qs, do_label l, t,
+              case by of
+                CaseSplit (proofs, facts) =>
+                CaseSplit (map (map do_step) proofs, facts)
+              | _ => by)
+      | do_step step = step
+  in map do_step proof end
+
+fun prefix_for_depth n = replicate_string (n + 1)
+
+val relabel_proof =
+  let
+    fun aux _ _ _ [] = []
+      | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) =
+        if l = no_label then
+          Assume (l, t) :: aux subst depth (next_assum, next_fact) proof
+        else
+          let val l' = (prefix_for_depth depth assum_prefix, next_assum) in
+            Assume (l', t) ::
+            aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof
+          end
+      | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) =
+        let
+          val (l', subst, next_fact) =
+            if l = no_label then
+              (l, subst, next_fact)
+            else
+              let
+                val l' = (prefix_for_depth depth fact_prefix, next_fact)
+              in (l', (l, l') :: subst, next_fact + 1) end
+          val relabel_facts =
+            apfst (map (fn l =>
+                           case AList.lookup (op =) subst l of
+                             SOME l' => l'
+                           | NONE => raise Fail ("unknown label " ^
+                                                 quote (string_for_label l))))
+          val by =
+            case by of
+              ByMetis facts => ByMetis (relabel_facts facts)
+            | CaseSplit (proofs, facts) =>
+              CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs,
+                         relabel_facts facts)
+        in
+          Have (qs, l', t, by) ::
+          aux subst depth (next_assum, next_fact) proof
+        end
+      | aux subst depth nextp (step :: proof) =
+        step :: aux subst depth nextp proof
+  in aux [] 0 (1, 1) end
+
+fun string_for_proof ctxt full_types i n =
+  let
+    fun fix_print_mode f x =
+      setmp_CRITICAL show_no_free_types true
+          (setmp_CRITICAL show_types true
+               (Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
+                                         (print_mode_value ())) f)) x
+    fun do_indent ind = replicate_string (ind * indent_size) " "
+    fun do_free (s, T) =
+      maybe_quote s ^ " :: " ^
+      maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T)
+    fun do_label l = if l = no_label then "" else string_for_label l ^ ": "
+    fun do_have qs =
+      (if member (op =) qs Moreover then "moreover " else "") ^
+      (if member (op =) qs Ultimately then "ultimately " else "") ^
+      (if member (op =) qs Then then
+         if member (op =) qs Show then "thus" else "hence"
+       else
+         if member (op =) qs Show then "show" else "have")
+    val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt)
+    fun do_facts (ls, ss) =
+      metis_command full_types 1 1
+                    (ls |> sort_distinct (prod_ord string_ord int_ord),
+                     ss |> sort_distinct string_ord)
+    and do_step ind (Fix xs) =
+        do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n"
+      | do_step ind (Let (t1, t2)) =
+        do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n"
+      | do_step ind (Assume (l, t)) =
+        do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n"
+      | do_step ind (Have (qs, l, t, ByMetis facts)) =
+        do_indent ind ^ do_have qs ^ " " ^
+        do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n"
+      | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) =
+        space_implode (do_indent ind ^ "moreover\n")
+                      (map (do_block ind) proofs) ^
+        do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^
+        do_facts facts ^ "\n"
+    and do_steps prefix suffix ind steps =
+      let val s = implode (map (do_step ind) steps) in
+        replicate_string (ind * indent_size - size prefix) " " ^ prefix ^
+        String.extract (s, ind * indent_size,
+                        SOME (size s - ind * indent_size - 1)) ^
+        suffix ^ "\n"
+      end
+    and do_block ind proof = do_steps "{ " " }" (ind + 1) proof
+    (* One-step proofs are pointless; better use the Metis one-liner
+       directly. *)
+    and do_proof [Have (_, _, _, ByMetis _)] = ""
+      | do_proof proof =
+        (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
+        do_indent 0 ^ "proof -\n" ^
+        do_steps "" "" 1 proof ^
+        do_indent 0 ^ (if n <> 1 then "next" else "qed")
+  in do_proof end
+
+fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
+                    (other_params as (full_types, _, atp_proof, axiom_names,
+                                      goal, i)) =
+  let
+    val (params, hyp_ts, concl_t) = strip_subgoal goal i
+    val frees = fold Term.add_frees (concl_t :: hyp_ts) []
+    val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) []
+    val n = Logic.count_prems (prop_of goal)
+    val (one_line_proof, lemma_names) = metis_proof_text other_params
+    fun isar_proof_for () =
+      case proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
+                                atp_proof conjecture_shape axiom_names params
+                                frees
+           |> redirect_proof conjecture_shape hyp_ts concl_t
+           |> kill_duplicate_assumptions_in_proof
+           |> then_chain_proof
+           |> kill_useless_labels_in_proof
+           |> relabel_proof
+           |> string_for_proof ctxt full_types i n of
+        "" => "\nNo structured proof available."
+      | proof => "\n\nStructured proof:\n" ^ Markup.markup Markup.sendback proof
+    val isar_proof =
+      if debug then
+        isar_proof_for ()
+      else
+        try isar_proof_for ()
+        |> the_default "\nWarning: The Isar proof construction failed."
+  in (one_line_proof ^ isar_proof, lemma_names) end
+
+fun proof_text isar_proof isar_params other_params =
+  (if isar_proof then isar_proof_text isar_params else metis_proof_text)
+      other_params
+
+end;