reconstruction framework for LEO-II's TPTP proofs;
authorsultana
Wed, 19 Feb 2014 15:57:02 +0000
changeset 55596 928b9f677165
parent 55595 2e2e9bc7c4c6
child 55597 25d7b485df81
reconstruction framework for LEO-II's TPTP proofs;
src/HOL/ROOT
src/HOL/TPTP/TPTP_Parser/tptp_reconstruct.ML
src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML
src/HOL/TPTP/TPTP_Proof_Reconstruction.thy
src/HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy
src/HOL/TPTP/TPTP_Proof_Reconstruction_Test_Units.thy
--- a/src/HOL/ROOT	Wed Feb 19 15:57:02 2014 +0000
+++ b/src/HOL/ROOT	Wed Feb 19 15:57:02 2014 +0000
@@ -668,6 +668,7 @@
     MaSh_Export
     TPTP_Interpret
     THF_Arith
+    TPTP_Proof_Reconstruction
   theories
     ATP_Problem_Import
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct.ML	Wed Feb 19 15:57:02 2014 +0000
@@ -0,0 +1,1925 @@
+(*  Title:      HOL/TPTP/TPTP_Parser/tptp_reconstruct.ML
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+Reconstructs TPTP proofs in Isabelle/HOL.
+Specialised to work with proofs produced by LEO-II.
+
+TODO
+ - Proof transformation to remove "copy" steps, and perhaps other dud inferences.
+*)
+
+signature TPTP_RECONSTRUCT =
+sig
+  (* Interface used by TPTP_Reconstruct.thy, to define LEO-II proof reconstruction. *)
+
+  datatype formula_kind =
+      Conjunctive of bool option
+    | Disjunctive of bool option
+    | Biimplicational of bool option
+    | Negative of bool option
+    | Existential of bool option * typ
+    | Universal of bool option * typ
+    | Equational of bool option * typ
+    | Atomic of bool option
+    | Implicational of bool option
+  type formula_meaning =
+    (string *
+     {role : TPTP_Syntax.role,
+      fmla : term,
+      source_inf_opt : TPTP_Proof.source_info option})
+  type proof_annotation =
+    {problem_name : TPTP_Problem_Name.problem_name,
+     skolem_defs : ((*skolem const name*)string * Binding.binding) list,
+     defs : ((*node name*)string * Binding.binding) list,
+     axs : ((*node name*)string * Binding.binding) list,
+     (*info for each node (for all lines in the TPTP proof)*)
+     meta : formula_meaning list}
+  type rule_info =
+    {inference_name : string, (*name of calculus rule*)
+     inference_fmla : term, (*the inference as a term*)
+     parents : string list}
+
+  exception UNPOLARISED of term
+
+  val remove_polarity : bool -> term -> term * bool
+  val interpret_bindings :
+     TPTP_Problem_Name.problem_name -> theory -> TPTP_Proof.parent_detail list -> (string * term) list -> (string * term) list
+  val diff_and_instantiate : Proof.context -> thm -> term -> term -> thm (*FIXME from library*)
+  val strip_top_all_vars : (string * typ) list -> term -> (string * typ) list * term
+  val strip_top_All_vars : term -> (string * typ) list * term
+  val strip_top_All_var : term -> (string * typ) * term
+  val new_consts_between : term -> term -> term list
+  val get_pannot_of_prob : theory -> TPTP_Problem_Name.problem_name -> proof_annotation
+  val inference_at_node : 'a -> TPTP_Problem_Name.problem_name -> formula_meaning list -> string -> rule_info option
+  val node_info : (string * 'a) list -> ('a -> 'b) -> string -> 'b
+
+  type step_id = string
+  datatype rolling_stock =
+      Step of step_id
+    | Assumed
+    | Unconjoin
+    | Split of step_id (*where split occurs*) *
+               step_id (*where split ends*) *
+               step_id list (*children of the split*)
+    | Synth_step of step_id (*A step which doesn't necessarily appear in
+      the original proof, or which has been modified slightly for better
+      handling by Isabelle*)
+    | Annotated_step of step_id * string (*Same interpretation as
+      "Step", except that additional information is attached. This is
+      currently used for debugging: Steps are mapped to Annotated_steps
+      and their rule names are included as strings*)
+    | Definition of step_id (*Mirrors TPTP role*)
+    | Axiom of step_id (*Mirrors TPTP role*)
+    | Caboose
+
+
+  (* Interface for using the proof reconstruction. *)
+
+  val import_thm : bool -> Path.T list -> Path.T -> (proof_annotation -> theory -> proof_annotation * theory) -> theory -> theory
+  val get_fmlas_of_prob : theory -> TPTP_Problem_Name.problem_name -> TPTP_Interpret.tptp_formula_meaning list
+  val structure_fmla_meaning : 'a * 'b * 'c * 'd -> 'a * {fmla: 'c, role: 'b, source_inf_opt: 'd}
+  val make_skeleton : Proof.context -> proof_annotation -> rolling_stock list
+  val naive_reconstruct_tacs :
+     (Proof.context -> TPTP_Problem_Name.problem_name -> step_id -> thm) ->
+     TPTP_Problem_Name.problem_name -> Proof.context -> (rolling_stock * term option * (thm * tactic) option) list
+  val naive_reconstruct_tac :
+     Proof.context -> (Proof.context -> TPTP_Problem_Name.problem_name -> step_id -> thm) -> TPTP_Problem_Name.problem_name -> tactic
+  val reconstruct : Proof.context -> (TPTP_Problem_Name.problem_name -> tactic) -> TPTP_Problem_Name.problem_name -> thm
+end
+
+structure TPTP_Reconstruct : TPTP_RECONSTRUCT =
+struct
+
+open TPTP_Reconstruct_Library
+open TPTP_Syntax
+
+(*FIXME move to more general struct*)
+(*Extract the formulas of an imported TPTP problem -- these formulas
+  may make up a proof*)
+fun get_fmlas_of_prob thy prob_name : TPTP_Interpret.tptp_formula_meaning list =
+  AList.lookup (op =) (TPTP_Interpret.get_manifests thy) prob_name
+  |> the |> #3 (*get formulas*);
+
+
+(** General **)
+
+(* Proof annotations *)
+
+(*FIXME modify TPTP_Interpret.tptp_formula_meaning into this type*)
+type formula_meaning =
+  (string *
+   {role : TPTP_Syntax.role,
+    fmla : term,
+    source_inf_opt : TPTP_Proof.source_info option})
+
+fun apply_to_parent_info f
+   (n, {role, fmla, source_inf_opt}) =
+  let
+    val source_inf_opt' =
+      case source_inf_opt of
+          NONE => NONE
+        | SOME (TPTP_Proof.Inference (inf_name, sinfos, pinfos)) =>
+            SOME (TPTP_Proof.Inference (inf_name, sinfos, f pinfos))
+  in
+   (n, {role = role, fmla = fmla, source_inf_opt = source_inf_opt'})
+  end
+
+fun structure_fmla_meaning (s, r, t, info) =
+  (s, {role = r, fmla = t, source_inf_opt = info})
+
+type proof_annotation =
+  {problem_name : TPTP_Problem_Name.problem_name,
+   skolem_defs : ((*skolem const name*)string * Binding.binding) list,
+   defs : ((*node name*)string * Binding.binding) list,
+   axs : ((*node name*)string * Binding.binding) list,
+   (*info for each node (for all lines in the TPTP proof)*)
+   meta : formula_meaning list}
+
+fun empty_pannot prob_name =
+  {problem_name = prob_name,
+   skolem_defs = [],
+   defs = [],
+   axs = [],
+   meta = []}
+
+
+(* Storage of proof data *)
+
+exception MANIFEST of TPTP_Problem_Name.problem_name * string (*FIXME move to TPTP_Interpret?*)
+
+type manifest = TPTP_Problem_Name.problem_name * proof_annotation
+
+(*manifest equality simply depends on problem name*)
+fun manifest_eq ((prob_name1, _), (prob_name2, _)) = prob_name1 = prob_name2
+
+structure TPTP_Reconstruction_Data = Theory_Data
+(
+  type T = manifest list
+  val empty = []
+  val extend = I
+  fun merge data : T = Library.merge manifest_eq data
+)
+val get_manifests : theory -> manifest list = TPTP_Reconstruction_Data.get
+
+fun update_manifest prob_name pannot thy =
+  let
+    val idx =
+      find_index
+        (fn (n, _) => n = prob_name)
+        (get_manifests thy)
+    val transf = (fn _ =>
+      (prob_name, pannot))
+  in
+    TPTP_Reconstruction_Data.map
+      (nth_map idx transf)
+      thy
+  end
+
+(*similar to get_fmlas_of_prob but for proofs*)
+fun get_pannot_of_prob thy prob_name : proof_annotation =
+  case AList.lookup (op =) (get_manifests thy) prob_name of
+      SOME pa => pa
+    | NONE => raise (MANIFEST (prob_name, "Could not find proof annotation"))
+
+
+(* Constants *)
+
+(*Prefix used for naming inferences which were added during proof
+transformation. (e.g., this is used to name "bind"-inference nodes
+described below)*)
+val inode_prefixK = "inode"
+
+(*New inference rule name, which is added to indicate that some
+variable has been instantiated. Additional proof metadata will
+indicate which variable, and how it was instantiated*)
+val bindK = "bind"
+
+(*New inference rule name, which is added to indicate that some
+(validity-preserving) preprocessing has been done to a (singleton)
+clause prior to it being split.*)
+val split_preprocessingK = "split_preprocessing"
+
+
+(* Storage of internal values *)
+
+type tptp_reconstruction_state = {next_int : int}
+structure TPTP_Reconstruction_Internal_Data = Theory_Data
+(
+  type T = tptp_reconstruction_state
+  val empty = {next_int = 0}
+  val extend = I
+  fun merge data : T = snd data
+)
+
+(*increment internal counter, and obtain the current next value*)
+fun get_next_int thy : int * theory =
+  let
+    val state = TPTP_Reconstruction_Internal_Data.get thy
+    val state' = {next_int = 1 + #next_int state}
+  in
+    (#next_int state,
+     TPTP_Reconstruction_Internal_Data.put state' thy)
+  end
+
+(*FIXME in some applications (e.g. where the name is used for an
+   inference node) need to check that the name is fresh, to avoid
+   collisions with other bits of the proof*)
+val get_next_name =
+  get_next_int
+  #> apfst (fn i => inode_prefixK ^ Int.toString i)
+
+
+(* Building the index *)
+
+(*thrown when we're expecting a TPTP_Proof.Bind annotation but find something else*)
+exception NON_BINDING
+(*given a list of pairs consisting of a variable name and
+  TPTP formula, returns the list consisting of the original
+  variable name and the interpreted HOL formula. Needs the
+  problem name to ensure use of correct interpretations for
+  constants and types.*)
+fun interpret_bindings (prob_name : TPTP_Problem_Name.problem_name) thy bindings acc =
+  if null bindings then acc
+  else
+    case hd bindings of
+        TPTP_Proof.Bind (v, fmla) =>
+          let
+            val (type_map, const_map) =
+                case AList.lookup (op =) (TPTP_Interpret.get_manifests thy) prob_name of
+                    NONE => raise (MANIFEST (prob_name, "Problem details not found in interpretation manifest"))
+                  | SOME (type_map, const_map, _) => (type_map, const_map)
+
+            (*FIXME get config from the envir or make it parameter*)
+            val config =
+              {cautious = true,
+               problem_name = SOME prob_name}
+            val result =
+              (v,
+               TPTP_Interpret.interpret_formula
+                config TPTP_Syntax.THF
+                const_map [] type_map fmla thy
+               |> fst)
+          in
+            interpret_bindings prob_name thy (tl bindings) (result :: acc)
+          end
+      | _ => raise NON_BINDING
+
+type rule_info =
+  {inference_name : string, (*name of calculus rule*)
+   inference_fmla : term, (*the inference as a term*)
+   parents : string list}
+
+(*Instantiates a binding in orig_parent_fmla. Used in a proof
+  transformation to factor out instantiations from inferences.*)
+fun apply_binding thy prob_name orig_parent_fmla target_fmla bindings =
+  let
+    val bindings' = interpret_bindings prob_name thy bindings []
+
+    (*capture selected free variables. these variables, and their
+      intended de Bruijn index, are included in "var_ctxt"*)
+    fun bind_free_vars var_ctxt t =
+      case t of
+          Const _ => t
+        | Var _ => t
+        | Bound _ => t
+        | Abs (x, ty, t') => Abs (x, ty, bind_free_vars (x :: var_ctxt) t')
+        | Free (x, ty) =>
+            let
+              val idx = find_index (fn y => y = x) var_ctxt
+            in
+              if idx > ~1 andalso
+                 ty = dummyT (*this check not really needed*) then
+                  Bound idx
+              else t
+            end
+        | t1 $ t2 => bind_free_vars var_ctxt t1 $ bind_free_vars var_ctxt t2
+
+    (*Instantiate specific quantified variables:
+      Look for subterms of form (! (% x. M)) where "x" appears as a "bound_var",
+      then replace "x" for "body" in "M".
+      Should only be applied at formula top level -- i.e., once past the quantifier
+      prefix we needn't bother with looking for bound_vars.
+      "var"_ctxt is used to keep track of lambda-bindings we encounter, to capture
+      free variables in "body" correctly (i.e., replace Free with Bound having the
+      right index)*)
+    fun instantiate_bound (binding as (bound_var, body)) (initial as (var_ctxt, t))  =
+      case t of
+          Const _ => initial
+        | Free _ => initial
+        | Var _ => initial
+        | Bound _ => initial
+        | Abs _ => initial
+        | t1 $ (t2 as Abs (x, ty, t')) =>
+            if is_Const t1 then
+              (*Could be fooled by shadowing, but if order matters
+                then should still be able to handle formulas like
+                (! X, X. F).*)
+              if x = bound_var andalso
+                 fst (dest_Const t1) = @{const_name All} then
+                  (*Body might contain free variables, so bind them using "var_ctxt".
+                    this involves replacing instances of Free with instances of Bound
+                    at the right index.*)
+                  let val body' = bind_free_vars var_ctxt body
+                  in
+                    (var_ctxt,
+                     betapply (t2, body'))
+                  end
+              else
+                  let
+                    val (var_ctxt', rest) = instantiate_bound binding (x :: var_ctxt, t')
+                  in
+                    (var_ctxt',
+                     t1 $ Abs (x, ty, rest))
+                  end
+            else initial
+        | t1 $ t2 =>
+            let
+              val (var_ctxt', rest) = instantiate_bound binding (var_ctxt, t2)
+            in
+              (var_ctxt', t1 $ rest)
+            end
+
+    (*Here we preempt the following problem:
+     if have (! X1, X2, X3. body), and X1 is instantiated to
+     "c X2 X3", then the current code will yield
+     (! X2, X3, X2a, X3a. body').
+     To avoid this, we must first push X1 in, before calling
+     instantiate_bound, to make sure that bound variables don't
+     get free.*)
+    fun safe_instantiate_bound (binding as (bound_var, body)) (var_ctxt, t) =
+       instantiate_bound binding
+         (var_ctxt, push_allvar_in bound_var t)
+
+    (*return true if one of the types is polymorphic*)
+    fun is_polymorphic tys =
+      if null tys then false
+      else
+        case hd tys of
+            Type (_, tys') => is_polymorphic (tl tys @ tys')
+          | TFree _ => true
+          | TVar _ => true
+
+    (*find the type of a quantified variable, at the "topmost" binding
+      occurrence*)
+    local
+      fun type_of_quantified_var' s ts =
+        if null ts then NONE
+        else
+          case hd ts of
+              Const _ => type_of_quantified_var' s (tl ts)
+            | Free _ => type_of_quantified_var' s (tl ts)
+            | Var _ => type_of_quantified_var' s (tl ts)
+            | Bound _ => type_of_quantified_var' s (tl ts)
+            | Abs (s', ty, t') =>
+                if s = s' then SOME ty
+                else type_of_quantified_var' s (t' :: tl ts)
+            | t1 $ t2 => type_of_quantified_var' s (t1 :: t2 :: tl ts)
+    in
+      fun type_of_quantified_var s =
+        single #> type_of_quantified_var' s
+    end
+
+    (*Form the universal closure of "t".
+      NOTE remark above "val frees" about ordering of quantified variables*)
+    fun close_formula t =
+      let
+          (*The ordering of Frees in this list affects the order in which variables appear
+            in the quantification prefix. Currently this is assumed not to matter.
+            This consists of a list of pairs: the first element consists of the "original"
+            free variable, and the latter consists of the monomorphised equivalent. The
+            two elements are identical if the original is already monomorphic.
+            This monomorphisation is needed since, owing to TPTP's lack of type annotations,
+            variables might not be constrained by type info. This results in them being
+            interpreted as polymorphic. E.g., this issue comes up in CSR148^1*)
+          val frees_monomorphised =
+            fold_aterms
+              (fn t => fn rest =>
+                 if is_Free t then
+                   let
+                     val (s, ty) = dest_Free t
+                     val ty' =
+                       if ty = dummyT orelse is_polymorphic [ty] then
+                         the (type_of_quantified_var s target_fmla)
+                       else ty
+                   in insert (op =) (t, Free (s, ty')) rest
+                   end
+                 else rest)
+              t []
+      in
+        Term.subst_free frees_monomorphised t
+        |> fold (fn (s, ty) => fn t =>
+                    HOLogic.mk_all (s, ty, t))
+              (map (snd #> dest_Free) frees_monomorphised)
+      end
+
+    (*FIXME currently assuming that we're only ever given a single binding each time this is called*)
+    val _ = @{assert} (length bindings' = 1)
+
+  in
+    fold safe_instantiate_bound bindings' ([], HOLogic.dest_Trueprop orig_parent_fmla)
+    |> snd (*discard var typing context*)
+    |> close_formula
+    |> single
+    |> Type_Infer_Context.infer_types (Context.proof_of (Context.Theory thy))
+    |> the_single
+    |> HOLogic.mk_Trueprop
+    |> rpair bindings'
+  end
+
+exception RECONSTRUCT of string
+
+(*Some of these may be redundant wrt the original aims of this
+  datatype, but it's useful to have a datatype to classify formulas
+  for use by other functions as well.*)
+datatype formula_kind =
+    Conjunctive of bool option
+  | Disjunctive of bool option
+  | Biimplicational of bool option
+  | Negative of bool option
+  | Existential of bool option * typ
+  | Universal of bool option * typ
+  | Equational of bool option * typ
+  | Atomic of bool option
+  | Implicational of bool option
+
+exception UNPOLARISED of term
+(*Remove "= $true" or "= $false$ from the edge
+  of a formula. Use "try" in case formula is not
+  polarised.*)
+fun remove_polarity strict formula =
+  case try HOLogic.dest_eq formula of
+      NONE => if strict then raise (UNPOLARISED formula)
+              else (formula, true)
+    | SOME (x, p as @{term True}) => (x, true)
+    | SOME (x, p as @{term False}) => (x, false)
+    | SOME (x, _) =>
+        if strict then raise (UNPOLARISED formula)
+        else (formula, true)
+
+(*flattens a formula wrt associative operators*)
+fun flatten formula_kind formula =
+  let
+    fun is_conj (Const (@{const_name HOL.conj}, _) $ _ $ _) = true
+      | is_conj _ = false
+    fun is_disj (Const (@{const_name HOL.disj}, _) $ _ $ _) = true
+      | is_disj _ = false
+    fun is_iff (Const (@{const_name HOL.eq}, ty) $ _ $ _) =
+          ty = ([HOLogic.boolT, HOLogic.boolT] ---> HOLogic.boolT)
+      | is_iff _ = false
+
+    fun flatten' formula acc =
+      case formula of
+          Const (@{const_name HOL.conj}, _) $ t1 $ t2 =>
+            (case formula_kind of
+                 Conjunctive _ =>
+                   let
+                     val left =
+                       if is_conj t1 then flatten' t1 acc else (t1 :: acc)
+                   in
+                       if is_conj t2 then flatten' t2 left else (t2 :: left)
+                   end
+               | _ => formula :: acc)
+        | Const (@{const_name HOL.disj}, _) $ t1 $ t2 =>
+            (case formula_kind of
+                 Disjunctive _ =>
+                   let
+                     val left =
+                       if is_disj t1 then flatten' t1 acc else (t1 :: acc)
+                   in
+                       if is_disj t2 then flatten' t2 left else (t2 :: left)
+                   end
+               | _ => formula :: acc)
+        | Const (@{const_name HOL.eq}, ty) $ t1 $ t2 =>
+            if ty = ([HOLogic.boolT, HOLogic.boolT] ---> HOLogic.boolT) then
+              case formula_kind of
+                   Biimplicational _ =>
+                     let
+                       val left =
+                         if is_iff t1 then flatten' t1 acc else (t1 :: acc)
+                     in
+                         if is_iff t2 then flatten' t2 left else (t2 :: left)
+                     end
+                 | _ => formula :: acc
+            else formula :: acc
+        | _ => [formula]
+
+    val formula' = try_dest_Trueprop formula
+  in
+    case formula_kind of
+        Conjunctive (SOME _) =>
+          remove_polarity false formula'
+          |> fst
+          |> (fn t => flatten' t [])
+      | Disjunctive (SOME _) =>
+          remove_polarity false formula'
+          |> fst
+          |> (fn t => flatten' t [])
+      | Biimplicational (SOME _) =>
+          remove_polarity false formula'
+          |> fst
+          |> (fn t => flatten' t [])
+      | _ => flatten' formula' []
+  end
+
+fun node_info fms projector node_name =
+  case AList.lookup (op =) fms node_name of
+      NONE =>
+        raise (RECONSTRUCT ("node " ^ node_name ^
+                            " doesn't exist"))
+    | SOME info => projector info
+
+(*Given a list of parent infos, extract the parent node names
+  and the additional info (e.g., if there was an instantiation
+  in addition to the inference).
+  if "filtered"=true then exclude axiom and definition parents*)
+fun dest_parent_infos filtered fms parent_infos : {name : string, details : TPTP_Proof.parent_detail list} list =
+  let
+    (*Removes "definition" dependencies since these play no
+      logical role -- i.e. they just give the expansions of
+      constants.
+      Removes "axiom" dependencies since these do not need to
+      be derived; the reconstruction handler in "leo2_tac" can
+      pick up the relevant axioms (using the info in the proof
+      annotation) and use them in its reconstruction.
+    *)
+    val filter_deps =
+      List.filter (fn {name, ...} =>
+        let
+          val role = node_info fms #role name
+        in role <> TPTP_Syntax.Role_Definition andalso
+            role <> TPTP_Syntax.Role_Axiom
+        end)
+    val parent_nodelist =
+      parent_infos
+      |> map (fn n =>
+                 case n of
+                     TPTP_Proof.Parent parent => {name = parent, details = []}
+                   | TPTP_Proof.ParentWithDetails (parent, details) =>
+                     {name = parent, details = details})
+  in
+    parent_nodelist
+    |> filtered ? filter_deps
+  end
+
+fun parents_of_node fms n =
+  case node_info fms #source_inf_opt n of
+      NONE => []
+    | SOME (TPTP_Proof.File _) => []
+    | SOME (TPTP_Proof.Inference (_, _ : TPTP_Proof.useful_info_as list, parent_infos)) =>
+        dest_parent_infos false fms parent_infos
+        |> map #name
+
+exception FIND_ANCESTOR_USING_RULE of string
+(*BFS for an ancestor inference involving a specific rule*)
+fun find_ancestor_using_rule pannot inference_rule (fringe : string list) : string =
+  if null fringe then
+    raise (FIND_ANCESTOR_USING_RULE inference_rule)
+  else
+    case node_info (#meta pannot) #source_inf_opt (hd fringe) of
+        NONE => find_ancestor_using_rule pannot inference_rule (tl fringe)
+      | SOME (TPTP_Proof.File _) => find_ancestor_using_rule pannot inference_rule (tl fringe)
+      | SOME (TPTP_Proof.Inference (rule_name, _ : TPTP_Proof.useful_info_as list, parent_infos)) =>
+          if rule_name = inference_rule then hd fringe
+          else
+            find_ancestor_using_rule pannot inference_rule
+             (tl fringe @
+              map #name (dest_parent_infos true (#meta pannot) parent_infos))
+
+(*Given a node in the proof, produce the term representing the inference
+  that took place in that step, the inference rule used, and which
+  other (non-axiom and non-definition) nodes participated in the
+  inference*)
+fun inference_at_node thy (prob_name : TPTP_Problem_Name.problem_name)
+     (fms : formula_meaning list) from : rule_info option =
+    let
+      exception INFERENCE_AT_NODE of string
+
+      (*lookup formula associated with a node*)
+      val fmla_of_node =
+          node_info fms #fmla
+          #> try_dest_Trueprop
+
+      fun build_inference_info rule_name parent_infos =
+        let
+          val _ = @{assert} (not (null parent_infos))
+
+          (*hypothesis formulas (with bindings already
+            instantiated during the proof-transformation
+            applied when loading the proof),
+            including any axioms or definitions*)
+          val parent_nodes =
+            dest_parent_infos false fms parent_infos
+            |> map #name
+
+          val parent_fmlas = map fmla_of_node (rev(*FIXME can do away with this? it matters because of order of conjunction. is there a matching rev elsewhere?*) parent_nodes)
+
+          val inference_term =
+            if null parent_fmlas then
+                fmla_of_node from
+                |> HOLogic.mk_Trueprop
+            else
+                Logic.mk_implies
+                 (fold
+                    (curry HOLogic.mk_conj)
+                    (tl parent_fmlas)
+                    (hd parent_fmlas)
+                  |> HOLogic.mk_Trueprop,
+                  fmla_of_node from |> HOLogic.mk_Trueprop)
+        in
+          SOME {inference_name = rule_name,
+                inference_fmla = inference_term,
+                parents = parent_nodes}
+        end
+    in
+      (*examine node's "source" annotation: we're only interested
+        if it's an inference*)
+      case node_info fms #source_inf_opt from of
+                NONE => NONE
+              | SOME (TPTP_Proof.File _) => NONE
+              | SOME (TPTP_Proof.Inference (rule_name, _ : TPTP_Proof.useful_info_as list, parent_infos)) =>
+                  if List.null parent_infos then
+                    raise (INFERENCE_AT_NODE
+                            ("empty parent list for node " ^
+                             from ^ ": check proof format"))
+                  else
+                    build_inference_info rule_name parent_infos
+    end
+
+
+(** Proof skeleton **)
+
+(* Emulating skeleton steps *)
+
+(*
+Builds a rule (thm) of the following form:
+
+
+                  prem1                   premn
+                   ...         ...         ...
+   major_prem     conc1                   concn
+  -----------------------------------------------
+                    conclusion
+
+where major_prem is a disjunction of prem1,...,premn.
+*)
+fun make_elimination_rule_t ctxt major_prem prems_and_concs conclusion =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    val minor_prems =
+      map (fn (v, conc) =>
+        Logic.mk_implies (v, HOLogic.mk_Trueprop conc))
+        prems_and_concs
+  in
+    (Logic.list_implies
+     (major_prem :: minor_prems,
+     conclusion))
+  end
+
+(*In summary, we emulate an n-way splitting rule via an n-way
+  disjunction elimination.
+
+  Given a split formula and conclusion, we prove a rule which
+  simulates the split. The conclusion is assumed to be a conjunction
+  of conclusions for each branch of the split. The
+  "minor_prem_assumptions" are the assumptions discharged in each
+  branch; they're passed to the function to make sure that the
+  generated rule is compatible with the skeleton (since the skeleton
+  fixes the "order" of the reconstruction, based on the proof's
+  structure).
+
+  Concretely, if P is "(_ & _) = $false" or "(_ | _) = $true" then
+  splitting behaves as follows:
+
+                     P
+      -------------------------------
+       _ = $false         _ = $false
+          ...       ...       ...
+           R1                  Rn
+      -------------------------------
+               R1 & ... & Rn
+
+  Splitting (binary) iffs works as follows:
+
+                  (A <=> B) = $false
+      ------------------------------------------
+       (A => B) = $false      (B => A) = $false
+             ...                     ...
+              R1                      R2
+      ------------------------------------------
+                        R1 & R2
+*)
+fun simulate_split ctxt split_fmla minor_prem_assumptions conclusion =
+  let
+    val prems_and_concs =
+      ListPair.zip (minor_prem_assumptions, flatten (Conjunctive NONE) conclusion)
+
+    val rule_t = make_elimination_rule_t ctxt split_fmla prems_and_concs conclusion
+
+    (*these are replaced by fresh variables in the abstract term*)
+    val abstraction_subterms =
+      (map (try_dest_Trueprop #> remove_polarity true #> fst)
+              minor_prem_assumptions)
+
+    (*generate an abstract rule as a term...*)
+    val abs_rule_t =
+      abstract
+        abstraction_subterms
+        rule_t
+      |> snd (*ignore mapping info. this is a bit wasteful*)
+             (*FIXME optimisation: instead on relying on diff
+                to regenerate this info, could use it directly*)
+
+    (*...and validate the abstract rule*)
+    val abs_rule_thm =
+      Goal.prove ctxt [] [] abs_rule_t
+       (fn pdata => HEADGOAL (blast_tac (#context pdata)))
+      |> Drule.export_without_context
+  in
+    (*Instantiate the abstract rule based on the contents of the
+      required instance*)
+    diff_and_instantiate ctxt abs_rule_thm (prop_of abs_rule_thm) rule_t
+  end
+
+
+(* Building the skeleton *)
+
+type step_id = string
+datatype rolling_stock =
+    Step of step_id
+  | Assumed
+  | Unconjoin
+  | Split of step_id (*where split occurs*) *
+             step_id (*where split ends*) *
+             step_id list (*children of the split*)
+  | Synth_step of step_id (*A step which doesn't necessarily appear in
+    the original proof, or which has been modified slightly for better
+    handling by Isabelle*) (*FIXME "inodes" should be made into Synth_steps*)
+  | Annotated_step of step_id * string (*Same interpretation as
+    "Step", except that additional information is attached. This is
+    currently used for debugging: Steps are mapped to Annotated_steps
+    and their rule names are included as strings*)
+  | Definition of step_id (*Mirrors TPTP role*)
+  | Axiom of step_id (*Mirrors TPTP role*)
+(*  | Derived of step_id -- to be used by memoization*)
+  | Caboose
+
+fun stock_to_string (Step n) = n
+  | stock_to_string (Annotated_step (n, anno)) = n ^ "(" ^ anno ^ ")"
+  | stock_to_string _ = error "Stock is not a step" (*FIXME more meaningful message*)
+
+fun filter_by_role tptp_role =
+  filter
+   (fn (_, info) =>
+       #role info = tptp_role)
+
+fun filter_by_name node_name =
+  filter
+   (fn (n, _) =>
+       n = node_name)
+
+exception NO_MARKER_NODE
+(*We fall back on node "1" in case the proof is not that of a theorem*)
+fun proof_beginning_node fms =
+  let
+    val result =
+      cascaded_filter_single true
+       [filter_by_role TPTP_Syntax.Role_Conjecture,
+        filter_by_name "1"] (*FIXME const*)
+       fms
+  in
+    case result of
+        SOME x => fst x (*get the node name*)
+      | NONE => raise NO_MARKER_NODE
+  end
+
+(*Get the name of the node where the proof ends*)
+fun proof_end_node fms =
+  (*FIXME this isn't very nice: we assume that the last line in the
+    proof file is the closing line of the proof. It would be nicer if
+    such a line is specially marked (with a role), since there is no
+    obvious ordering on names, since they can be strings.
+    Another way would be to run an analysis on the graph to find
+    this node, since it has properties which should make it unique
+    in a graph*)
+  fms
+  |> hd (*since proof has been reversed prior*)
+  |> fst (*get node name*)
+
+(*Generate list of (possibly reconstructed) inferences which can be
+  composed together to reconstruct the whole proof.*)
+fun make_skeleton ctxt (pannot : proof_annotation) : rolling_stock list =
+  let
+    val thy = Proof_Context.theory_of ctxt
+
+    fun stock_is_ax_or_def (Axiom _) = true
+      | stock_is_ax_or_def (Definition _) = true
+      | stock_is_ax_or_def _ = false
+
+    fun stock_of n =
+      case node_info (#meta pannot) #role n of
+          TPTP_Syntax.Role_Definition => (true, Definition n)
+        | TPTP_Syntax.Role_Axiom => (true, Axiom n)
+        | _ => (false, Step n)
+
+    fun n_is_split_conjecture (inference_info : rule_info option) =
+      case inference_info of
+          NONE => false
+        | SOME inference_info => #inference_name inference_info = "split_conjecture"
+
+    (*Different kinds of inference sequences:
+        - Linear: (just add a step to the skeleton)
+           ---...---
+
+        - Fan-in: (treat each in-path as conjoined with the others. Linearise all the paths, and concatenate them.)
+                  /---...
+           ------<
+                  \---...
+
+        - Real split: Instead of treating as a conjunction, as in
+           normal fan-ins, we need to treat specially by looking
+           at the location where the split occurs, and turn the
+           split inference into a validity-preserving subproof.
+           As with fan-ins, we handle each fan-in path, and
+           concatenate.
+                  /---...---\
+           ------<           >------
+                  \---...---/
+
+        - Fake split: (treat like linear, since there isn't a split-node)
+           ------<---...----------
+
+      Different kinds of sequences endings:
+        - "Stop before": Non-decreasing list of nodes where should terminate.
+                         This starts off with the end node, and the split_nodes
+                         will be added dynamically as the skeleton is built.
+        - Axiom/Definition
+     *)
+
+    (*The following functions build the skeleton for the reconstruction starting
+      from the node labelled "n" and stopping just before an element in stop_just_befores*)
+    (*FIXME could throw exception if none of stop_just_befores is ever encountered*)
+
+    (*This approach below is naive because it linearises the proof DAG, and this would
+      duplicate some effort if the DAG isn't already linear.*)
+    exception SKELETON
+
+    fun check_parents stop_just_befores n =
+      let
+        val parents = parents_of_node (#meta pannot) n
+      in
+        if length parents = 1 then
+          AList.lookup (op =) stop_just_befores (the_single parents)
+        else
+          NONE
+      end
+
+    fun naive_skeleton' stop_just_befores n =
+      case check_parents stop_just_befores n of
+          SOME skel => skel
+        | NONE =>
+            let
+              val inference_info = inference_at_node thy (#problem_name pannot) (#meta pannot) n
+            in
+                if is_none inference_info then
+                  (*this is the case for the conjecture, definitions and axioms*)
+                    if node_info (#meta pannot) #role n = TPTP_Syntax.Role_Definition then
+                      [(Definition n), Assumed]
+                    else if node_info (#meta pannot) #role n = TPTP_Syntax.Role_Axiom then
+                      [Axiom n]
+                    else raise SKELETON
+                else
+                  let
+                    val inference_info = the inference_info
+                    val parents = #parents inference_info
+                  in
+                    (*FIXME memoize antecedent_steps?*)
+                    if #inference_name inference_info = "solved_all_splits" andalso length parents > 1 then
+                      (*splitting involves fanning out then in; this is to be
+                        treated different than other fan-out-ins.*)
+                      let
+                        (*find where the proofs fanned-out: pick some antecedent,
+                          then find ancestor to use a "split_conjecture" inference.*)
+                        (*NOTE we assume that splits can't be nested*)
+                        val split_node =
+                          find_ancestor_using_rule pannot "split_conjecture" [hd parents]
+                          |> parents_of_node (#meta pannot)
+                          |> the_single
+
+                        (*compute the skeletons starting at parents to either the split_node
+                          if the antecedent is descended from the split_node, or the
+                          stop_just_before otherwise*)
+                        val skeletons_up =
+                          map (naive_skeleton' ((split_node, [Assumed]) :: stop_just_befores)) parents
+                      in
+                        (*point to the split node, so that custom rule can be built later on*)
+                        Step n :: (Split (split_node, n, parents)) :: (*this will create the elimination rule*)
+                         naive_skeleton' stop_just_befores split_node @ (*this will discharge the major premise*)
+                         List.concat skeletons_up @ [Assumed] (*this will discharge the minor premises*)
+                      end
+                    else if length parents > 1 then
+                      (*Handle fan-in nodes which aren't split-sinks by
+                        enclosing each branch but one in conjI-assumption invocations*)
+                        let
+                          val skeletons_up =
+                            map (naive_skeleton' stop_just_befores) parents
+                        in
+                          Step n :: concat_between skeletons_up (SOME Unconjoin, NONE) @ [Assumed]
+                        end
+                    else
+                      Step n :: naive_skeleton' stop_just_befores (the_single parents)
+                  end
+            end
+  in
+    if List.null (#meta pannot) then [] (*in case "proof" file is empty*)
+    else
+      naive_skeleton'
+       [(proof_beginning_node (#meta pannot), [Assumed])]
+       (proof_end_node (#meta pannot))
+      (*make last step the Caboose*)
+      |> rev |> tl |> cons Caboose |> rev (*FIXME hacky*)
+  end
+
+
+(* Using the skeleton *)
+
+exception SKELETON
+local
+    (*Change the negated assumption (which is output by the contradiction rule) into
+      a form familiar to Leo2*)
+    val neg_eq_false =
+      @{lemma "!! P. (~ P) ==> (P = False)" by auto}
+
+    (*FIXME this is just a dummy thm to annotate the assumption tac "atac"*)
+    val solved_all_splits =
+      @{lemma "False = True ==> False" by auto}
+
+    fun skel_to_naive_tactic ctxt prover_tac prob_name skel memo = fn st =>
+      let
+        val thy = Proof_Context.theory_of ctxt
+        val pannot = get_pannot_of_prob thy prob_name
+        fun tac_and_memo node memo =
+          case AList.lookup (op =) memo node of
+              NONE =>
+                let
+                  val tac =
+                    (*FIXME formula_sizelimit not being
+                            checked here*)
+                    prover_tac ctxt prob_name node
+                in (tac, (node, tac) :: memo) end
+            | SOME tac => (tac, memo)
+        fun rest skel' memo =
+          skel_to_naive_tactic ctxt prover_tac prob_name skel' memo
+
+        val tactic =
+          if null skel then
+            raise SKELETON (*FIXME or classify it as a Caboose: TRY (HEADGOAL atac) *)
+          else
+            case hd skel of
+                Assumed => TRY (HEADGOAL atac) THEN rest (tl skel) memo
+              | Caboose => TRY (HEADGOAL atac)
+              | Unconjoin => rtac @{thm conjI} 1 THEN rest (tl skel) memo
+              | Split (split_node, solved_node, antes) =>
+                  let
+                    val split_fmla = node_info (#meta pannot) #fmla split_node
+                    val conclusion =
+                      (inference_at_node thy prob_name (#meta pannot) solved_node
+                       |> the
+                       |> #inference_fmla)
+                      |> Logic.dest_implies (*FIXME there might be !!-variables?*)
+                      |> #1
+                    val minor_prems_assumps =
+                      map (fn ante => find_ancestor_using_rule pannot "split_conjecture" [ante]) antes
+                      |> map (node_info (#meta pannot) #fmla)
+                    val split_thm =
+                      simulate_split ctxt split_fmla minor_prems_assumps conclusion
+                  in
+                    rtac split_thm 1 THEN rest (tl skel) memo
+                  end
+              | Step s =>
+                  let
+                    val (tac, memo') = tac_and_memo s memo
+                  in
+                    rtac tac 1 THEN rest (tl skel) memo'
+                  end
+              | Definition n =>
+                  let
+                    val def_thm =
+                      case AList.lookup (op =) (#defs pannot) n of
+                          NONE => error ("Did not find definition: " ^ n)
+                        | SOME binding =>
+                            Binding.dest binding
+                            |> #3
+                            |> Global_Theory.get_thm thy
+                  in
+                    rtac def_thm 1 THEN rest (tl skel) memo
+                  end
+              | Axiom n =>
+                  let
+                    val ax_thm =
+                      case AList.lookup (op =) (#axs pannot) n of
+                          NONE => error ("Did not find axiom: " ^ n)
+                        | SOME binding =>
+                            Binding.dest binding
+                            |> #3
+                            |> Global_Theory.get_thm thy
+                  in
+                    rtac ax_thm 1 THEN rest (tl skel) memo
+                  end
+              | _ => raise SKELETON
+      in tactic st end
+(*FIXME fuse these*)
+    (*As above, but creates debug-friendly tactic.
+      This is also used for "partial proof reconstruction"*)
+    fun skel_to_naive_tactic_dbg prover_tac ctxt prob_name skel (memo : (string * (thm * tactic) option) list) =
+      let
+        val thy = Proof_Context.theory_of ctxt
+        val pannot = get_pannot_of_prob thy prob_name
+
+        fun rtac_wrap thm_f i = fn st =>
+          let
+            val thy = Thm.theory_of_thm st
+          in
+            rtac (thm_f thy) i st
+          end
+
+        (*Some nodes don't have an inference name, such as the conjecture,
+          definitions and axioms. Such nodes shouldn't appear in the
+          skeleton.*)
+        fun inference_name_of_node node =
+           case AList.lookup (op =) (#meta pannot) node of
+               NONE => (warning "Inference step lacks an inference name"; "(Shouldn't be here)")
+             | SOME info =>
+                 case #source_inf_opt info of
+                     SOME (TPTP_Proof.Inference (infname, _, _)) =>
+                       infname
+                   | _ => (warning "Inference step lacks an inference name"; "(Shouldn't be here)")
+
+        fun inference_fmla node =
+          case inference_at_node thy prob_name (#meta pannot) node of
+              NONE => NONE
+            | SOME {inference_fmla, ...} => SOME inference_fmla
+
+        fun rest memo' ctxt' = skel_to_naive_tactic_dbg prover_tac ctxt' prob_name (tl skel) memo'
+        (*reconstruct the inference. also set timeout in case
+          tactic takes too long*)
+        val try_make_step =
+          (*FIXME const timeout*)
+          (* TimeLimit.timeLimit (Time.fromSeconds 5) *)
+          (fn ctxt' =>
+             let
+               fun thm ctxt'' = prover_tac ctxt'' prob_name (hd skel |> stock_to_string)
+               val reconstructed_inference = thm ctxt'
+               val rec_inf_tac = fn st =>
+                 let
+                   val ctxt =
+                     Thm.theory_of_thm st
+                     |> Proof_Context.init_global
+                 in
+                   HEADGOAL (rtac (thm ctxt)) st
+                 end
+             in (reconstructed_inference,
+                 rec_inf_tac)
+             end)
+        fun ignore_interpretation_exn f x = SOME (f x)
+          handle
+              INTERPRET_INFERENCE => NONE
+            | exn => reraise exn
+      in
+        if List.null skel then
+          raise SKELETON
+        (*FIXME or classify it as follows:
+          [(Caboose,
+            prop_of @{thm asm_rl}
+            |> SOME,
+            SOME (@{thm asm_rl}, TRY (HEADGOAL atac)))]
+         *)
+        else
+          case hd skel of
+              Assumed =>
+                (hd skel,
+                 prop_of @{thm asm_rl}
+                 |> SOME,
+                 SOME (@{thm asm_rl}, TRY (HEADGOAL atac))) :: rest memo ctxt
+            | Caboose =>
+                [(Caboose,
+                  prop_of @{thm asm_rl}
+                  |> SOME,
+                  SOME (@{thm asm_rl}, TRY (HEADGOAL atac)))]
+            | Unconjoin =>
+                (hd skel,
+                 prop_of @{thm conjI}
+                 |> SOME,
+                 SOME (@{thm conjI}, rtac @{thm conjI} 1)) :: rest memo ctxt
+            | Split (split_node, solved_node, antes) =>
+                let
+                  val split_fmla = node_info (#meta pannot) #fmla split_node
+                  val conclusion =
+                        (inference_at_node thy prob_name (#meta pannot) solved_node
+                         |> the
+                         |> #inference_fmla)
+                        |> Logic.dest_implies (*FIXME there might be !!-variables?*)
+                        |> #1
+                  val minor_prems_assumps =
+                      map (fn ante => find_ancestor_using_rule pannot "split_conjecture" [ante]) antes
+                      |> map (node_info (#meta pannot) #fmla)
+                  val split_thm =
+                      simulate_split ctxt split_fmla minor_prems_assumps conclusion
+                in
+                  (hd skel,
+                   prop_of split_thm
+                   |> SOME,
+                   SOME (split_thm, rtac split_thm 1)) :: rest memo ctxt
+                end
+            | Step node =>
+                let
+                  val inference_name = inference_name_of_node node
+                  val inference_fmla = inference_fmla node
+
+                  (*FIXME debugging code
+                  val _ =
+                    if Config.get ctxt tptp_trace_reconstruction then
+                       (tracing ("handling node " ^ node);
+                        tracing ("inference " ^ inference_name);
+                        if is_some inference_fmla then
+                          tracing ("formula size " ^ Int.toString (Term.size_of_term (the inference_fmla)))
+                        else ()(*;
+                        tracing ("formula " ^ @{make_string inference_fmla}) *))
+                    else ()*)
+
+                  val (inference_instance_thm, memo', ctxt') =
+                    case AList.lookup (op =) memo node of
+                        NONE =>
+                          let
+                            val (thm, ctxt') =
+                              (*Instead of NONE could have another value indicating that the formula was too big*)
+                                if is_some inference_fmla andalso
+                                   (*FIXME could have different inference rules have different sizelimits*)
+                                   exceeds_tptp_max_term_size ctxt (Term.size_of_term (the inference_fmla)) then
+                                    (
+                                     warning ("Gave up on node " ^ node ^ " because of fmla size " ^
+                                              Int.toString (Term.size_of_term (the inference_fmla)));
+                                     (NONE, ctxt)
+                                    )
+                                else
+                                  let
+                                    val maybe_thm = ignore_interpretation_exn try_make_step ctxt
+                                    val ctxt' =
+                                      if is_some maybe_thm then
+                                        the maybe_thm
+                                        |> #1
+                                        |> Thm.theory_of_thm |> Proof_Context.init_global
+                                      else ctxt
+                                  in
+                                    (maybe_thm, ctxt')
+                                  end
+                          in (thm, (node, thm) :: memo, ctxt') end
+                      | SOME maybe_thm => (maybe_thm, memo, ctxt)
+                in
+                  (Annotated_step (node, inference_name),
+                   inference_fmla,
+                   inference_instance_thm) :: rest memo' ctxt'
+                end
+            | Definition n =>
+                let
+                  fun def_thm thy =
+                    case AList.lookup (op =) (#defs pannot) n of
+                        NONE => error ("Did not find definition: " ^ n)
+                      | SOME binding =>
+                          Binding.dest binding
+                          |> #3
+                          |> Global_Theory.get_thm thy
+                in
+                  (hd skel,
+                   prop_of (def_thm thy)
+                   |> SOME,
+                   SOME (def_thm thy,
+                         HEADGOAL (rtac_wrap def_thm))) :: rest memo ctxt
+                end
+            | Axiom n =>
+                let
+                  val ax_thm =
+                    case AList.lookup (op =) (#axs pannot) n of
+                        NONE => error ("Did not find axiom: " ^ n)
+                      | SOME binding =>
+                          Binding.dest binding
+                          |> #3
+                          |> Global_Theory.get_thm thy
+                in
+                  (hd skel,
+                   prop_of ax_thm
+                   |> SOME,
+                   SOME (ax_thm, rtac ax_thm 1)) :: rest memo ctxt
+                end
+      end
+
+    (*The next function handles cases where Leo2 doesn't include the solved_all_splits
+      step at the end (e.g. because there wouldn't be a split -- the proof
+      would be linear*)
+    fun sas_if_needed_tac ctxt prob_name =
+      let
+        val thy = Proof_Context.theory_of ctxt
+        val pannot = get_pannot_of_prob thy prob_name
+        val last_inference_info_opt =
+          find_first
+           (fn (_, info) => #role info = TPTP_Syntax.Role_Plain)
+           (#meta pannot)
+        val last_inference_info =
+          case last_inference_info_opt of
+              NONE => NONE
+            | SOME (_, info) => #source_inf_opt info
+      in
+        if is_some last_inference_info andalso
+         TPTP_Proof.is_inference_called "solved_all_splits"
+          (the last_inference_info)
+        then (@{thm asm_rl}, all_tac)
+        else (solved_all_splits, TRY (rtac solved_all_splits 1))
+      end
+in
+  (*Build a tactic from a skeleton. This is naive because it uses the naive skeleton.
+    The inference interpretation ("prover_tac") is a parameter -- it would usually be
+    different for different provers.*)
+  fun naive_reconstruct_tac ctxt prover_tac prob_name =
+    let
+      val thy = Proof_Context.theory_of ctxt
+    in
+      rtac @{thm ccontr} 1
+      THEN dtac neg_eq_false 1
+      THEN (sas_if_needed_tac ctxt prob_name |> #2)
+      THEN skel_to_naive_tactic ctxt prover_tac prob_name
+       (make_skeleton ctxt
+        (get_pannot_of_prob thy prob_name)) []
+    end
+
+  (*As above, but generates a list of tactics. This is useful for debugging, to apply
+    the tactics one by one manually.*)
+  fun naive_reconstruct_tacs prover_tac prob_name ctxt =
+    let
+      val thy = Proof_Context.theory_of ctxt
+    in
+      (Synth_step "ccontr", prop_of @{thm ccontr} |> SOME,
+       SOME (@{thm ccontr}, rtac @{thm ccontr} 1)) ::
+      (Synth_step "neg_eq_false", prop_of neg_eq_false |> SOME,
+       SOME (neg_eq_false, dtac neg_eq_false 1)) ::
+      (Synth_step "sas_if_needed_tac", prop_of @{thm asm_rl}(*FIXME *) |> SOME,
+       SOME (sas_if_needed_tac ctxt prob_name)) ::
+      skel_to_naive_tactic_dbg prover_tac ctxt prob_name
+       (make_skeleton ctxt
+        (get_pannot_of_prob thy prob_name)) []
+    end
+end
+
+(*Produces a theorem given a tactic and a parsed proof. This function is handy
+to test reconstruction, since it automates the interpretation and proving of the
+parsed proof's goal.*)
+fun reconstruct ctxt tactic prob_name =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    val pannot = get_pannot_of_prob thy prob_name
+    val goal =
+      #meta pannot
+      |> List.filter (fn (_, info) =>
+          #role info = TPTP_Syntax.Role_Conjecture)
+  in
+    if null (#meta pannot) then
+      (*since the proof is empty, return a trivial result.*)
+      @{thm TrueI}
+    else if null goal then
+      raise (RECONSTRUCT "Proof lacks conjecture")
+    else
+      the_single goal
+      |> snd |> #fmla
+      |> (fn fmla => Goal.prove ctxt [] [] fmla (fn _ => tactic prob_name))
+  end
+
+
+(** Skolemisation setup **)
+
+(*Ignore these constants if they appear in the conclusion but not the hypothesis*)
+(*FIXME possibly incomplete*)
+val ignore_consts =
+  [HOLogic.conj, HOLogic.disj, HOLogic.imp, HOLogic.Not]
+
+(*Difference between the constants appearing between two terms, minus "ignore_consts"*)
+fun new_consts_between t1 t2 =
+  List.filter
+   (fn n => not (List.exists (fn n' => n' = n) ignore_consts))
+   (list_diff (consts_in t2) (consts_in t1))
+
+(*Generate definition binding for an equation*)
+fun mk_bind_eq prob_name params ((n, ty), t) =
+  let
+    val bnd =
+      Binding.name (List.last (space_explode "." n) ^ "_def")
+      |> Binding.qualify false (TPTP_Problem_Name.mangle_problem_name prob_name)
+    val t' =
+      Term.list_comb (Const (n, ty), params)
+      |> rpair t
+      |> HOLogic.mk_eq
+      |> HOLogic.mk_Trueprop
+      |> fold Logic.all params
+  in
+    (bnd, t')
+  end
+
+(*Generate binding for an axiom. Similar to "mk_bind_eq"*)
+fun mk_bind_ax prob_name node t =
+  let
+    val bnd =
+      Binding.name node
+      (*FIXME add suffix? e.g. ^ "_ax"*)
+      |> Binding.qualify false (TPTP_Problem_Name.mangle_problem_name prob_name)
+  in
+    (bnd, t)
+  end
+
+(*Extract the constant name, type, and its definition*)
+fun get_defn_components
+  (Const (@{const_name HOL.Trueprop}, _) $
+    (Const (@{const_name HOL.eq}, _) $
+      Const (name, ty) $ t)) = ((name, ty), t)
+
+
+(*** Proof transformations ***)
+
+(*Transforms a proof_annotation value.
+  Argument "f" is the proof transformer*)
+fun transf_pannot f (pannot : proof_annotation) : (theory * proof_annotation) =
+  let
+    val (thy', fms') = f (#meta pannot)
+  in
+    (thy',
+     {problem_name = #problem_name pannot,
+      skolem_defs = #skolem_defs pannot,
+      defs = #defs pannot,
+      axs = #axs pannot,
+      meta = fms'})
+  end
+
+
+(** Proof transformer to add virtual inference steps
+    encoding "bind" annotations in Leo-II proofs **)
+
+(*
+Involves finding an inference of this form:
+
+       (!x1 ... xn. F)   ...   Cn
+  ------------------------------------ (Rule name)
+          G[t1/x1, ..., tn/xn]
+
+and turn it into this:
+
+
+     (!x1 ... xn. F)
+  ---------------------- bind
+   F[t1/x1, ..., tn/xn]           ...   Cn
+  -------------------------------------------- (Rule name)
+                    G
+
+where "bind" is an inference rule (distinct from any rule name used
+by Leo2) to indicate such inferences.  This transformation is used
+to factor out instantiations, thus allowing the reconstruction to
+focus on (Rule name) rather than "(Rule name) + instantiations".
+*)
+fun interpolate_binds prob_name thy fms : theory * formula_meaning list =
+  let
+    fun factor_out_bind target_node pinfo intermediate_thy =
+      case pinfo of
+         TPTP_Proof.ParentWithDetails (n, pdetails) =>
+           (*create new node which contains the "bind" inference,
+             to be added to graph*)
+           let
+             val (new_node_name, thy') = get_next_name intermediate_thy
+             val orig_fmla = node_info fms #fmla n
+             val target_fmla = node_info fms #fmla target_node
+             val new_node =
+              (new_node_name,
+               {role = TPTP_Syntax.Role_Plain,
+                fmla = apply_binding thy' prob_name orig_fmla target_fmla pdetails |> fst,
+                source_inf_opt =
+                  SOME (TPTP_Proof.Inference (bindK, [], [pinfo]))})
+           in
+             ((TPTP_Proof.Parent new_node_name, SOME new_node), thy')
+           end
+       | _ => ((pinfo, NONE), intermediate_thy)
+    fun process_nodes (step as (n, data)) (intermediate_thy, rest) =
+      case #source_inf_opt data of
+          SOME (TPTP_Proof.Inference (inf_name, sinfos, pinfos)) =>
+            let
+              val ((pinfos', parent_nodes), thy') =
+                fold_map (factor_out_bind n) pinfos intermediate_thy
+                |> apfst ListPair.unzip
+              val step' =
+                (n, {role = #role data, fmla = #fmla data,
+                 source_inf_opt = SOME (TPTP_Proof.Inference (inf_name, sinfos, pinfos'))})
+            in (thy', fold_options parent_nodes @ step' :: rest) end
+        | _ => (intermediate_thy, step :: rest)
+  in
+    fold process_nodes fms (thy, [])
+    (*new_nodes must come at the beginning, since we assume that the last line in a proof is the closing line*)
+    |> apsnd rev
+  end
+
+
+(** Proof transformer to add virtual inference steps
+    encoding any transformation done immediately prior
+    to a splitting step **)
+
+(*
+Involves finding an inference of this form:
+
+                   F = $false
+  ----------------------------------- split_conjecture
+    (F1 = $false) ... (Fn = $false)
+
+where F doesn't have an "and" or "iff" at the top level,
+and turn it into this:
+
+                   F = $false
+  ----------------------------------- split_preprocessing
+            (F1 % ... % Fn) = $false
+  ----------------------------------- split_conjecture
+    (F1 = $false) ... (Fn = $false)
+
+where "%" is either an "and" or an "iff" connective.
+This transformation is used to clarify the clause structure, to
+make it immediately "obvious" how splitting is taking place
+(by factoring out the other syntactic transformations -- e.g.
+related to quantifiers -- performed by Leo2). Having the clause
+in this "clearer" form makes the inference amenable to handling
+using the "abstraction" technique, which allows us to validate
+large inferences.
+*)
+exception PREPROCESS_SPLITS
+fun preprocess_splits prob_name thy fms : theory * formula_meaning list =
+  let
+    (*Simulate the transformation done by Leo2's preprocessing
+      step during splitting.
+      NOTE: we assume that the clause is a singleton
+
+      This transformation does the following:
+       - miniscopes !-quantifiers (and recurs)
+       - removes redundant ?-quantifiers (and recurs)
+       - eliminates double negation (and recurs)
+       - breaks up conjunction (and recurs)
+       - expands iff (and doesn't recur)*)
+    fun transform_fmla i fmla_t =
+      case fmla_t of
+          Const (@{const_name "HOL.All"}, ty) $ Abs (s, ty', t') =>
+            let
+              val (i', fmla_ts) = transform_fmla i t'
+            in
+              if i' > i then
+                (i' + 1,
+                 map (fn t =>
+                  Const (@{const_name "HOL.All"}, ty) $ Abs (s, ty', t))
+                fmla_ts)
+              else (i, [fmla_t])
+            end
+        | Const (@{const_name "HOL.Ex"}, ty) $ Abs (s, ty', t') =>
+            if loose_bvar (t', 0) then
+              (i, [fmla_t])
+            else transform_fmla (i + 1) t'
+        | @{term HOL.Not} $ (@{term HOL.Not} $ t') =>
+            transform_fmla (i + 1) t'
+        | @{term HOL.conj} $ t1 $ t2 =>
+            let
+              val (i1, fmla_t1s) = transform_fmla (i + 1) t1
+              val (i2, fmla_t2s) = transform_fmla (i + 1) t2
+            in
+              (i1 + i2 - i, fmla_t1s @ fmla_t2s)
+            end
+        | Const (@{const_name HOL.eq}, ty) $ t1 $ t2 =>
+            let
+              val (T1, (T2, res)) =
+                dest_funT ty
+                |> apsnd dest_funT
+            in
+              if T1 = HOLogic.boolT andalso T2 = HOLogic.boolT andalso
+                 res = HOLogic.boolT then
+                (i + 1,
+                  [HOLogic.mk_imp (t1, t2),
+                   HOLogic.mk_imp (t2, t1)])
+              else (i, [fmla_t])
+            end
+        | _ => (i, [fmla_t])
+
+    fun preprocess_split thy split_node_name fmla_t =
+      (*create new node which contains the new inference,
+        to be added to graph*)
+      let
+        val (node_name, thy') = get_next_name thy
+        val (changes, fmla_conjs) =
+          transform_fmla 0 fmla_t
+          |> apsnd rev (*otherwise we run into problems because
+                         of commutativity of conjunction*)
+        val target_fmla =
+          fold (curry HOLogic.mk_conj) (tl fmla_conjs) (hd fmla_conjs)
+        val new_node =
+         (node_name,
+          {role = TPTP_Syntax.Role_Plain,
+           fmla =
+             HOLogic.mk_eq (target_fmla, @{term False}) (*polarise*)
+             |> HOLogic.mk_Trueprop,
+           source_inf_opt =
+             SOME (TPTP_Proof.Inference (split_preprocessingK, [], [TPTP_Proof.Parent split_node_name]))})
+      in
+        if changes = 0 then NONE
+        else SOME (TPTP_Proof.Parent node_name, new_node, thy')
+      end
+  in
+    fold
+     (fn step as (n, data) => fn (intermediate_thy, redirections, rest) =>
+       case #source_inf_opt data of
+            SOME (TPTP_Proof.Inference
+                   (inf_name, sinfos, pinfos)) =>
+              if inf_name <> "split_conjecture" then
+                (intermediate_thy, redirections, step :: rest)
+              else
+                let
+                  (*
+                   NOTE: here we assume that the node only has one
+                         parent, and that there is no additional
+                         parent info.
+                   *)
+                  val split_node_name =
+                    case pinfos of
+                        [TPTP_Proof.Parent n] => n
+                      | _ => raise PREPROCESS_SPLITS
+                (*check if we've already handled that already node*)
+                in
+                  case AList.lookup (op =) redirections split_node_name of
+                      SOME preprocessed_split_node_name =>
+                        let
+                          val step' =
+                            apply_to_parent_info (fn _ => [TPTP_Proof.Parent preprocessed_split_node_name]) step
+                        in (intermediate_thy, redirections, step' :: rest) end
+                    | NONE =>
+                        let
+                          (*we know the polarity to be $false, from knowing Leo2*)
+                          val split_fmla =
+                            try_dest_Trueprop (node_info fms #fmla split_node_name)
+                            |> remove_polarity true
+                            |> fst
+
+                          val preprocess_result =
+                            preprocess_split intermediate_thy
+                              split_node_name
+                              split_fmla
+                        in
+                          if is_none preprocess_result then
+                            (*no preprocessing done by Leo2, so no need to introduce
+                              a virtual inference. cache this result by
+                              redirecting the split_node to itself*)
+                            (intermediate_thy,
+                             (split_node_name, split_node_name) :: redirections,
+                             step :: rest)
+                          else
+                            let
+                              val (new_parent_info, new_parent_node, thy') = the preprocess_result
+                              val step' =
+                                (n, {role = #role data, fmla = #fmla data,
+                                 source_inf_opt = SOME (TPTP_Proof.Inference (inf_name, sinfos, [new_parent_info]))})
+                            in
+                              (thy',
+                               (split_node_name, fst new_parent_node) :: redirections,
+                               step' :: new_parent_node :: rest)
+                            end
+                        end
+                end
+          | _ => (intermediate_thy, redirections, step :: rest))
+     (rev fms) (*this allows us to put new inferences before other inferences which use them*)
+     (thy, [], [])
+    |> (fn (x, _, z) => (x, z)) (*discard redirection info*)
+  end
+
+
+(** Proof transformer to remove repeated quantification **)
+
+exception DROP_REPEATED_QUANTIFICATION
+fun drop_repeated_quantification thy (fms : formula_meaning list) : theory * formula_meaning list =
+  let
+    (*In case of repeated quantification, removes outer quantification.
+      Only need to look at top-level, since the repeated quantification
+      generally occurs at clause level*)
+    fun remove_repeated_quantification seen t =
+      case t of
+          (*NOTE we're assuming that variables having the same name, have the same type throughout*)
+          Const (@{const_name "HOL.All"}, ty) $ Abs (s, ty', t') =>
+            let
+              val (seen_so_far, seen') =
+                case AList.lookup (op =) seen s of
+                    NONE => (0, (s, 0) :: seen)
+                  | SOME n => (n + 1, AList.update (op =) (s, n + 1) seen)
+              val (pre_final_t, final_seen) = remove_repeated_quantification seen' t'
+              val final_t =
+                case AList.lookup (op =) final_seen s of
+                    NONE => raise DROP_REPEATED_QUANTIFICATION
+                  | SOME n =>
+                      if n > seen_so_far then pre_final_t
+                      else Const (@{const_name "HOL.All"}, ty) $ Abs (s, ty', pre_final_t)
+            in (final_t, final_seen) end
+        | _ => (t, seen)
+
+    fun remove_repeated_quantification' (n, {role, fmla, source_inf_opt}) =
+      (n,
+       {role = role,
+        fmla =
+          try_dest_Trueprop fmla
+          |> remove_repeated_quantification []
+          |> fst
+          |> HOLogic.mk_Trueprop,
+        source_inf_opt = source_inf_opt})
+  in
+    (thy, map remove_repeated_quantification' fms)
+  end
+
+
+(** Proof transformer to detect a redundant splitting and remove
+    the redundant branch. **)
+
+fun node_is_inference fms rule_name node_name =
+  case node_info fms #source_inf_opt node_name of
+      NONE => false
+    | SOME (TPTP_Proof.File _) => false
+    | SOME (TPTP_Proof.Inference (rule_name', _, _)) => rule_name' = rule_name
+
+(*In this analysis we're interested if there exists a split-free
+  path between the end of the proof and the negated conjecture.
+  If so, then this path (or the shortest such path) could be
+  retained, and the rest of the proof erased.*)
+datatype branch_info =
+    Split_free (*Path is not part of a split. This is only used when path reaches the negated conjecture.*)
+  | Split_present (*Path is one of a number of splits. Such paths are excluded.*)
+  | Coinconsistent of int (*Path leads to a clause which is inconsistent with nodes concluded by other paths.
+                            Therefore this path should be kept if the others are kept
+                            (i.e., unless one of them results from a split)*)
+  | No_info (*Analysis hasn't come across anything definite yet, though it still hasn't completed.*)
+(*A "paths" value consist of every way of reaching the destination,
+  including information come across it so far. Taking the head of
+  each way gives the fringe. All paths should share the same source
+  and sink.*)
+type path = (branch_info * string list)
+exception PRUNE_REDUNDANT_SPLITS
+fun prune_redundant_splits prob_name thy fms : theory * formula_meaning list =
+  let
+    (*All paths start at the contradiction*)
+    val initial_path = (No_info, [proof_end_node fms])
+    (*All paths should end at the proof's beginning*)
+    val end_node = proof_beginning_node fms
+
+    fun compute_path (path as ((info,
+                       (n :: ns)) : path))(*i.e. node list can't be empty*)
+        intermediate_thy =
+      case info of
+          Split_free => (([path], []), intermediate_thy)
+        | Coinconsistent branch_id =>
+            (*If this branch has a split_conjecture parent then all "sibling" branches get erased.*)
+            (*This branch can't lead to yet another coinconsistent branch (in the case of Leo2).*)
+            let
+              val parent_nodes = parents_of_node fms n
+            in
+              if List.exists (node_is_inference fms "split_conjecture") parent_nodes then
+                (([], [branch_id]), intermediate_thy) (*all related branches are to be deleted*)
+              else
+                list_prod [] parent_nodes (n :: ns)
+                |> map (fn ns' => (Coinconsistent branch_id, ns'))
+                |> (fn x => ((x, []), intermediate_thy))
+            end
+
+        | No_info =>
+            let
+              val parent_nodes = parents_of_node fms n
+
+              (*if this node is a consistency checking node then parent nodes will be marked as coinconsistent*)
+              val (thy', new_branch_info) =
+                if node_is_inference fms "fo_atp_e" n orelse
+                   node_is_inference fms "res" n then
+                  let
+                    val (i', intermediate_thy') = get_next_int intermediate_thy
+                  in
+                    (intermediate_thy', SOME (Coinconsistent i'))
+                  end
+                else (intermediate_thy, NONE)
+            in
+              if List.exists (node_is_inference fms "split_conjecture") parent_nodes then
+                (([], []), thy')
+              else
+                list_prod [] parent_nodes (n :: ns)
+                |> map (fn ns' =>
+                          let
+                            val info =
+                              if is_some new_branch_info then the new_branch_info
+                              else
+                                if hd ns' = end_node then Split_free else No_info
+                          in (info, ns') end)
+                |> (fn x => ((x, []), thy'))
+            end
+        | _ => raise PRUNE_REDUNDANT_SPLITS
+
+    fun compute_paths intermediate_thy (paths : path list) =
+      if filter (fn (_, ns) => ns <> [] andalso hd ns = end_node) paths = paths then
+        (*fixpoint reached when all paths are at the head position*)
+        (intermediate_thy, paths)
+      else
+        let
+          val filtered_paths = filter (fn (info, _) : path => info <> Split_present) paths (*not interested in paths containing a split*)
+          val (paths', thy') =
+            fold_map compute_path filtered_paths intermediate_thy
+        in
+          paths'
+          |> ListPair.unzip (*we get a list of pairs of lists. we want a pair of lists*)
+          |> (fn (paths, branch_ids) =>
+               (List.concat paths,
+                (*remove duplicate branch_ids*)
+                fold (Library.insert (op =)) (List.concat branch_ids) []))
+          (*filter paths having branch_ids appearing in the second list*)
+          |> (fn (paths, branch_ids) =>
+              filter (fn (info, _) =>
+                        case info of
+                            Coinconsistent branch_id => List.exists (fn x => x = branch_id) branch_ids
+                          | _ => true) paths)
+          |> compute_paths thy'
+        end
+
+    val (thy', paths) =
+      compute_paths thy [initial_path]
+      |> apsnd
+          (filter (fn (branch_info, _) =>
+                  case branch_info of
+                      Split_free => true
+                    | Coinconsistent _ => true
+                    | _ => false))
+    (*Extract subset of fms which is used in a path.
+      Also, remove references (in parent info annotations) to erased nodes.*)
+    fun path_to_fms ((_, nodes) : path) =
+      fold
+       (fn n => fn fms' =>
+          case AList.lookup (op =) fms' n of
+              SOME _ => fms'
+            | NONE =>
+               let
+                 val node_info = the (AList.lookup (op =) fms n)
+
+                 val source_info' =
+                   case #source_inf_opt node_info of
+                       NONE => error "Only the conjecture is an orphan"
+                     | SOME (source_info as TPTP_Proof.File _) => source_info
+                     | SOME (source_info as
+                             TPTP_Proof.Inference (inference_name,
+                                                   useful_infos : TPTP_Proof.useful_info_as list,
+                                                   parent_infos)) =>
+                         let
+                           fun is_node_in_fms' parent_info =
+                             let
+                               val parent_nodename =
+                                 case parent_info of
+                                     TPTP_Proof.Parent n => n
+                                   | TPTP_Proof.ParentWithDetails (n, _) => n
+                             in
+                               case AList.lookup (op =) fms' parent_nodename of
+                                   NONE => false
+                                 | SOME _ => true
+                             end
+                         in
+                           TPTP_Proof.Inference (inference_name,
+                                                 useful_infos,
+                                                 filter is_node_in_fms' parent_infos)
+                         end
+               in
+                   (n,
+                    {role = #role node_info,
+                     fmla = #fmla node_info,
+                     source_inf_opt = SOME source_info'}) :: fms'
+               end)
+       nodes
+       []
+  in
+    if null paths then (thy', fms) else
+      (thy',
+       hd(*FIXME could pick path based on length, or some notion of "difficulty"*) paths
+       |> path_to_fms)
+  end
+
+
+(*** Main functions ***)
+
+(*interpret proof*)
+fun import_thm cautious path_prefixes file_name
+ (on_load : proof_annotation -> theory -> (proof_annotation * theory)) thy =
+  let
+    val prob_name =
+      Path.base file_name
+      |> Path.implode
+      |> TPTP_Problem_Name.parse_problem_name
+    val thy1 = TPTP_Interpret.import_file cautious path_prefixes file_name [] [] thy
+    val fms = get_fmlas_of_prob thy1 prob_name
+  in
+    if List.null fms then
+      (warning ("File " ^ Path.implode file_name ^ " appears empty!");
+       TPTP_Reconstruction_Data.map (cons ((prob_name, empty_pannot prob_name))) thy1)
+    else
+      let
+        val defn_equations =
+          List.filter (fn (_, role, _, _) => role = TPTP_Syntax.Role_Definition) fms
+          |> map (fn (node, _, t, _) =>
+               (node,
+                get_defn_components t
+                |> mk_bind_eq prob_name []))
+        val axioms =
+          List.filter (fn (_, role, _, _) => role = TPTP_Syntax.Role_Axiom) fms
+          |> map (fn (node, _, t, _) =>
+               (node,
+                mk_bind_ax prob_name node t))
+
+        (*add definitions and axioms to the theory*)
+        val thy2 =
+          fold
+           (fn bnd => fn thy =>
+              let
+                val ((name, thm), thy') = Thm.add_axiom_global bnd thy
+              in Global_Theory.add_thm ((#1 bnd, thm), []) thy' |> #2 end)
+           (map snd defn_equations @ map snd axioms)
+          thy1
+
+        (*apply global proof transformations*)
+        val (thy3, pre_pannot) : theory * proof_annotation =
+          transf_pannot
+           (prune_redundant_splits prob_name thy2
+            #-> interpolate_binds prob_name
+            #-> preprocess_splits prob_name
+            #-> drop_repeated_quantification)
+           {problem_name = prob_name,
+            skolem_defs = [],
+            defs = map (apsnd fst) defn_equations,
+            axs = map (apsnd fst) axioms,
+            meta = map (fn (n, r, t, info) => (n, {role=r, fmla=t, source_inf_opt=info})) fms}
+
+        (*store pannot*)
+        val thy4 = TPTP_Reconstruction_Data.map (cons ((prob_name, pre_pannot))) thy3
+
+        (*run hook, which might result in changed pannot and theory*)
+        val (pannot, thy5) = on_load pre_pannot thy4
+
+      (*store the most recent pannot*)
+      in TPTP_Reconstruction_Data.map (cons ((prob_name, pannot))) thy5 end
+  end
+
+(*This has been disabled since it requires a hook to be specified to use "import_thm"
+val _ =
+  Outer_Syntax.improper_command @{command_spec "import_leo2_proof"} "import TPTP proof"
+    (Parse.path >> (fn name =>
+      Toplevel.theory (fn thy =>
+       let val path = Path.explode name
+       in import_thm true [Path.dir path, Path.explode "$TPTP"] path (*FIXME hook needs to be given here*)
+thy end)))
+*)
+
+
+(** Archive **)
+(*FIXME move elsewhere*)
+(*This contains currently unused, but possibly useful, functions written
+  during experimentation, in case they are useful later on*)
+
+(*given a list of rules and a node, return
+  SOME (rule name) if that node's rule name
+  belongs to the list of rules*)
+fun match_rules_of_current (pannot : proof_annotation) rules n =
+  case node_info (#meta pannot) #source_inf_opt n of
+      NONE => NONE
+    | SOME (TPTP_Proof.File _) => NONE
+    | SOME (TPTP_Proof.Inference (rule_name, _ : TPTP_Proof.useful_info_as list, _)) =>
+        if member (op =) rules rule_name then SOME rule_name else NONE
+
+(*given a node and a list of rules, determine
+  whether all the rules can be matched to
+  parent nodes. If nonstrict then there may be
+  more parents than given rules.*)
+fun match_rules_of_immediate_previous (pannot : proof_annotation) strict rules n =
+  case node_info (#meta pannot) #source_inf_opt n of
+      NONE => null rules
+    | SOME (TPTP_Proof.File _) => null rules
+    | SOME (TPTP_Proof.Inference (rule_name, _ : TPTP_Proof.useful_info_as list, parent_infos)) =>
+        let
+          val matched_rules : string option list =
+            map (match_rules_of_current pannot rules)
+                (dest_parent_infos true (#meta pannot) parent_infos |> map #name)
+        in
+          if strict andalso member (op =) matched_rules NONE then false
+          else
+            (*check that all the rules were matched*)
+            fold
+              (fn (rule : string) => fn (st, matches : string option list) =>
+                 if not st then (st, matches)
+                 else
+                   let
+                     val idx = find_index (fn match => SOME rule = match) matches
+                   in
+                     if idx < 0 then (false, matches)
+                     else
+                       (st, nth_drop idx matches)
+                   end)
+             rules
+             (true, matched_rules)
+            |> #1 (*discard the other info*)
+        end
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML	Wed Feb 19 15:57:02 2014 +0000
@@ -0,0 +1,795 @@
+(*  Title:      HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+Collection of general functions used in the reconstruction module.
+*)
+
+signature TPTP_RECONSTRUCT_LIBRARY =
+sig
+  exception BREAK_LIST
+  val break_list : 'a list -> 'a * 'a list
+  val break_seq : 'a Seq.seq -> 'a * 'a Seq.seq
+  exception MULTI_ELEMENT_LIST
+  val cascaded_filter_single : bool -> ('a list -> 'a list) list -> 'a list -> 'a option
+  val concat_between : 'a list list -> ('a option * 'a option) -> 'a list
+  exception DIFF_TYPE of typ * typ
+  exception DIFF of term * term
+  val diff :
+     theory ->
+     term * term -> (term * term) list * (typ * typ) list
+  exception DISPLACE_KV
+  val displace_kv : ''a -> (''a * 'b) list -> (''a * 'b) list
+  val enumerate : int -> 'a list -> (int * 'a) list
+  val fold_options : 'a option list -> 'a list
+  val find_and_remove : ('a -> bool) -> 'a list -> 'a * 'a list
+  val lift_option : ('a -> 'b) -> 'a option -> 'b option
+  val list_diff : ''a list -> ''a list -> ''a list
+  val list_prod : 'a list list -> 'a list -> 'a list -> 'a list list
+  val permute : ''a list -> ''a list list
+  val prefix_intersection_list :
+     ''a list -> ''a list -> ''a list
+  val repeat_until_fixpoint : (''a -> ''a) -> ''a -> ''a
+  val switch : ('a -> 'b -> 'c) -> 'b -> 'a -> 'c
+  val zip_amap :
+       'a list ->
+       'b list ->
+       ('a * 'b) list -> ('a * 'b) list * ('a list * 'b list)
+
+  val consts_in : term -> term list
+  val head_quantified_variable :
+     int -> thm -> (string * typ) option
+  val push_allvar_in : string -> term -> term
+  val strip_top_All_var : term -> (string * typ) * term
+  val strip_top_All_vars : term -> (string * typ) list * term
+  val strip_top_all_vars :
+     (string * typ) list -> term -> (string * typ) list * term
+  val trace_tac' :
+     string ->
+     ('a -> thm -> 'b Seq.seq) -> 'a -> thm -> 'b Seq.seq
+  val try_dest_Trueprop : term -> term
+
+  val type_devar : ((indexname * sort) * typ) list -> term -> term
+  val diff_and_instantiate : Proof.context -> thm -> term -> term -> thm
+
+  val batter : int -> tactic
+  val break_hypotheses : int -> tactic
+  val clause_breaker : int -> tactic
+  (* val dist_all_and_tac : Proof.context -> int -> tactic *)(*FIXME unused*)
+  val reassociate_conjs_tac : Proof.context -> int -> tactic
+
+  val ASAP : (int -> tactic) -> (int -> tactic) -> int -> tactic
+  val COND' :
+     ('a -> thm -> bool) ->
+     ('a -> tactic) -> ('a -> tactic) -> 'a -> tactic
+
+  val TERMFUN :
+     (term list * term -> 'a) -> int option -> thm -> 'a list
+  val TERMPRED :
+     (term -> bool) ->
+     (term -> bool) -> int option -> thm -> bool
+
+  val guided_abstract :
+     bool -> term -> term -> ((string * typ) * term) * term list
+  val abstract :
+     term list -> term -> ((string * typ) * term) list * term
+end
+
+structure TPTP_Reconstruct_Library : TPTP_RECONSTRUCT_LIBRARY =
+struct
+
+(*zip as much as possible*)
+fun zip_amap [] ys acc = (acc, ([], ys))
+  | zip_amap xs [] acc = (acc, (xs, []))
+  | zip_amap (x :: xs) (y :: ys) acc =
+      zip_amap xs ys ((x, y) :: acc);
+
+(*Pair a list up with the position number of each element,
+  starting from n*)
+fun enumerate n ls =
+  let
+    fun enumerate' [] _ acc = acc
+      | enumerate' (x :: xs) n acc = enumerate' xs (n + 1) ((n, x) :: acc)
+  in
+    enumerate' ls n []
+    |> rev
+  end
+
+(*
+enumerate 0 [];
+enumerate 0 ["a", "b", "c"];
+*)
+
+(*List subtraction*)
+fun list_diff l1 l2 =
+  List.filter (fn x => List.all (fn y => x <> y) l2) l1
+
+val _ = @{assert}
+  (list_diff [1,2,3] [2,4] = [1, 3])
+
+(* [a,b] times_list [c,d] gives [[a,c,d], [b,c,d]] *)
+fun list_prod acc [] _ = rev acc
+  | list_prod acc (x :: xs) ys = list_prod ((x :: ys) :: acc) xs ys
+
+fun repeat_until_fixpoint f x =
+  let
+    val x' = f x
+  in
+    if x = x' then x else repeat_until_fixpoint f x'
+  end
+
+(*compute all permutations of a list*)
+fun permute l =
+  let
+    fun permute' (l, []) = [(l, [])]
+      | permute' (l, xs) =
+          map (fn x => (x :: l, filter (fn y => y <> x) xs)) xs
+          |> map permute'
+          |> List.concat
+  in
+    permute' ([], l)
+    |> map fst
+  end
+(*
+permute [1,2,3];
+permute ["A", "B"]
+*)
+
+(*this exception is raised when the pair we wish to displace
+  isn't found in the association list*)
+exception DISPLACE_KV;
+(*move a key-value pair, determined by the k, to the beginning of
+  an association list. it moves the first occurrence of a pair
+  keyed by "k"*)
+local
+  fun fold_fun k (kv as (k', v)) (l, buff) =
+    if is_some buff then (kv :: l, buff)
+    else
+      if k = k' then
+        (l, SOME kv)
+      else
+        (kv :: l, buff)
+in
+  (*"k" is the key value of the pair we wish to displace*)
+  fun displace_kv k alist =
+    let
+      val (pre_alist, kv) = fold (fold_fun k) alist ([], NONE)
+    in
+      if is_some kv then
+        the kv :: rev pre_alist
+      else raise DISPLACE_KV
+    end
+end
+
+(*Given two lists, it generates a new list where
+  the intersection of the lists forms the prefix
+  of the new list.*)
+local
+  fun prefix_intersection_list' (acc_pre, acc_pro) l1 l2 =
+    if null l1 then
+      List.rev acc_pre @ List.rev acc_pro
+    else if null l2 then
+      List.rev acc_pre @ l1 @ List.rev acc_pro
+    else
+      let val l1_hd = hd l1
+      in
+        prefix_intersection_list'
+         (if member (op =) l2 l1_hd then
+            (l1_hd :: acc_pre, acc_pro)
+          else
+           (acc_pre, l1_hd :: acc_pro))
+         (tl l1) l2
+      end
+in
+  fun prefix_intersection_list l1 l2 = prefix_intersection_list' ([], []) l1 l2
+end;
+
+val _ = @{assert}
+  (prefix_intersection_list [1,2,3,4,5] [1,3,5] = [1, 3, 5, 2, 4]);
+
+val _ = @{assert}
+  (prefix_intersection_list [1,2,3,4,5] [] = [1,2,3,4,5]);
+
+val _ = @{assert}
+  (prefix_intersection_list [] [1,3,5] = [])
+
+fun switch f y x = f x y
+
+(*Given a value of type "'a option list", produce
+  a value of type "'a list" by dropping the NONE elements
+  and projecting the SOME elements.*)
+fun fold_options opt_list =
+  fold
+   (fn x => fn l => if is_some x then the x :: l else l)
+   opt_list
+   [];
+
+val _ = @{assert}
+  ([2,0,1] =
+   fold_options [NONE, SOME 1, NONE, SOME 0, NONE, NONE, SOME 2]);
+
+fun lift_option (f : 'a -> 'b) (x_opt : 'a option) : 'b option =
+  case x_opt of
+      NONE => NONE
+    | SOME x => SOME (f x)
+
+fun break_seq x = (Seq.hd x, Seq.tl x)
+
+exception BREAK_LIST
+fun break_list (x :: xs) = (x, xs)
+  | break_list _ = raise BREAK_LIST
+
+exception MULTI_ELEMENT_LIST
+(*Try a number of predicates, in order, to find a single element.
+  Predicates are expected to either return an empty list or a
+  singleton list. If strict=true and list has more than one element,
+  then raise an exception. Otherwise try a new predicate.*)
+fun cascaded_filter_single strict preds l =
+  case preds of
+      [] => NONE
+    | (p :: ps) =>
+      case p l of
+          [] => cascaded_filter_single strict ps l
+        | [x] => SOME x
+        | l =>
+            if strict then raise MULTI_ELEMENT_LIST
+            else cascaded_filter_single strict ps l
+
+(*concat but with optional before-and-after delimiters*)
+fun concat_between [] _ = []
+  | concat_between [l] _ = l
+  | concat_between (l :: ls) (seps as (bef, aft)) =
+    let
+      val pre = if is_some bef then the bef :: l else l
+      val mid = if is_some aft then [the aft] else []
+      val post = concat_between ls seps
+    in
+      pre @ mid @ post
+    end
+
+(*Given a list, find an element satisfying pred, and return
+  a pair consisting of that element and the list minus the element.*)
+fun find_and_remove pred l =
+  find_index pred l
+  |> switch chop l
+  |> apsnd break_list
+  |> (fn (xs, (y, ys)) => (y, xs @ ys))
+
+val _ = @{assert} (find_and_remove (curry (op =) 3) [0,1,2,3,4,5] = (3, [0,1,2,4,5]))
+
+
+(** Functions on terms **)
+
+(*Extract the forall-prefix of a term, and return a pair consisting of the prefix
+  and the body*)
+local
+  (*Strip off HOL's All combinator if it's at the toplevel*)
+  fun try_dest_All (Const (@{const_name HOL.All}, _) $ t) = t
+    | try_dest_All (Const (@{const_name HOL.Trueprop}, _) $ t) = try_dest_All t
+    | try_dest_All t = t
+
+  val _ = @{assert}
+    ((@{term "! x. (! y. P) = True"}
+      |> try_dest_All
+      |> Term.strip_abs_vars)
+     = [("x", @{typ "'a"})])
+
+  val _ = @{assert}
+    ((@{prop "! x. (! y. P) = True"}
+      |> try_dest_All
+      |> Term.strip_abs_vars)
+     = [("x", @{typ "'a"})])
+
+  fun strip_top_All_vars' once acc t =
+    let
+      val t' = try_dest_All t
+      val var =
+        try (Term.strip_abs_vars #> hd) t'
+
+      fun strip v t =
+        (v, subst_bounds ([Free v], Term.strip_abs_body t))
+    in
+      if t' = t orelse is_none var then (acc, t)
+      else
+        let
+          val (v, t) = strip (the var) t'
+          val acc' = v :: acc
+        in
+          if once then (acc', t)
+          else strip_top_All_vars' once acc' t
+        end
+    end
+in
+  fun strip_top_All_vars t = strip_top_All_vars' false [] t
+
+val _ =
+  let
+    val answer =
+      ([("x", @{typ "'a"})],
+       HOLogic.all_const @{typ "'a"} $
+        (HOLogic.eq_const @{typ "'a"} $
+         Free ("x", @{typ "'a"})))
+  in
+    @{assert}
+      ((@{term "! x. All (op = x)"}
+        |> strip_top_All_vars)
+       = answer)
+  end
+
+  (*like strip_top_All_vars, but peels a single variable off, instead of all of them*)
+  fun strip_top_All_var t =
+    strip_top_All_vars' true [] t
+    |> apfst the_single
+end
+
+(*like strip_top_All_vars but for "all" instead of "All"*)
+fun strip_top_all_vars acc t =
+  if Logic.is_all t then
+    let
+      val (v, t') = Logic.dest_all t
+      (*bound instances in t' are replaced with free vars*)
+    in
+      strip_top_all_vars (v :: acc) t'
+    end
+  else (acc, (*variables are returned in FILO order*)
+        t)
+
+(*given a term "t"
+    ! X Y Z. t'
+  then then "push_allvar_in "X" t" will give
+    ! Y Z X. t'
+*)
+fun push_allvar_in v t =
+  let
+    val (vs, t') = strip_top_All_vars t
+    val vs' = displace_kv v vs
+  in
+    fold (fn (v, ty) => fn t =>
+      HOLogic.mk_all (v, ty, t)) vs' t'
+  end
+
+(*Lists all consts in a term, uniquely*)
+fun consts_in (Const c) = [Const c]
+  | consts_in (Free _) = []
+  | consts_in (Var _) = []
+  | consts_in (Bound _) = []
+  | consts_in (Abs (_, _, t)) = consts_in t
+  | consts_in (t1 $ t2) = union (op =) (consts_in t1) (consts_in t2);
+
+exception DIFF of term * term
+exception DIFF_TYPE of typ * typ
+(*This carries out naive form of matching.  It "diffs" two formulas,
+  to create a function which maps (schematic or non-schematic)
+  variables to terms.  The first argument is the more "general" term.
+  The second argument is used to find the "image" for the variables in
+  the first argument which don't appear in the second argument.
+
+  Note that the list that is returned might have duplicate entries.
+  It's not checked to see if the same variable maps to different
+  values -- that should be regarded as an error.*)
+fun diff thy (initial as (t_gen, t)) =
+  let
+    fun diff_ty acc [] = acc
+      | diff_ty acc ((pair as (ty_gen, ty)) :: ts) =
+          case pair of
+              (Type (s1, ty_gens1), Type (s2, ty_gens2)) =>
+                if s1 <> s2 orelse
+                 length ty_gens1 <> length ty_gens2 then
+                  raise (DIFF (t_gen, t))
+                else
+                  diff_ty acc
+                   (ts @ ListPair.zip (ty_gens1, ty_gens2))
+            | (TFree (s1, so1), TFree (s2, so2)) =>
+                if s1 <> s2 orelse
+                 not (Sign.subsort thy (so2, so1)) then
+                  raise (DIFF (t_gen, t))
+                else
+                  diff_ty acc ts
+            | (TVar (idx1, so1), TVar (idx2, so2)) =>
+                if idx1 <> idx2 orelse
+                 not (Sign.subsort thy (so2, so1)) then
+                  raise (DIFF (t_gen, t))
+                else
+                  diff_ty acc ts
+            | (TFree _, _) => diff_ty (pair :: acc) ts
+            | (TVar _, _) => diff_ty (pair :: acc) ts
+            | _ => raise (DIFF_TYPE pair)
+
+    fun diff' (acc as (acc_t, acc_ty)) (pair as (t_gen, t)) ts =
+      case pair of
+          (Const (s1, ty1), Const (s2, ty2)) =>
+            if s1 <> s2 orelse
+             not (Sign.typ_instance thy (ty2, ty1)) then
+              raise (DIFF (t_gen, t))
+            else
+              diff_probs acc ts
+        | (Free (s1, ty1), Free (s2, ty2)) =>
+            if s1 <> s2 orelse
+             not (Sign.typ_instance thy (ty2, ty1)) then
+              raise (DIFF (t_gen, t))
+            else
+              diff_probs acc ts
+        | (Var (idx1, ty1), Var (idx2, ty2)) =>
+            if idx1 <> idx2 orelse
+             not (Sign.typ_instance thy (ty2, ty1)) then
+              raise (DIFF (t_gen, t))
+            else
+              diff_probs acc ts
+        | (Bound i1, Bound i2) =>
+            if i1 <> i2 then
+              raise (DIFF (t_gen, t))
+            else
+              diff_probs acc ts
+        | (Abs (s1, ty1, t1), Abs (s2, ty2, t2)) =>
+            if s1 <> s2 orelse
+             not (Sign.typ_instance thy (ty2, ty1)) then
+              raise (DIFF (t_gen, t))
+            else
+              diff' acc (t1, t2) ts
+        | (ta1 $ ta2, tb1 $ tb2) =>
+            diff_probs acc ((ta1, tb1) :: (ta2, tb2) :: ts)
+
+        (*the particularly important bit*)
+        | (Free (_, ty), _) =>
+            diff_probs
+             (pair :: acc_t,
+              diff_ty acc_ty [(ty, Term.fastype_of t)])
+             ts
+        | (Var (_, ty), _) =>
+            diff_probs
+             (pair :: acc_t,
+              diff_ty acc_ty [(ty, Term.fastype_of t)])
+             ts
+
+        (*everything else is problematic*)
+        | _ => raise (DIFF (t_gen, t))
+
+    and diff_probs acc ts =
+      case ts of
+          [] => acc
+        | (pair :: ts') => diff' acc pair ts'
+  in
+    diff_probs ([], []) [initial]
+  end
+
+(*Abstracts occurrences of "t_sub" in "t", returning a list of
+  abstractions of "t" with a Var at each occurrence of "t_sub".
+  If "strong=true" then it uses strong abstraction (i.e., replaces
+   all occurrnces of "t_sub"), otherwise it uses weak abstraction
+   (i.e., replaces the occurrences one at a time).
+  NOTE there are many more possibilities between strong and week.
+    These can be enumerated by abstracting based on the powerset
+    of occurrences (minus the null element, which would correspond
+    to "t").
+*)
+fun guided_abstract strong t_sub t =
+  let
+    val varnames = Term.add_frees t [] |> map #1
+    val prefixK = "v"
+    val freshvar =
+      let
+        fun find_fresh i =
+          let
+            val varname = prefixK ^ Int.toString i
+          in
+            if member (op =) varnames varname then
+              find_fresh (i + 1)
+            else
+              (varname, fastype_of t_sub)
+          end
+      in
+        find_fresh 0
+      end
+
+    fun guided_abstract' t =
+      case t of
+          Abs (s, ty, t') =>
+            if t = t_sub then [Free freshvar]
+            else
+              (map (fn t' => Abs (s, ty, t'))
+               (guided_abstract' t'))
+        | t1 $ t2 =>
+            if t = t_sub then [Free freshvar]
+            else
+                (map (fn t' => t' $ t2)
+                  (guided_abstract' t1)) @
+                (map (fn t' => t1 $ t')
+                  (guided_abstract' t2))
+        | _ =>
+            if t = t_sub then [Free freshvar]
+            else [t]
+
+    fun guided_abstract_strong' t =
+      let
+        fun continue t = guided_abstract_strong' t
+          |> (fn x => if null x then t
+                else the_single x)
+      in
+        case t of
+            Abs (s, ty, t') =>
+              if t = t_sub then [Free freshvar]
+              else
+                [Abs (s, ty, continue t')]
+          | t1 $ t2 =>
+              if t = t_sub then [Free freshvar]
+              else
+                [continue t1 $ continue t2]
+          | _ =>
+              if t = t_sub then [Free freshvar]
+              else [t]
+      end
+
+  in
+    ((freshvar, t_sub),
+     if strong then guided_abstract_strong' t
+     else guided_abstract' t)
+  end
+
+(*Carries out strong abstraction of a term guided by a list of
+  other terms.
+  In case some of the latter terms happen to be the same, it
+  only abstracts them once.
+  It returns the abstracted term, together with a map from
+   the fresh names to the terms.*)
+fun abstract ts t =
+  fold_map (apsnd the_single oo (guided_abstract true)) ts t
+  |> (fn (v_and_ts, t') =>
+       let
+         val (vs, ts) = ListPair.unzip v_and_ts
+         val vs' =
+           (* list_diff vs (list_diff (Term.add_frees t' []) vs) *)
+           Term.add_frees t' []
+           |> list_diff vs
+           |> list_diff vs
+         val v'_and_ts =
+           map (fn v =>
+             (v, AList.lookup (op =) v_and_ts v |> the))
+            vs'
+       in
+         (v'_and_ts, t')
+       end)
+
+(*Instantiate type variables in a term, based on a type environment*)
+fun type_devar (tyenv : ((indexname * sort) * typ) list) (t : term) : term =
+  case t of
+      Const (s, ty) => Const (s, Term_Subst.instantiateT tyenv ty)
+    | Free (s, ty) => Free (s, Term_Subst.instantiateT tyenv ty)
+    | Var (idx, ty) => Var (idx, Term_Subst.instantiateT tyenv ty)
+    | Bound _ => t
+    | Abs (s, ty, t') =>
+        Abs (s, Term_Subst.instantiateT tyenv ty, type_devar tyenv t')
+    | t1 $ t2 => type_devar tyenv t1 $ type_devar tyenv t2
+
+(*Take a "diff" between an (abstract) thm's term, and another term
+  (the latter is an instance of the form), then instantiate the
+  abstract theorem. This is a way of turning the latter term into
+  a theorem, but without exposing the proof-search functions to
+  complex terms.
+  In addition to the abstract thm ("scheme_thm"), this function is
+  also supplied with the (sub)term of the abstract thm ("scheme_t")
+  we want to use in the diff, in case only part of "scheme_t"
+  might be needed (not the whole "prop_of scheme_thm")*)
+fun diff_and_instantiate ctxt scheme_thm scheme_t instance_t =
+  let
+    val thy = Proof_Context.theory_of ctxt
+
+    val (term_pairing, type_pairing) =
+      diff thy (scheme_t, instance_t)
+
+    (*valuation of type variables*)
+    val typeval = map (pairself (ctyp_of thy)) type_pairing
+
+    val typeval_env =
+      map (apfst dest_TVar) type_pairing
+    (*valuation of term variables*)
+    val termval =
+      map (apfst (type_devar typeval_env)) term_pairing
+      |> map (pairself (cterm_of thy))
+  in
+    Thm.instantiate (typeval, termval) scheme_thm
+  end
+
+(*FIXME this is bad form?*)
+val try_dest_Trueprop = perhaps (try HOLogic.dest_Trueprop)
+
+
+(** Some tacticals **)
+
+(*Lift COND to be parametrised by subgoal number*)
+fun COND' sat' tac'1 tac'2 i =
+  COND (sat' i) (tac'1 i) (tac'2 i)
+
+(*Apply simplification ("wittler") as few times as possible
+  before being able to apply a tactic ("tac").
+  This is like a lazy version of REPEAT, since it attempts
+  to REPEAT a tactic the smallest number times as possible,
+  to make some other tactic succeed subsequently.*)
+fun ASAP wittler (tac : int -> tactic) (i : int) = fn st =>
+  let
+    val tac_result = tac i st
+    val pulled_tac_result = Seq.pull tac_result
+    val tac_failed =
+      is_none pulled_tac_result orelse
+       not (has_fewer_prems 1 (fst (the pulled_tac_result)))
+  in
+    if tac_failed then (wittler THEN' ASAP wittler tac) i st
+    else tac_result
+  end
+
+
+(** Some tactics **)
+
+val break_hypotheses =
+ ((REPEAT_DETERM o etac @{thm conjE})
+  THEN' (REPEAT_DETERM o etac @{thm disjE})
+ ) #> CHANGED
+
+(*Prove subgoals of form A ==> B1 | ... | A | ... | Bn*)
+val clause_breaker =
+  (REPEAT o (resolve_tac [@{thm "disjI1"}, @{thm "disjI2"}, @{thm "conjI"}]))
+  THEN'  atac
+
+(*
+  Refines a subgoal have the form:
+    A1 ... An ==> B1 | ... | Aj | ... | Bi | ... | Ak | ...
+  into multiple subgoals of the form:
+    A'1 ==> B1 | ... | Aj | ... | Bi | ... | Ak | ...
+     :
+    A'm ==> B1 | ... | Aj | ... | Bi | ... | Ak | ...
+  where {A'1 .. A'm} is disjoint from {B1, ..., Aj, ..., Bi, ..., Ak, ...}
+  (and solves the subgoal completely if the first set is empty)
+*)
+val batter =
+  break_hypotheses
+  THEN' K (ALLGOALS (TRY o clause_breaker))
+
+(*Same idiom as ex_expander_tac*)
+fun dist_all_and_tac ctxt i =
+   let
+     val simpset =
+       empty_simpset ctxt
+       |> Simplifier.add_simp
+           @{lemma "! x. P x & Q x \<equiv> (! x. P x) & (! x. Q x)"
+              by (rule eq_reflection, auto)}
+   in
+     CHANGED (asm_full_simp_tac simpset i)
+   end
+
+fun reassociate_conjs_tac ctxt =
+  asm_full_simp_tac
+   (Simplifier.add_simp
+    @{lemma "(A & B) & C == A & B & C" by auto} (*FIXME duplicates @{thm simp_meta(3)}*)
+    (Raw_Simplifier.empty_simpset ctxt))
+  #> CHANGED
+  #> REPEAT_DETERM
+
+
+(** Subgoal analysis **)
+
+(*Given an inference
+        C
+      -----
+        D
+  This function returns "SOME X" if C = "! X. C'".
+  If C has no quantification prefix, then returns NONE.*)
+fun head_quantified_variable i = fn st =>
+  let
+    val thy = Thm.theory_of_thm st
+    val ctxt = Proof_Context.init_global thy
+
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+
+    val hypos =
+      if null gls then []
+      else
+        rpair (i - 1) gls
+        |> uncurry nth
+        |> strip_top_all_vars []
+        |> snd
+        |> Logic.strip_horn
+        |> fst
+
+    fun foralls_of_hd_hypos () =
+      hd hypos
+      |> try_dest_Trueprop
+      |> strip_top_All_vars
+      |> #1
+      |> rev
+
+    val quantified_variables = foralls_of_hd_hypos ()
+  in
+    if null hypos orelse null quantified_variables then NONE
+    else SOME (hd quantified_variables)
+  end
+
+
+(** Builders for goal analysers or transformers **)
+
+(*Lifts function over terms to apply it to subgoals.
+  "fun_over_terms" has type (term list * term -> 'a), where
+  (term list * term) will be the term representations of the
+  hypotheses and conclusion.
+  if i_opt=SOME i then applies fun_over_terms to that
+   subgoal and returns singleton result.
+  otherwise applies fun_over_terms to all subgoals and return
+   list of results.*)
+fun TERMFUN
+ (fun_over_terms : term list * term -> 'a)
+ (i_opt : int option) : thm -> 'a list = fn st =>
+  let
+    val t_raws =
+        Thm.rep_thm st
+        |> #prop
+        |> strip_top_all_vars []
+        |> snd
+        |> Logic.strip_horn
+        |> fst
+  in
+    if null t_raws then []
+    else
+      let
+        val ts =
+          let
+            val stripper =
+              strip_top_all_vars []
+              #> snd
+              #> Logic.strip_horn
+              #> apsnd try_dest_Trueprop
+              #> apfst (map try_dest_Trueprop)
+          in
+            map stripper t_raws
+          end
+      in
+        case i_opt of
+            NONE =>
+              map fun_over_terms ts
+          | SOME i =>
+              nth ts (i - 1)
+              |> fun_over_terms
+              |> single
+      end
+  end
+
+(*Applies a predicate to subgoal(s) conclusion(s)*)
+fun TERMPRED
+ (hyp_pred_over_terms : term -> bool)
+ (conc_pred_over_terms : term -> bool)
+ (i_opt : int option) : thm -> bool = fn st =>
+    let
+      val hyp_results =
+        TERMFUN (fst (*discard hypotheses*)
+                 #> map hyp_pred_over_terms) i_opt st
+      val conc_results =
+        TERMFUN (snd (*discard hypotheses*)
+                 #> conc_pred_over_terms) i_opt st
+      val _ = @{assert} (length hyp_results = length conc_results)
+    in
+      if null hyp_results then true
+      else
+        let
+          val hyps_conjoined =
+            fold (fn a => fn b =>
+              b andalso (List.all (fn x => x) a)) hyp_results true
+          val concs_conjoined =
+            fold (fn a => fn b =>
+              b andalso a) conc_results true
+        in hyps_conjoined andalso concs_conjoined end
+    end
+
+
+(** Tracing **)
+(*If "tac i st" succeeds then msg is printed to "trace" channel*)
+fun trace_tac' msg tac i st =
+  let
+    val thy = Thm.theory_of_thm st
+    val ctxt = Proof_Context.init_global thy
+    val result = tac i st
+  in
+    if Config.get ctxt tptp_trace_reconstruction andalso
+     not (is_none (Seq.pull result)) then
+      (tracing msg; result)
+    else result
+  end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy	Wed Feb 19 15:57:02 2014 +0000
@@ -0,0 +1,2223 @@
+(*  Title:      HOL/TPTP/TPTP_Proof_Reconstruction.thy
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+Proof reconstruction for Leo-II.
+
+TODO:
+  use RemoveRedundantQuantifications instead of the ad hoc use of
+   remove_redundant_quantification_in_lit and remove_redundant_quantification
+*)
+
+theory TPTP_Proof_Reconstruction
+imports TPTP_Parser TPTP_Interpret
+(* keywords "import_leo2_proof" :: thy_decl *) (*FIXME currently unused*)
+begin
+
+
+section "Setup"
+
+ML {*
+  val tptp_unexceptional_reconstruction = Attrib.setup_config_bool @{binding tptp_unexceptional_reconstruction} (K false)
+  fun unexceptional_reconstruction ctxt = Config.get ctxt tptp_unexceptional_reconstruction
+  val tptp_informative_failure = Attrib.setup_config_bool @{binding tptp_informative_failure} (K false)
+  fun informative_failure ctxt = Config.get ctxt tptp_informative_failure
+  val tptp_trace_reconstruction = Attrib.setup_config_bool @{binding tptp_trace_reconstruction} (K false)
+  val tptp_max_term_size = Attrib.setup_config_int @{binding tptp_max_term_size} (K 0) (*0=infinity*)
+
+  fun exceeds_tptp_max_term_size ctxt size =
+    let
+      val max = Config.get ctxt tptp_max_term_size
+    in
+      if max = 0 then false
+      else size > max
+    end
+*}
+
+(*FIXME move to TPTP_Proof_Reconstruction_Test_Units*)
+declare [[
+  tptp_unexceptional_reconstruction = false, (*NOTE should be "false" while testing*)
+  tptp_informative_failure = true
+]]
+
+ML_file "TPTP_Parser/tptp_reconstruct_library.ML"
+ML "open TPTP_Reconstruct_Library"
+ML_file "TPTP_Parser/tptp_reconstruct.ML"
+
+(*FIXME fudge*)
+declare [[
+  blast_depth_limit = 10,
+  unify_search_bound = 5
+]]
+
+
+section "Proof reconstruction"
+text {*There are two parts to proof reconstruction:
+\begin{itemize}
+  \item interpreting the inferences
+  \item building the skeleton, which indicates how to compose
+    individual inferences into subproofs, and then compose the
+    subproofs to give the proof).
+\end{itemize}
+
+One step detects unsound inferences, and the other step detects
+unsound composition of inferences.  The two parts can be weakly
+coupled. They rely on a "proof index" which maps nodes to the
+inference information. This information consists of the (usually
+prover-specific) name of the inference step, and the Isabelle
+formalisation of the inference as a term. The inference interpretation
+then maps these terms into meta-theorems, and the skeleton is used to
+compose the inference-level steps into a proof.
+
+Leo2 operates on conjunctions of clauses. Each Leo2 inference has the
+following form, where Cx are clauses:
+
+           C1 && ... && Cn
+          -----------------
+          C'1 && ... && C'n
+
+Clauses consist of disjunctions of literals (shown as Px below), and might
+have a prefix of !-bound variables, as shown below.
+
+  ! X... { P1 || ... || Pn}
+
+Literals are usually assigned a polarity, but this isn't always the
+case; you can come across inferences looking like this (where A is an
+object-level formula):
+
+             F
+          --------
+          F = true
+
+The symbol "||" represents literal-level disjunction and "&&" is
+clause-level conjunction. Rules will typically lift formula-level
+conjunctions; for instance the following rule lifts object-level
+disjunction:
+
+          {    (A | B) = true    || ... } && ...
+          --------------------------------------
+          { A = true || B = true || ... } && ...
+
+
+Using this setup, efficiency might be gained by only interpreting
+inferences once, merging identical inference steps, and merging
+identical subproofs into single inferences thus avoiding some effort.
+We can also attempt to minimising proof search when interpreting
+inferences.
+
+It is hoped that this setup can target other provers by modifying the
+clause representation to fit them, and adapting the inference
+interpretation to handle the rules used by the prover. It should also
+facilitate composing together proofs found by different provers.
+*}
+
+
+subsection "Instantiation"
+
+lemma polar_allE [rule_format]:
+  "\<lbrakk>(\<forall>x. P x) = True; (P x) = True \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+  "\<lbrakk>(\<exists>x. P x) = False; (P x) = False \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by auto
+
+lemma polar_exE [rule_format]:
+  "\<lbrakk>(\<exists>x. P x) = True; \<And>x. (P x) = True \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+  "\<lbrakk>(\<forall>x. P x) = False; \<And>x. (P x) = False \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by auto
+
+ML {*
+(*This carries out an allE-like rule but on (polarised) literals.
+ Instead of yielding a free variable (which is a hell for the
+ matcher) it seeks to use one of the subgoals' parameters.
+ This ought to be sufficient for emulating extcnf_combined,
+ but note that the complexity of the problem can be enormous.*)
+fun inst_parametermatch_tac ctxt thms i = fn st =>
+  let
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+
+    val parameters =
+      if null gls then []
+      else
+        rpair (i - 1) gls
+        |> uncurry nth
+        |> strip_top_all_vars []
+        |> fst
+        |> map fst (*just get the parameter names*)
+  in
+    if null parameters then no_tac st
+    else
+      let
+        fun instantiate param =
+           (map (eres_inst_tac ctxt [(("x", 0), param)]) thms
+                   |> FIRST')
+        val attempts = map instantiate parameters
+      in
+        (fold (curry (op APPEND')) attempts (K no_tac)) i st
+      end
+  end
+
+(*Attempts to use the polar_allE theorems on a specific subgoal.*)
+fun forall_pos_tac ctxt = inst_parametermatch_tac ctxt @{thms polar_allE}
+*}
+
+ML {*
+(*This is similar to inst_parametermatch_tac, but prefers to
+  match variables having identical names. Logically, this is
+  a hack. But it reduces the complexity of the problem.*)
+fun nominal_inst_parametermatch_tac ctxt thm i = fn st =>
+  let
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+
+    val parameters =
+      if null gls then []
+      else
+        rpair (i - 1) gls
+        |> uncurry nth
+        |> strip_top_all_vars []
+        |> fst
+        |> map fst (*just get the parameter names*)
+  in
+    if null parameters then no_tac st
+    else
+      let
+        fun instantiates param =
+           eres_inst_tac ctxt [(("x", 0), param)] thm
+
+        val quantified_var = head_quantified_variable i st
+      in
+        if is_none quantified_var then no_tac st
+        else
+          if member (op =) parameters (the quantified_var |> fst) then
+            instantiates (the quantified_var |> fst) i st
+          else
+            K no_tac i st
+      end
+  end
+*}
+
+
+subsection "Prefix massaging"
+
+ML {*
+exception NO_GOALS
+
+(*Get quantifier prefix of the hypothesis and conclusion, reorder
+  the hypothesis' quantifiers to have the ones appearing in the
+  conclusion first.*)
+fun canonicalise_qtfr_order ctxt i = fn st =>
+  let
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+  in
+    if null gls then raise NO_GOALS
+    else
+      let
+        val (params, (hyp_clause, conc_clause)) =
+          rpair (i - 1) gls
+          |> uncurry nth
+          |> strip_top_all_vars []
+          |> apsnd Logic.dest_implies
+
+        val (hyp_quants, hyp_body) =
+          HOLogic.dest_Trueprop hyp_clause
+          |> strip_top_All_vars
+          |> apfst rev
+
+        val conc_quants =
+          HOLogic.dest_Trueprop conc_clause
+          |> strip_top_All_vars
+          |> fst
+
+        val new_hyp =
+          (* fold absfree new_hyp_prefix hyp_body *)
+          (*HOLogic.list_all*)
+          fold_rev (fn (v, ty) => fn t => HOLogic.mk_all (v, ty, t))
+           (prefix_intersection_list
+             hyp_quants conc_quants)
+           hyp_body
+          |> HOLogic.mk_Trueprop
+
+         val thm = Goal.prove ctxt [] []
+           (Logic.mk_implies (hyp_clause, new_hyp))
+           (fn _ =>
+              (REPEAT_DETERM (HEADGOAL (rtac @{thm allI})))
+              THEN (REPEAT_DETERM
+                    (HEADGOAL
+                     (nominal_inst_parametermatch_tac ctxt @{thm allE})))
+              THEN HEADGOAL atac)
+      in
+        dtac thm i st
+      end
+    end
+*}
+
+
+subsection "Some general rules and congruences"
+
+(*this isn't an actual rule used in Leo2, but it seems to be
+  applied implicitly during some Leo2 inferences.*)
+lemma polarise: "P ==> P = True" by auto
+
+ML {*
+fun is_polarised t =
+  (TPTP_Reconstruct.remove_polarity true t; true)
+  handle TPTP_Reconstruct.UNPOLARISED _ => false
+
+val polarise_subgoal_hyps =
+  COND' (SOME #> TERMPRED is_polarised (fn _ => true)) (K no_tac) (dtac @{thm polarise})
+*}
+
+lemma simp_meta [rule_format]:
+  "(A --> B) == (~A | B)"
+  "(A | B) | C == A | B | C"
+  "(A & B) & C == A & B & C"
+  "(~ (~ A)) == A"
+  (* "(A & B) == (~ (~A | ~B))" *)
+  "~ (A & B) == (~A | ~B)"
+  "~(A | B) == (~A) & (~B)"
+by auto
+
+
+subsection "Emulation of Leo2's inference rules"
+
+(*this is not included in simp_meta since it would make a mess of the polarities*)
+lemma expand_iff [rule_format]:
+ "((A :: bool) = B) \<equiv> (~ A | B) & (~ B | A)"
+by (rule eq_reflection, auto)
+
+lemma polarity_switch [rule_format]:
+  "(\<not> P) = True \<Longrightarrow> P = False"
+  "(\<not> P) = False \<Longrightarrow> P = True"
+  "P = False \<Longrightarrow> (\<not> P) = True"
+  "P = True \<Longrightarrow> (\<not> P) = False"
+by auto
+
+lemma solved_all_splits: "False = True \<Longrightarrow> False" by simp
+ML {*
+val solved_all_splits_tac =
+  TRY (etac @{thm conjE} 1)
+  THEN rtac @{thm solved_all_splits} 1
+  THEN atac 1
+*}
+
+lemma lots_of_logic_expansions_meta [rule_format]:
+  "(((A :: bool) = B) = True) == (((A \<longrightarrow> B) = True) & ((B \<longrightarrow> A) = True))"
+  "((A :: bool) = B) = False == (((~A) | B) = False) | (((~B) | A) = False)"
+
+  "((F = G) = True) == (! x. (F x = G x)) = True"
+  "((F = G) = False) == (! x. (F x = G x)) = False"
+
+  "(A | B) = True == (A = True) | (B = True)"
+  "(A & B) = False == (A = False) | (B = False)"
+  "(A | B) = False == (A = False) & (B = False)"
+  "(A & B) = True == (A = True) & (B = True)"
+  "(~ A) = True == A = False"
+  "(~ A) = False == A = True"
+  "~ (A = True) == A = False"
+  "~ (A = False) == A = True"
+by (rule eq_reflection, auto)+
+
+(*this is used in extcnf_combined handler*)
+lemma eq_neg_bool: "((A :: bool) = B) = False ==> ((~ (A | B)) | ~ ((~ A) | (~ B))) = False"
+by auto
+
+lemma eq_pos_bool:
+  "((A :: bool) = B) = True ==> ((~ (A | B)) | ~ (~ A | ~ B)) = True"
+  "(A = B) = True \<Longrightarrow> A = True \<or> B = False"
+  "(A = B) = True \<Longrightarrow> A = False \<or> B = True"
+by auto
+
+(*next formula is more versatile than
+    "(F = G) = True \<Longrightarrow> \<forall>x. ((F x = G x) = True)"
+  since it doesn't assume that clause is singleton. After splitqtfr,
+  and after applying allI exhaustively to the conclusion, we can
+  use the existing functions to find the "(F x = G x) = True"
+  disjunct in the conclusion*)
+lemma eq_pos_func: "\<And> x. (F = G) = True \<Longrightarrow> (F x = G x) = True"
+by auto
+
+(*make sure the conclusion consists of just "False"*)
+lemma flip:
+  "((A = True) ==> False) ==> A = False"
+  "((A = False) ==> False) ==> A = True"
+by auto
+
+(*FIXME try to use Drule.equal_elim_rule1 directly for this*)
+lemma equal_elim_rule1: "(A \<equiv> B) \<Longrightarrow> A \<Longrightarrow> B" by auto
+lemmas leo2_rules =
+ lots_of_logic_expansions_meta[THEN equal_elim_rule1]
+
+(*FIXME is there any overlap with lots_of_logic_expansions_meta or leo2_rules?*)
+lemma extuni_bool2 [rule_format]: "(A = B) = False \<Longrightarrow> (A = True) | (B = True)" by auto
+lemma extuni_bool1 [rule_format]: "(A = B) = False \<Longrightarrow> (A = False) | (B = False)" by auto
+lemma extuni_triv [rule_format]: "(A = A) = False \<Longrightarrow> R" by auto
+
+(*Order (of A, B, C, D) matters*)
+lemma dec_commut_eq [rule_format]:
+  "((A = B) = (C = D)) = False \<Longrightarrow> (B = C) = False | (A = D) = False"
+  "((A = B) = (C = D)) = False \<Longrightarrow> (B = D) = False | (A = C) = False"
+by auto
+lemma dec_commut_disj [rule_format]:
+  "((A \<or> B) = (C \<or> D)) = False \<Longrightarrow> (B = C) = False \<or> (A = D) = False"
+by auto
+
+lemma extuni_func [rule_format]: "(F = G) = False \<Longrightarrow> (! X. (F X = G X)) = False" by auto
+
+
+subsection "Emulation: tactics"
+
+ML {*
+(*Instantiate a variable according to the info given in the
+  proof annotation. Through this we avoid having to come up
+  with instantiations during reconstruction.*)
+fun bind_tac ctxt prob_name ordered_binds =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    fun term_to_string t =
+        Print_Mode.with_modes [""]
+          (fn () => Output.output (Syntax.string_of_term ctxt t)) ()
+    val ordered_instances =
+      TPTP_Reconstruct.interpret_bindings prob_name thy ordered_binds []
+      |> map (snd #> term_to_string)
+      |> permute
+
+    (*instantiate a list of variables, order matters*)
+    fun instantiate_vars ctxt vars : tactic =
+      map (fn var =>
+            Rule_Insts.eres_inst_tac ctxt
+             [(("x", 0), var)] @{thm allE} 1)
+          vars
+      |> EVERY
+
+    fun instantiate_tac vars =
+      instantiate_vars ctxt vars
+      THEN (HEADGOAL atac)
+  in
+    HEADGOAL (canonicalise_qtfr_order ctxt)
+    THEN (REPEAT_DETERM (HEADGOAL (rtac @{thm allI})))
+    THEN REPEAT_DETERM (HEADGOAL (nominal_inst_parametermatch_tac ctxt @{thm allE}))
+    (*now only the variable to instantiate should be left*)
+    THEN FIRST (map instantiate_tac ordered_instances)
+  end
+*}
+
+ML {*
+(*Simplification tactics*)
+local
+  fun rew_goal_tac thms ctxt i =
+    rewrite_goal_tac ctxt thms i
+    |> CHANGED
+in
+  val expander_animal =
+    rew_goal_tac (@{thms simp_meta} @ @{thms lots_of_logic_expansions_meta})
+
+  val simper_animal =
+    rew_goal_tac @{thms simp_meta}
+end
+*}
+
+lemma prop_normalise [rule_format]:
+  "(A | B) | C == A | B | C"
+  "(A & B) & C == A & B & C"
+  "A | B == ~(~A & ~B)"
+  "~~ A == A"
+by auto
+ML {*
+(*i.e., break_conclusion*)
+fun flip_conclusion_tac ctxt =
+  let
+    val default_tac =
+      (TRY o CHANGED o (rewrite_goal_tac ctxt @{thms prop_normalise}))
+      THEN' rtac @{thm notI}
+      THEN' (REPEAT_DETERM o etac @{thm conjE})
+      THEN' (TRY o (expander_animal ctxt))
+  in
+    default_tac ORELSE' resolve_tac @{thms flip}
+  end
+*}
+
+
+subsection "Skolemisation"
+
+lemma skolemise [rule_format]:
+  "\<forall> P. (~ (! x. P x)) \<longrightarrow> ~ (P (SOME x. ~ P x))"
+proof -
+  have "\<And> P. (~ (! x. P x)) \<Longrightarrow> ~ (P (SOME x. ~ P x))"
+  proof -
+    fix P
+    assume ption: "~ (! x. P x)"
+    hence a: "? x. ~ P x" by force
+
+    have hilbert : "\<And> P. (? x. P x) \<Longrightarrow> (P (SOME x. P x))"
+    proof -
+      fix P
+      assume "(? x. P x)"
+      thus "(P (SOME x. P x))"
+        apply auto
+        apply (rule someI)
+        apply auto
+        done
+    qed
+
+    from a show "~ P (SOME x. ~ P x)"
+    proof -
+      assume "? x. ~ P x"
+      hence "\<not> P (SOME x. \<not> P x)" by (rule hilbert)
+      thus ?thesis .
+    qed
+  qed
+  thus ?thesis by blast
+qed
+
+lemma polar_skolemise [rule_format]:
+  "\<forall> P. (! x. P x) = False \<longrightarrow> (P (SOME x. ~ P x)) = False"
+proof -
+  have "\<And> P. (! x. P x) = False \<Longrightarrow> (P (SOME x. ~ P x)) = False"
+  proof -
+    fix P
+    assume ption: "(! x. P x) = False"
+    hence "\<not> (\<forall> x. P x)" by force
+    hence "\<not> All P" by force
+    hence "\<not> (P (SOME x. \<not> P x))" by (rule skolemise)
+    thus "(P (SOME x. \<not> P x)) = False" by force
+  qed
+  thus ?thesis by blast
+qed
+
+lemma leo2_skolemise [rule_format]:
+  "\<forall> P sk. (! x. P x) = False \<longrightarrow> (sk = (SOME x. ~ P x)) \<longrightarrow> (P sk) = False"
+by (clarify, rule polar_skolemise)
+
+lemma lift_forall [rule_format]:
+  "!! x. (! x. A x) = True ==> (A x) = True"
+  "!! x. (? x. A x) = False ==> (A x) = False"
+by auto
+lemma lift_exists [rule_format]:
+  "\<lbrakk>(All P) = False; sk = (SOME x. \<not> P x)\<rbrakk> \<Longrightarrow> P sk = False"
+  "\<lbrakk>(Ex P) = True; sk = (SOME x. P x)\<rbrakk> \<Longrightarrow> P sk = True"
+apply (drule polar_skolemise, simp)
+apply (simp, drule someI_ex, simp)
+done
+
+ML {*
+(*FIXME LHS should be constant. Currently allow variables for testing. Probably should still allow Vars (but not Frees) since they'll act as intermediate values*)
+fun conc_is_skolem_def t =
+  case t of
+      Const (@{const_name HOL.eq}, _) $ t' $ (Const (@{const_name Hilbert_Choice.Eps}, _) $ _) =>
+      let
+        val (h, args) =
+          strip_comb t'
+          |> apfst (strip_abs #> snd #> strip_comb #> fst)
+        val get_const_name = dest_Const #> fst
+        val h_property =
+          is_Free h orelse
+          is_Var h orelse
+          (is_Const h
+           andalso (get_const_name h <> get_const_name @{term HOL.Ex})
+           andalso (get_const_name h <> get_const_name @{term HOL.All})
+           andalso (h <> @{term Hilbert_Choice.Eps})
+           andalso (h <> @{term HOL.conj})
+           andalso (h <> @{term HOL.disj})
+           andalso (h <> @{term HOL.eq})
+           andalso (h <> @{term HOL.implies})
+           andalso (h <> @{term HOL.The})
+           andalso (h <> @{term HOL.Ex1})
+           andalso (h <> @{term HOL.Not})
+           andalso (h <> @{term HOL.iff})
+           andalso (h <> @{term HOL.not_equal}))
+        val args_property =
+          fold (fn t => fn b =>
+           b andalso is_Free t) args true
+      in
+        h_property andalso args_property
+      end
+    | _ => false
+*}
+
+ML {*
+(*Hack used to detect if a Skolem definition, with an LHS Var, has had the LHS instantiated into an unacceptable term.*)
+fun conc_is_bad_skolem_def t =
+  case t of
+      Const (@{const_name HOL.eq}, _) $ t' $ (Const (@{const_name Hilbert_Choice.Eps}, _) $ _) =>
+      let
+        val (h, args) = strip_comb t'
+        val get_const_name = dest_Const #> fst
+        val const_h_test =
+          if is_Const h then
+            (get_const_name h = get_const_name @{term HOL.Ex})
+             orelse (get_const_name h = get_const_name @{term HOL.All})
+             orelse (h = @{term Hilbert_Choice.Eps})
+             orelse (h = @{term HOL.conj})
+             orelse (h = @{term HOL.disj})
+             orelse (h = @{term HOL.eq})
+             orelse (h = @{term HOL.implies})
+             orelse (h = @{term HOL.The})
+             orelse (h = @{term HOL.Ex1})
+             orelse (h = @{term HOL.Not})
+             orelse (h = @{term HOL.iff})
+             orelse (h = @{term HOL.not_equal})
+          else true
+        val h_property =
+          not (is_Free h) andalso
+          not (is_Var h) andalso
+          const_h_test
+        val args_property =
+          fold (fn t => fn b =>
+           b andalso is_Free t) args true
+      in
+        h_property andalso args_property
+      end
+    | _ => false
+*}
+
+ML {*
+fun get_skolem_conc t =
+  let
+    val t' =
+      strip_top_all_vars [] t
+      |> snd
+      |> try_dest_Trueprop
+  in
+    case t' of
+        Const (@{const_name HOL.eq}, _) $ t' $ (Const (@{const_name Hilbert_Choice.Eps}, _) $ _) => SOME t'
+      | _ => NONE
+  end
+
+fun get_skolem_conc_const t =
+  lift_option
+   (fn t' =>
+     head_of t'
+     |> strip_abs_body
+     |> head_of
+     |> dest_Const)
+   (get_skolem_conc t)
+*}
+
+(*
+Technique for handling quantifiers:
+  Principles:
+  * allE should always match with a !!
+  * exE should match with a constant,
+     or bind a fresh !! -- currently not doing the latter since it never seems to arised in normal Leo2 proofs.
+*)
+
+ML {*
+fun forall_neg_tac candidate_consts ctxt i = fn st =>
+  let
+    val thy = Proof_Context.theory_of ctxt
+
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+
+    val parameters =
+      if null gls then ""
+      else
+        rpair (i - 1) gls
+        |> uncurry nth
+        |> strip_top_all_vars []
+        |> fst
+        |> map fst (*just get the parameter names*)
+        |> (fn l =>
+              if null l then ""
+              else
+                space_implode " " l
+                |> pair " "
+                |> op ^)
+
+  in
+    if null gls orelse null candidate_consts then no_tac st
+    else
+      let
+        fun instantiate const_name =
+          dres_inst_tac ctxt [(("sk", 0), const_name ^ parameters)] @{thm leo2_skolemise}
+        val attempts = map instantiate candidate_consts
+      in
+        (fold (curry (op APPEND')) attempts (K no_tac)) i st
+      end
+  end
+*}
+
+ML {*
+exception SKOLEM_DEF of term (*The tactic wasn't pointed at a skolem definition*)
+exception NO_SKOLEM_DEF of (*skolem const name*)string * Binding.binding * term (*The tactic could not find a skolem definition in the theory*)
+fun absorb_skolem_def ctxt prob_name_opt i = fn st =>
+  let
+    val thy = Proof_Context.theory_of ctxt
+
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+
+    val conclusion =
+      if null gls then
+        (*this should never be thrown*)
+        raise NO_GOALS
+      else
+        rpair (i - 1) gls
+        |> uncurry nth
+        |> strip_top_all_vars []
+        |> snd
+        |> Logic.strip_horn
+        |> snd
+
+    fun skolem_const_info_of t =
+      case t of
+          Const (@{const_name HOL.Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ t' $ (Const (@{const_name Hilbert_Choice.Eps}, _) $ _)) =>
+          head_of t'
+          |> strip_abs_body (*since in general might have a skolem term, so we want to rip out the prefixing lambdas to get to the constant (which should be at head position)*)
+          |> head_of
+          |> dest_Const
+        | _ => raise SKOLEM_DEF t
+
+    val const_name =
+      skolem_const_info_of conclusion
+      |> fst
+
+    val def_name = const_name ^ "_def"
+
+    val bnd_def = (*FIXME consts*)
+      const_name
+      |> space_implode "." o tl o space_explode "." (*FIXME hack to drop theory-name prefix*)
+      |> Binding.qualified_name
+      |> Binding.suffix_name "_def"
+
+    val bnd_name =
+      case prob_name_opt of
+          NONE => bnd_def
+        | SOME prob_name =>
+(*            Binding.qualify false
+             (TPTP_Problem_Name.mangle_problem_name prob_name)
+*)
+             bnd_def
+
+    val thm =
+      if Name_Space.defined_entry (Theory.axiom_space thy) def_name then
+        Thm.axiom thy def_name
+      else
+        if is_none prob_name_opt then
+          (*This mode is for testing, so we can be a bit
+            looser with theories*)
+          Thm.add_axiom_global (bnd_name, conclusion) thy
+          |> fst |> snd
+        else
+          raise (NO_SKOLEM_DEF (def_name, bnd_name, conclusion))
+  in
+    rtac (Drule.export_without_context thm) i st
+  end
+  handle SKOLEM_DEF _ => no_tac st
+*}
+
+ML {*
+(*
+In current system, there should only be 2 subgoals: the one where
+the skolem definition is being built (with a Var in the LHS), and the other subgoal using Var.
+*)
+(*arity must be greater than 0. if arity=0 then
+  there's no need to use this expensive matching.*)
+fun find_skolem_term ctxt consts_candidate arity = fn st =>
+  let
+    val _ = @{assert} (arity > 0)
+
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+
+    (*extract the conclusion of each subgoal*)
+    val conclusions =
+      if null gls then
+        raise NO_GOALS
+      else
+        map (strip_top_all_vars [] #> snd #> Logic.strip_horn #> snd) gls
+        (*Remove skolem-definition conclusion, to avoid wasting time analysing it*)
+        |> List.filter (try_dest_Trueprop #> conc_is_skolem_def #> not)
+        (*There should only be a single goal*) (*FIXME this might not always be the case, in practice*)
+        (* |> tap (fn x => @{assert} (is_some (try the_single x))) *)
+
+    (*look for subterms headed by a skolem constant, and whose
+      arguments are all parameter Vars*)
+    fun get_skolem_terms args (acc : term list) t =
+      case t of
+          (c as Const _) $ (v as Free _) =>
+            if c = consts_candidate andalso
+             arity = length args + 1 then
+              (list_comb (c, v :: args)) :: acc
+            else acc
+        | t1 $ (v as Free _) =>
+            get_skolem_terms (v :: args) acc t1 @
+             get_skolem_terms [] acc t1
+        | t1 $ t2 =>
+            get_skolem_terms [] acc t1 @
+             get_skolem_terms [] acc t2
+        | Abs (_, _, t') => get_skolem_terms [] acc t'
+        | _ => acc
+  in
+    map (strip_top_All_vars #> snd) conclusions
+    |> map (get_skolem_terms [] [])
+    |> List.concat
+    |> distinct (op =)
+  end
+*}
+
+ML {*
+fun instantiate_skols ctxt consts_candidates i = fn st =>
+  let
+    val thy = Proof_Context.theory_of ctxt
+
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+
+    val (params, conclusion) =
+      if null gls then
+        raise NO_GOALS
+      else
+        rpair (i - 1) gls
+        |> uncurry nth
+        |> strip_top_all_vars []
+        |> apsnd (Logic.strip_horn #> snd)
+
+    fun skolem_const_info_of t =
+      case t of
+          Const (@{const_name HOL.Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ lhs $ (Const (@{const_name Hilbert_Choice.Eps}, _) $ rhs)) =>
+          let
+            (*the parameters we will concern ourselves with*)
+            val params' =
+              Term.add_frees lhs []
+              |> distinct (op =)
+            (*check to make sure that params' <= params*)
+            val _ = @{assert} (List.all (member (op =) params) params')
+            val skolem_const_ty =
+              let
+                val (skolem_const_prety, no_params) =
+                  Term.strip_comb lhs
+                  |> apfst (dest_Var #> snd) (*head of lhs consists of a logical variable. we just want its type.*)
+                  |> apsnd length
+
+                val _ = @{assert} (length params = no_params)
+
+                (*get value type of a function type after n arguments have been supplied*)
+                fun get_val_ty n ty =
+                  if n = 0 then ty
+                  else get_val_ty (n - 1) (dest_funT ty |> snd)
+              in
+                get_val_ty no_params skolem_const_prety
+              end
+
+          in
+            (skolem_const_ty, params')
+          end
+        | _ => raise (SKOLEM_DEF t)
+
+(*
+find skolem const candidates which, after applying distinct members of params' we end up with, give us something of type skolem_const_ty.
+
+given a candidate's type, skolem_const_ty, and params', we get some pemutations of params' (i.e. the order in which they can be given to the candidate in order to get skolem_const_ty). If the list of permutations is empty, then we cannot use that candidate.
+*)
+(*
+only returns a single matching -- since terms are linear, and variable arguments are Vars, order shouldn't matter, so we can ignore permutations.
+doesn't work with polymorphism (for which we'd need to use type unification) -- this is OK since no terms should be polymorphic, since Leo2 proofs aren't.
+*)
+    fun use_candidate target_ty params acc cur_ty =
+      if null params then
+        if Type.eq_type Vartab.empty (cur_ty, target_ty) then
+          SOME (rev acc)
+        else NONE
+      else
+        let
+          val (arg_ty, val_ty) = Term.dest_funT cur_ty
+          (*now find a param of type arg_ty*)
+          val (candidate_param, params') =
+            find_and_remove
+             (snd #> pair arg_ty #> Type.eq_type Vartab.empty)
+             params
+        in
+          use_candidate target_ty params' (candidate_param :: acc) val_ty
+        end
+        handle TYPE ("dest_funT", _, _) => NONE
+             | DEST_LIST => NONE
+
+    val (skolem_const_ty, params') = skolem_const_info_of conclusion
+
+(*
+For each candidate, build a term and pass it to Thm.instantiate, whic in turn is chained with PRIMITIVE to give us this_tactic.
+
+Big picture:
+  we run the following:
+    drule leo2_skolemise THEN' this_tactic
+
+NOTE: remember to APPEND' instead of ORELSE' the two tactics relating to skolemisation
+*)
+
+    val filtered_candidates =
+      map (dest_Const
+           #> snd
+           #> use_candidate skolem_const_ty params' [])
+       consts_candidates (* prefiltered_candidates *)
+      |> pair consts_candidates (* prefiltered_candidates *)
+      |> ListPair.zip
+      |> filter (snd #> is_none #> not)
+      |> map (apsnd the)
+
+    val skolem_terms =
+      let
+        fun make_result_t (t, args) =
+          (* list_comb (t, map Free args) *)
+          if length args > 0 then
+            hd (find_skolem_term ctxt t (length args) st)
+          else t
+      in
+        map make_result_t filtered_candidates
+      end
+
+    (*prefix a skolem term with bindings for the parameters*)
+    (* val contextualise = fold absdummy (map snd params) *)
+    val contextualise = fold absfree params
+
+    val skolem_cts = map (contextualise #> cterm_of thy) skolem_terms
+
+
+(*now the instantiation code*)
+
+    (*there should only be one Var -- that is from the previous application of drule leo2_skolemise. We look for it at the head position in some equation at a conclusion of a subgoal.*)
+    val var_opt =
+      let
+        val pre_var =
+          gls
+          |> map
+              (strip_top_all_vars [] #> snd #>
+               Logic.strip_horn #> snd #>
+               get_skolem_conc)
+          |> switch (fold (fn x => fn l => if is_some x then the x :: l else l)) []
+          |> map (switch Term.add_vars [])
+          |> List.concat
+
+        fun make_var pre_var =
+          the_single pre_var
+          |> Var
+          |> cterm_of thy
+          |> SOME
+      in
+        if null pre_var then NONE
+        else make_var pre_var
+     end
+
+    fun instantiate_tac from to =
+      Thm.instantiate ([], [(from, to)])
+      |> PRIMITIVE
+
+    val tectic =
+      if is_none var_opt then no_tac
+      else
+        fold (curry (op APPEND)) (map (instantiate_tac (the var_opt)) skolem_cts) no_tac
+
+  in
+    tectic st
+  end
+*}
+
+ML {*
+fun new_skolem_tac ctxt consts_candidates =
+  let
+    fun tec thm =
+      dtac thm
+      THEN' instantiate_skols ctxt consts_candidates
+  in
+    if null consts_candidates then K no_tac
+    else FIRST' (map tec @{thms lift_exists})
+  end
+*}
+
+(*
+need a tactic to expand "? x . P" to "~ ! x. ~ P"
+*)
+ML {*
+fun ex_expander_tac ctxt i =
+   let
+     val simpset =
+       empty_simpset ctxt (*NOTE for some reason, Bind exception gets raised if ctxt's simpset isn't emptied*)
+       |> Simplifier.add_simp @{lemma "Ex P == (~ (! x. ~ P x))" by auto}
+   in
+     CHANGED (asm_full_simp_tac simpset i)
+   end
+*}
+
+
+subsubsection "extuni_dec"
+
+ML {*
+(*n-ary decomposition. Code is based on the n-ary arg_cong generator*)
+fun extuni_dec_n ctxt arity =
+  let
+    val _ = @{assert} (arity > 0)
+    val is =
+      upto (1, arity)
+      |> map Int.toString
+    val arg_tys = map (fn i => TFree ("arg" ^ i ^ "_ty", HOLogic.typeS)) is
+    val res_ty = TFree ("res" ^ "_ty", HOLogic.typeS)
+    val f_ty = arg_tys ---> res_ty
+    val f = Free ("f", f_ty)
+    val xs = map (fn i =>
+      Free ("x" ^ i, TFree ("arg" ^ i ^ "_ty", HOLogic.typeS))) is
+    (*FIXME DRY principle*)
+    val ys = map (fn i =>
+      Free ("y" ^ i, TFree ("arg" ^ i ^ "_ty", HOLogic.typeS))) is
+
+    val hyp_lhs = list_comb (f, xs)
+    val hyp_rhs = list_comb (f, ys)
+    val hyp_eq =
+      HOLogic.eq_const res_ty $ hyp_lhs $ hyp_rhs
+    val hyp =
+      HOLogic.eq_const HOLogic.boolT $ hyp_eq $ @{term False}
+      |> HOLogic.mk_Trueprop
+    fun conc_eq i =
+      let
+        val ty = TFree ("arg" ^ i ^ "_ty", HOLogic.typeS)
+        val x = Free ("x" ^ i, ty)
+        val y = Free ("y" ^ i, ty)
+        val eq = HOLogic.eq_const ty $ x $ y
+      in
+        HOLogic.eq_const HOLogic.boolT $ eq $ @{term False}
+      end
+
+    val conc_disjs = map conc_eq is
+
+    val conc =
+      if length conc_disjs = 1 then
+        the_single conc_disjs
+      else
+        fold
+         (fn t => fn t_conc => HOLogic.mk_disj (t_conc, t))
+         (tl conc_disjs) (hd conc_disjs)
+
+    val t =
+      Logic.mk_implies (hyp, HOLogic.mk_Trueprop conc)
+  in
+    Goal.prove ctxt [] [] t (fn _ => auto_tac ctxt)
+    |> Drule.export_without_context
+  end
+*}
+
+ML {*
+(*Determine the arity of a function which the "dec"
+  unification rule is about to be applied.
+  NOTE:
+    * Assumes that there is a single hypothesis
+*)
+fun find_dec_arity i = fn st =>
+  let
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+  in
+    if null gls then raise NO_GOALS
+    else
+      let
+        val (params, (literal, conc_clause)) =
+          rpair (i - 1) gls
+          |> uncurry nth
+          |> strip_top_all_vars []
+          |> apsnd Logic.strip_horn
+          |> apsnd (apfst the_single)
+
+        val get_ty =
+          HOLogic.dest_Trueprop
+          #> strip_top_All_vars
+          #> snd
+          #> HOLogic.dest_eq (*polarity's "="*)
+          #> fst
+          #> HOLogic.dest_eq (*the unification constraint's "="*)
+          #> fst
+          #> head_of
+          #> dest_Const
+          #> snd
+
+       fun arity_of ty =
+         let
+           val (_, res_ty) = dest_funT ty
+
+         in
+           1 + arity_of res_ty
+         end
+         handle (TYPE ("dest_funT", _, _)) => 0
+
+      in
+        arity_of (get_ty literal)
+      end
+  end
+
+(*given an inference, it returns the parameters (i.e., we've already matched the leading & shared quantification in the hypothesis & conclusion clauses), and the "raw" inference*)
+fun breakdown_inference i = fn st =>
+  let
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+  in
+    if null gls then raise NO_GOALS
+    else
+      rpair (i - 1) gls
+      |> uncurry nth
+      |> strip_top_all_vars []
+  end
+
+(*build a custom elimination rule for extuni_dec, and instantiate it to match a specific subgoal*)
+fun extuni_dec_elim_rule ctxt arity i = fn st =>
+  let
+    val rule = extuni_dec_n ctxt arity
+
+    val rule_hyp =
+      prop_of rule
+      |> Logic.dest_implies
+      |> fst (*assuming that rule has single hypothesis*)
+
+    (*having run break_hypothesis earlier, we know that the hypothesis
+      now consists of a single literal. We can (and should)
+      disregard the conclusion, since it hasn't been "broken",
+      and it might include some unwanted literals -- the latter
+      could cause "diff" to fail (since they won't agree with the
+      rule we have generated.*)
+
+    val inference_hyp =
+      snd (breakdown_inference i st)
+      |> Logic.dest_implies
+      |> fst (*assuming that inference has single hypothesis,
+               as explained above.*)
+  in
+    TPTP_Reconstruct_Library.diff_and_instantiate ctxt rule rule_hyp inference_hyp
+  end
+
+fun extuni_dec_tac ctxt i = fn st =>
+  let
+    val arity = find_dec_arity i st
+
+    fun elim_tac i st =
+      let
+        val rule =
+          extuni_dec_elim_rule ctxt arity i st
+          (*in case we itroduced free variables during
+            instantiation, we generalise the rule to make
+            those free variables into logical variables.*)
+          |> Thm.forall_intr_frees
+          |> Drule.export_without_context
+      in dtac rule i st end
+      handle NO_GOALS => no_tac st
+
+    fun closure tac =
+     (*batter fails if there's no toplevel disjunction in the
+       hypothesis, so we also try atac*)
+      SOLVE o (tac THEN' (batter ORELSE' atac))
+    val search_tac =
+      ASAP
+        (rtac @{thm disjI1} APPEND' rtac @{thm disjI2})
+        (FIRST' (map closure
+                  [dresolve_tac @{thms dec_commut_eq},
+                   dtac @{thm dec_commut_disj},
+                   elim_tac]))
+  in
+    (CHANGED o search_tac) i st
+  end
+*}
+
+
+subsubsection "standard_cnf"
+(*Given a standard_cnf inference, normalise it
+     e.g. ((A & B) & C \<longrightarrow> D & E \<longrightarrow> F \<longrightarrow> G) = False
+     is changed to
+          (A & B & C & D & E & F \<longrightarrow> G) = False
+ then custom-build a metatheorem which validates this:
+          (A & B & C & D & E & F \<longrightarrow> G) = False
+       -------------------------------------------
+          (A = True) & (B = True) & (C = True) &
+          (D = True) & (E = True) & (F = True) & (G = False)
+ and apply this metatheorem.
+
+There aren't any "positive" standard_cnfs in Leo2's calculus:
+  e.g.,  "(A \<longrightarrow> B) = True \<Longrightarrow> A = False | (A = True & B = True)"
+since "standard_cnf" seems to be applied at the preprocessing
+stage, together with splitting.
+*)
+
+ML {*
+(*Conjunctive counterparts to Term.disjuncts_aux and Term.disjuncts*)
+fun conjuncts_aux (Const (@{const_name HOL.conj}, _) $ t $ t') conjs =
+     conjuncts_aux t (conjuncts_aux t' conjs)
+  | conjuncts_aux t conjs = t :: conjs
+
+fun conjuncts t = conjuncts_aux t []
+
+(*HOL equivalent of Logic.strip_horn*)
+local
+  fun imp_strip_horn' acc (Const (@{const_name HOL.implies}, _) $ A $ B) =
+        imp_strip_horn' (A :: acc) B
+    | imp_strip_horn' acc t = (acc, t)
+in
+  fun imp_strip_horn t =
+    imp_strip_horn' [] t
+    |> apfst rev
+end
+*}
+
+ML {*
+(*Returns whether the antecedents are separated by conjunctions
+  or implications; the number of antecedents; and the polarity
+  of the original clause -- I think this will always be "false".*)
+fun standard_cnf_type ctxt i : thm -> (TPTP_Reconstruct.formula_kind * int * bool) option = fn st =>
+  let
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+
+    val hypos =
+      if null gls then raise NO_GOALS
+      else
+        rpair (i - 1) gls
+        |> uncurry nth
+        |> TPTP_Reconstruct.strip_top_all_vars []
+        |> snd
+        |> Logic.strip_horn
+        |> fst
+
+    (*hypothesis clause should be singleton*)
+    val _ = @{assert} (length hypos = 1)
+
+    val (t, pol) = the_single hypos
+      |> try_dest_Trueprop
+      |> TPTP_Reconstruct.strip_top_All_vars
+      |> snd
+      |> TPTP_Reconstruct.remove_polarity true
+
+    (*literal is negative*)
+    val _ = @{assert} (not pol)
+
+    val (antes, conc) = imp_strip_horn t
+
+    val (ante_type, antes') =
+      if length antes = 1 then
+        let
+          val conjunctive_antes =
+            the_single antes
+            |> conjuncts
+        in
+          if length conjunctive_antes > 1 then
+            (TPTP_Reconstruct.Conjunctive NONE,
+             conjunctive_antes)
+          else
+            (TPTP_Reconstruct.Implicational NONE,
+             antes)
+        end
+      else
+        (TPTP_Reconstruct.Implicational NONE,
+         antes)
+  in
+    if null antes then NONE
+    else SOME (ante_type, length antes', pol)
+  end
+*}
+
+ML {*
+(*Given a certain standard_cnf type, build a metatheorem that would
+  validate it*)
+fun mk_standard_cnf ctxt kind arity =
+  let
+    val _ = @{assert} (arity > 0)
+    val vars =
+      upto (1, arity + 1)
+      |> map (fn i => Free ("x" ^ Int.toString i, HOLogic.boolT))
+
+    val consequent = hd vars
+    val antecedents = tl vars
+
+    val conc =
+      fold
+       (curry HOLogic.mk_conj)
+       (map (fn var => HOLogic.mk_eq (var, @{term True})) antecedents)
+       (HOLogic.mk_eq (consequent, @{term False}))
+
+    val pre_hyp =
+      case kind of
+          TPTP_Reconstruct.Conjunctive NONE =>
+            curry HOLogic.mk_imp
+             (if length antecedents = 1 then the_single antecedents
+              else
+                fold (curry HOLogic.mk_conj) (tl antecedents) (hd antecedents))
+             (hd vars)
+        | TPTP_Reconstruct.Implicational NONE =>
+            fold (curry HOLogic.mk_imp) antecedents consequent
+
+    val hyp = HOLogic.mk_eq (pre_hyp, @{term False})
+
+    val t =
+      Logic.mk_implies (HOLogic.mk_Trueprop  hyp, HOLogic.mk_Trueprop conc)
+  in
+    Goal.prove ctxt [] [] t (fn _ => HEADGOAL (blast_tac ctxt))
+    |> Drule.export_without_context
+  end
+*}
+
+ML {*
+(*Applies a d-tactic, then breaks it up conjunctively.
+  This can be used to transform subgoals as follows:
+     (A \<longrightarrow> B) = False  \<Longrightarrow> R
+              |
+              v
+  \<lbrakk>A = True; B = False\<rbrakk> \<Longrightarrow> R
+*)
+fun weak_conj_tac drule =
+  dtac drule THEN' (REPEAT_DETERM o etac @{thm conjE})
+*}
+
+ML {*
+val uncurry_lit_neg_tac =
+  dtac @{lemma "(A \<longrightarrow> B \<longrightarrow> C) = False \<Longrightarrow> (A & B \<longrightarrow> C) = False" by auto}
+  #> REPEAT_DETERM
+*}
+
+ML {*
+fun standard_cnf_tac ctxt i = fn st =>
+  let
+    fun core_tactic i = fn st =>
+      case standard_cnf_type ctxt i st of
+          NONE => no_tac st
+        | SOME (kind, arity, _) =>
+            let
+              val rule = mk_standard_cnf ctxt kind arity;
+            in
+              (weak_conj_tac rule THEN' atac) i st
+            end
+  in
+    (uncurry_lit_neg_tac
+     THEN' TPTP_Reconstruct_Library.reassociate_conjs_tac ctxt
+     THEN' core_tactic) i st
+  end
+*}
+
+
+subsubsection "Emulator prep"
+
+ML {*
+datatype cleanup_feature =
+    RemoveHypothesesFromSkolemDefs
+  | RemoveDuplicates
+
+datatype loop_feature =
+    Close_Branch
+  | ConjI
+  | King_Cong
+  | Break_Hypotheses
+  | Donkey_Cong (*simper_animal + ex_expander_tac*)
+  | RemoveRedundantQuantifications
+  | Assumption
+
+  (*Closely based on Leo2 calculus*)
+  | Existential_Free
+  | Existential_Var
+  | Universal
+  | Not_pos
+  | Not_neg
+  | Or_pos
+  | Or_neg
+  | Equal_pos
+  | Equal_neg
+  | Extuni_Bool2
+  | Extuni_Bool1
+  | Extuni_Dec
+  | Extuni_Bind
+  | Extuni_Triv
+  | Extuni_FlexRigid
+  | Extuni_Func
+  | Polarity_switch
+  | Forall_special_pos
+
+datatype feature =
+    ConstsDiff
+  | StripQuantifiers
+  | Flip_Conclusion
+  | Loop of loop_feature list
+  | LoopOnce of loop_feature list
+  | InnerLoopOnce of loop_feature list
+  | CleanUp of cleanup_feature list
+  | AbsorbSkolemDefs
+*}
+
+ML {*
+fun can_feature x l =
+  let
+    fun sublist_of_clean_up el =
+      case el of
+          CleanUp l'' => SOME l''
+        | _ => NONE
+    fun sublist_of_loop el =
+      case el of
+          Loop l'' => SOME l''
+        | _ => NONE
+    fun sublist_of_loop_once el =
+      case el of
+          LoopOnce l'' => SOME l''
+        | _ => NONE
+    fun sublist_of_inner_loop_once el =
+      case el of
+          InnerLoopOnce l'' => SOME l''
+        | _ => NONE
+
+    fun check_sublist sought_sublist opt_list =
+      if List.all is_none opt_list then false
+      else
+        fold_options opt_list
+        |> List.concat
+        |> pair sought_sublist
+        |> subset (op =)
+  in
+    case x of
+        CleanUp l' =>
+          map sublist_of_clean_up l
+          |> check_sublist l'
+      | Loop l' =>
+          map sublist_of_loop l
+          |> check_sublist l'
+      | LoopOnce l' =>
+          map sublist_of_loop_once l
+          |> check_sublist l'
+      | InnerLoopOnce l' =>
+          map sublist_of_inner_loop_once l
+          |> check_sublist l'
+      | _ => List.exists (curry (op =) x) l
+  end;
+
+fun loop_can_feature loop_feats l =
+  can_feature (Loop loop_feats) l orelse
+  can_feature (LoopOnce loop_feats) l orelse
+  can_feature (InnerLoopOnce loop_feats) l;
+
+@{assert} (can_feature ConstsDiff [StripQuantifiers, ConstsDiff]);
+
+@{assert}
+  (can_feature (CleanUp [RemoveHypothesesFromSkolemDefs])
+    [CleanUp [RemoveHypothesesFromSkolemDefs, RemoveDuplicates]]);
+
+@{assert}
+  (can_feature (Loop []) [Loop [Existential_Var]]);
+
+@{assert}
+  (not (can_feature (Loop []) [InnerLoopOnce [Existential_Var]]));
+*}
+
+ML {*
+exception NO_LOOP_FEATS
+fun get_loop_feats (feats : feature list) =
+  let
+    val loop_find =
+      fold (fn x => fn loop_feats_acc =>
+        if is_some loop_feats_acc then loop_feats_acc
+        else
+          case x of
+              Loop loop_feats => SOME loop_feats
+            | LoopOnce loop_feats => SOME loop_feats
+            | InnerLoopOnce loop_feats => SOME loop_feats
+            | _ => NONE)
+       feats
+       NONE
+  in
+    if is_some loop_find then the loop_find
+    else raise NO_LOOP_FEATS
+  end;
+
+@{assert}
+  (get_loop_feats [Loop [King_Cong, Break_Hypotheses, Existential_Free, Existential_Var, Universal]] =
+   [King_Cong, Break_Hypotheses, Existential_Free, Existential_Var, Universal])
+*}
+
+(*use as elim rule to remove premises*)
+lemma insa_prems: "\<lbrakk>Q; P\<rbrakk> \<Longrightarrow> P" by auto
+ML {*
+fun cleanup_skolem_defs feats =
+  let
+    (*remove hypotheses from skolem defs,
+     after testing that they look like skolem defs*)
+    val dehypothesise_skolem_defs =
+      COND' (SOME #> TERMPRED (fn _ => true) conc_is_skolem_def)
+        (REPEAT_DETERM o etac @{thm insa_prems})
+        (K no_tac)
+  in
+    if can_feature (CleanUp [RemoveHypothesesFromSkolemDefs]) feats then
+      ALLGOALS (TRY o dehypothesise_skolem_defs)
+    else all_tac
+  end
+*}
+
+ML {*
+fun remove_duplicates_tac feats =
+  (if can_feature (CleanUp [RemoveDuplicates]) feats then
+     ALLGOALS distinct_subgoal_tac
+   else all_tac)
+*}
+
+ML {*
+(*given a goal state, indicates the skolem constants committed-to in it (i.e. appearing in LHS of a skolem definition)*)
+val which_skolem_concs_used = fn st =>
+  let
+    val feats = [CleanUp [RemoveHypothesesFromSkolemDefs, RemoveDuplicates]]
+    val scrubup_tac =
+      cleanup_skolem_defs feats
+      THEN remove_duplicates_tac feats
+  in
+    scrubup_tac st
+    |> break_seq
+    |> tap (fn (_, rest) => @{assert} (null (Seq.list_of rest)))
+    |> fst
+    |> TERMFUN (snd (*discard hypotheses*)
+                 #> get_skolem_conc_const) NONE
+    |> switch (fold (fn x => fn l => if is_some x then the x :: l else l)) []
+    |> map Const
+  end
+*}
+
+ML {*
+fun exists_tac ctxt feats consts_diff =
+  let
+    val ex_var =
+      if loop_can_feature [Existential_Var] feats andalso consts_diff <> [] then
+        new_skolem_tac ctxt consts_diff
+        (*We're making sure that each skolem constant is used once in instantiations.*)
+      else K no_tac
+
+    val ex_free =
+      if loop_can_feature [Existential_Free] feats andalso consts_diff = [] then
+        eresolve_tac @{thms polar_exE}
+      else K no_tac
+  in
+    ex_var APPEND' ex_free
+  end
+
+fun forall_tac ctxt feats =
+  if loop_can_feature [Universal] feats then
+    forall_pos_tac ctxt
+  else K no_tac
+*}
+
+
+subsubsection "Finite types"
+(*lift quantification from a singleton literal to a singleton clause*)
+lemma forall_pos_lift:
+"\<lbrakk>(! X. P X) = True; ! X. (P X = True) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" by auto
+
+(*predicate over the type of the leading quantified variable*)
+
+ML {*
+val extcnf_forall_special_pos_tac =
+  let
+    val bool =
+      ["True", "False"]
+
+    val bool_to_bool =
+      ["% _ . True", "% _ . False", "% x . x", "Not"]
+
+    val tecs =
+      map (fn t_s =>
+       eres_inst_tac @{context} [(("x", 0), t_s)] @{thm allE}
+       THEN' atac)
+  in
+    (TRY o etac @{thm forall_pos_lift})
+    THEN' (atac
+           ORELSE' FIRST'
+            (*FIXME could check the type of the leading quantified variable, instead of trying everything*)
+            (tecs (bool @ bool_to_bool)))
+  end
+*}
+
+
+subsubsection "Emulator"
+
+lemma efq: "[|A = True; A = False|] ==> R" by auto
+ML {*
+val efq_tac =
+  (etac @{thm efq} THEN' atac)
+  ORELSE' atac
+*}
+
+ML {*
+(*This is applied to all subgoals, repeatedly*)
+fun extcnf_combined_main ctxt feats consts_diff =
+  let
+    (*This is applied to subgoals which don't have a conclusion
+      consisting of a Skolem definition*)
+    fun extcnf_combined_tac' ctxt i = fn st =>
+      let
+        val skolem_consts_used_so_far = which_skolem_concs_used st
+        val consts_diff' = subtract (op =) skolem_consts_used_so_far consts_diff
+
+        fun feat_to_tac feat =
+          case feat of
+              Close_Branch => trace_tac' "mark: closer" efq_tac
+            | ConjI => trace_tac' "mark: conjI" (rtac @{thm conjI})
+            | King_Cong => trace_tac' "mark: expander_animal" (expander_animal ctxt)
+            | Break_Hypotheses => trace_tac' "mark: break_hypotheses" break_hypotheses
+            | RemoveRedundantQuantifications => K all_tac
+(*
+FIXME Building this into the loop instead.. maybe not the ideal choice
+            | RemoveRedundantQuantifications =>
+                trace_tac' "mark: strip_unused_variable_hyp"
+                 (REPEAT_DETERM o remove_redundant_quantification_in_lit)
+*)
+
+            | Assumption => atac
+(*FIXME both Existential_Free and Existential_Var run same code*)
+            | Existential_Free => trace_tac' "mark: forall_neg" (exists_tac ctxt feats consts_diff')
+            | Existential_Var => trace_tac' "mark: forall_neg" (exists_tac ctxt feats consts_diff')
+            | Universal => trace_tac' "mark: forall_pos" (forall_tac ctxt feats)
+            | Not_pos => trace_tac' "mark: not_pos" (dtac @{thm leo2_rules(9)})
+            | Not_neg => trace_tac' "mark: not_neg" (dtac @{thm leo2_rules(10)})
+            | Or_pos => trace_tac' "mark: or_pos" (dtac @{thm leo2_rules(5)}) (*could add (6) for negated conjunction*)
+            | Or_neg => trace_tac' "mark: or_neg" (dtac @{thm leo2_rules(7)})
+            | Equal_pos => trace_tac' "mark: equal_pos" (dresolve_tac (@{thms eq_pos_bool} @ [@{thm leo2_rules(3)}, @{thm eq_pos_func}]))
+            | Equal_neg => trace_tac' "mark: equal_neg" (dresolve_tac [@{thm eq_neg_bool}, @{thm leo2_rules(4)}])
+            | Donkey_Cong => trace_tac' "mark: donkey_cong" (simper_animal ctxt THEN' ex_expander_tac ctxt)
+
+            | Extuni_Bool2 => trace_tac' "mark: extuni_bool2" (dtac @{thm extuni_bool2})
+            | Extuni_Bool1 => trace_tac' "mark: extuni_bool1" (dtac @{thm extuni_bool1})
+            | Extuni_Bind => trace_tac' "mark: extuni_triv" (etac @{thm extuni_triv})
+            | Extuni_Triv => trace_tac' "mark: extuni_triv" (etac @{thm extuni_triv})
+            | Extuni_Dec => trace_tac' "mark: extuni_dec_tac" (extuni_dec_tac ctxt)
+            | Extuni_FlexRigid => trace_tac' "mark: extuni_flex_rigid" (atac ORELSE' asm_full_simp_tac ctxt)
+            | Extuni_Func => trace_tac' "mark: extuni_func" (dtac @{thm extuni_func})
+            | Polarity_switch => trace_tac' "mark: polarity_switch" (eresolve_tac @{thms polarity_switch})
+            | Forall_special_pos => trace_tac' "mark: dorall_special_pos" extcnf_forall_special_pos_tac
+
+        val core_tac =
+          get_loop_feats feats
+          |> map feat_to_tac
+          |> FIRST'
+      in
+        core_tac i st
+      end
+
+    (*This is applied to all subgoals, repeatedly*)
+    fun extcnf_combined_tac ctxt i =
+      COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i))
+        no_tac
+        (extcnf_combined_tac' ctxt i)
+
+    val core_tac = CHANGED (ALLGOALS (IF_UNSOLVED o TRY o extcnf_combined_tac ctxt))
+
+    val full_tac = REPEAT core_tac
+
+  in
+    CHANGED
+      (if can_feature (InnerLoopOnce []) feats then
+         core_tac
+       else full_tac)
+  end
+
+val interpreted_consts =
+  [@{const_name HOL.All}, @{const_name HOL.Ex},
+   @{const_name Hilbert_Choice.Eps},
+   @{const_name HOL.conj},
+   @{const_name HOL.disj},
+   @{const_name HOL.eq},
+   @{const_name HOL.implies},
+   @{const_name HOL.The},
+   @{const_name HOL.Ex1},
+   @{const_name HOL.Not},
+   (* @{const_name HOL.iff}, *) (*FIXME do these exist?*)
+   (* @{const_name HOL.not_equal}, *)
+   @{const_name HOL.False},
+   @{const_name HOL.True},
+   @{const_name "==>"}]
+
+fun strip_qtfrs_tac ctxt =
+  REPEAT_DETERM (HEADGOAL (rtac @{thm allI}))
+  THEN REPEAT_DETERM (HEADGOAL (etac @{thm exE}))
+  THEN HEADGOAL (canonicalise_qtfr_order ctxt)
+  THEN
+    ((REPEAT (HEADGOAL (nominal_inst_parametermatch_tac ctxt @{thm allE})))
+     APPEND (REPEAT (HEADGOAL (inst_parametermatch_tac ctxt [@{thm allE}]))))
+  (*FIXME need to handle "@{thm exI}"?*)
+
+(*difference in constants between the hypothesis clause and the conclusion clause*)
+fun clause_consts_diff thm =
+  let
+    val t =
+      prop_of thm
+      |> Logic.dest_implies
+      |> fst
+
+      (*This bit should not be needed, since Leo2 inferences don't have parameters*)
+      |> TPTP_Reconstruct.strip_top_all_vars []
+      |> snd
+
+    val do_diff =
+      Logic.dest_implies
+      #> uncurry TPTP_Reconstruct.new_consts_between
+      #> filter
+           (fn Const (n, _) =>
+             not (member (op =) interpreted_consts n))
+  in
+    if head_of t = Logic.implies then do_diff t
+    else []
+  end
+*}
+
+ML {*
+(*remove quantification in hypothesis clause (! X. t), if
+  X not free in t*)
+fun remove_redundant_quantification ctxt i = fn st =>
+  let
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+  in
+    if null gls then raise NO_GOALS
+    else
+      let
+        val (params, (hyp_clauses, conc_clause)) =
+          rpair (i - 1) gls
+          |> uncurry nth
+          |> TPTP_Reconstruct.strip_top_all_vars []
+          |> apsnd Logic.strip_horn
+      in
+        (*this is to fail gracefully in case this tactic is applied to a goal which doesn't have a single hypothesis*)
+        if length hyp_clauses > 1 then no_tac st
+        else
+          let
+            val hyp_clause = the_single hyp_clauses
+            val sep_prefix =
+              HOLogic.dest_Trueprop
+              #> TPTP_Reconstruct.strip_top_All_vars
+              #> apfst rev
+            val (hyp_prefix, hyp_body) = sep_prefix hyp_clause
+            val (conc_prefix, conc_body) = sep_prefix conc_clause
+          in
+            if null hyp_prefix orelse
+              member (op =) conc_prefix (hd hyp_prefix) orelse
+              member (op =)  (Term.add_frees hyp_body []) (hd hyp_prefix) then
+              no_tac st
+            else
+              eres_inst_tac ctxt [(("x", 0), "(@X. False)")] @{thm allE} i st
+          end
+     end
+  end
+*}
+
+ML {*
+fun remove_redundant_quantification_ignore_skolems ctxt i =
+  COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i))
+    no_tac
+    (remove_redundant_quantification ctxt i)
+*}
+
+lemma drop_redundant_literal_qtfr:
+  "(! X. P) = True \<Longrightarrow> P = True"
+  "(? X. P) = True \<Longrightarrow> P = True"
+  "(! X. P) = False \<Longrightarrow> P = False"
+  "(? X. P) = False \<Longrightarrow> P = False"
+by auto
+
+ML {*
+(*remove quantification in the literal "(! X. t) = True/False"
+  in the singleton hypothesis clause, if X not free in t*)
+fun remove_redundant_quantification_in_lit ctxt i = fn st =>
+  let
+    val gls =
+      prop_of st
+      |> Logic.strip_horn
+      |> fst
+  in
+    if null gls then raise NO_GOALS
+    else
+      let
+        val (params, (hyp_clauses, conc_clause)) =
+          rpair (i - 1) gls
+          |> uncurry nth
+          |> TPTP_Reconstruct.strip_top_all_vars []
+          |> apsnd Logic.strip_horn
+      in
+        (*this is to fail gracefully in case this tactic is applied to a goal which doesn't have a single hypothesis*)
+        if length hyp_clauses > 1 then no_tac st
+        else
+          let
+            fun literal_content (Const (@{const_name HOL.eq}, _) $ lhs $ (rhs as @{term True})) = SOME (lhs, rhs)
+              | literal_content (Const (@{const_name HOL.eq}, _) $ lhs $ (rhs as @{term False})) = SOME (lhs, rhs)
+              | literal_content t = NONE
+
+            val hyp_clause =
+              the_single hyp_clauses
+              |> HOLogic.dest_Trueprop
+              |> literal_content
+
+          in
+            if is_none hyp_clause then
+              no_tac st
+            else
+              let
+                val (hyp_lit_prefix, hyp_lit_body) =
+                  the hyp_clause
+                  |> (fn (t, polarity) =>
+                       TPTP_Reconstruct.strip_top_All_vars t
+                       |> apfst rev)
+              in
+                if null hyp_lit_prefix orelse
+                  member (op =)  (Term.add_frees hyp_lit_body []) (hd hyp_lit_prefix) then
+                  no_tac st
+                else
+                  dresolve_tac @{thms drop_redundant_literal_qtfr} i st
+              end
+          end
+     end
+  end
+*}
+
+ML {*
+fun remove_redundant_quantification_in_lit_ignore_skolems ctxt i =
+  COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i))
+    no_tac
+    (remove_redundant_quantification_in_lit ctxt i)
+*}
+
+ML {*
+fun extcnf_combined_tac ctxt prob_name_opt feats skolem_consts = fn st =>
+  let
+    val thy = Proof_Context.theory_of ctxt
+
+    (*Initially, st consists of a single goal, showing the
+      hypothesis clause implying the conclusion clause.
+      There are no parameters.*)
+    val consts_diff =
+      union (op =) skolem_consts
+       (if can_feature ConstsDiff feats then
+          clause_consts_diff st
+        else [])
+
+    val main_tac =
+      if can_feature (LoopOnce []) feats orelse can_feature (InnerLoopOnce []) feats then
+        extcnf_combined_main ctxt feats consts_diff
+      else if can_feature (Loop []) feats then
+        BEST_FIRST (TERMPRED (fn _ => true) conc_is_skolem_def NONE, size_of_thm)
+(*FIXME maybe need to weaken predicate to include "solved form"?*)
+         (extcnf_combined_main ctxt feats consts_diff)
+      else all_tac (*to allow us to use the cleaning features*)
+
+    (*Remove hypotheses from Skolem definitions,
+      then remove duplicate subgoals,
+      then we should be left with skolem definitions:
+        absorb them as axioms into the theory.*)
+    val cleanup =
+      cleanup_skolem_defs feats
+      THEN remove_duplicates_tac feats
+      THEN (if can_feature AbsorbSkolemDefs feats then
+              ALLGOALS (absorb_skolem_def ctxt prob_name_opt)
+            else all_tac)
+
+    val have_loop_feats =
+      (get_loop_feats feats; true)
+      handle NO_LOOP_FEATS => false
+
+    val tec =
+      (if can_feature StripQuantifiers feats then
+         (REPEAT (CHANGED (strip_qtfrs_tac ctxt)))
+       else all_tac)
+      THEN (if can_feature Flip_Conclusion feats then
+             HEADGOAL (flip_conclusion_tac ctxt)
+           else all_tac)
+
+      (*after stripping the quantifiers any remaining quantifiers
+        can be simply eliminated -- they're redundant*)
+      (*FIXME instead of just using allE, instantiate to a silly
+         term, to remove opportunities for unification.*)
+      THEN (REPEAT_DETERM (etac @{thm allE} 1))
+
+      THEN (REPEAT_DETERM (rtac @{thm allI} 1))
+
+      THEN (if have_loop_feats then
+              REPEAT (CHANGED
+              ((ALLGOALS (TRY o clause_breaker)) (*brush away literals which don't change*)
+               THEN
+                (*FIXME move this to a different level?*)
+                (if loop_can_feature [Polarity_switch] feats then
+                   all_tac
+                 else
+                   (TRY (IF_UNSOLVED (HEADGOAL (remove_redundant_quantification_ignore_skolems ctxt))))
+                   THEN (TRY (IF_UNSOLVED (HEADGOAL (remove_redundant_quantification_in_lit_ignore_skolems ctxt)))))
+               THEN (TRY main_tac)))
+            else
+              all_tac)
+      THEN IF_UNSOLVED cleanup
+
+  in
+    DEPTH_SOLVE (CHANGED tec) st
+  end
+*}
+
+
+subsubsection "unfold_def"
+
+(*this is used when handling unfold_tac, because the skeleton includes the definitions conjoined with the goal. it turns out that, for my tactic, the definitions are harmful. instead of modifying the skeleton (which may be nontrivial) i'm just dropping the information using this lemma. obviously, and from the name, order matters here.*)
+lemma drop_first_hypothesis [rule_format]: "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B" by auto
+
+(*Unfold_def works by reducing the goal to a meta equation,
+  then working on it until it can be discharged by atac,
+  or reflexive, or else turned back into an object equation
+  and broken down further.*)
+lemma un_meta_polarise: "(X \<equiv> True) \<Longrightarrow> X" by auto
+lemma meta_polarise: "X \<Longrightarrow> X \<equiv> True" by auto
+
+ML {*
+fun unfold_def_tac ctxt depends_on_defs = fn st =>
+  let
+    (*This is used when we end up with something like
+        (A & B) \<equiv> True \<Longrightarrow> (B & A) \<equiv> True.
+      It breaks down this subgoal until it can be trivially
+      discharged.
+     *)
+    val kill_meta_eqs_tac =
+      dtac @{thm un_meta_polarise}
+      THEN' rtac @{thm meta_polarise}
+      THEN' (REPEAT_DETERM o (etac @{thm conjE}))
+      THEN' (REPEAT_DETERM o (rtac @{thm conjI} ORELSE' atac))
+
+    val continue_reducing_tac =
+      rtac @{thm meta_eq_to_obj_eq} 1
+      THEN (REPEAT_DETERM (ex_expander_tac ctxt 1))
+      THEN TRY (polarise_subgoal_hyps 1) (*no need to REPEAT_DETERM here, since there should only be one hypothesis*)
+      THEN TRY (dtac @{thm eq_reflection} 1)
+      THEN (TRY ((CHANGED o rewrite_goal_tac ctxt
+              (@{thm expand_iff} :: @{thms simp_meta})) 1))
+      THEN HEADGOAL (rtac @{thm reflexive}
+                     ORELSE' atac
+                     ORELSE' kill_meta_eqs_tac)
+
+    val tectic =
+      (rtac @{thm polarise} 1 THEN atac 1)
+      ORELSE
+        (REPEAT_DETERM (etac @{thm conjE} 1 THEN etac @{thm drop_first_hypothesis} 1)
+         THEN PRIMITIVE (Conv.fconv_rule Drule.eta_long_conversion)
+         THEN (REPEAT_DETERM (ex_expander_tac ctxt 1))
+         THEN (TRY ((CHANGED o rewrite_goal_tac ctxt @{thms simp_meta}) 1))
+         THEN PRIMITIVE (Conv.fconv_rule Drule.eta_long_conversion)
+         THEN
+           (HEADGOAL atac
+           ORELSE
+            (unfold_tac ctxt depends_on_defs
+             THEN IF_UNSOLVED continue_reducing_tac)))
+  in
+    tectic st
+  end
+*}
+
+
+subsection "Handling split 'preprocessing'"
+
+lemma split_tranfs:
+  "! x. P x & Q x \<equiv> (! x. P x) & (! x. Q x)"
+  "~ (~ A) \<equiv> A"
+  "? x. A \<equiv> A"
+  "(A & B) & C \<equiv> A & B & C"
+  "A = B \<equiv> (A --> B) & (B --> A)"
+by (rule eq_reflection, auto)+
+
+(*Same idiom as ex_expander_tac*)
+ML {*
+fun split_simp_tac (ctxt : Proof.context) i =
+   let
+     val simpset =
+       fold Simplifier.add_simp @{thms split_tranfs} (empty_simpset ctxt)
+   in
+     CHANGED (asm_full_simp_tac simpset i)
+   end
+*}
+
+
+subsection "Alternative reconstruction tactics"
+ML {*
+(*An "auto"-based proof reconstruction, where we attempt to reconstruct each inference
+  using auto_tac. A realistic tactic would inspect the inference name and act
+  accordingly.*)
+fun auto_based_reconstruction_tac ctxt prob_name n =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name
+  in
+    TPTP_Reconstruct.inference_at_node
+     thy
+     prob_name (#meta pannot) n
+      |> the
+      |> (fn {inference_fmla, ...} =>
+          Goal.prove ctxt [] [] inference_fmla
+           (fn pdata => auto_tac (#context pdata)))
+  end
+*}
+
+(*An oracle-based reconstruction, which is only used to test the shunting part of the system*)
+oracle oracle_iinterp = "fn t => t"
+ML {*
+fun oracle_based_reconstruction_tac ctxt prob_name n =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name
+  in
+    TPTP_Reconstruct.inference_at_node
+     thy
+     prob_name (#meta pannot) n
+      |> the
+      |> (fn {inference_fmla, ...} => cterm_of thy inference_fmla)
+      |> oracle_iinterp
+  end
+*}
+
+
+subsection "Leo2 reconstruction tactic"
+
+ML {*
+exception UNSUPPORTED_ROLE
+exception INTERPRET_INFERENCE
+
+(*Failure reports can be adjusted to avoid interrupting
+  an overall reconstruction process*)
+fun fail ctxt x =
+  if unexceptional_reconstruction ctxt then
+    (warning x; raise INTERPRET_INFERENCE)
+  else error x
+
+fun interpret_leo2_inference_tac ctxt prob_name node =
+  let
+    val thy = Proof_Context.theory_of ctxt
+
+    val _ =
+      if Config.get ctxt tptp_trace_reconstruction then
+        tracing ("interpret_inference reconstructing node" ^ node ^ " of " ^ TPTP_Problem_Name.mangle_problem_name prob_name)
+      else ()
+
+    val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name
+
+    fun nonfull_extcnf_combined_tac feats =
+      extcnf_combined_tac ctxt (SOME prob_name)
+       [ConstsDiff,
+        StripQuantifiers,
+        InnerLoopOnce (Break_Hypotheses :: (*FIXME RemoveRedundantQuantifications :: *) feats),
+        AbsorbSkolemDefs]
+       []
+
+    val source_inf_opt =
+      AList.lookup (op =) (#meta pannot)
+      #> the
+      #> #source_inf_opt
+
+    (*FIXME integrate this with other lookup code, or in the early analysis*)
+    local
+      fun node_is_of_role role node =
+        AList.lookup (op =) (#meta pannot) node |> the
+        |> #role
+        |> (fn role' => role = role')
+
+      fun roled_dependencies_names role =
+        let
+          fun values () =
+            case role of
+                TPTP_Syntax.Role_Definition =>
+                  map (apsnd Binding.dest) (#defs pannot)
+              | TPTP_Syntax.Role_Axiom =>
+                  map (apsnd Binding.dest) (#axs pannot)
+              | _ => raise UNSUPPORTED_ROLE
+          in
+            if is_none (source_inf_opt node) then []
+            else
+              case the (source_inf_opt node) of
+                  TPTP_Proof.Inference (_, _, parent_inf) =>
+                    List.map TPTP_Proof.parent_name parent_inf
+                    |> List.filter (node_is_of_role role)
+                    |> (*FIXME currently definitions are not
+                         included in the proof annotations, so
+                         i'm using all the definitions available
+                         in the proof. ideally i should only
+                         use the ones in the proof annotation.*)
+                       (fn x =>
+                         if role = TPTP_Syntax.Role_Definition then
+                           let fun values () = map (apsnd Binding.dest) (#defs pannot)
+                           in
+                             map snd (values ())
+                           end
+                         else
+                         map (fn node => AList.lookup (op =) (values ()) node |> the) x)
+                | _ => []
+         end
+
+      val roled_dependencies =
+        roled_dependencies_names
+        #> map (#3 #> Global_Theory.get_thm thy)
+    in
+      val depends_on_defs = roled_dependencies TPTP_Syntax.Role_Definition
+      val depends_on_axs = roled_dependencies TPTP_Syntax.Role_Axiom
+      val depends_on_defs_names = roled_dependencies_names TPTP_Syntax.Role_Definition
+    end
+
+    fun get_binds source_inf_opt =
+      case the source_inf_opt of
+          TPTP_Proof.Inference (_, _, parent_inf) =>
+            List.map
+              (fn TPTP_Proof.Parent _ => []
+                | TPTP_Proof.ParentWithDetails (_, parent_details) => parent_details)
+              parent_inf
+            |> List.concat
+        | _ => []
+
+    val inference_name =
+      case TPTP_Reconstruct.inference_at_node thy prob_name (#meta pannot) node of
+          NONE => fail ctxt "Cannot reconstruct rule: no information"
+        | SOME {inference_name, ...} => inference_name
+    val default_tac = HEADGOAL (blast_tac ctxt)
+  in
+    case inference_name of
+      "fo_atp_e" =>
+        HEADGOAL (etac (oracle_based_reconstruction_tac ctxt prob_name node))
+    | "copy" =>
+         HEADGOAL
+          (atac
+           ORELSE'
+              (rtac @{thm polarise}
+               THEN' atac))
+    | "polarity_switch" => nonfull_extcnf_combined_tac [Polarity_switch]
+    | "solved_all_splits" => solved_all_splits_tac
+    | "extcnf_not_pos" => nonfull_extcnf_combined_tac [Not_pos]
+    | "extcnf_forall_pos" => nonfull_extcnf_combined_tac [Universal]
+    | "negate_conjecture" => fail ctxt "Should not handle negate_conjecture here"
+    | "unfold_def" => unfold_def_tac ctxt depends_on_defs
+    | "extcnf_not_neg" => nonfull_extcnf_combined_tac [Not_neg]
+    | "extcnf_or_neg" => nonfull_extcnf_combined_tac [Or_neg]
+    | "extcnf_equal_pos" => nonfull_extcnf_combined_tac [Equal_pos]
+    | "extcnf_equal_neg" => nonfull_extcnf_combined_tac [Equal_neg]
+    | "extcnf_forall_special_pos" =>
+         nonfull_extcnf_combined_tac [Forall_special_pos]
+         ORELSE HEADGOAL (blast_tac ctxt)
+    | "extcnf_or_pos" => nonfull_extcnf_combined_tac [Or_pos]
+    | "extuni_bool2" => nonfull_extcnf_combined_tac [Extuni_Bool2]
+    | "extuni_bool1" => nonfull_extcnf_combined_tac [Extuni_Bool1]
+    | "extuni_dec" =>
+        HEADGOAL atac
+        ORELSE nonfull_extcnf_combined_tac [Extuni_Dec]
+    | "extuni_bind" => nonfull_extcnf_combined_tac [Extuni_Bind]
+    | "extuni_triv" => nonfull_extcnf_combined_tac [Extuni_Triv]
+    | "extuni_flex_rigid" => nonfull_extcnf_combined_tac [Extuni_FlexRigid]
+    | "prim_subst" => nonfull_extcnf_combined_tac [Assumption]
+    | "bind" =>
+        let
+          val ordered_binds = get_binds (source_inf_opt node)
+        in
+          bind_tac ctxt prob_name ordered_binds
+        end
+    | "standard_cnf" => HEADGOAL (standard_cnf_tac ctxt)
+    | "extcnf_forall_neg" =>
+        nonfull_extcnf_combined_tac
+         [Existential_Var(* , RemoveRedundantQuantifications *)] (*FIXME RemoveRedundantQuantifications*)
+    | "extuni_func" =>
+        nonfull_extcnf_combined_tac [Extuni_Func, Existential_Var]
+    | "replace_leibnizEQ" => nonfull_extcnf_combined_tac [Assumption]
+    | "replace_andrewsEQ" => nonfull_extcnf_combined_tac [Assumption]
+    | "split_preprocessing" =>
+         (REPEAT (HEADGOAL (split_simp_tac ctxt)))
+         THEN TRY (PRIMITIVE (Conv.fconv_rule Drule.eta_long_conversion))
+         THEN HEADGOAL atac
+
+    (*FIXME some of these could eventually be handled specially*)
+    | "fac_restr" => default_tac
+    | "sim" => default_tac
+    | "res" => default_tac
+    | "rename" => default_tac
+    | "flexflex" => default_tac
+    | other => fail ctxt ("Unknown inference rule: " ^ other)
+  end
+*}
+
+ML {*
+fun interpret_leo2_inference ctxt prob_name node =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name
+
+    val (inference_name, inference_fmla) =
+      case TPTP_Reconstruct.inference_at_node thy prob_name (#meta pannot) node of
+          NONE => fail ctxt "Cannot reconstruct rule: no information"
+        | SOME {inference_name, inference_fmla, ...} =>
+            (inference_name, inference_fmla)
+
+    val proof_outcome =
+      let
+        fun prove () =
+          Goal.prove ctxt [] [] inference_fmla
+           (fn pdata => interpret_leo2_inference_tac
+            (#context pdata) prob_name node)
+      in
+        if informative_failure ctxt then SOME (prove ())
+        else try prove ()
+      end
+
+  in case proof_outcome of
+      NONE => fail ctxt (Pretty.string_of
+        (Pretty.block
+          [Pretty.str ("Failed inference reconstruction for '" ^
+            inference_name ^ "' at node " ^ node ^ ":\n"),
+           Syntax.pretty_term ctxt inference_fmla]))
+    | SOME thm => thm
+  end
+*}
+
+ML {*
+(*filter a set of nodes based on which inference rule was used to
+  derive a node*)
+fun nodes_by_inference (fms : TPTP_Reconstruct.formula_meaning list) inference_rule =
+  let
+    fun fold_fun n l =
+      case TPTP_Reconstruct.node_info fms #source_inf_opt n of
+          NONE => l
+        | SOME (TPTP_Proof.File _) => l
+        | SOME (TPTP_Proof.Inference (rule_name, _, _)) =>
+            if rule_name = inference_rule then n :: l
+            else l
+  in
+    fold fold_fun (map fst fms) []
+  end
+*}
+
+ML {*
+fun leo2_on_load (pannot : TPTP_Reconstruct.proof_annotation) thy =
+  let
+    val ctxt = Proof_Context.init_global thy
+    val dud = ("", Binding.empty, @{term False})
+    val pre_skolem_defs =
+      nodes_by_inference (#meta pannot) "extcnf_forall_neg" @
+       nodes_by_inference (#meta pannot) "extuni_func"
+      |> map (fn x =>
+              (interpret_leo2_inference ctxt (#problem_name pannot) x; dud)
+               handle NO_SKOLEM_DEF (s, bnd, t) => (s, bnd, t))
+      |> filter (fn (x, _, _) => x <> "") (*In case no skolem constants were introduced in that inference*)
+    val skolem_defs = map (fn (x, y, _) => (x, y)) pre_skolem_defs
+    val thy' =
+      fold (fn skolem_def => fn thy =>
+             let
+               val ((s, thm), thy') = Thm.add_axiom_global skolem_def thy
+               (* val _ = warning ("Added skolem definition " ^ s ^ ": " ^  @{make_string thm}) *) (*FIXME use of make_string*)
+             in thy' end)
+       (map (fn (_, y, z) => (y, z)) pre_skolem_defs)
+       thy
+  in
+    ({problem_name = #problem_name pannot,
+      skolem_defs = skolem_defs,
+      defs = #defs pannot,
+      axs = #axs pannot,
+      meta = #meta pannot},
+     thy')
+  end
+*}
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy	Wed Feb 19 15:57:02 2014 +0000
@@ -0,0 +1,822 @@
+(*  Title:      HOL/TPTP/TPTP_Proof_Reconstruction.thy
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+Various tests for the proof reconstruction module.
+
+NOTE
+  - Makes use of the PolyML structure.
+  - looks for THF proofs in the path indicated by $THF_PROOFS
+*)
+
+theory TPTP_Proof_Reconstruction_Test
+imports TPTP_Test TPTP_Proof_Reconstruction
+begin
+
+declare [[
+  tptp_trace_reconstruction = false,
+  tptp_test_all = false,
+  (* tptp_test_all = true, *)
+  tptp_test_timeout = 30,
+  (* tptp_max_term_size = 200 *)
+  tptp_max_term_size = 0
+]]
+
+ML {*
+  if test_all @{context} then ()
+  else
+    (Toplevel.debug := true;
+     PolyML.print_depth 200;
+     PolyML.Compiler.maxInlineSize := 0)
+*}
+
+
+section "Importing proofs"
+
+ML {*
+val probs =
+  (* "$THF_PROOFS/SYN991^1.p.out" *) (*lacks conjecture*)
+  (* "$THF_PROOFS/SYO040^2.p.out" *)
+  (* "$THF_PROOFS/NUM640^1.p.out" *)
+  (* "$THF_PROOFS/SEU553^2.p.out" *)
+  (* "$THF_PROOFS/NUM665^1.p.out" *)
+  (* "$THF_PROOFS/SEV161^5.p.out" *)
+  (* "$THF_PROOFS/SET014^4.p.out" *)
+  "$THF_PROOFS/NUM667^1.p.out"
+  |> Path.explode
+  |> single
+
+val prob_names =
+  probs
+  |> map (Path.base #> Path.implode #> TPTP_Problem_Name.Nonstandard)
+*}
+
+setup {*
+  if test_all @{context} then I
+  else
+    fold
+     (fn path =>
+       TPTP_Reconstruct.import_thm true [Path.dir path, Path.explode "$THF_PROOFS"] path leo2_on_load)
+     probs
+*}
+
+text "Display nicely."
+ML {*
+fun display_nicely ctxt (fms : TPTP_Reconstruct.formula_meaning list) =
+  List.app (fn ((n, data) : TPTP_Reconstruct.formula_meaning) =>
+    Pretty.writeln
+      (Pretty.block
+        [Pretty.str (n ^ " "),
+         Syntax.pretty_term ctxt (#fmla data),
+         Pretty.str (
+          if is_none (#source_inf_opt data) then ""
+          else ("\n\tannotation: " ^
+           PolyML.makestring (the (#source_inf_opt data : TPTP_Proof.source_info option))))])
+    ) (rev fms);
+
+(*FIXME hack for testing*)
+fun test_fmla thy =
+  TPTP_Reconstruct.get_fmlas_of_prob thy (hd prob_names);
+
+fun test_pannot thy =
+  TPTP_Reconstruct.get_pannot_of_prob thy (hd prob_names);
+
+if test_all @{context} orelse prob_names = [] then ()
+else
+  display_nicely @{context}
+  (#meta (test_pannot @{theory}))
+(* To look at the original proof (i.e., before the proof transformations applied
+   when the proof is loaded) replace previous line with:
+   (test_fmla @{theory}
+    |> map TPTP_Reconstruct.structure_fmla_meaning)
+*)
+*}
+
+ML {*
+fun step_range_tester f_x f_exn ctxt prob_name from until =
+  let
+    val max =
+      case until of
+          SOME x => x
+        | NONE =>
+            if is_some Int.maxInt then the Int.maxInt else 999999
+    fun test_step x =
+      if x > max then ()
+      else
+        (f_x x;
+         (interpret_leo2_inference ctxt prob_name (Int.toString x); ())
+         handle e => f_exn e; (*FIXME naive. should let Interrupt through*)
+         (*assumes that inferences are numbered consecutively*)
+         test_step (x + 1))
+  in
+    test_step from
+  end
+
+val step_range_tester_tracing =
+  step_range_tester
+   (fn x => tracing ("@step " ^ Int.toString x))
+   (fn e => tracing ("!!" ^ PolyML.makestring e))
+*}
+
+ML {*
+(*try to reconstruct each inference step*)
+if test_all @{context} orelse prob_names = []
+orelse true (*NOTE currently disabled*)
+then ()
+else
+  let
+    (*FIXME not guaranteed to be the right nodes*)
+    val heur_start = 3
+    val heur_end =
+      hd (#meta (test_pannot @{theory}))
+      |> #1
+      |> Int.fromString
+  in
+    step_range_tester_tracing @{context} (hd prob_names) heur_start heur_end
+  end
+*}
+
+
+section "Building metadata and tactics"
+
+subsection "Building the skeleton"
+ML {*
+if test_all @{context} orelse prob_names = [] then []
+else TPTP_Reconstruct.make_skeleton @{context} (test_pannot @{theory});
+
+length it
+*}
+
+
+subsection "The 'one shot' tactic approach"
+ML {*
+val the_tactic =
+  if test_all @{context} then []
+  else
+    map (fn prob_name =>
+      (TPTP_Reconstruct.naive_reconstruct_tac @{context} interpret_leo2_inference (* auto_based_reconstruction_tac *) (* oracle_based_reconstruction_tac *) prob_name))
+     prob_names;
+*}
+
+
+subsection "The 'piecemeal' approach"
+ML {*
+val the_tactics =
+  if test_all @{context} then []
+  else
+    map (fn prob_name =>
+      TPTP_Reconstruct.naive_reconstruct_tacs interpret_leo2_inference (* auto_based_reconstruction_tac *) (* oracle_based_reconstruction_tac *) prob_name @{context})
+     prob_names;
+*}
+
+ML {*
+print_depth 2000;
+the_tactics
+|> map (filter (fn (_, _, x) => is_none x)
+        #> map (fn (x, SOME y, _) => (x, cterm_of @{theory} y)))
+*}
+
+
+section "Using metadata and tactics"
+text "There are various ways of testing the two ways (whole tactics or lists of tactics) of representing 'reconstructors'."
+
+
+subsection "The 'one shot' tactic approach"
+text "First we test whole tactics."
+ML {*
+(*produce thm*)
+if test_all @{context} then []
+else
+  map (
+    (* try *) (TPTP_Reconstruct.reconstruct @{context}
+     (fn prob_name =>
+       TPTP_Reconstruct.naive_reconstruct_tac @{context} interpret_leo2_inference prob_name
+     (* oracle_based_reconstruction_tac *))))
+   prob_names
+*}
+
+
+subsection "The 'piecemeal' approach"
+ML {*
+fun attac n = List.nth (List.nth (the_tactics, 0), n) |> #3 |> the |> snd
+fun attac_to n 0 = attac n
+  | attac_to n m = attac n THEN attac_to (n + 1) (m - 1)
+fun shotac n = List.nth (List.nth (the_tactics, 0), n) |> #3 |> the |> fst
+*}
+
+ML {*
+(*Given a list of reconstructed inferences (as in "the_tactics" above,
+  count the number of failures and successes, and list the failed inference
+  reconstructions.*)
+fun evaluate_the_tactics [] acc = acc
+  | evaluate_the_tactics ((node_no, (inf_name, inf_fmla, NONE)) :: xs) ((fai, suc), inf_list) =
+      let
+        val score = (fai + 1, suc)
+        val index_info = get_index (fn (x, _) => if x = node_no then SOME true else NONE) inf_list
+        val inf_list' =
+          case index_info of
+              NONE => (node_no, (inf_name, inf_fmla, 1)) :: inf_list
+            | SOME (idx, _) =>
+                nth_map idx (fn (node_no, (inf_name, inf_fmla, count)) => (node_no, (inf_name, inf_fmla, count + 1))) inf_list
+      in
+        evaluate_the_tactics xs (score, inf_list')
+      end
+  | evaluate_the_tactics ((_, (_, _, SOME _)) :: xs) ((fai, suc), inf_list) =
+      evaluate_the_tactics xs ((fai, suc + 1), inf_list)
+*}
+
+
+text "Now we build a tactic by combining lists of tactics"
+ML {*
+(*given a list of tactics to be applied in sequence (i.e., they
+  follow a skeleton), we build a single tactic, interleaving
+  some tracing info to help with debugging.*)
+fun step_by_step_tacs verbose (thm_tacs : (thm * tactic) list) : tactic =
+    let
+      fun interleave_tacs [] [] = all_tac
+        | interleave_tacs (tac1 :: tacs1) (tac2 :: tacs2) =
+            EVERY [tac1, tac2]
+            THEN interleave_tacs tacs1 tacs2
+      val thms_to_traceprint =
+        map (fn thm => fn st =>
+              (*FIXME uses makestring*)
+              print_tac (PolyML.makestring thm) st)
+
+    in
+      if verbose then
+        ListPair.unzip thm_tacs
+        |> apfst (fn thms => enumerate 1 thms)
+        |> apfst thms_to_traceprint
+        |> uncurry interleave_tacs
+      else EVERY (map #2 thm_tacs)
+    end
+*}
+
+ML {*
+(*apply step_by_step_tacs to all problems under test*)
+val narrated_tactics =
+ map (map (#3 #> the)
+      #> step_by_step_tacs false)
+   the_tactics;
+
+(*produce thm*)
+(*use narrated_tactics to reconstruct all problems under test*)
+if test_all @{context} then []
+else
+  map (fn (prob_name, tac) =>
+         TPTP_Reconstruct.reconstruct @{context}
+           (fn _ => tac) prob_name)
+    (ListPair.zip (prob_names, narrated_tactics))
+*}
+
+
+subsection "Manually using 'piecemeal' approach"
+text "Another testing possibility involves manually creating a lemma
+and running through the list of tactics generating to prove that lemma. The following code shows the goal of each problem under test, and then for each problem returns the list of tactics which can be invoked individually as shown below."
+ML {*
+fun show_goal ctxt prob_name =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name
+  in
+    #meta pannot
+    |> List.filter (fn (_, info) =>
+        #role info = TPTP_Syntax.Role_Conjecture)
+    |> hd
+    |> snd |> #fmla
+    |> cterm_of thy
+  end;
+
+if test_all @{context} then []
+else
+  map (show_goal @{context}) prob_names;
+*}
+
+ML {*
+(*project out the list of tactics from "the_tactics"*)
+val just_the_tacs  =
+ map (map (#3 #> the #> #2))
+   the_tactics;
+
+map length just_the_tacs
+*}
+
+ML {*
+(*like just_the_tacs, but extract the thms, to inspect their thys*)
+val just_the_thms  =
+ map (map (#3 #> the #> #1))
+   the_tactics;
+
+map length just_the_thms;
+*}
+
+ML {*
+(*given a thm, show us the axioms in its thy*)
+val axms_of_thy_of_thm =
+  Thm.theory_of_thm
+  #> ` Theory.axioms_of
+  #> apsnd cterm_of
+  #> swap
+  #> apsnd (map snd)
+  #> uncurry map
+*}
+
+ML {*
+(*Show the skeleton-level inference which is done by each element of just_the_tacs. This is useful when debugging using the technique shown next*)
+if test_all @{context} orelse prob_names = [] then ()
+else
+  the_tactics
+  |> hd
+  |> map #1
+  |> TPTP_Reconstruct_Library.enumerate 0
+  |> List.app (PolyML.makestring #> writeln)
+  *}
+
+ML {*
+fun leo2_tac_wrap prob_name step i = fn st =>
+  let
+    val ctxt =
+      Thm.theory_of_thm st
+      |> Proof_Context.init_global
+  in
+    rtac (interpret_leo2_inference ctxt prob_name step) i st
+  end
+*}
+
+(*FIXME move these examples elsewhere*)
+(*
+lemma "\<forall>(Xj\<Colon>TPTP_Interpret.ind) Xk\<Colon>TPTP_Interpret.ind.
+        bnd_cCKB6_BLACK Xj Xk \<longrightarrow>
+        bnd_cCKB6_BLACK (bnd_s (bnd_s (bnd_s Xj))) (bnd_s Xk)"
+apply (tactic {*nth (nth just_the_tacs 0) 0*})
+apply (tactic {*nth (nth just_the_tacs 0) 1*})
+apply (tactic {*nth (nth just_the_tacs 0) 2*})
+apply (tactic {*nth (nth just_the_tacs 0) 3*})
+apply (tactic {*nth (nth just_the_tacs 0) 4*})
+apply (tactic {*nth (nth just_the_tacs 0) 5*})
+ML_prf "nth (hd the_tactics) 6"
+apply (tactic {*nth (nth just_the_tacs 0) 6*})
+apply (tactic {*nth (nth just_the_tacs 0) 7*})
+apply (tactic {*nth (nth just_the_tacs 0) 8*})
+apply (tactic {*nth (nth just_the_tacs 0) 9*})
+apply (tactic {*nth (nth just_the_tacs 0) 10*})
+apply (tactic {*nth (nth just_the_tacs 0) 11*})
+apply (tactic {*nth (nth just_the_tacs 0) 12*})
+apply (tactic {*nth (nth just_the_tacs 0) 13*})
+apply (tactic {*nth (nth just_the_tacs 0) 14*})
+apply (tactic {*nth (nth just_the_tacs 0) 15*})
+
+apply (tactic {*nth (nth just_the_tacs 0) 16*})
+
+(*
+apply (tactic {*
+rtac (interpret_leo2_inference @{context} (hd prob_names) "8") 1
+*})
+apply (tactic {*
+rtac (interpret_leo2_inference @{context} (hd prob_names) "7") 1
+*})
+apply (tactic {*
+rtac (interpret_leo2_inference @{context} (hd prob_names) "6") 1
+*})
+(*
+apply (tactic {*
+rtac (interpret_leo2_inference @{context} (hd prob_names) "4") 1
+*})
+*)
+*)
+
+apply (tactic {*nth (nth just_the_tacs 0) 17*})
+apply (tactic {*nth (nth just_the_tacs 0) 18*})
+apply (tactic {*nth (nth just_the_tacs 0) 19*})
+apply (tactic {*nth (nth just_the_tacs 0) 20*})
+apply (tactic {*nth (nth just_the_tacs 0) 21*})
+
+ML_prf "nth (hd the_tactics) 21"
+ML_prf "nth (hd the_tactics) 22"
+
+apply (tactic {*nth (nth just_the_tacs 0) 22*})
+apply (tactic {*nth (nth just_the_tacs 0) 23*})
+apply (tactic {*nth (nth just_the_tacs 0) 24*})
+apply (tactic {*nth (nth just_the_tacs 0) 25*})
+
+
+ML_prf "nth (hd the_tactics) 19"
+
+apply (tactic {*
+interpret_leo2_inference_wrap (hd prob_names) "8" 1
+*})
+apply (tactic {*
+interpret_leo2_inference_wrap (hd prob_names) "7" 1
+*})
+apply (tactic {*
+interpret_leo2_inference_wrap (hd prob_names) "6" 1
+*})
+apply (tactic {*
+interpret_leo2_inference_wrap (hd prob_names) "4" 1
+*})
+
+ML_prf "nth (hd the_tactics) 20"
+ML_prf "nth (hd the_tactics) 21"
+ML_prf "nth (hd the_tactics) 22"
+*)
+
+(*
+lemma "bnd_powersetE1 \<longrightarrow>
+     bnd_sepInPowerset \<longrightarrow>
+     (\<forall>A Xphi. bnd_subset (bnd_dsetconstr A Xphi) A)"
+apply (tactic {*nth (nth just_the_tacs 0) 0*})
+apply (tactic {*nth (nth just_the_tacs 0) 1*})
+apply (tactic {*nth (nth just_the_tacs 0) 2*})
+apply (tactic {*nth (nth just_the_tacs 0) 3*})
+apply (tactic {*nth (nth just_the_tacs 0) 4*})
+apply (tactic {*nth (nth just_the_tacs 0) 5*})
+ML_prf "nth (hd the_tactics) 6"
+apply (tactic {*nth (nth just_the_tacs 0) 6*})
+apply (tactic {*nth (nth just_the_tacs 0) 7*})
+apply (tactic {*nth (nth just_the_tacs 0) 8*})
+apply (tactic {*nth (nth just_the_tacs 0) 9*})
+apply (tactic {*nth (nth just_the_tacs 0) 10*})
+apply (tactic {*nth (nth just_the_tacs 0) 11*})
+apply (tactic {*nth (nth just_the_tacs 0) 12*})
+apply (tactic {*nth (nth just_the_tacs 0) 13*})
+apply (tactic {*nth (nth just_the_tacs 0) 14*})
+apply (tactic {*nth (nth just_the_tacs 0) 15*})
+apply (tactic {*nth (nth just_the_tacs 0) 16*})
+apply (tactic {*nth (nth just_the_tacs 0) 17*})
+apply (tactic {*nth (nth just_the_tacs 0) 18*})
+apply (tactic {*nth (nth just_the_tacs 0) 19*})
+apply (tactic {*nth (nth just_the_tacs 0) 20*})
+apply (tactic {*nth (nth just_the_tacs 0) 21*})
+apply (tactic {*nth (nth just_the_tacs 0) 22*})
+apply (tactic {*nth (nth just_the_tacs 0) 23*})
+apply (tactic {*nth (nth just_the_tacs 0) 24*})
+apply (tactic {*nth (nth just_the_tacs 0) 25*})
+(* apply (tactic {*nth (nth just_the_tacs 0) 26*}) *)
+ML_prf "nth (hd the_tactics) 26"
+apply (subgoal_tac "(\<not> (\<forall>A Xphi. bnd_subset (bnd_dsetconstr A Xphi) A)) =
+       True \<Longrightarrow>
+       (\<not> bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) =
+       True")
+prefer 2
+apply (thin_tac "(bnd_powersetE1 \<longrightarrow>
+      bnd_sepInPowerset \<longrightarrow>
+      (\<forall>A Xphi. bnd_subset (bnd_dsetconstr A Xphi) A)) =
+     False")
+apply (tactic {*extcnf_combined_simulator_tac (hd prob_names) 1*})
+apply (tactic {*extcnf_combined_simulator_tac (hd prob_names) 1*})
+apply (tactic {*extcnf_combined_simulator_tac (hd prob_names) 1*})
+apply (tactic {*extcnf_combined_simulator_tac (hd prob_names) 1*})
+
+apply simp
+
+(* apply (tactic {*nth (nth just_the_tacs 0) 26*}) *)
+apply (tactic {*nth (nth just_the_tacs 0) 27*})
+apply (tactic {*nth (nth just_the_tacs 0) 28*})
+apply (tactic {*nth (nth just_the_tacs 0) 29*})
+apply (tactic {*nth (nth just_the_tacs 0) 30*})
+apply (tactic {*nth (nth just_the_tacs 0) 31*})
+apply (tactic {*nth (nth just_the_tacs 0) 32*})
+apply (tactic {*nth (nth just_the_tacs 0) 33*})
+apply (tactic {*nth (nth just_the_tacs 0) 34*})
+apply (tactic {*nth (nth just_the_tacs 0) 35*})
+apply (tactic {*nth (nth just_the_tacs 0) 36*})
+apply (tactic {*nth (nth just_the_tacs 0) 37*})
+apply (tactic {*nth (nth just_the_tacs 0) 38*})
+apply (tactic {*nth (nth just_the_tacs 0) 39*})
+apply (tactic {*nth (nth just_the_tacs 0) 40*})
+apply (tactic {*nth (nth just_the_tacs 0) 41*})
+apply (tactic {*nth (nth just_the_tacs 0) 42*})
+apply (tactic {*nth (nth just_the_tacs 0) 43*})
+apply (tactic {*nth (nth just_the_tacs 0) 44*})
+apply (tactic {*nth (nth just_the_tacs 0) 45*})
+apply (tactic {*nth (nth just_the_tacs 0) 46*})
+apply (tactic {*nth (nth just_the_tacs 0) 47*})
+apply (tactic {*nth (nth just_the_tacs 0) 48*})
+apply (tactic {*nth (nth just_the_tacs 0) 49*})
+apply (tactic {*nth (nth just_the_tacs 0) 50*})
+apply (tactic {*nth (nth just_the_tacs 0) 51*})
+done
+*)
+
+(*
+We can use just_the_tacs as follows:
+
+(this is from SEV012^5.p.out)
+lemma "((\<forall>(Xx :: bool) (Xy :: bool). True \<longrightarrow> True) \<and>
+      (\<forall>(Xx :: bool) (Xy :: bool) (Xz :: bool). True \<and> True \<longrightarrow> True)) \<and>
+     (\<lambda>(Xx :: bool) (Xy :: bool). True) = (\<lambda>Xx Xy. True)"
+apply (tactic {*nth (nth just_the_tacs 0) 0*})
+apply (tactic {*nth (nth just_the_tacs 0) 1*})
+apply (tactic {*nth (nth just_the_tacs 0) 2*})
+apply (tactic {*nth (nth just_the_tacs 0) 3*})
+apply (tactic {*nth (nth just_the_tacs 0) 4*})
+apply (tactic {*nth (nth just_the_tacs 0) 5*})
+ML_prf "nth (hd the_tactics) 6"
+apply (tactic {*nth (nth just_the_tacs 0) 6*})
+apply (tactic {*nth (nth just_the_tacs 0) 7*})
+apply (tactic {*nth (nth just_the_tacs 0) 8*})
+apply (tactic {*nth (nth just_the_tacs 0) 9*})
+apply (tactic {*nth (nth just_the_tacs 0) 10*})
+apply (tactic {*nth (nth just_the_tacs 0) 11*})
+apply (tactic {*nth (nth just_the_tacs 0) 12*})
+apply (tactic {*nth (nth just_the_tacs 0) 13*})
+apply (tactic {*nth (nth just_the_tacs 0) 14*})
+apply (tactic {*nth (nth just_the_tacs 0) 15*})
+apply (tactic {*nth (nth just_the_tacs 0) 16*})
+apply (tactic {*nth (nth just_the_tacs 0) 17*})
+apply (tactic {*nth (nth just_the_tacs 0) 18*})
+apply (tactic {*nth (nth just_the_tacs 0) 19*})
+apply (tactic {*nth (nth just_the_tacs 0) 20*})
+apply (tactic {*nth (nth just_the_tacs 0) 21*})
+apply (tactic {*nth (nth just_the_tacs 0) 22*})
+done
+
+(*
+We could also use previous definitions directly,
+e.g. the following should prove the goal at a go:
+- apply (tactic {*narrated_tactics |> hd |> hd*})
+- apply (tactic {*
+    TPTP_Reconstruct.naive_reconstruct_tac
+     interpret_leo2_inference
+     (hd prob_names)
+     @{context}*})
+(Note that the previous two methods don't work in this
+ "lemma" testing mode, not sure why. The previous methods
+ (producing the thm values directly) should work though.)
+*)
+*)
+
+
+section "Testing against benchmark"
+
+ML {*
+(*if reconstruction_info value is NONE then a big error must have occurred*)
+type reconstruction_info =
+  ((int(*no of failures*) * int(*no of successes*)) *
+  (TPTP_Reconstruct.rolling_stock * term option(*inference formula*) * int (*number of times the inference occurs in the skeleton*)) list) option
+
+datatype proof_contents =
+    No_info
+  | Empty
+  | Nonempty of reconstruction_info
+
+(*To make output less cluttered in whole-run tests*)
+fun erase_inference_fmlas (Nonempty (SOME (outline, inf_info))) =
+      Nonempty (SOME (outline, map (fn (inf_name, _, count) => (inf_name, NONE, count)) inf_info))
+  | erase_inference_fmlas x = x
+*}
+
+ML {*
+(*Report on how many inferences in a proof are reconstructed, and give some
+  info about the inferences for which reconstruction failed.*)
+fun test_partial_reconstruction thy prob_file =
+  let
+    val prob_name =
+      (Path.base #> Path.implode #> TPTP_Problem_Name.Nonstandard) prob_file
+
+    val thy' =
+      try
+       (TPTP_Reconstruct.import_thm
+        true
+        [Path.dir prob_file, Path.explode "$TPTP"]
+        prob_file leo2_on_load)
+       thy
+
+    val ctxt' =
+      if is_some thy' then SOME (Proof_Context.init_global (the thy')) else NONE
+
+    (*to test if proof is empty*)
+    val fms =
+      if is_some thy'
+      then SOME (TPTP_Reconstruct.get_fmlas_of_prob (the thy') prob_name)
+      else NONE
+
+    val the_tactics =
+      if is_some thy' then
+        SOME (TPTP_Reconstruct.naive_reconstruct_tacs (* metis_based_reconstruction_tac *)
+interpret_leo2_inference (* auto_based_reconstruction_tac *) (* oracle_based_reconstruction_tac *) prob_name (the ctxt'))
+      else NONE
+
+(* val _ = tracing ("tt=" ^ PolyML.makestring the_tactics) *)
+
+    val skeleton =
+      if is_some thy' then
+        SOME (TPTP_Reconstruct.make_skeleton (the ctxt')
+              (TPTP_Reconstruct.get_pannot_of_prob (the thy') prob_name))
+      else NONE
+
+    val skeleton_and_tactics =
+      if is_some thy' then
+        SOME (ListPair.zip (the skeleton, the the_tactics))
+      else NONE
+
+    val result =
+      if is_some thy' then
+        SOME (evaluate_the_tactics (the skeleton_and_tactics)
+              ((0, 0), []))
+      else NONE
+
+    (*strip node names*)
+    val result' =
+      if is_some result then SOME (apsnd (map #2) (the result)) else NONE
+  in
+    if is_some fms andalso List.null (the fms) then Empty
+    else Nonempty result'
+  end
+*}
+
+ML {*
+  (*default timeout is 1 min*)
+  fun reconstruct timeout light_output file thy =
+    let
+      val timer = Timer.startRealTimer ()
+    in
+      TimeLimit.timeLimit (Time.fromSeconds (if timeout = 0 then 60 else timeout))
+       (test_partial_reconstruction thy
+        #> light_output ? erase_inference_fmlas
+        #> PolyML.makestring (* FIXME *)
+        #> (fn s => report (Proof_Context.init_global thy) (PolyML.makestring file ^ " === " ^ s ^
+             " t=" ^ (Timer.checkRealTimer timer |> Time.toMilliseconds |> PolyML.makestring))))
+       file
+    end
+*}
+
+ML {*
+  (*this version of "reconstruct" builds theorems, instead of lists of reconstructed inferences*)
+  (*default timeout is 1 min*)
+  fun reconstruct timeout file thy =
+    let
+      val timer = Timer.startRealTimer ()
+      val thy' =
+        TPTP_Reconstruct.import_thm true
+         [Path.dir file, Path.explode "$TPTP"]
+         file leo2_on_load thy
+
+      val ctxt = Proof_Context.init_global thy' (*FIXME pass ctxt instead of thy*)
+      val prob_name =
+        file
+        |> Path.base
+        |> Path.implode
+        |> TPTP_Problem_Name.Nonstandard
+    in
+      TimeLimit.timeLimit (Time.fromSeconds (if timeout = 0 then 60 else timeout))
+       (fn prob_name =>
+        (can
+          (TPTP_Reconstruct.reconstruct ctxt (fn prob_name =>
+            TPTP_Reconstruct.naive_reconstruct_tac ctxt interpret_leo2_inference prob_name (* oracle_based_reconstruction_tac *))) prob_name )
+       |> (fn s => report ctxt (Path.print file ^ " === " ^ Bool.toString s ^
+             " t=" ^ (Timer.checkRealTimer timer |> Time.toMilliseconds |> PolyML.makestring))))
+       prob_name
+    end
+*}
+
+ML {*
+  fun reconstruction_test timeout ctxt =
+    test_fn ctxt
+     (fn file => reconstruct timeout file (Proof_Context.theory_of ctxt))
+     "reconstructor"
+     ()
+*}
+
+ML {*
+datatype examination_results =
+    Whole_proof of string(*filename*) * proof_contents
+  | Specific_rule of string(*filename*) * string(*inference rule*) * term option list
+
+(*Look out for failures reconstructing a particular inference rule*)
+fun filter_failures inference_name (Whole_proof (filename, results)) =
+  let
+    val filtered_results =
+      case results of
+          Nonempty (SOME results') =>
+            #2 results'
+            |> map (fn (stock as TPTP_Reconstruct.Annotated_step (_, inf_name), inf_fmla, _) =>
+                 if inf_name = inference_name then [inf_fmla] else [])
+            |> List.concat
+        | _ => []
+  in Specific_rule (filename, inference_name, filtered_results) end
+
+(*Returns detailed info about a proof-reconstruction attempt.
+  If rule_name is specified then the related failed inferences
+  are returned, otherwise all failed inferences are returned.*)
+fun examine_failed_inferences ctxt filename rule_name =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    val prob_file = Path.explode filename
+    val results =
+      if test_all ctxt then No_info
+      else test_partial_reconstruction thy prob_file
+  in
+    Whole_proof (filename, results)
+    |> is_some rule_name ? (fn x =>
+                             filter_failures (the rule_name) x)
+  end
+*}
+
+ML {*
+exception NONSENSE
+
+fun annotation_or_id (TPTP_Reconstruct.Step n) = n
+  | annotation_or_id (TPTP_Reconstruct.Annotated_step (n, anno)) = anno
+  | annotation_or_id TPTP_Reconstruct.Assumed = "assumption"
+  | annotation_or_id TPTP_Reconstruct.Unconjoin = "conjI"
+  | annotation_or_id TPTP_Reconstruct.Caboose = "(end)"
+  | annotation_or_id (TPTP_Reconstruct.Synth_step s) = s
+  | annotation_or_id (TPTP_Reconstruct.Split (split_node, soln_node, _)) = "split_at " ^ split_node ^ " " ^ soln_node;
+
+fun count_failures (Whole_proof (_, No_info)) = raise NONSENSE
+  | count_failures (Whole_proof (_, Empty)) = raise NONSENSE
+  | count_failures (Whole_proof (_, Nonempty NONE)) = raise NONSENSE
+  | count_failures (Whole_proof (_, Nonempty (SOME (((n, _), _))))) = n
+  | count_failures (Specific_rule (_, _, t)) = length t
+
+fun pre_classify_failures [] alist = alist
+  | pre_classify_failures ((stock, _, _) :: xs) alist =
+      let
+        val inf = annotation_or_id stock
+        val count = AList.lookup (op =) alist inf
+      in
+        if is_none count
+        then pre_classify_failures xs ((inf, 1) :: alist)
+        else
+          pre_classify_failures xs
+           (AList.update (op =) (inf, the count + 1) alist)
+      end
+
+fun classify_failures (Whole_proof (_, Nonempty (SOME (((_, _), inferences))))) = pre_classify_failures inferences []
+  | classify_failures (Specific_rule (_, rule, t)) = [(rule, length t)]
+  | classify_failures _ = raise NONSENSE
+*}
+
+ML {*
+val regressions = map (fn s => "$THF_PROOFS/" ^ s)
+  ["SEV405^5.p.out",
+   (*"SYO377^5.p.out", Always seems to raise Interrupt on my laptop -- probably because node 475 has lots of premises*)
+   "PUZ031^5.p.out",
+   "ALG001^5.p.out",
+   "SYO238^5.p.out",
+   (*"SEV158^5.p.out", This is big*)
+   "SYO285^5.p.out",
+   "../SYO285^5.p.out_reduced",
+   (* "SYO225^5.p.out", This is big*)
+   "SYO291^5.p.out",
+   "SET669^3.p.out",
+   "SEV233^5.p.out",
+   (*"SEU511^1.p.out", This is big*)
+   "SEV161^5.p.out",
+   "SEV012^5.p.out",
+   "SYO035^1.p.out",
+   "SYO291^5.p.out",
+   "SET741^4.p.out", (*involves both definitions and contorted splitting. has nice graph.*)
+   "SEU548^2.p.out",
+   "SEU513^2.p.out",
+   "SYO006^1.p.out",
+   "SYO371^5.p.out" (*has contorted splitting, like SYO006^1.p.out, but doesn't involve definitions*)
+  ]
+*}
+
+ML {*
+val experiment = examine_failed_inferences @{context}
+  (List.last regressions) NONE;
+
+(*
+val experiment_focus =
+  filter_failures "extcnf_combined" experiment;
+*)
+
+(*
+count_failures experiment_focus
+classify_failures experiment
+*)
+*}
+
+text "Run reconstruction on all problems in a benchmark (provided via a script)
+and report on partial success."
+
+declare [[
+  tptp_test_all = true,
+  tptp_test_timeout = 10
+]]
+
+ML {*
+  (*problem source*)
+  val tptp_probs_dir =
+    Path.explode "$THF_PROOFS"
+    |> Path.expand;
+*}
+
+ML {*
+  if test_all @{context} then
+    (report @{context} "Reconstructing proofs";
+    S timed_test (reconstruction_test (get_timeout @{context})) @{context} (TPTP_Syntax.get_file_list tptp_probs_dir))
+  else ()
+*}
+
+(*
+Debugging strategy:
+  1) get list of all proofs
+  2) order by size
+  3) try to construct each in turn, given some timeout
+
+Use this to find the smallest failure, then debug that.
+*)
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test_Units.thy	Wed Feb 19 15:57:02 2014 +0000
@@ -0,0 +1,2314 @@
+(*  Title:      HOL/TPTP/TPTP_Proof_Reconstruction.thy
+    Author:     Nik Sultana, Cambridge University Computer Laboratory
+
+Unit tests for proof reconstruction module.
+
+NOTE
+  - Makes use of the PolyML structure.
+*)
+
+theory TPTP_Proof_Reconstruction_Test
+imports TPTP_Test TPTP_Proof_Reconstruction
+begin
+
+ML {*
+print_depth 200;
+Toplevel.debug := true;
+PolyML.Compiler.maxInlineSize := 0;
+(* FIXME doesn't work with Isabelle?
+   PolyML.Compiler.debug := true *)
+*}
+
+declare [[
+  tptp_trace_reconstruction = true
+]]
+
+lemma "! (X1 :: bool) (X2 :: bool) (X3 :: bool) (X4 :: bool) (X5 :: bool). P \<Longrightarrow> ! (X1 :: bool) (X3 :: bool) (X5 :: bool). P"
+apply (tactic {*canonicalise_qtfr_order @{context} 1*})
+oops
+
+lemma "! (X1 :: bool) (X2 :: bool) (X3 :: bool) (X4 :: bool) (X5 :: bool). P \<Longrightarrow> ! (X1 :: bool) (X3 :: bool) (X5 :: bool). P"
+apply (tactic {*canonicalise_qtfr_order @{context} 1*})
+apply (rule allI)+
+apply (tactic {*nominal_inst_parametermatch_tac @{context} @{thm allE} 1*})+
+oops
+
+(*Could test bind_tac further with NUM667^1 inode43*)
+
+(*
+  (* SEU581^2.p_nux *)
+     (* (Annotated_step ("inode1", "bind"), *)
+lemma "\<forall>(SV5\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+            SV6\<Colon>TPTP_Interpret.ind.
+            (bnd_in (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15)
+              (bnd_powerset bnd_sK1_A) =
+             bnd_in (bnd_dsetconstr SV6 SV5)
+              (bnd_powerset SV6)) =
+            False \<Longrightarrow>
+         (bnd_in (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15)
+           (bnd_powerset bnd_sK1_A) =
+          bnd_in (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15)
+           (bnd_powerset bnd_sK1_A)) =
+         False"
+ML_prf {*
+open TPTP_Syntax;
+open TPTP_Proof;
+
+
+val binds =
+[Bind ("SV6", Atom (THF_Atom_term (Term_Func (Uninterpreted "sK1_A", [])))), Bind ("SV5", Quant (Lambda, [("SX0", SOME (Fmla_type (Atom (THF_Atom_term (Term_Func (TypeSymbol Type_Ind, []))))))], Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "sK2_SY15", []))), Atom (THF_Atom_term (Term_Var "SX0"))])))]
+(* |> TPTP_Reconstruct.permute *)
+
+(*
+val binds =
+[Bind ("SV5", Quant (Lambda, [("SX0", SOME (Fmla_type (Atom (THF_Atom_term (Term_Func (TypeSymbol Type_Ind, []))))))], Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "sK2_SY15", []))), Atom (THF_Atom_term (Term_Var "SX0"))]))),
+Bind ("SV6", Atom (THF_Atom_term (Term_Func (Uninterpreted "sK1_A", []))))
+]
+*)
+
+val tec =
+(*
+  map (bind_tac @{context} (hd prob_names)) binds
+  |> FIRST
+*)
+  bind_tac @{context} (hd prob_names) binds
+*}
+apply (tactic {*tec*})
+done
+
+     (* (Annotated_step ("inode2", "bind"), *)
+lemma "\<forall>(SV7\<Colon>TPTP_Interpret.ind) SV8\<Colon>TPTP_Interpret.ind.
+            (bnd_subset SV8 SV7 =
+             bnd_subset (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15)
+              bnd_sK1_A) =
+            False \<or>
+            bnd_in SV8 (bnd_powerset SV7) = False \<Longrightarrow>
+         (bnd_subset (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15)
+           bnd_sK1_A =
+          bnd_subset (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15)
+           bnd_sK1_A) =
+         False \<or>
+         bnd_in (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15)
+          (bnd_powerset bnd_sK1_A) =
+         False"
+ML_prf {*
+open TPTP_Syntax;
+open TPTP_Proof;
+
+
+val binds =
+[Bind ("SV8", Fmla (Interpreted_ExtraLogic Apply, [Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "dsetconstr", []))), Atom (THF_Atom_term (Term_Func (Uninterpreted "sK1_A", [])))]), Quant (Lambda, [("SX0", SOME (Fmla_type (Atom (THF_Atom_term (Term_Func (TypeSymbol Type_Ind, []))))))], Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "sK2_SY15", []))), Atom (THF_Atom_term (Term_Var "SX0"))]))])), Bind ("SV7", Atom (THF_Atom_term (Term_Func (Uninterpreted "sK1_A", []))))]
+(* |> TPTP_Reconstruct.permute *)
+
+val tec =
+(*
+  map (bind_tac @{context} (hd prob_names)) binds
+  |> FIRST
+*)
+  bind_tac @{context} (hd prob_names) binds
+*}
+apply (tactic {*tec*})
+done
+*)
+
+(*
+from SEU897^5
+lemma "
+\<forall>SV9 SV10 SV11 SV12 SV13 SV14.
+   (((((bnd_sK5_SY14 SV14 SV13 SV12 = SV11) = False \<or>
+       (bnd_sK4_SX0 = SV10 (bnd_sK5_SY14 SV9 SV10 SV11)) =
+       False) \<or>
+      bnd_cR SV14 = False) \<or>
+     (SV12 = SV13 SV14) = False) \<or>
+    bnd_cR SV9 = False) \<or>
+   (SV11 = SV10 SV9) = False \<Longrightarrow>
+\<forall>SV14 SV13 SV12 SV10 SV9.
+   (((((bnd_sK5_SY14 SV14 SV13 SV12 =
+        bnd_sK5_SY14 SV14 SV13 SV12) =
+       False \<or>
+       (bnd_sK4_SX0 =
+        SV10
+         (bnd_sK5_SY14 SV9 SV10
+           (bnd_sK5_SY14 SV14 SV13 SV12))) =
+       False) \<or>
+      bnd_cR SV14 = False) \<or>
+     (SV12 = SV13 SV14) = False) \<or>
+    bnd_cR SV9 = False) \<or>
+   (bnd_sK5_SY14 SV14 SV13 SV12 = SV10 SV9) = False"
+ML_prf {*
+open TPTP_Syntax;
+open TPTP_Proof;
+
+val binds =
+[Bind ("SV11", Fmla (Interpreted_ExtraLogic Apply, [Fmla (Interpreted_ExtraLogic Apply, [Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "sK5_SY14", []))), Atom (THF_Atom_term (Term_Var "SV14"))]), Atom (THF_Atom_term (Term_Var "SV13"))]), Atom (THF_Atom_term (Term_Var "SV12"))]))]
+
+val tec = bind_tac @{context} (hd prob_names) binds
+*}
+apply (tactic {*tec*})
+done
+*)
+
+
+subsection "Interpreting the inferences"
+
+(*from SET598^5
+lemma "(bnd_sK1_X = (\<lambda>SY17. bnd_sK2_Y SY17 \<and> bnd_sK3_Z SY17) \<longrightarrow>
+   ((\<forall>SY25. bnd_sK1_X SY25 \<longrightarrow> bnd_sK2_Y SY25) \<and>
+    (\<forall>SY26. bnd_sK1_X SY26 \<longrightarrow> bnd_sK3_Z SY26)) \<and>
+   (\<forall>SY27.
+       (\<forall>SY21. SY27 SY21 \<longrightarrow> bnd_sK2_Y SY21) \<and>
+       (\<forall>SY15. SY27 SY15 \<longrightarrow> bnd_sK3_Z SY15) \<longrightarrow>
+       (\<forall>SY30. SY27 SY30 \<longrightarrow> bnd_sK1_X SY30))) =
+  False \<Longrightarrow>
+  (\<not> (bnd_sK1_X = (\<lambda>SY17. bnd_sK2_Y SY17 \<and> bnd_sK3_Z SY17) \<longrightarrow>
+      ((\<forall>SY25. bnd_sK1_X SY25 \<longrightarrow> bnd_sK2_Y SY25) \<and>
+       (\<forall>SY26. bnd_sK1_X SY26 \<longrightarrow> bnd_sK3_Z SY26)) \<and>
+      (\<forall>SY27.
+          (\<forall>SY21. SY27 SY21 \<longrightarrow> bnd_sK2_Y SY21) \<and>
+          (\<forall>SY15. SY27 SY15 \<longrightarrow> bnd_sK3_Z SY15) \<longrightarrow>
+          (\<forall>SY30. SY27 SY30 \<longrightarrow> bnd_sK1_X SY30)))) =
+  True"
+apply (tactic {*polarity_switch_tac @{context}*})
+done
+lemma "
+  (((\<forall>SY25. bnd_sK1_X SY25 \<longrightarrow> bnd_sK2_Y SY25) \<and>
+    (\<forall>SY26. bnd_sK1_X SY26 \<longrightarrow> bnd_sK3_Z SY26)) \<and>
+   (\<forall>SY27.
+       (\<forall>SY21. SY27 SY21 \<longrightarrow> bnd_sK2_Y SY21) \<and>
+       (\<forall>SY15. SY27 SY15 \<longrightarrow> bnd_sK3_Z SY15) \<longrightarrow>
+       (\<forall>SY30. SY27 SY30 \<longrightarrow> bnd_sK1_X SY30)) \<longrightarrow>
+   bnd_sK1_X = (\<lambda>SY17. bnd_sK2_Y SY17 \<and> bnd_sK3_Z SY17)) =
+  False \<Longrightarrow>
+  (\<not> (((\<forall>SY25. bnd_sK1_X SY25 \<longrightarrow> bnd_sK2_Y SY25) \<and>
+       (\<forall>SY26. bnd_sK1_X SY26 \<longrightarrow> bnd_sK3_Z SY26)) \<and>
+      (\<forall>SY27.
+          (\<forall>SY21. SY27 SY21 \<longrightarrow> bnd_sK2_Y SY21) \<and>
+          (\<forall>SY15. SY27 SY15 \<longrightarrow> bnd_sK3_Z SY15) \<longrightarrow>
+          (\<forall>SY30. SY27 SY30 \<longrightarrow> bnd_sK1_X SY30)) \<longrightarrow>
+      bnd_sK1_X = (\<lambda>SY17. bnd_sK2_Y SY17 \<and> bnd_sK3_Z SY17))) =
+  True"
+apply (tactic {*polarity_switch_tac @{context}*})
+done
+*)
+
+(* beware lack of type annotations
+(* lemma "!!x. (A x = B x) = False ==> (B x = A x) = False" *)
+(* lemma "!!x. (A x = B x) = True ==> (B x = A x) = True" *)
+(* lemma "((A x) = (B x)) = True ==> ((B x) = (A x)) = True" *)
+lemma "(A = B) = True ==> (B = A) = True"
+*)
+lemma "!!x. ((A x :: bool) = B x) = False ==> (B x = A x) = False"
+apply (tactic {*expander_animal @{context} 1*})
+oops
+
+lemma "(A & B) ==> ~(~A | ~B)"
+by (tactic {*expander_animal @{context} 1*})
+
+lemma "(A | B) ==> ~(~A & ~B)"
+by (tactic {*expander_animal @{context} 1*})
+
+lemma "(A | B) | C ==> A | (B | C)"
+by (tactic {*expander_animal @{context} 1*})
+
+lemma "(~~A) = B ==> A = B"
+by (tactic {*expander_animal @{context} 1*})
+
+lemma "~ ~ (A = True) ==> A = True"
+by (tactic {*expander_animal @{context} 1*})
+
+(*This might not be a goal which might realistically arise:
+lemma "((~~A) = B) & (B = (~~A)) ==> ~(~(A = B) | ~(B = A))" *)
+lemma "((~~A) = True) ==> ~(~(A = True) | ~(True = A))"
+apply (tactic {*expander_animal @{context} 1*})+
+apply (rule conjI)
+apply assumption
+apply (rule sym, assumption)
+done
+
+lemma "A = B ==> ((~~A) = B) & (B = (~~A)) ==> ~(~(A = B) | ~(B = A))"
+by (tactic {*expander_animal @{context} 1*})+
+
+(*some lemmas assume constants in the signature of PUZ114^5*)
+consts
+  PUZ114_5_bnd_sK1 :: "TPTP_Interpret.ind"
+  PUZ114_5_bnd_sK2 :: "TPTP_Interpret.ind"
+  PUZ114_5_bnd_sK3 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+  PUZ114_5_bnd_sK4 :: "(TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind"
+  PUZ114_5_bnd_sK5 :: "(TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind"
+  PUZ114_5_bnd_s :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind"
+  PUZ114_5_bnd_c1 :: TPTP_Interpret.ind
+
+(*testing logical expansion*)
+lemma "!! SY30. (SY30 PUZ114_5_bnd_c1 PUZ114_5_bnd_c1 \<and>
+       (\<forall>Xj Xk.
+           SY30 Xj Xk \<longrightarrow>
+           SY30 (PUZ114_5_bnd_s (PUZ114_5_bnd_s Xj)) Xk \<and>
+           SY30 (PUZ114_5_bnd_s Xj) (PUZ114_5_bnd_s Xk)) \<longrightarrow>
+       SY30 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2)
+==> (
+       (~ SY30 PUZ114_5_bnd_c1 PUZ114_5_bnd_c1)
+     | (~ (\<forall>Xj Xk.
+           SY30 Xj Xk \<longrightarrow>
+           SY30 (PUZ114_5_bnd_s (PUZ114_5_bnd_s Xj)) Xk \<and>
+           SY30 (PUZ114_5_bnd_s Xj) (PUZ114_5_bnd_s Xk)))
+     | SY30 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2
+)"
+by (tactic {*expander_animal @{context} 1*})+
+
+(*
+extcnf_forall_pos:
+
+     (! X. L1) | ... | Ln
+ ----------------------------   X' fresh
+  ! X'. (L1[X'/X] | ... | Ln)
+
+After elimination rule has been applied we'll have a subgoal which looks like this:
+            (! X. L1)
+ ----------------------------   X' fresh
+  ! X'. (L1[X'/X] | ... | Ln)
+and we need to transform it so that, in Isabelle, we go from
+ (! X. L1) ==> ! X'. (L1[X'/X] | ... | Ln)
+to
+ \<And> X'. L1[X'/X] ==> (L1[X'/X] | ... | Ln)
+(where X' is fresh, or renamings are done suitably).*)
+
+lemma "A | B \<Longrightarrow> A | B | C"
+apply (tactic {*flip_conclusion_tac @{context} 1*})+
+apply (tactic {*break_hypotheses 1*})+
+oops
+
+consts
+  CSR122_1_bnd_lBill_THFTYPE_i :: TPTP_Interpret.ind
+  CSR122_1_bnd_lMary_THFTYPE_i :: TPTP_Interpret.ind
+  CSR122_1_bnd_lSue_THFTYPE_i :: TPTP_Interpret.ind
+  CSR122_1_bnd_n2009_THFTYPE_i :: TPTP_Interpret.ind
+  CSR122_1_bnd_lYearFn_THFTYPE_IiiI :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind"
+  CSR122_1_bnd_holdsDuring_THFTYPE_IiooI ::
+    "TPTP_Interpret.ind \<Rightarrow> bool \<Rightarrow> bool"
+  CSR122_1_bnd_likes_THFTYPE_IiioI ::
+    "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+
+lemma "\<forall>SV2. (CSR122_1_bnd_holdsDuring_THFTYPE_IiooI
+                 (CSR122_1_bnd_lYearFn_THFTYPE_IiiI CSR122_1_bnd_n2009_THFTYPE_i)
+                 (\<not> (\<not> CSR122_1_bnd_likes_THFTYPE_IiioI
+                        CSR122_1_bnd_lMary_THFTYPE_i
+                        CSR122_1_bnd_lBill_THFTYPE_i \<or>
+                     \<not> CSR122_1_bnd_likes_THFTYPE_IiioI
+                        CSR122_1_bnd_lSue_THFTYPE_i
+                        CSR122_1_bnd_lBill_THFTYPE_i)) =
+                CSR122_1_bnd_holdsDuring_THFTYPE_IiooI SV2 True) =
+               False \<Longrightarrow>
+         \<forall>SV2. (CSR122_1_bnd_lYearFn_THFTYPE_IiiI CSR122_1_bnd_n2009_THFTYPE_i =
+                SV2) =
+               False \<or>
+               ((\<not> (\<not> CSR122_1_bnd_likes_THFTYPE_IiioI
+                       CSR122_1_bnd_lMary_THFTYPE_i CSR122_1_bnd_lBill_THFTYPE_i \<or>
+                    \<not> CSR122_1_bnd_likes_THFTYPE_IiioI CSR122_1_bnd_lSue_THFTYPE_i
+                       CSR122_1_bnd_lBill_THFTYPE_i)) =
+                True) =
+               False"
+apply (rule allI, erule_tac x = "SV2" in allE)
+apply (tactic {*extuni_dec_tac @{context} 1*})
+done
+
+(*SEU882^5*)
+(*
+lemma
+ "\<forall>(SV2\<Colon>TPTP_Interpret.ind)
+        SV1\<Colon>TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind.
+        (SV1 SV2 = bnd_sK1_Xy) =
+        False
+   \<Longrightarrow>
+   \<forall>SV2\<Colon>TPTP_Interpret.ind.
+            (bnd_sK1_Xy = bnd_sK1_Xy) =
+            False"
+ML_prf {*
+open TPTP_Syntax;
+open TPTP_Proof;
+
+val binds =
+[Bind ("SV1", Quant (Lambda, [("SX0", SOME (Fmla_type (Atom (THF_Atom_term (Term_Func (TypeSymbol Type_Ind, []))))))], Atom (THF_Atom_term (Term_Func (Uninterpreted "sK1_Xy", [])))))]
+
+val tec = bind_tac @{context} (hd prob_names) binds
+*}
+(*
+apply (tactic {*strip_qtfrs
+                (* THEN tec *)*})
+*)
+apply (tactic {*tec*})
+done
+*)
+
+lemma "A | B \<Longrightarrow> C1 | A | C2 | B | C3"
+apply (erule disjE)
+apply (tactic {*clause_breaker 1*})
+apply (tactic {*clause_breaker 1*})
+done
+
+lemma "A \<Longrightarrow> A"
+apply (tactic {*clause_breaker 1*})
+done
+
+typedecl NUM667_1_bnd_nat
+consts
+  NUM667_1_bnd_less :: "NUM667_1_bnd_nat \<Rightarrow> NUM667_1_bnd_nat \<Rightarrow> bool"
+  NUM667_1_bnd_x :: NUM667_1_bnd_nat
+  NUM667_1_bnd_y :: NUM667_1_bnd_nat
+
+(*NUM667^1 node 302 -- dec*)
+lemma "\<forall>SV12 SV13 SV14 SV9 SV10 SV11.
+       ((((NUM667_1_bnd_less SV12 SV13 = NUM667_1_bnd_less SV11 SV10) = False \<or>
+          (SV14 = SV13) = False) \<or>
+         NUM667_1_bnd_less SV12 SV14 = False) \<or>
+        NUM667_1_bnd_less SV9 SV10 = True) \<or>
+       (SV9 = SV11) =
+       False \<Longrightarrow>
+       \<forall>SV9 SV14 SV10 SV13 SV11 SV12.
+       (((((SV12 = SV11) = False \<or> (SV13 = SV10) = False) \<or>
+          (SV14 = SV13) = False) \<or>
+         NUM667_1_bnd_less SV12 SV14 = False) \<or>
+        NUM667_1_bnd_less SV9 SV10 = True) \<or>
+       (SV9 = SV11) =
+       False"
+apply (tactic {*strip_qtfrs_tac @{context}*})
+apply (tactic {*break_hypotheses 1*})
+apply (tactic {*ALLGOALS (TRY o clause_breaker)*})
+apply (tactic {*extuni_dec_tac @{context} 1*})
+done
+
+ML {*
+extuni_dec_n @{context} 2;
+*}
+
+(*NUM667^1, node 202*)
+lemma "\<forall>SV9 SV10 SV11.
+       ((((SV9 = SV11) = (NUM667_1_bnd_x = NUM667_1_bnd_y)) = False \<or>
+         NUM667_1_bnd_less SV11 SV10 = False) \<or>
+        NUM667_1_bnd_less SV9 SV10 = True) \<or>
+       NUM667_1_bnd_less NUM667_1_bnd_x NUM667_1_bnd_y =
+       True \<Longrightarrow>
+       \<forall>SV10 SV9 SV11.
+       ((((SV11 = NUM667_1_bnd_x) = False \<or> (SV9 = NUM667_1_bnd_y) = False) \<or>
+         NUM667_1_bnd_less SV11 SV10 = False) \<or>
+        NUM667_1_bnd_less SV9 SV10 = True) \<or>
+       NUM667_1_bnd_less NUM667_1_bnd_x NUM667_1_bnd_y =
+       True"
+apply (tactic {*strip_qtfrs_tac @{context}*})
+apply (tactic {*break_hypotheses 1*})
+apply (tactic {*ALLGOALS (TRY o clause_breaker)*})
+apply (tactic {*extuni_dec_tac @{context} 1*})
+done
+
+(*NUM667^1 node 141*)
+(*
+lemma "((bnd_x = bnd_x) = False \<or> (bnd_y = bnd_z) = False) \<or>
+         bnd_less bnd_x bnd_y = True \<Longrightarrow>
+         (bnd_y = bnd_z) = False \<or> bnd_less bnd_x bnd_y = True"
+apply (tactic {*strip_qtfrs*})
+apply (tactic {*break_hypotheses 1*})
+apply (tactic {*ALLGOALS (TRY o clause_breaker)*})
+apply (erule extuni_triv)
+done
+*)
+
+ML {*
+fun full_extcnf_combined_tac ctxt =
+  extcnf_combined_tac ctxt NONE
+   [ConstsDiff,
+    StripQuantifiers,
+    Flip_Conclusion,
+    Loop [
+     Close_Branch,
+     ConjI,
+     King_Cong,
+     Break_Hypotheses,
+     Existential_Free,
+     Existential_Var,
+     Universal,
+     RemoveRedundantQuantifications],
+    CleanUp [RemoveHypothesesFromSkolemDefs, RemoveDuplicates],
+    AbsorbSkolemDefs]
+   []
+*}
+
+ML {*
+fun nonfull_extcnf_combined_tac ctxt feats =
+  extcnf_combined_tac ctxt NONE
+   [ConstsDiff,
+    StripQuantifiers,
+    InnerLoopOnce (Break_Hypotheses :: feats),
+    AbsorbSkolemDefs]
+   []
+*}
+
+consts SEU882_5_bnd_sK1_Xy :: TPTP_Interpret.ind
+lemma
+  "\<forall>SV2. (SEU882_5_bnd_sK1_Xy = SEU882_5_bnd_sK1_Xy) = False \<Longrightarrow>
+   (SEU882_5_bnd_sK1_Xy = SEU882_5_bnd_sK1_Xy) = False"
+(* apply (erule_tac x = "(@X. False)" in allE) *)
+(* apply (tactic {*remove_redundant_quantification 1*}) *)
+(* apply assumption *)
+by (tactic {*nonfull_extcnf_combined_tac @{context} [RemoveRedundantQuantifications, Extuni_FlexRigid]*})
+
+(*NUM667^1*)
+(*
+     (* (Annotated_step ("153", "extuni_triv"), *)
+lemma "((bnd_y = bnd_x) = False \<or> (bnd_z = bnd_z) = False) \<or>
+         (bnd_y = bnd_z) = True \<Longrightarrow>
+         (bnd_y = bnd_x) = False \<or> (bnd_y = bnd_z) = True"
+apply (tactic {*nonfull_extcnf_combined_tac [Extuni_Triv]*})
+done
+     (* (Annotated_step ("162", "extuni_triv"), *)
+lemma "((bnd_y = bnd_x) = False \<or> (bnd_z = bnd_z) = False) \<or>
+         bnd_less bnd_y bnd_z = True \<Longrightarrow>
+         (bnd_y = bnd_x) = False \<or> bnd_less bnd_y bnd_z = True"
+apply (tactic {*nonfull_extcnf_combined_tac [Extuni_Triv]*})
+done
+*)
+
+(* SEU602^2 *)
+consts
+  SEU602_2_bnd_sK7_E :: "(TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind"
+  SEU602_2_bnd_sK2_SY17 :: TPTP_Interpret.ind
+  SEU602_2_bnd_in :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+
+     (* (Annotated_step ("113", "extuni_func"), *)
+lemma "\<forall>SV49\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+            (SV49 =
+             (\<lambda>SY23\<Colon>TPTP_Interpret.ind.
+                 \<not> SEU602_2_bnd_in SY23 SEU602_2_bnd_sK2_SY17)) =
+            False \<Longrightarrow>
+         \<forall>SV49\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+            (SV49 (SEU602_2_bnd_sK7_E SV49) =
+             (\<not> SEU602_2_bnd_in (SEU602_2_bnd_sK7_E SV49) SEU602_2_bnd_sK2_SY17)) =
+            False"
+(*FIXME this (and similar) tests are getting the "Bad background theory of goal state" error since upgrading to Isabelle2013-2.*)
+by (tactic {*fn thm =>
+  let
+    val ctxt =
+      theory_of_thm thm
+      |> Context.Theory
+      |> Context.proof_of
+  in nonfull_extcnf_combined_tac ctxt [Extuni_Func, Existential_Var] thm
+  end*})
+(*by (tactic {*nonfull_extcnf_combined_tac @{context} [Extuni_Func, Existential_Var]*})*)
+oops
+
+consts
+  SEV405_5_bnd_sK1_SY2 :: "(TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind"
+  SEV405_5_bnd_cA :: bool
+
+lemma "\<forall>SV1\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+            (\<forall>SY2\<Colon>TPTP_Interpret.ind.
+                \<not> (\<not> (\<not> SV1 SY2 \<or> SEV405_5_bnd_cA) \<or>
+                   \<not> (\<not> SEV405_5_bnd_cA \<or> SV1 SY2))) =
+            False \<Longrightarrow>
+         \<forall>SV1\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+            (\<not> (\<not> (\<not> SV1 (SEV405_5_bnd_sK1_SY2 SV1) \<or> SEV405_5_bnd_cA) \<or>
+                \<not> (\<not> SEV405_5_bnd_cA \<or> SV1 (SEV405_5_bnd_sK1_SY2 SV1)))) =
+            False"
+by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+(*
+strip quantifiers -- creating a space of permutations; from shallowest to deepest (iterative deepening)
+flip the conclusion -- giving us branch
+apply some collection of rules, in some order, until the space has been explored completely. advantage of not having extcnf_combined: search space is shallow -- particularly if the collection of rules is small.
+*)
+
+consts
+  SEU581_2_bnd_sK1 :: "TPTP_Interpret.ind"
+  SEU581_2_bnd_sK2 :: "TPTP_Interpret.ind \<Rightarrow> bool"
+  SEU581_2_bnd_subset :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> HOL.bool"
+  SEU581_2_bnd_dsetconstr ::  "TPTP_Interpret.ind \<Rightarrow> (TPTP_Interpret.ind \<Rightarrow> HOL.bool) \<Rightarrow> TPTP_Interpret.ind"
+
+(*testing parameters*)
+lemma "! X :: TPTP_Interpret.ind . (\<forall>A B. SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True
+\<Longrightarrow> ! X :: TPTP_Interpret.ind . (\<forall>A B. \<not> SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "(A & B) = True ==> (A | B) = True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "(\<not> bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = True \<Longrightarrow> (bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = False"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+(*testing goals with parameters*)
+lemma "(\<not> bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = True \<Longrightarrow> ! X. (bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = False"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "(A & B) = True ==> (B & A) = True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+(*appreciating differences between THEN, REPEAT, and APPEND*)
+lemma "A & B ==> B & A"
+apply (tactic {*
+  TRY (etac @{thm conjE} 1)
+  THEN TRY (rtac @{thm conjI} 1)*})
+by assumption+
+
+lemma "A & B ==> B & A"
+by (tactic {*
+  etac @{thm conjE} 1
+  THEN rtac @{thm conjI} 1
+  THEN REPEAT (atac 1)*})
+
+lemma "A & B ==> B & A"
+apply (tactic {*
+  rtac @{thm conjI} 1
+  APPEND etac @{thm conjE} 1*})+
+back
+by assumption+
+
+consts
+  SEU581_2_bnd_sK3 :: "TPTP_Interpret.ind"
+  SEU581_2_bnd_sK4 :: "TPTP_Interpret.ind"
+  SEU581_2_bnd_sK5 :: "(TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind"
+  SEU581_2_bnd_powerset :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind"
+  SEU581_2_bnd_in :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+
+consts
+  bnd_c1 :: TPTP_Interpret.ind
+  bnd_s :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind"
+
+lemma "(\<forall>SX0. (\<not> (\<not> SX0 (PUZ114_5_bnd_sK4 SX0) (PUZ114_5_bnd_sK5 SX0) \<or>
+              \<not> (\<not> SX0 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SX0)))
+                    (PUZ114_5_bnd_sK5 SX0) \<or>
+                 \<not> SX0 (bnd_s (PUZ114_5_bnd_sK4 SX0))
+                    (bnd_s (PUZ114_5_bnd_sK5 SX0)))) \<or>
+           \<not> SX0 bnd_c1 bnd_c1) \<or>
+          SX0 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2) =
+   True ==> \<forall>SV1. ((\<not> (\<not> SV1 (PUZ114_5_bnd_sK4 SV1) (PUZ114_5_bnd_sK5 SV1) \<or>
+              \<not> (\<not> SV1 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SV1)))
+                    (PUZ114_5_bnd_sK5 SV1) \<or>
+                 \<not> SV1 (bnd_s (PUZ114_5_bnd_sK4 SV1))
+                    (bnd_s (PUZ114_5_bnd_sK5 SV1)))) \<or>
+           \<not> SV1 bnd_c1 bnd_c1) \<or>
+          SV1 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2) =
+         True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "(\<not> SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1) = True \<Longrightarrow> (SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1) = False"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+(*testing repeated application of simulator*)
+lemma "(\<not> \<not> False) = True \<Longrightarrow>
+    SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1 = True \<or>
+    False = True \<or> False = True \<or> False = True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+(*Testing non-normal conclusion. Ideally we should be able to apply
+  the tactic to arbitrary chains of extcnf steps -- where it's not
+  generally the case that the conclusions are normal.*)
+lemma "(\<not> \<not> False) = True \<Longrightarrow>
+    SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1 = True \<or>
+    (\<not> False) = False"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+(*testing repeated application of simulator, involving different extcnf rules*)
+lemma "(\<not> \<not> (False | False)) = True \<Longrightarrow>
+    SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1 = True \<or>
+    False = True \<or> False = True \<or> False = True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+(*testing logical expansion*)
+lemma "(\<forall>A B. SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True
+\<Longrightarrow> (\<forall>A B. \<not> SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+(*testing extcnf_forall_pos*)
+lemma "(\<forall>A Xphi. SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr A Xphi) (SEU581_2_bnd_powerset A)) = True \<Longrightarrow> \<forall>SV1. (\<forall>SY14.
+             SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr SV1 SY14)
+              (SEU581_2_bnd_powerset SV1)) = True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "((\<forall>A Xphi. SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr A Xphi) (SEU581_2_bnd_powerset A)) = True) | ((~ False) = False) \<Longrightarrow>
+\<forall>SV1. ((\<forall>SY14. SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr SV1 SY14) (SEU581_2_bnd_powerset SV1)) = True) | ((~ False) = False)"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+(*testing parameters*)
+lemma "(\<forall>A B. SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True
+\<Longrightarrow> ! X. (\<forall>A B. \<not> SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "((? A .P1 A) = False) | P2 = True \<Longrightarrow> !X. ((P1 X) = False | P2 = True)"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "((!A . (P1a A | P1b A)) = True) | (P2 = True) \<Longrightarrow> !X. (P1a X = True | P1b X = True | P2 = True)"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "! Y. (((!A .(P1a A | P1b A)) = True) | P2 = True) \<Longrightarrow> ! Y. (!X. (P1a X = True | P1b X = True | P2 = True))"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "! Y. (((!A .(P1a A | P1b A)) = True) | P2 = True) \<Longrightarrow> ! Y. (!X. (P1a X = True | P1b X = True | P2 = True))"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "! Y. (((!A .(P1a A | P1b A)) = True) | P2 = True) \<Longrightarrow> ! Y. (!X. (P1a X = True | P1b X = True | P2 = True))"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+consts dud_bnd_s :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind"
+
+(*this lemma kills blast*)
+lemma "(\<not> (\<forall>SX0 SX1.
+          \<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or> PUZ114_5_bnd_sK3 (dud_bnd_s (dud_bnd_s SX0)) SX1) \<or>
+    \<not> (\<forall>SX0 SX1.
+          \<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or>
+          PUZ114_5_bnd_sK3 (dud_bnd_s SX0) (dud_bnd_s SX1))) =
+   False \<Longrightarrow> (\<not> (\<forall>SX0 SX1.
+          \<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or>
+          PUZ114_5_bnd_sK3 (dud_bnd_s SX0) (dud_bnd_s SX1))) =
+   False"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+(*testing logical expansion -- this should be done by blast*)
+lemma "(\<forall>A B. bnd_in B (bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True
+\<Longrightarrow> (\<forall>A B. \<not> bnd_in B (bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+(*testing related to PUZ114^5.p.out*)
+lemma "\<forall>SV1. ((\<not> (\<not> SV1 (PUZ114_5_bnd_sK4 SV1) (PUZ114_5_bnd_sK5 SV1) \<or>
+                    \<not> (\<not> SV1 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SV1)))
+                          (PUZ114_5_bnd_sK5 SV1) \<or>
+                       \<not> SV1 (bnd_s (PUZ114_5_bnd_sK4 SV1))
+                          (bnd_s (PUZ114_5_bnd_sK5 SV1))))) =
+                True \<or>
+                (\<not> SV1 bnd_c1 bnd_c1) = True) \<or>
+               SV1 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2 = True \<Longrightarrow>
+         \<forall>SV1. (SV1 bnd_c1 bnd_c1 = False \<or>
+                (\<not> (\<not> SV1 (PUZ114_5_bnd_sK4 SV1) (PUZ114_5_bnd_sK5 SV1) \<or>
+                    \<not> (\<not> SV1 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SV1)))
+                          (PUZ114_5_bnd_sK5 SV1) \<or>
+                       \<not> SV1 (bnd_s (PUZ114_5_bnd_sK4 SV1))
+                          (bnd_s (PUZ114_5_bnd_sK5 SV1))))) =
+                True) \<or>
+               SV1 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2 = True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "\<forall>SV2. (\<forall>SY41.
+                   \<not> PUZ114_5_bnd_sK3 SV2 SY41 \<or>
+                   PUZ114_5_bnd_sK3 (dud_bnd_s (dud_bnd_s SV2)) SY41) =
+               True \<Longrightarrow>
+         \<forall>SV4 SV2.
+            (\<not> PUZ114_5_bnd_sK3 SV2 SV4 \<or>
+             PUZ114_5_bnd_sK3 (dud_bnd_s (dud_bnd_s SV2)) SV4) =
+            True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+lemma "\<forall>SV3. (\<forall>SY42.
+                   \<not> PUZ114_5_bnd_sK3 SV3 SY42 \<or>
+                   PUZ114_5_bnd_sK3 (dud_bnd_s SV3) (dud_bnd_s SY42)) =
+               True \<Longrightarrow>
+         \<forall>SV5 SV3.
+            (\<not> PUZ114_5_bnd_sK3 SV3 SV5 \<or>
+             PUZ114_5_bnd_sK3 (dud_bnd_s SV3) (dud_bnd_s SV5)) =
+            True"
+by (tactic {*full_extcnf_combined_tac @{context}*})
+
+
+subsection "unfold_def"
+     (* (Annotated_step ("9", "unfold_def"), *)
+lemma "bnd_kpairiskpair =
+             (ALL Xx Xy.
+                 bnd_iskpair
+                  (bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset)
+                    (bnd_setadjoin (bnd_setadjoin Xx (bnd_setadjoin Xy bnd_emptyset))
+                      bnd_emptyset))) &
+             bnd_kpair =
+             (%Xx Xy.
+                 bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset)
+                  (bnd_setadjoin (bnd_setadjoin Xx (bnd_setadjoin Xy bnd_emptyset))
+                    bnd_emptyset)) &
+             bnd_iskpair =
+             (%A. EX Xx. bnd_in Xx (bnd_setunion A) &
+                         (EX Xy. bnd_in Xy (bnd_setunion A) &
+                                 A = bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset)
+                                      (bnd_setadjoin
+                                        (bnd_setadjoin Xx
+(bnd_setadjoin Xy bnd_emptyset))
+                                        bnd_emptyset))) &
+             (~ (ALL SY0 SY1.
+                    EX SY3.
+                       bnd_in SY3
+                        (bnd_setunion
+                          (bnd_setadjoin (bnd_setadjoin SY0 bnd_emptyset)
+                            (bnd_setadjoin
+                              (bnd_setadjoin SY0 (bnd_setadjoin SY1 bnd_emptyset))
+                              bnd_emptyset))) &
+                       (EX SY4.
+                           bnd_in SY4
+                            (bnd_setunion
+                              (bnd_setadjoin (bnd_setadjoin SY0 bnd_emptyset)
+                                (bnd_setadjoin
+                                  (bnd_setadjoin SY0
+                                    (bnd_setadjoin SY1 bnd_emptyset))
+                                  bnd_emptyset))) &
+                           bnd_setadjoin (bnd_setadjoin SY0 bnd_emptyset)
+                            (bnd_setadjoin
+                              (bnd_setadjoin SY0 (bnd_setadjoin SY1 bnd_emptyset))
+                              bnd_emptyset) =
+                           bnd_setadjoin (bnd_setadjoin SY3 bnd_emptyset)
+                            (bnd_setadjoin
+                              (bnd_setadjoin SY3 (bnd_setadjoin SY4 bnd_emptyset))
+                              bnd_emptyset)))) =
+             True
+             ==> (~ (ALL SX0 SX1.
+                        ~ (ALL SX2.
+                              ~ ~ (~ bnd_in SX2
+                                      (bnd_setunion
+                                        (bnd_setadjoin
+(bnd_setadjoin SX0 bnd_emptyset)
+(bnd_setadjoin (bnd_setadjoin SX0 (bnd_setadjoin SX1 bnd_emptyset)) bnd_emptyset))) |
+                                   ~ ~ (ALL SX3.
+ ~ ~ (~ bnd_in SX3
+         (bnd_setunion
+           (bnd_setadjoin (bnd_setadjoin SX0 bnd_emptyset)
+             (bnd_setadjoin (bnd_setadjoin SX0 (bnd_setadjoin SX1 bnd_emptyset))
+               bnd_emptyset))) |
+      bnd_setadjoin (bnd_setadjoin SX0 bnd_emptyset)
+       (bnd_setadjoin (bnd_setadjoin SX0 (bnd_setadjoin SX1 bnd_emptyset))
+         bnd_emptyset) ~=
+      bnd_setadjoin (bnd_setadjoin SX2 bnd_emptyset)
+       (bnd_setadjoin (bnd_setadjoin SX2 (bnd_setadjoin SX3 bnd_emptyset))
+         bnd_emptyset))))))) =
+                 True"
+by (tactic {*unfold_def_tac @{context} []*})
+
+     (* (Annotated_step ("10", "unfold_def"), *)
+lemma "bnd_kpairiskpair =
+             (ALL Xx Xy.
+                 bnd_iskpair
+                  (bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset)
+                    (bnd_setadjoin (bnd_setadjoin Xx (bnd_setadjoin Xy bnd_emptyset))
+                      bnd_emptyset))) &
+             bnd_kpair =
+             (%Xx Xy.
+                 bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset)
+                  (bnd_setadjoin (bnd_setadjoin Xx (bnd_setadjoin Xy bnd_emptyset))
+                    bnd_emptyset)) &
+             bnd_iskpair =
+             (%A. EX Xx. bnd_in Xx (bnd_setunion A) &
+                         (EX Xy. bnd_in Xy (bnd_setunion A) &
+                                 A = bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset)
+                                      (bnd_setadjoin
+                                        (bnd_setadjoin Xx
+(bnd_setadjoin Xy bnd_emptyset))
+                                        bnd_emptyset))) &
+             (ALL SY5 SY6.
+                 EX SY7.
+                    bnd_in SY7
+                     (bnd_setunion
+                       (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset)
+                         (bnd_setadjoin
+                           (bnd_setadjoin SY5 (bnd_setadjoin SY6 bnd_emptyset))
+                           bnd_emptyset))) &
+                    (EX SY8.
+                        bnd_in SY8
+                         (bnd_setunion
+                           (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset)
+                             (bnd_setadjoin
+                               (bnd_setadjoin SY5 (bnd_setadjoin SY6 bnd_emptyset))
+                               bnd_emptyset))) &
+                        bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset)
+                         (bnd_setadjoin
+                           (bnd_setadjoin SY5 (bnd_setadjoin SY6 bnd_emptyset))
+                           bnd_emptyset) =
+                        bnd_setadjoin (bnd_setadjoin SY7 bnd_emptyset)
+                         (bnd_setadjoin
+                           (bnd_setadjoin SY7 (bnd_setadjoin SY8 bnd_emptyset))
+                           bnd_emptyset))) =
+             True
+             ==> (ALL SX0 SX1.
+                     ~ (ALL SX2.
+                           ~ ~ (~ bnd_in SX2
+                                   (bnd_setunion
+                                     (bnd_setadjoin (bnd_setadjoin SX0 bnd_emptyset)
+                                       (bnd_setadjoin
+                                         (bnd_setadjoin SX0
+ (bnd_setadjoin SX1 bnd_emptyset))
+                                         bnd_emptyset))) |
+                                ~ ~ (ALL SX3.
+                                        ~ ~ (~ bnd_in SX3
+      (bnd_setunion
+        (bnd_setadjoin (bnd_setadjoin SX0 bnd_emptyset)
+          (bnd_setadjoin (bnd_setadjoin SX0 (bnd_setadjoin SX1 bnd_emptyset))
+            bnd_emptyset))) |
+   bnd_setadjoin (bnd_setadjoin SX0 bnd_emptyset)
+    (bnd_setadjoin (bnd_setadjoin SX0 (bnd_setadjoin SX1 bnd_emptyset))
+      bnd_emptyset) ~=
+   bnd_setadjoin (bnd_setadjoin SX2 bnd_emptyset)
+    (bnd_setadjoin (bnd_setadjoin SX2 (bnd_setadjoin SX3 bnd_emptyset))
+      bnd_emptyset)))))) =
+                 True"
+by (tactic {*unfold_def_tac @{context} []*})
+
+     (* (Annotated_step ("12", "unfold_def"), *)
+lemma "bnd_cCKB6_BLACK =
+         (\<lambda>Xu Xv.
+             \<forall>Xw. Xw bnd_c1 bnd_c1 \<and>
+                  (\<forall>Xj Xk.
+                      Xw Xj Xk \<longrightarrow>
+                      Xw (bnd_s (bnd_s Xj)) Xk \<and>
+                      Xw (bnd_s Xj) (bnd_s Xk)) \<longrightarrow>
+                  Xw Xu Xv) \<and>
+         ((((\<forall>SY36 SY37.
+                \<not> PUZ114_5_bnd_sK3 SY36 SY37 \<or>
+                PUZ114_5_bnd_sK3 (bnd_s (bnd_s SY36)) SY37) \<and>
+            (\<forall>SY38 SY39.
+                \<not> PUZ114_5_bnd_sK3 SY38 SY39 \<or>
+                PUZ114_5_bnd_sK3 (bnd_s SY38) (bnd_s SY39))) \<and>
+           PUZ114_5_bnd_sK3 bnd_c1 bnd_c1) \<and>
+          \<not> PUZ114_5_bnd_sK3 (bnd_s (bnd_s (bnd_s PUZ114_5_bnd_sK1)))
+             (bnd_s PUZ114_5_bnd_sK2)) =
+         True \<Longrightarrow>
+         (\<not> (\<not> \<not> (\<not> \<not> (\<not> (\<forall>SX0 SX1.
+                             \<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or>
+                             PUZ114_5_bnd_sK3 (bnd_s (bnd_s SX0)) SX1) \<or>
+                       \<not> (\<forall>SX0 SX1.
+                             \<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or>
+                             PUZ114_5_bnd_sK3 (bnd_s SX0) (bnd_s SX1))) \<or>
+                  \<not> PUZ114_5_bnd_sK3 bnd_c1 bnd_c1) \<or>
+             \<not> \<not> PUZ114_5_bnd_sK3 (bnd_s (bnd_s (bnd_s PUZ114_5_bnd_sK1)))
+                  (bnd_s PUZ114_5_bnd_sK2))) =
+         True"
+(*
+apply (erule conjE)+
+apply (erule subst)+
+apply (tactic {*log_expander 1*})+
+apply (rule refl)
+*)
+by (tactic {*unfold_def_tac @{context} []*})
+
+     (* (Annotated_step ("13", "unfold_def"), *)
+lemma "bnd_cCKB6_BLACK =
+         (\<lambda>Xu Xv.
+             \<forall>Xw. Xw bnd_c1 bnd_c1 \<and>
+                  (\<forall>Xj Xk.
+                      Xw Xj Xk \<longrightarrow>
+                      Xw (bnd_s (bnd_s Xj)) Xk \<and>
+                      Xw (bnd_s Xj) (bnd_s Xk)) \<longrightarrow>
+                  Xw Xu Xv) \<and>
+         (\<forall>SY30.
+             (SY30 (PUZ114_5_bnd_sK4 SY30) (PUZ114_5_bnd_sK5 SY30) \<and>
+              (\<not> SY30 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SY30)))
+                  (PUZ114_5_bnd_sK5 SY30) \<or>
+               \<not> SY30 (bnd_s (PUZ114_5_bnd_sK4 SY30))
+                  (bnd_s (PUZ114_5_bnd_sK5 SY30))) \<or>
+              \<not> SY30 bnd_c1 bnd_c1) \<or>
+             SY30 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2) =
+         True \<Longrightarrow>
+         (\<forall>SX0. (\<not> (\<not> SX0 (PUZ114_5_bnd_sK4 SX0) (PUZ114_5_bnd_sK5 SX0) \<or>
+                    \<not> (\<not> SX0 (bnd_s (bnd_s (PUZ114_5_bnd_sK4 SX0)))
+                          (PUZ114_5_bnd_sK5 SX0) \<or>
+                       \<not> SX0 (bnd_s (PUZ114_5_bnd_sK4 SX0))
+                          (bnd_s (PUZ114_5_bnd_sK5 SX0)))) \<or>
+                 \<not> SX0 bnd_c1 bnd_c1) \<or>
+                SX0 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2) =
+         True"
+(*
+apply (erule conjE)+
+apply (tactic {*expander_animal 1*})+
+apply assumption
+*)
+by (tactic {*unfold_def_tac @{context} []*})
+
+(*FIXME move this heuristic elsewhere*)
+ML {*
+(*Other than the list (which must not be empty) this function
+  expects a parameter indicating the smallest integer.
+  (Using Int.minInt isn't always viable).*)
+fun max_int_floored min l =
+  if null l then raise List.Empty
+  else fold (curry Int.max) l min;
+
+val _ = @{assert} (max_int_floored ~101002 [1]  = 1)
+val _ = @{assert} (max_int_floored 0 [1, 3, 5] = 5)
+
+fun max_index_floored min l =
+  let
+    val max = max_int_floored min l
+  in find_index (pair max #> op =) l end
+*}
+
+ML {*
+max_index_floored 0 [1, 3, 5]
+*}
+
+ML {*
+(*
+Given argument ([h_1, ..., h_n], conc),
+obtained from term of form
+  h_1 ==> ... ==> h_n ==> conclusion,
+this function indicates which h_i is biggest,
+or NONE if h_n = 0.
+*)
+fun biggest_hypothesis (hypos, _) =
+  if null hypos then NONE
+  else
+    map size_of_term hypos
+    |> max_index_floored 0
+    |> SOME
+*}
+
+ML {*
+fun biggest_hyp_first_tac i = fn st =>
+  let
+    val results = TERMFUN biggest_hypothesis (SOME i) st
+val _ = tracing ("result=" ^ PolyML.makestring results)
+  in
+    if null results then no_tac st
+    else
+      let
+        val result = the_single results
+      in
+        case result of
+            NONE => no_tac st
+          | SOME n =>
+              if n > 0 then rotate_tac n i st else no_tac st
+      end
+  end
+*}
+
+     (* (Annotated_step ("6", "unfold_def"), *)
+lemma  "(\<not> (\<exists>U :: TPTP_Interpret.ind \<Rightarrow> bool. \<forall>V. U V = SEV405_5_bnd_cA)) = True \<Longrightarrow>
+         (\<not> \<not> (\<forall>SX0 :: TPTP_Interpret.ind \<Rightarrow> bool. \<not> (\<forall>SX1. \<not> (\<not> (\<not> SX0 SX1 \<or> SEV405_5_bnd_cA) \<or>
+ \<not> (\<not> SEV405_5_bnd_cA \<or> SX0 SX1))))) =
+         True"
+(* by (tactic {*unfold_def_tac []*}) *)
+oops
+
+subsection "Using leo2_tac"
+(*these require PUZ114^5's proof to be loaded
+
+ML {*leo2_tac @{context} (hd prob_names) "50"*}
+
+ML {*leo2_tac @{context} (hd prob_names) "4"*}
+
+ML {*leo2_tac @{context} (hd prob_names) "9"*}
+
+     (* (Annotated_step ("9", "extcnf_combined"), *)
+lemma "(\<forall>SY30.
+             SY30 bnd_c1 bnd_c1 \<and>
+             (\<forall>Xj Xk.
+                 SY30 Xj Xk \<longrightarrow>
+                 SY30 (bnd_s (bnd_s Xj)) Xk \<and>
+                 SY30 (bnd_s Xj) (bnd_s Xk)) \<longrightarrow>
+             SY30 bnd_sK1 bnd_sK2) =
+         True \<Longrightarrow>
+         (\<forall>SY30.
+             (SY30 (bnd_sK4 SY30) (bnd_sK5 SY30) \<and>
+              (\<not> SY30 (bnd_s (bnd_s (bnd_sK4 SY30)))
+                  (bnd_sK5 SY30) \<or>
+               \<not> SY30 (bnd_s (bnd_sK4 SY30))
+                  (bnd_s (bnd_sK5 SY30))) \<or>
+              \<not> SY30 bnd_c1 bnd_c1) \<or>
+             SY30 bnd_sK1 bnd_sK2) =
+         True"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "9") 1*})
+*)
+
+
+
+typedecl GEG007_1_bnd_reg
+consts
+  GEG007_1_bnd_sK7_SX2 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> GEG007_1_bnd_reg"
+  GEG007_1_bnd_sK6_SX2 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> GEG007_1_bnd_reg"
+  GEG007_1_bnd_a :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+  GEG007_1_bnd_catalunya  :: "GEG007_1_bnd_reg"
+  GEG007_1_bnd_spain :: "GEG007_1_bnd_reg"
+  GEG007_1_bnd_c :: "GEG007_1_bnd_reg \<Rightarrow> GEG007_1_bnd_reg \<Rightarrow> bool"
+
+     (* (Annotated_step ("147", "extcnf_forall_neg"), *)
+lemma "\<forall>SV13 SV6.
+            (\<forall>SX2. \<not> GEG007_1_bnd_c SX2 GEG007_1_bnd_spain \<or>
+                   GEG007_1_bnd_c SX2 GEG007_1_bnd_catalunya) =
+            False \<or>
+            GEG007_1_bnd_a SV6 SV13 = False \<Longrightarrow>
+         \<forall>SV6 SV13.
+            (\<not> GEG007_1_bnd_c (GEG007_1_bnd_sK7_SX2 SV13 SV6) GEG007_1_bnd_spain \<or>
+             GEG007_1_bnd_c (GEG007_1_bnd_sK7_SX2 SV13 SV6) GEG007_1_bnd_catalunya) =
+            False \<or>
+            GEG007_1_bnd_a SV6 SV13 = False"
+by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+
+     (* (Annotated_step ("116", "extcnf_forall_neg"), *)
+lemma "\<forall>SV13 SV6.
+            (\<forall>SX2. \<not> \<not> (\<not> \<not> (\<not> GEG007_1_bnd_c SX2 GEG007_1_bnd_catalunya \<or>
+                             \<not> \<not> \<not> (\<forall>SX3.
+       \<not> \<not> (\<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SX3 \<or> GEG007_1_bnd_c SX4 SX2) \<or>
+            \<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SX3 \<or>
+                     GEG007_1_bnd_c SX4 GEG007_1_bnd_catalunya)))) \<or>
+                        \<not> \<not> (\<not> GEG007_1_bnd_c SX2 GEG007_1_bnd_spain \<or>
+                             \<not> \<not> \<not> (\<forall>SX3.
+       \<not> \<not> (\<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SX3 \<or> GEG007_1_bnd_c SX4 SX2) \<or>
+            \<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SX3 \<or>
+                     GEG007_1_bnd_c SX4 GEG007_1_bnd_spain)))))) =
+            False \<or>
+            GEG007_1_bnd_a SV6 SV13 = False \<Longrightarrow>
+         \<forall>SV6 SV13.
+            (\<not> \<not> (\<not> \<not> (\<not> GEG007_1_bnd_c (GEG007_1_bnd_sK6_SX2 SV13 SV6)
+                          GEG007_1_bnd_catalunya \<or>
+                       \<not> \<not> \<not> (\<forall>SY68.
+ \<not> \<not> (\<not> (\<forall>SY69.
+            \<not> GEG007_1_bnd_c SY69 SY68 \<or>
+            GEG007_1_bnd_c SY69 (GEG007_1_bnd_sK6_SX2 SV13 SV6)) \<or>
+      \<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SY68 \<or> GEG007_1_bnd_c SX4 GEG007_1_bnd_catalunya)))) \<or>
+                  \<not> \<not> (\<not> GEG007_1_bnd_c (GEG007_1_bnd_sK6_SX2 SV13 SV6)
+                          GEG007_1_bnd_spain \<or>
+                       \<not> \<not> \<not> (\<forall>SY71.
+ \<not> \<not> (\<not> (\<forall>SY72.
+            \<not> GEG007_1_bnd_c SY72 SY71 \<or>
+            GEG007_1_bnd_c SY72 (GEG007_1_bnd_sK6_SX2 SV13 SV6)) \<or>
+      \<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SY71 \<or> GEG007_1_bnd_c SX4 GEG007_1_bnd_spain)))))) =
+            False \<or>
+            GEG007_1_bnd_a SV6 SV13 = False"
+by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+
+consts PUZ107_5_bnd_sK1_SX0 ::
+  "TPTP_Interpret.ind
+      \<Rightarrow> TPTP_Interpret.ind
+        \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+
+lemma "\<forall>(SV4\<Colon>TPTP_Interpret.ind) (SV8\<Colon>TPTP_Interpret.ind)
+   (SV6\<Colon>TPTP_Interpret.ind) (SV2\<Colon>TPTP_Interpret.ind)
+   (SV3\<Colon>TPTP_Interpret.ind) SV1\<Colon>TPTP_Interpret.ind.
+   ((SV1 \<noteq> SV3) = False \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or>
+   PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False \<Longrightarrow>
+\<forall>(SV4\<Colon>TPTP_Interpret.ind) (SV8\<Colon>TPTP_Interpret.ind)
+   (SV6\<Colon>TPTP_Interpret.ind) (SV2\<Colon>TPTP_Interpret.ind)
+   (SV3\<Colon>TPTP_Interpret.ind) SV1\<Colon>TPTP_Interpret.ind.
+   ((SV1 = SV3) = True \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or>
+   PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False"
+by (tactic {*nonfull_extcnf_combined_tac @{context} [Not_neg]*})
+
+lemma "
+\<forall>(SV8\<Colon>TPTP_Interpret.ind) (SV6\<Colon>TPTP_Interpret.ind)
+   (SV4\<Colon>TPTP_Interpret.ind) (SV2\<Colon>TPTP_Interpret.ind)
+   (SV3\<Colon>TPTP_Interpret.ind) SV1\<Colon>TPTP_Interpret.ind.
+   ((SV1 \<noteq> SV3 \<or> SV2 \<noteq> SV4) = False \<or>
+    PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or>
+   PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False \<Longrightarrow>
+\<forall>(SV4\<Colon>TPTP_Interpret.ind) (SV8\<Colon>TPTP_Interpret.ind)
+   (SV6\<Colon>TPTP_Interpret.ind) (SV2\<Colon>TPTP_Interpret.ind)
+   (SV3\<Colon>TPTP_Interpret.ind) SV1\<Colon>TPTP_Interpret.ind.
+   ((SV1 \<noteq> SV3) = False \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or>
+   PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False"
+by (tactic {*nonfull_extcnf_combined_tac @{context} [Or_neg]*})
+
+consts
+  NUM016_5_bnd_a :: TPTP_Interpret.ind
+  NUM016_5_bnd_prime :: "TPTP_Interpret.ind \<Rightarrow> bool"
+  NUM016_5_bnd_factorial_plus_one :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind"
+  NUM016_5_bnd_prime_divisor :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind"
+  NUM016_5_bnd_divides :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+  NUM016_5_bnd_less :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+
+     (* (Annotated_step ("6", "unfold_def"), *)
+lemma "((((((((((((\<forall>X\<Colon>TPTP_Interpret.ind. \<not> NUM016_5_bnd_less X X) \<and>
+                    (\<forall>(X\<Colon>TPTP_Interpret.ind)
+                        Y\<Colon>TPTP_Interpret.ind.
+                        \<not> NUM016_5_bnd_less X Y \<or> \<not> NUM016_5_bnd_less Y X)) \<and>
+                   (\<forall>X\<Colon>TPTP_Interpret.ind. NUM016_5_bnd_divides X X)) \<and>
+                  (\<forall>(X\<Colon>TPTP_Interpret.ind)
+                      (Y\<Colon>TPTP_Interpret.ind)
+                      Z\<Colon>TPTP_Interpret.ind.
+                      (\<not> NUM016_5_bnd_divides X Y \<or> \<not> NUM016_5_bnd_divides Y Z) \<or>
+                      NUM016_5_bnd_divides X Z)) \<and>
+                 (\<forall>(X\<Colon>TPTP_Interpret.ind) Y\<Colon>TPTP_Interpret.ind.
+                     \<not> NUM016_5_bnd_divides X Y \<or> \<not> NUM016_5_bnd_less Y X)) \<and>
+                (\<forall>X\<Colon>TPTP_Interpret.ind.
+                    NUM016_5_bnd_less X (NUM016_5_bnd_factorial_plus_one X))) \<and>
+               (\<forall>(X\<Colon>TPTP_Interpret.ind) Y\<Colon>TPTP_Interpret.ind.
+                   \<not> NUM016_5_bnd_divides X (NUM016_5_bnd_factorial_plus_one Y) \<or>
+                   NUM016_5_bnd_less Y X)) \<and>
+              (\<forall>X\<Colon>TPTP_Interpret.ind.
+                  NUM016_5_bnd_prime X \<or>
+                  NUM016_5_bnd_divides (NUM016_5_bnd_prime_divisor X) X)) \<and>
+             (\<forall>X\<Colon>TPTP_Interpret.ind.
+                 NUM016_5_bnd_prime X \<or>
+                 NUM016_5_bnd_prime (NUM016_5_bnd_prime_divisor X))) \<and>
+            (\<forall>X\<Colon>TPTP_Interpret.ind.
+                NUM016_5_bnd_prime X \<or>
+                NUM016_5_bnd_less (NUM016_5_bnd_prime_divisor X) X)) \<and>
+           NUM016_5_bnd_prime NUM016_5_bnd_a) \<and>
+          (\<forall>X\<Colon>TPTP_Interpret.ind.
+              (\<not> NUM016_5_bnd_prime X \<or> \<not> NUM016_5_bnd_less NUM016_5_bnd_a X) \<or>
+              NUM016_5_bnd_less (NUM016_5_bnd_factorial_plus_one NUM016_5_bnd_a) X)) =
+         True \<Longrightarrow>
+         (\<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+     \<not> NUM016_5_bnd_less SX0 SX0) \<or>
+                               \<not> (\<forall>(SX0\<Colon>TPTP_Interpret.ind)
+     SX1\<Colon>TPTP_Interpret.ind.
+     \<not> NUM016_5_bnd_less SX0 SX1 \<or> \<not> NUM016_5_bnd_less SX1 SX0)) \<or>
+                          \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+NUM016_5_bnd_divides SX0 SX0)) \<or>
+                     \<not> (\<forall>(SX0\<Colon>TPTP_Interpret.ind)
+                           (SX1\<Colon>TPTP_Interpret.ind)
+                           SX2\<Colon>TPTP_Interpret.ind.
+                           (\<not> NUM016_5_bnd_divides SX0 SX1 \<or>
+                            \<not> NUM016_5_bnd_divides SX1 SX2) \<or>
+                           NUM016_5_bnd_divides SX0 SX2)) \<or>
+                \<not> (\<forall>(SX0\<Colon>TPTP_Interpret.ind)
+                      SX1\<Colon>TPTP_Interpret.ind.
+                      \<not> NUM016_5_bnd_divides SX0 SX1 \<or>
+                      \<not> NUM016_5_bnd_less SX1 SX0)) \<or>
+           \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+                 NUM016_5_bnd_less SX0 (NUM016_5_bnd_factorial_plus_one SX0))) \<or>
+      \<not> (\<forall>(SX0\<Colon>TPTP_Interpret.ind) SX1\<Colon>TPTP_Interpret.ind.
+            \<not> NUM016_5_bnd_divides SX0 (NUM016_5_bnd_factorial_plus_one SX1) \<or>
+            NUM016_5_bnd_less SX1 SX0)) \<or>
+ \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+       NUM016_5_bnd_prime SX0 \<or>
+       NUM016_5_bnd_divides (NUM016_5_bnd_prime_divisor SX0) SX0)) \<or>
+                            \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+  NUM016_5_bnd_prime SX0 \<or> NUM016_5_bnd_prime (NUM016_5_bnd_prime_divisor SX0))) \<or>
+                       \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+                             NUM016_5_bnd_prime SX0 \<or>
+                             NUM016_5_bnd_less (NUM016_5_bnd_prime_divisor SX0)
+                              SX0)) \<or>
+                  \<not> NUM016_5_bnd_prime NUM016_5_bnd_a) \<or>
+             \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+                   (\<not> NUM016_5_bnd_prime SX0 \<or> \<not> NUM016_5_bnd_less NUM016_5_bnd_a SX0) \<or>
+                   NUM016_5_bnd_less (NUM016_5_bnd_factorial_plus_one NUM016_5_bnd_a)
+                    SX0))) =
+         True"
+by (tactic {*unfold_def_tac @{context} []*})
+
+     (* (Annotated_step ("3", "unfold_def"), *)
+lemma "(~ ((((((((((((ALL X. ~ bnd_less X X) &
+                           (ALL X Y. ~ bnd_less X Y | ~ bnd_less Y X)) &
+                          (ALL X. bnd_divides X X)) &
+                         (ALL X Y Z.
+                             (~ bnd_divides X Y | ~ bnd_divides Y Z) |
+                             bnd_divides X Z)) &
+                        (ALL X Y. ~ bnd_divides X Y | ~ bnd_less Y X)) &
+                       (ALL X. bnd_less X (bnd_factorial_plus_one X))) &
+                      (ALL X Y.
+                          ~ bnd_divides X (bnd_factorial_plus_one Y) |
+                          bnd_less Y X)) &
+                     (ALL X. bnd_prime X | bnd_divides (bnd_prime_divisor X) X)) &
+                    (ALL X. bnd_prime X | bnd_prime (bnd_prime_divisor X))) &
+                   (ALL X. bnd_prime X | bnd_less (bnd_prime_divisor X) X)) &
+                  bnd_prime bnd_a) &
+                 (ALL X. (~ bnd_prime X | ~ bnd_less bnd_a X) |
+                         bnd_less (bnd_factorial_plus_one bnd_a) X))) =
+             False
+             ==> (~ ((((((((((((ALL X. ~ bnd_less X X) &
+                               (ALL X Y. ~ bnd_less X Y | ~ bnd_less Y X)) &
+                              (ALL X. bnd_divides X X)) &
+                             (ALL X Y Z.
+                                 (~ bnd_divides X Y | ~ bnd_divides Y Z) |
+                                 bnd_divides X Z)) &
+                            (ALL X Y. ~ bnd_divides X Y | ~ bnd_less Y X)) &
+                           (ALL X. bnd_less X (bnd_factorial_plus_one X))) &
+                          (ALL X Y.
+                              ~ bnd_divides X (bnd_factorial_plus_one Y) |
+                              bnd_less Y X)) &
+                         (ALL X. bnd_prime X |
+                                 bnd_divides (bnd_prime_divisor X) X)) &
+                        (ALL X. bnd_prime X | bnd_prime (bnd_prime_divisor X))) &
+                       (ALL X. bnd_prime X | bnd_less (bnd_prime_divisor X) X)) &
+                      bnd_prime bnd_a) &
+                     (ALL X. (~ bnd_prime X | ~ bnd_less bnd_a X) |
+                             bnd_less (bnd_factorial_plus_one bnd_a) X))) =
+                 False"
+by (tactic {*unfold_def_tac @{context} []*})
+
+(* SET062^6.p.out
+      [[(Annotated_step ("3", "unfold_def"), *)
+lemma "(\<forall>Z3. False \<longrightarrow> bnd_cA Z3) = False \<Longrightarrow>
+         (\<forall>Z3. False \<longrightarrow> bnd_cA Z3) = False"
+by (tactic {*unfold_def_tac @{context} []*})
+
+(*
+(* SEU559^2.p.out *)
+   (* [[(Annotated_step ("3", "unfold_def"), *)
+lemma "bnd_subset = (\<lambda>A B. \<forall>Xx. bnd_in Xx A \<longrightarrow> bnd_in Xx B) \<and>
+         (\<forall>A B. (\<forall>Xx. bnd_in Xx A \<longrightarrow> bnd_in Xx B) \<longrightarrow>
+                bnd_subset A B) =
+         False \<Longrightarrow>
+         (\<forall>SY0 SY1.
+             (\<forall>Xx. bnd_in Xx SY0 \<longrightarrow> bnd_in Xx SY1) \<longrightarrow>
+             (\<forall>SY5. bnd_in SY5 SY0 \<longrightarrow> bnd_in SY5 SY1)) =
+         False"
+by (tactic {*unfold_def_tac [@{thm bnd_subset_def}]*})
+
+(* SEU559^2.p.out
+    [[(Annotated_step ("6", "unfold_def"), *)
+lemma "(\<not> (\<exists>Xx. \<forall>Xy. Xx \<longrightarrow> Xy)) = True \<Longrightarrow>
+         (\<not> \<not> (\<forall>SX0. \<not> (\<forall>SX1. \<not> SX0 \<or> SX1))) = True"
+by (tactic {*unfold_def_tac []*})
+
+(* SEU502^2.p.out
+    [[(Annotated_step ("3", "unfold_def"), *)
+lemma "bnd_emptysetE =
+         (\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<and>
+         (bnd_emptysetE \<longrightarrow>
+          (\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> False)) =
+         False \<Longrightarrow>
+         ((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow>
+          (\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> False)) =
+         False"
+by (tactic {*unfold_def_tac [@{thm bnd_emptysetE_def}]*})
+*)
+
+typedecl AGT037_2_bnd_mu
+consts
+  AGT037_2_bnd_sK1_SX0 :: TPTP_Interpret.ind
+  AGT037_2_bnd_cola :: AGT037_2_bnd_mu
+  AGT037_2_bnd_jan :: AGT037_2_bnd_mu
+  AGT037_2_bnd_possibly_likes :: "AGT037_2_bnd_mu \<Rightarrow> AGT037_2_bnd_mu \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+  AGT037_2_bnd_sK5_SY68 ::
+    "TPTP_Interpret.ind
+     \<Rightarrow> AGT037_2_bnd_mu
+       \<Rightarrow> AGT037_2_bnd_mu
+         \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind"
+  AGT037_2_bnd_likes :: "AGT037_2_bnd_mu \<Rightarrow> AGT037_2_bnd_mu \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+  AGT037_2_bnd_very_much_likes :: "AGT037_2_bnd_mu \<Rightarrow> AGT037_2_bnd_mu \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+  AGT037_2_bnd_a1 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+  AGT037_2_bnd_a2 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+  AGT037_2_bnd_a3 :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
+
+(*test that nullary skolem terms are OK*)
+     (* (Annotated_step ("79", "extcnf_forall_neg"), *)
+lemma "(\<forall>SX0\<Colon>TPTP_Interpret.ind.
+             AGT037_2_bnd_possibly_likes AGT037_2_bnd_jan AGT037_2_bnd_cola SX0) =
+         False \<Longrightarrow>
+         AGT037_2_bnd_possibly_likes AGT037_2_bnd_jan AGT037_2_bnd_cola AGT037_2_bnd_sK1_SX0 =
+         False"
+by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+
+     (* (Annotated_step ("202", "extcnf_forall_neg"), *)
+lemma "\<forall>(SV13\<Colon>TPTP_Interpret.ind) (SV39\<Colon>AGT037_2_bnd_mu) (SV29\<Colon>AGT037_2_bnd_mu)
+            SV45\<Colon>TPTP_Interpret.ind.
+            ((((\<forall>SY68\<Colon>TPTP_Interpret.ind.
+                   \<not> AGT037_2_bnd_a1 SV45 SY68 \<or>
+                   AGT037_2_bnd_likes SV29 SV39 SY68) =
+               False \<or>
+               (\<not> (\<forall>SY69\<Colon>TPTP_Interpret.ind.
+                      \<not> AGT037_2_bnd_a2 SV45 SY69 \<or>
+                      AGT037_2_bnd_likes SV29 SV39 SY69)) =
+               True) \<or>
+              AGT037_2_bnd_likes SV29 SV39 SV45 = False) \<or>
+             AGT037_2_bnd_very_much_likes SV29 SV39 SV45 = True) \<or>
+            AGT037_2_bnd_a3 SV13 SV45 = False \<Longrightarrow>
+         \<forall>(SV29\<Colon>AGT037_2_bnd_mu) (SV39\<Colon>AGT037_2_bnd_mu) (SV13\<Colon>TPTP_Interpret.ind)
+            SV45\<Colon>TPTP_Interpret.ind.
+            ((((\<not> AGT037_2_bnd_a1 SV45
+                   (AGT037_2_bnd_sK5_SY68 SV13 SV39 SV29 SV45) \<or>
+                AGT037_2_bnd_likes SV29 SV39
+                 (AGT037_2_bnd_sK5_SY68 SV13 SV39 SV29 SV45)) =
+               False \<or>
+               (\<not> (\<forall>SY69\<Colon>TPTP_Interpret.ind.
+                      \<not> AGT037_2_bnd_a2 SV45 SY69 \<or>
+                      AGT037_2_bnd_likes SV29 SV39 SY69)) =
+               True) \<or>
+              AGT037_2_bnd_likes SV29 SV39 SV45 = False) \<or>
+             AGT037_2_bnd_very_much_likes SV29 SV39 SV45 = True) \<or>
+            AGT037_2_bnd_a3 SV13 SV45 = False"
+(*
+apply (rule allI)+
+apply (erule_tac x = "SV13" in allE)
+apply (erule_tac x = "SV39" in allE)
+apply (erule_tac x = "SV29" in allE)
+apply (erule_tac x = "SV45" in allE)
+apply (erule disjE)+
+defer
+apply (tactic {*clause_breaker 1*})+
+apply (drule_tac sk = "bnd_sK5_SY68 SV13 SV39 SV29 SV45" in leo2_skolemise)
+defer
+apply (tactic {*clause_breaker 1*})
+apply (tactic {*nonfull_extcnf_combined_tac []*})
+*)
+by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+
+(*(*NUM667^1*)
+lemma "\<forall>SV12 SV13 SV14 SV9 SV10 SV11.
+   ((((bnd_less SV12 SV13 = bnd_less SV11 SV10) = False \<or>
+      (SV14 = SV13) = False) \<or>
+     bnd_less SV12 SV14 = False) \<or>
+    bnd_less SV9 SV10 = True) \<or>
+   (SV9 = SV11) = False \<Longrightarrow>
+\<forall>SV9 SV14 SV10 SV11 SV13 SV12.
+   ((((bnd_less SV12 SV13 = False \<or>
+       bnd_less SV11 SV10 = False) \<or>
+      (SV14 = SV13) = False) \<or>
+     bnd_less SV12 SV14 = False) \<or>
+    bnd_less SV9 SV10 = True) \<or>
+   (SV9 = SV11) = False"
+(*
+apply (tactic {*
+  extcnf_combined_tac NONE
+   [ConstsDiff,
+    StripQuantifiers]
+   []*})
+*)
+(*
+apply (rule allI)+
+apply (erule_tac x = "SV12" in allE)
+apply (erule_tac x = "SV13" in allE)
+apply (erule_tac x = "SV14" in allE)
+apply (erule_tac x = "SV9" in allE)
+apply (erule_tac x = "SV10" in allE)
+apply (erule_tac x = "SV11" in allE)
+*)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "300") 1*})
+
+
+(*NUM667^1 node 302 -- dec*)
+lemma "\<forall>SV12 SV13 SV14 SV9 SV10 SV11.
+       ((((bnd_less SV12 SV13 = bnd_less SV11 SV10) = False \<or>
+          (SV14 = SV13) = False) \<or>
+         bnd_less SV12 SV14 = False) \<or>
+        bnd_less SV9 SV10 = True) \<or>
+       (SV9 = SV11) =
+       False \<Longrightarrow>
+       \<forall>SV9 SV14 SV10 SV13 SV11 SV12.
+       (((((SV12 = SV11) = False \<or> (SV13 = SV10) = False) \<or>
+          (SV14 = SV13) = False) \<or>
+         bnd_less SV12 SV14 = False) \<or>
+        bnd_less SV9 SV10 = True) \<or>
+       (SV9 = SV11) =
+       False"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "302") 1*})
+*)
+
+
+(*
+(*CSR122^2*)
+     (* (Annotated_step ("23", "extuni_bool2"), *)
+lemma "(bnd_holdsDuring_THFTYPE_IiooI
+           (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+           (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+                  bnd_lBill_THFTYPE_i \<or>
+               \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                  bnd_lBill_THFTYPE_i)) =
+          bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+           bnd_lBill_THFTYPE_i) =
+         False \<Longrightarrow>
+         bnd_holdsDuring_THFTYPE_IiooI
+          (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+          (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+                 bnd_lBill_THFTYPE_i \<or>
+              \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                 bnd_lBill_THFTYPE_i)) =
+         True \<or>
+         bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+          bnd_lBill_THFTYPE_i =
+         True"
+(* apply (erule extuni_bool2) *)
+(* done *)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "23") 1*})
+
+     (* (Annotated_step ("24", "extuni_bool1"), *)
+lemma "(bnd_holdsDuring_THFTYPE_IiooI
+           (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+           (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+                  bnd_lBill_THFTYPE_i \<or>
+               \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                  bnd_lBill_THFTYPE_i)) =
+          bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+           bnd_lBill_THFTYPE_i) =
+         False \<Longrightarrow>
+         bnd_holdsDuring_THFTYPE_IiooI
+          (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+          (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+                 bnd_lBill_THFTYPE_i \<or>
+              \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                 bnd_lBill_THFTYPE_i)) =
+         False \<or>
+         bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+          bnd_lBill_THFTYPE_i =
+         False"
+(* apply (erule extuni_bool1) *)
+(* done *)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "24") 1*})
+
+     (* (Annotated_step ("25", "extuni_bool2"), *)
+lemma "(bnd_holdsDuring_THFTYPE_IiooI
+           (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+           (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+                  bnd_lBill_THFTYPE_i \<or>
+               \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                  bnd_lBill_THFTYPE_i)) =
+          bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+           bnd_lBill_THFTYPE_i) =
+         False \<Longrightarrow>
+         bnd_holdsDuring_THFTYPE_IiooI
+          (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+          (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+                 bnd_lBill_THFTYPE_i \<or>
+              \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                 bnd_lBill_THFTYPE_i)) =
+         True \<or>
+         bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+          bnd_lBill_THFTYPE_i =
+         True"
+(* apply (erule extuni_bool2) *)
+(* done *)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "25") 1*})
+
+     (* (Annotated_step ("26", "extuni_bool1"), *)
+lemma "\<forall>SV2. (bnd_holdsDuring_THFTYPE_IiooI
+                 (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+                 (\<not> (\<not> bnd_likes_THFTYPE_IiioI
+                        bnd_lMary_THFTYPE_i
+                        bnd_lBill_THFTYPE_i \<or>
+                     \<not> bnd_likes_THFTYPE_IiioI
+                        bnd_lSue_THFTYPE_i
+                        bnd_lBill_THFTYPE_i)) =
+                bnd_holdsDuring_THFTYPE_IiooI SV2 True) =
+               False \<Longrightarrow>
+         \<forall>SV2. bnd_holdsDuring_THFTYPE_IiooI
+                (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+                (\<not> (\<not> bnd_likes_THFTYPE_IiioI
+                       bnd_lMary_THFTYPE_i bnd_lBill_THFTYPE_i \<or>
+                    \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                       bnd_lBill_THFTYPE_i)) =
+               False \<or>
+               bnd_holdsDuring_THFTYPE_IiooI SV2 True = False"
+(* apply (rule allI, erule allE) *)
+(* apply (erule extuni_bool1) *)
+(* done *)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "26") 1*})
+
+     (* (Annotated_step ("27", "extuni_bool2"), *)
+lemma "\<forall>SV2. (bnd_holdsDuring_THFTYPE_IiooI
+                 (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+                 (\<not> (\<not> bnd_likes_THFTYPE_IiioI
+                        bnd_lMary_THFTYPE_i
+                        bnd_lBill_THFTYPE_i \<or>
+                     \<not> bnd_likes_THFTYPE_IiioI
+                        bnd_lSue_THFTYPE_i
+                        bnd_lBill_THFTYPE_i)) =
+                bnd_holdsDuring_THFTYPE_IiooI SV2 True) =
+               False \<Longrightarrow>
+         \<forall>SV2. bnd_holdsDuring_THFTYPE_IiooI
+                (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+                (\<not> (\<not> bnd_likes_THFTYPE_IiioI
+                       bnd_lMary_THFTYPE_i bnd_lBill_THFTYPE_i \<or>
+                    \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                       bnd_lBill_THFTYPE_i)) =
+               True \<or>
+               bnd_holdsDuring_THFTYPE_IiooI SV2 True = True"
+(* apply (rule allI, erule allE) *)
+(* apply (erule extuni_bool2) *)
+(* done *)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "27") 1*})
+
+     (* (Annotated_step ("30", "extuni_bool1"), *)
+lemma "((\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+                 bnd_lBill_THFTYPE_i \<or>
+              \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                 bnd_lBill_THFTYPE_i)) =
+          True) =
+         False \<Longrightarrow>
+         (\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+                bnd_lBill_THFTYPE_i \<or>
+             \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                bnd_lBill_THFTYPE_i)) =
+         False \<or>
+         True = False"
+(* apply (erule extuni_bool1) *)
+(* done *)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "30") 1*})
+
+     (* (Annotated_step ("29", "extuni_bind"), *)
+lemma "(bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i =
+          bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i) =
+         False \<or>
+         ((\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+                 bnd_lBill_THFTYPE_i \<or>
+              \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                 bnd_lBill_THFTYPE_i)) =
+          True) =
+         False \<Longrightarrow>
+         ((\<not> (\<not> bnd_likes_THFTYPE_IiioI bnd_lMary_THFTYPE_i
+                 bnd_lBill_THFTYPE_i \<or>
+              \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                 bnd_lBill_THFTYPE_i)) =
+          True) =
+         False"
+(* apply (tactic {*break_hypotheses 1*}) *)
+(* apply (erule extuni_bind) *)
+(* apply (tactic {*clause_breaker 1*}) *)
+(* done *)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "29") 1*})
+
+     (* (Annotated_step ("28", "extuni_dec"), *)
+lemma "\<forall>SV2. (bnd_holdsDuring_THFTYPE_IiooI
+                 (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i)
+                 (\<not> (\<not> bnd_likes_THFTYPE_IiioI
+                        bnd_lMary_THFTYPE_i
+                        bnd_lBill_THFTYPE_i \<or>
+                     \<not> bnd_likes_THFTYPE_IiioI
+                        bnd_lSue_THFTYPE_i
+                        bnd_lBill_THFTYPE_i)) =
+                bnd_holdsDuring_THFTYPE_IiooI SV2 True) =
+               False \<Longrightarrow>
+         \<forall>SV2. (bnd_lYearFn_THFTYPE_IiiI bnd_n2009_THFTYPE_i =
+                SV2) =
+               False \<or>
+               ((\<not> (\<not> bnd_likes_THFTYPE_IiioI
+                       bnd_lMary_THFTYPE_i bnd_lBill_THFTYPE_i \<or>
+                    \<not> bnd_likes_THFTYPE_IiioI bnd_lSue_THFTYPE_i
+                       bnd_lBill_THFTYPE_i)) =
+                True) =
+               False"
+(* apply (rule allI) *)
+(* apply (erule_tac x = "SV2" in allE) *)
+(* apply (erule extuni_dec_2) *)
+(* done *)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "28") 1*})
+*)
+
+(* QUA002^1
+   (* [[(Annotated_step ("49", "extuni_dec"), *)
+lemma "((bnd_sK3_E = bnd_sK1_X1 \<or> bnd_sK3_E = bnd_sK2_X2) =
+          (bnd_sK3_E = bnd_sK2_X2 \<or> bnd_sK3_E = bnd_sK1_X1)) =
+         False \<Longrightarrow>
+         ((bnd_sK3_E = bnd_sK2_X2) = (bnd_sK3_E = bnd_sK2_X2)) =
+         False \<or>
+         ((bnd_sK3_E = bnd_sK1_X1) = (bnd_sK3_E = bnd_sK1_X1)) =
+         False"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "49") 1*})
+
+     (* (Annotated_step ("20", "unfold_def"), *)
+lemma "(bnd_addition bnd_sK1_X1 bnd_sK2_X2 \<noteq>
+          bnd_addition bnd_sK2_X2 bnd_sK1_X1) =
+         True \<Longrightarrow>
+         (bnd_sup
+           (\<lambda>SX0\<Colon>TPTP_Interpret.ind.
+               SX0 = bnd_sK1_X1 \<or> SX0 = bnd_sK2_X2) \<noteq>
+          bnd_sup
+           (\<lambda>SX0\<Colon>TPTP_Interpret.ind.
+               SX0 = bnd_sK2_X2 \<or> SX0 = bnd_sK1_X1)) =
+         True"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "20") 1*})
+*)
+
+(*
+(*SEU620^2*)
+     (* (Annotated_step ("11", "unfold_def"), *)
+lemma "bnd_kpairiskpair =
+         (\<forall>Xx Xy.
+             bnd_iskpair
+              (bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset)
+                (bnd_setadjoin
+                  (bnd_setadjoin Xx
+                    (bnd_setadjoin Xy bnd_emptyset))
+                  bnd_emptyset))) \<and>
+         bnd_kpair =
+         (\<lambda>Xx Xy.
+             bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset)
+              (bnd_setadjoin
+                (bnd_setadjoin Xx
+                  (bnd_setadjoin Xy bnd_emptyset))
+                bnd_emptyset)) \<and>
+         bnd_iskpair =
+         (\<lambda>A. \<exists>Xx. bnd_in Xx (bnd_setunion A) \<and>
+                   (\<exists>Xy. bnd_in Xy (bnd_setunion A) \<and>
+                         A =
+                         bnd_setadjoin
+                          (bnd_setadjoin Xx bnd_emptyset)
+                          (bnd_setadjoin
+                            (bnd_setadjoin Xx
+                              (bnd_setadjoin Xy bnd_emptyset))
+                            bnd_emptyset))) \<and>
+         (\<forall>SY5 SY6.
+             (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset)
+               (bnd_setadjoin
+                 (bnd_setadjoin SY5
+                   (bnd_setadjoin SY6 bnd_emptyset))
+                 bnd_emptyset) =
+              bnd_setadjoin
+               (bnd_setadjoin (bnd_sK3 SY6 SY5) bnd_emptyset)
+               (bnd_setadjoin
+                 (bnd_setadjoin (bnd_sK3 SY6 SY5)
+                   (bnd_setadjoin (bnd_sK4 SY6 SY5)
+                     bnd_emptyset))
+                 bnd_emptyset) \<and>
+              bnd_in (bnd_sK4 SY6 SY5)
+               (bnd_setunion
+                 (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset)
+                   (bnd_setadjoin
+                     (bnd_setadjoin SY5
+                       (bnd_setadjoin SY6 bnd_emptyset))
+                     bnd_emptyset)))) \<and>
+             bnd_in (bnd_sK3 SY6 SY5)
+              (bnd_setunion
+                (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset)
+                  (bnd_setadjoin
+                    (bnd_setadjoin SY5
+                      (bnd_setadjoin SY6 bnd_emptyset))
+                    bnd_emptyset)))) =
+         True \<Longrightarrow>
+         (\<forall>SX0 SX1.
+             \<not> (\<not> \<not> (bnd_setadjoin
+                      (bnd_setadjoin SX0 bnd_emptyset)
+                      (bnd_setadjoin
+                        (bnd_setadjoin SX0
+                          (bnd_setadjoin SX1 bnd_emptyset))
+                        bnd_emptyset) \<noteq>
+                     bnd_setadjoin
+                      (bnd_setadjoin (bnd_sK3 SX1 SX0)
+                        bnd_emptyset)
+                      (bnd_setadjoin
+                        (bnd_setadjoin (bnd_sK3 SX1 SX0)
+                          (bnd_setadjoin (bnd_sK4 SX1 SX0)
+                            bnd_emptyset))
+                        bnd_emptyset) \<or>
+                     \<not> bnd_in (bnd_sK4 SX1 SX0)
+                        (bnd_setunion
+                          (bnd_setadjoin
+                            (bnd_setadjoin SX0 bnd_emptyset)
+                            (bnd_setadjoin
+                              (bnd_setadjoin SX0
+(bnd_setadjoin SX1 bnd_emptyset))
+                              bnd_emptyset)))) \<or>
+                \<not> bnd_in (bnd_sK3 SX1 SX0)
+                   (bnd_setunion
+                     (bnd_setadjoin
+                       (bnd_setadjoin SX0 bnd_emptyset)
+                       (bnd_setadjoin
+                         (bnd_setadjoin SX0
+                           (bnd_setadjoin SX1 bnd_emptyset))
+                         bnd_emptyset))))) =
+         True"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "11") 1*})
+
+     (* (Annotated_step ("3", "unfold_def"), *)
+lemma "bnd_kpairiskpair =
+         (\<forall>Xx Xy.
+             bnd_iskpair
+              (bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset)
+                (bnd_setadjoin
+                  (bnd_setadjoin Xx
+                    (bnd_setadjoin Xy bnd_emptyset))
+                  bnd_emptyset))) \<and>
+         bnd_kpair =
+         (\<lambda>Xx Xy.
+             bnd_setadjoin (bnd_setadjoin Xx bnd_emptyset)
+              (bnd_setadjoin
+                (bnd_setadjoin Xx
+                  (bnd_setadjoin Xy bnd_emptyset))
+                bnd_emptyset)) \<and>
+         bnd_iskpair =
+         (\<lambda>A. \<exists>Xx. bnd_in Xx (bnd_setunion A) \<and>
+                   (\<exists>Xy. bnd_in Xy (bnd_setunion A) \<and>
+                         A =
+                         bnd_setadjoin
+                          (bnd_setadjoin Xx bnd_emptyset)
+                          (bnd_setadjoin
+                            (bnd_setadjoin Xx
+                              (bnd_setadjoin Xy bnd_emptyset))
+                            bnd_emptyset))) \<and>
+         (bnd_kpairiskpair \<longrightarrow>
+          (\<forall>Xx Xy. bnd_iskpair (bnd_kpair Xx Xy))) =
+         False \<Longrightarrow>
+         ((\<forall>SY5 SY6.
+              \<exists>SY7. bnd_in SY7
+                     (bnd_setunion
+                       (bnd_setadjoin
+                         (bnd_setadjoin SY5 bnd_emptyset)
+                         (bnd_setadjoin
+                           (bnd_setadjoin SY5
+                             (bnd_setadjoin SY6 bnd_emptyset))
+                           bnd_emptyset))) \<and>
+                    (\<exists>SY8. bnd_in SY8
+                            (bnd_setunion
+                              (bnd_setadjoin
+(bnd_setadjoin SY5 bnd_emptyset)
+(bnd_setadjoin
+  (bnd_setadjoin SY5 (bnd_setadjoin SY6 bnd_emptyset))
+  bnd_emptyset))) \<and>
+                           bnd_setadjoin
+                            (bnd_setadjoin SY5 bnd_emptyset)
+                            (bnd_setadjoin
+                              (bnd_setadjoin SY5
+(bnd_setadjoin SY6 bnd_emptyset))
+                              bnd_emptyset) =
+                           bnd_setadjoin
+                            (bnd_setadjoin SY7 bnd_emptyset)
+                            (bnd_setadjoin
+                              (bnd_setadjoin SY7
+(bnd_setadjoin SY8 bnd_emptyset))
+                              bnd_emptyset))) \<longrightarrow>
+          (\<forall>SY0 SY1.
+              \<exists>SY3. bnd_in SY3
+                     (bnd_setunion
+                       (bnd_setadjoin
+                         (bnd_setadjoin SY0 bnd_emptyset)
+                         (bnd_setadjoin
+                           (bnd_setadjoin SY0
+                             (bnd_setadjoin SY1 bnd_emptyset))
+                           bnd_emptyset))) \<and>
+                    (\<exists>SY4. bnd_in SY4
+                            (bnd_setunion
+                              (bnd_setadjoin
+(bnd_setadjoin SY0 bnd_emptyset)
+(bnd_setadjoin
+  (bnd_setadjoin SY0 (bnd_setadjoin SY1 bnd_emptyset))
+  bnd_emptyset))) \<and>
+                           bnd_setadjoin
+                            (bnd_setadjoin SY0 bnd_emptyset)
+                            (bnd_setadjoin
+                              (bnd_setadjoin SY0
+(bnd_setadjoin SY1 bnd_emptyset))
+                              bnd_emptyset) =
+                           bnd_setadjoin
+                            (bnd_setadjoin SY3 bnd_emptyset)
+                            (bnd_setadjoin
+                              (bnd_setadjoin SY3
+(bnd_setadjoin SY4 bnd_emptyset))
+                              bnd_emptyset)))) =
+         False"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "3") 1*})
+
+     (* (Annotated_step ("8", "extcnf_combined"), *)
+lemma "(\<forall>SY5 SY6.
+             \<exists>SY7. bnd_in SY7
+                    (bnd_setunion
+                      (bnd_setadjoin
+                        (bnd_setadjoin SY5 bnd_emptyset)
+                        (bnd_setadjoin
+                          (bnd_setadjoin SY5
+                            (bnd_setadjoin SY6 bnd_emptyset))
+                          bnd_emptyset))) \<and>
+                   (\<exists>SY8. bnd_in SY8
+                           (bnd_setunion
+                             (bnd_setadjoin
+                               (bnd_setadjoin SY5 bnd_emptyset)
+                               (bnd_setadjoin
+ (bnd_setadjoin SY5 (bnd_setadjoin SY6 bnd_emptyset))
+ bnd_emptyset))) \<and>
+                          bnd_setadjoin
+                           (bnd_setadjoin SY5 bnd_emptyset)
+                           (bnd_setadjoin
+                             (bnd_setadjoin SY5
+                               (bnd_setadjoin SY6 bnd_emptyset))
+                             bnd_emptyset) =
+                          bnd_setadjoin
+                           (bnd_setadjoin SY7 bnd_emptyset)
+                           (bnd_setadjoin
+                             (bnd_setadjoin SY7
+                               (bnd_setadjoin SY8 bnd_emptyset))
+                             bnd_emptyset))) =
+         True \<Longrightarrow>
+         (\<forall>SY5 SY6.
+             (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset)
+               (bnd_setadjoin
+                 (bnd_setadjoin SY5
+                   (bnd_setadjoin SY6 bnd_emptyset))
+                 bnd_emptyset) =
+              bnd_setadjoin
+               (bnd_setadjoin (bnd_sK3 SY6 SY5) bnd_emptyset)
+               (bnd_setadjoin
+                 (bnd_setadjoin (bnd_sK3 SY6 SY5)
+                   (bnd_setadjoin (bnd_sK4 SY6 SY5)
+                     bnd_emptyset))
+                 bnd_emptyset) \<and>
+              bnd_in (bnd_sK4 SY6 SY5)
+               (bnd_setunion
+                 (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset)
+                   (bnd_setadjoin
+                     (bnd_setadjoin SY5
+                       (bnd_setadjoin SY6 bnd_emptyset))
+                     bnd_emptyset)))) \<and>
+             bnd_in (bnd_sK3 SY6 SY5)
+              (bnd_setunion
+                (bnd_setadjoin (bnd_setadjoin SY5 bnd_emptyset)
+                  (bnd_setadjoin
+                    (bnd_setadjoin SY5
+                      (bnd_setadjoin SY6 bnd_emptyset))
+                    bnd_emptyset)))) =
+         True"
+by (tactic {*
+HEADGOAL (extcnf_combined_tac Full false (hd prob_names))
+*})
+
+     (* (Annotated_step ("7", "extcnf_combined"), *)
+lemma "(\<not> (\<forall>SY0 SY1.
+                \<exists>SY3. bnd_in SY3
+                       (bnd_setunion
+                         (bnd_setadjoin
+                           (bnd_setadjoin SY0 bnd_emptyset)
+                           (bnd_setadjoin
+                             (bnd_setadjoin SY0
+                               (bnd_setadjoin SY1 bnd_emptyset))
+                             bnd_emptyset))) \<and>
+                      (\<exists>SY4. bnd_in SY4
+                              (bnd_setunion
+(bnd_setadjoin (bnd_setadjoin SY0 bnd_emptyset)
+  (bnd_setadjoin
+    (bnd_setadjoin SY0 (bnd_setadjoin SY1 bnd_emptyset))
+    bnd_emptyset))) \<and>
+                             bnd_setadjoin
+                              (bnd_setadjoin SY0 bnd_emptyset)
+                              (bnd_setadjoin
+(bnd_setadjoin SY0 (bnd_setadjoin SY1 bnd_emptyset))
+bnd_emptyset) =
+                             bnd_setadjoin
+                              (bnd_setadjoin SY3 bnd_emptyset)
+                              (bnd_setadjoin
+(bnd_setadjoin SY3 (bnd_setadjoin SY4 bnd_emptyset))
+bnd_emptyset)))) =
+         True \<Longrightarrow>
+         (\<forall>SY24.
+             (\<forall>SY25.
+                 bnd_setadjoin
+                  (bnd_setadjoin bnd_sK1 bnd_emptyset)
+                  (bnd_setadjoin
+                    (bnd_setadjoin bnd_sK1
+                      (bnd_setadjoin bnd_sK2 bnd_emptyset))
+                    bnd_emptyset) \<noteq>
+                 bnd_setadjoin (bnd_setadjoin SY24 bnd_emptyset)
+                  (bnd_setadjoin
+                    (bnd_setadjoin SY24
+                      (bnd_setadjoin SY25 bnd_emptyset))
+                    bnd_emptyset) \<or>
+                 \<not> bnd_in SY25
+                    (bnd_setunion
+                      (bnd_setadjoin
+                        (bnd_setadjoin bnd_sK1 bnd_emptyset)
+                        (bnd_setadjoin
+                          (bnd_setadjoin bnd_sK1
+                            (bnd_setadjoin bnd_sK2
+                              bnd_emptyset))
+                          bnd_emptyset)))) \<or>
+             \<not> bnd_in SY24
+                (bnd_setunion
+                  (bnd_setadjoin
+                    (bnd_setadjoin bnd_sK1 bnd_emptyset)
+                    (bnd_setadjoin
+                      (bnd_setadjoin bnd_sK1
+                        (bnd_setadjoin bnd_sK2 bnd_emptyset))
+                      bnd_emptyset)))) =
+         True"
+by (tactic {*HEADGOAL (extcnf_combined_tac Full false (hd prob_names))*})
+*)
+
+(*PUZ081^2*)
+(*
+     (* (Annotated_step ("9", "unfold_def"), *)
+lemma "bnd_says bnd_mel
+          (\<not> bnd_knave bnd_zoey \<and> \<not> bnd_knave bnd_mel) \<Longrightarrow>
+         bnd_says bnd_mel
+          (\<not> bnd_knave bnd_zoey \<and> \<not> bnd_knave bnd_mel) =
+         True"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "9") 1*})
+
+     (* (Annotated_step ("10", "unfold_def"), *)
+lemma "bnd_says bnd_zoey (bnd_knave bnd_mel) \<Longrightarrow>
+         bnd_says bnd_zoey (bnd_knave bnd_mel) = True"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "10") 1*})
+
+     (* (Annotated_step ("11", "unfold_def"), *)
+lemma "\<forall>P S. bnd_knave P \<and> bnd_says P S \<longrightarrow> \<not> S \<Longrightarrow>
+         (\<forall>P S. bnd_knave P \<and> bnd_says P S \<longrightarrow> \<not> S) = True"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "11") 1*})
+
+     (* (Annotated_step ("12", "unfold_def"), *)
+lemma "\<forall>P S. bnd_knight P \<and> bnd_says P S \<longrightarrow> S \<Longrightarrow>
+         (\<forall>P S. bnd_knight P \<and> bnd_says P S \<longrightarrow> S) = True"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "12") 1*})
+
+     (* (Annotated_step ("13", "unfold_def"), *)
+lemma "\<forall>P. bnd_knight P \<noteq> bnd_knave P \<Longrightarrow>
+         (\<forall>P. bnd_knight P \<noteq> bnd_knave P) = True"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "13") 1*})
+
+     (* (Annotated_step ("15", "extcnf_combined"), *)
+lemma "(\<not> (\<exists>TZ TM. TZ bnd_zoey \<and> TM bnd_mel)) = True \<Longrightarrow>
+         ((\<forall>TM. \<not> TM bnd_mel) \<or> (\<forall>TZ. \<not> TZ bnd_zoey)) = True"
+by (tactic {*extcnf_combined_tac Full false (hd prob_names) 1*})
+
+     (* (Annotated_step ("18", "extcnf_combined"), *)
+lemma "(\<forall>P. bnd_knight P \<noteq> bnd_knave P) = True \<Longrightarrow>
+         ((\<forall>P. \<not> bnd_knave P \<or> \<not> bnd_knight P) \<and>
+          (\<forall>P. bnd_knave P \<or> bnd_knight P)) =
+         True"
+by (tactic {*extcnf_combined_tac Full false (hd prob_names) 1*})
+*)
+
+(*
+(*from SEU667^2.p.out*)
+     (* (Annotated_step ("10", "unfold_def"), *)
+lemma "bnd_dpsetconstrSub =
+         (\<forall>A B Xphi.
+             bnd_subset (bnd_dpsetconstr A B Xphi)
+              (bnd_cartprod A B)) \<and>
+         bnd_dpsetconstr =
+         (\<lambda>A B Xphi.
+             bnd_dsetconstr (bnd_cartprod A B)
+              (\<lambda>Xu. \<exists>Xx. bnd_in Xx A \<and>
+                         (\<exists>Xy. (bnd_in Xy B \<and> Xphi Xx Xy) \<and>
+                               Xu = bnd_kpair Xx Xy))) \<and>
+         bnd_breln = (\<lambda>A B C. bnd_subset C (bnd_cartprod A B)) \<and>
+         (\<not> bnd_subset
+             (bnd_dsetconstr (bnd_cartprod bnd_sK1 bnd_sK2)
+               (\<lambda>SY66.
+                   \<exists>SY67.
+                      bnd_in SY67 bnd_sK1 \<and>
+                      (\<exists>SY68.
+                          (bnd_in SY68 bnd_sK2 \<and>
+                           bnd_sK3 SY67 SY68) \<and>
+                          SY66 = bnd_kpair SY67 SY68)))
+             (bnd_cartprod bnd_sK1 bnd_sK2)) =
+         True \<Longrightarrow>
+         (\<not> bnd_subset
+             (bnd_dsetconstr (bnd_cartprod bnd_sK1 bnd_sK2)
+               (\<lambda>SX0. \<not> (\<forall>SX1. \<not> \<not> (\<not> bnd_in SX1 bnd_sK1 \<or>
+    \<not> \<not> (\<forall>SX2. \<not> \<not> (\<not> \<not> (\<not> bnd_in SX2 bnd_sK2 \<or>
+                         \<not> bnd_sK3 SX1 SX2) \<or>
+                    SX0 \<noteq> bnd_kpair SX1 SX2))))))
+             (bnd_cartprod bnd_sK1 bnd_sK2)) =
+         True"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "10") 1*})
+
+
+     (* (Annotated_step ("11", "unfold_def"), *)
+lemma "bnd_dpsetconstrSub =
+         (\<forall>A B Xphi.
+             bnd_subset (bnd_dpsetconstr A B Xphi)
+              (bnd_cartprod A B)) \<and>
+         bnd_dpsetconstr =
+         (\<lambda>A B Xphi.
+             bnd_dsetconstr (bnd_cartprod A B)
+              (\<lambda>Xu. \<exists>Xx. bnd_in Xx A \<and>
+                         (\<exists>Xy. (bnd_in Xy B \<and> Xphi Xx Xy) \<and>
+                               Xu = bnd_kpair Xx Xy))) \<and>
+         bnd_breln = (\<lambda>A B C. bnd_subset C (bnd_cartprod A B)) \<and>
+         (\<forall>SY21 SY22 SY23.
+             bnd_subset
+              (bnd_dsetconstr (bnd_cartprod SY21 SY22)
+                (\<lambda>SY35.
+                    \<exists>SY36.
+                       bnd_in SY36 SY21 \<and>
+                       (\<exists>SY37.
+                           (bnd_in SY37 SY22 \<and> SY23 SY36 SY37) \<and>
+                           SY35 = bnd_kpair SY36 SY37)))
+              (bnd_cartprod SY21 SY22)) =
+         True \<Longrightarrow>
+         (\<forall>SX0 SX1 SX2.
+             bnd_subset
+              (bnd_dsetconstr (bnd_cartprod SX0 SX1)
+                (\<lambda>SX3. \<not> (\<forall>SX4. \<not> \<not> (\<not> bnd_in SX4 SX0 \<or>
+     \<not> \<not> (\<forall>SX5. \<not> \<not> (\<not> \<not> (\<not> bnd_in SX5 SX1 \<or> \<not> SX2 SX4 SX5) \<or>
+                     SX3 \<noteq> bnd_kpair SX4 SX5))))))
+              (bnd_cartprod SX0 SX1)) =
+         True"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "11") 1*})
+*)
+
+(*
+(*from ALG001^5*)
+      (* (Annotated_step ("4", "extcnf_forall_neg"), *)
+lemma "(\<forall>(Xh1 :: bnd_g \<Rightarrow> bnd_b) (Xh2 :: bnd_b \<Rightarrow> bnd_a) (Xf1 :: bnd_g \<Rightarrow> bnd_g \<Rightarrow> bnd_g) (Xf2 :: bnd_b \<Rightarrow> bnd_b \<Rightarrow> bnd_b) (Xf3 :: bnd_a \<Rightarrow> bnd_a \<Rightarrow> bnd_a).
+             (\<forall>Xx Xy. Xh1 (Xf1 Xx Xy) = Xf2 (Xh1 Xx) (Xh1 Xy)) \<and>
+             (\<forall>Xx Xy.
+                 Xh2 (Xf2 Xx Xy) = Xf3 (Xh2 Xx) (Xh2 Xy)) \<longrightarrow>
+             (\<forall>Xx Xy.
+                 Xh2 (Xh1 (Xf1 Xx Xy)) =
+                 Xf3 (Xh2 (Xh1 Xx)) (Xh2 (Xh1 Xy)))) =
+         False \<Longrightarrow>
+         ((\<forall>SY35 SY36.
+              bnd_sK1 (bnd_sK3 SY35 SY36) =
+              bnd_sK4 (bnd_sK1 SY35) (bnd_sK1 SY36)) \<and>
+          (\<forall>SY31 SY32.
+              bnd_sK2 (bnd_sK4 SY31 SY32) =
+              bnd_sK5 (bnd_sK2 SY31) (bnd_sK2 SY32)) \<longrightarrow>
+          (\<forall>SY39 SY40.
+              bnd_sK2 (bnd_sK1 (bnd_sK3 SY39 SY40)) =
+              bnd_sK5 (bnd_sK2 (bnd_sK1 SY39))
+               (bnd_sK2 (bnd_sK1 SY40)))) =
+         False"
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "4") 1*})
+*)
+
+(*SYN044^4*)
+(*
+ML {*
+print_depth 1400;
+(* the_tactics *)
+*}
+
+     (* (Annotated_step ("12", "unfold_def"), *)
+lemma "bnd_mor =
+         (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             (Y\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+             X U \<or> Y U) \<and>
+         bnd_mnot =
+         (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+             \<not> X U) \<and>
+         bnd_mimplies =
+         (\<lambda>U\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. bnd_mor (bnd_mnot U)) \<and>
+         bnd_mbox_s4 =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) X\<Colon>TPTP_Interpret.ind.
+             \<forall>Y\<Colon>TPTP_Interpret.ind. bnd_irel X Y \<longrightarrow> P Y) \<and>
+         bnd_mand =
+         (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             (Y\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+             X U \<and> Y U) \<and>
+         bnd_ixor =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             bnd_inot (bnd_iequiv P Q)) \<and>
+         bnd_ivalid = All \<and>
+         bnd_itrue = (\<lambda>W\<Colon>TPTP_Interpret.ind. True) \<and>
+         bnd_isatisfiable = Ex \<and>
+         bnd_ior =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             bnd_mor (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and>
+         bnd_inot =
+         (\<lambda>P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             bnd_mnot (bnd_mbox_s4 P)) \<and>
+         bnd_iinvalid =
+         (\<lambda>Phi\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             \<forall>W\<Colon>TPTP_Interpret.ind. \<not> Phi W) \<and>
+         bnd_iimplies =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             bnd_mimplies (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and>
+         bnd_iimplied =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. bnd_iimplies Q P) \<and>
+         bnd_ifalse = bnd_inot bnd_itrue \<and>
+         bnd_iequiv =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             bnd_iand (bnd_iimplies P Q) (bnd_iimplies Q P)) \<and>
+         bnd_icountersatisfiable =
+         (\<lambda>Phi\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             \<exists>W\<Colon>TPTP_Interpret.ind. \<not> Phi W) \<and>
+         bnd_iatom = (\<lambda>P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. P) \<and>
+         bnd_iand = bnd_mand \<and>
+         (\<forall>(X\<Colon>TPTP_Interpret.ind) (Y\<Colon>TPTP_Interpret.ind)
+             Z\<Colon>TPTP_Interpret.ind.
+             bnd_irel X Y \<and> bnd_irel Y Z \<longrightarrow> bnd_irel X Z) \<Longrightarrow>
+         (\<forall>(X\<Colon>TPTP_Interpret.ind) (Y\<Colon>TPTP_Interpret.ind)
+             Z\<Colon>TPTP_Interpret.ind.
+             bnd_irel X Y \<and> bnd_irel Y Z \<longrightarrow> bnd_irel X Z) =
+         True"
+(* by (tactic {*tectoc @{context}*}) *)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "12") 1*})
+
+     (* (Annotated_step ("11", "unfold_def"), *)
+lemma "bnd_mor =
+         (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             (Y\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+             X U \<or> Y U) \<and>
+         bnd_mnot =
+         (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+             \<not> X U) \<and>
+         bnd_mimplies =
+         (\<lambda>U\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. bnd_mor (bnd_mnot U)) \<and>
+         bnd_mbox_s4 =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) X\<Colon>TPTP_Interpret.ind.
+             \<forall>Y\<Colon>TPTP_Interpret.ind. bnd_irel X Y \<longrightarrow> P Y) \<and>
+         bnd_mand =
+         (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             (Y\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+             X U \<and> Y U) \<and>
+         bnd_ixor =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             bnd_inot (bnd_iequiv P Q)) \<and>
+         bnd_ivalid = All \<and>
+         bnd_itrue = (\<lambda>W\<Colon>TPTP_Interpret.ind. True) \<and>
+         bnd_isatisfiable = Ex \<and>
+         bnd_ior =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             bnd_mor (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and>
+         bnd_inot =
+         (\<lambda>P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             bnd_mnot (bnd_mbox_s4 P)) \<and>
+         bnd_iinvalid =
+         (\<lambda>Phi\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             \<forall>W\<Colon>TPTP_Interpret.ind. \<not> Phi W) \<and>
+         bnd_iimplies =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             bnd_mimplies (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and>
+         bnd_iimplied =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. bnd_iimplies Q P) \<and>
+         bnd_ifalse = bnd_inot bnd_itrue \<and>
+         bnd_iequiv =
+         (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
+             Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             bnd_iand (bnd_iimplies P Q) (bnd_iimplies Q P)) \<and>
+         bnd_icountersatisfiable =
+         (\<lambda>Phi\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+             \<exists>W\<Colon>TPTP_Interpret.ind. \<not> Phi W) \<and>
+         bnd_iatom = (\<lambda>P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. P) \<and>
+         bnd_iand = bnd_mand \<and>
+         bnd_ivalid
+          (bnd_iimplies (bnd_iatom bnd_q) (bnd_iatom bnd_r)) \<Longrightarrow>
+         (\<forall>SY161\<Colon>TPTP_Interpret.ind.
+             \<not> (\<forall>SY162\<Colon>TPTP_Interpret.ind.
+                   bnd_irel SY161 SY162 \<longrightarrow> bnd_q SY162) \<or>
+             (\<forall>SY163\<Colon>TPTP_Interpret.ind.
+                 bnd_irel SY161 SY163 \<longrightarrow> bnd_r SY163)) =
+         True"
+(* by (tactic {*tectoc @{context}*}) *)
+by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "11") 1*})
+
+lemma "
+(\<forall>SY136.
+    \<not> (\<forall>SY137. bnd_irel SY136 SY137 \<longrightarrow> bnd_r SY137) \<or>
+    (\<forall>SY138.
+        bnd_irel SY136 SY138 \<longrightarrow> bnd_p SY138 \<and> bnd_q SY138)) =
+True \<Longrightarrow>
+(\<forall>SY136.
+    bnd_irel SY136 (bnd_sK5 SY136) \<and> \<not> bnd_r (bnd_sK5 SY136) \<or>
+    (\<forall>SY138. \<not> bnd_irel SY136 SY138 \<or> bnd_p SY138) \<and>
+    (\<forall>SY138. \<not> bnd_irel SY136 SY138 \<or> bnd_q SY138)) =
+True"
+by (tactic {*HEADGOAL (extcnf_combined_tac Full false (hd prob_names))*})
+*)
+
+     (* (Annotated_step ("11", "extcnf_forall_neg"), *)
+lemma "\<forall>SV1\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+            (\<forall>SY2\<Colon>TPTP_Interpret.ind.
+                \<not> (\<not> (\<not> SV1 SY2 \<or> SEV405_5_bnd_cA) \<or>
+                   \<not> (\<not> SEV405_5_bnd_cA \<or> SV1 SY2))) =
+            False \<Longrightarrow>
+         \<forall>SV1\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+            (\<not> (\<not> (\<not> SV1 (SEV405_5_bnd_sK1_SY2 SV1) \<or> SEV405_5_bnd_cA) \<or>
+                \<not> (\<not> SEV405_5_bnd_cA \<or> SV1 (SEV405_5_bnd_sK1_SY2 SV1)))) =
+            False"
+(* apply (tactic {*full_extcnf_combined_tac*}) *)
+by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+
+     (* (Annotated_step ("13", "extcnf_forall_pos"), *)
+lemma "(\<forall>SY31 SY32.
+             bnd_sK2 (bnd_sK4 SY31 SY32) =
+             bnd_sK5 (bnd_sK2 SY31) (bnd_sK2 SY32)) =
+         True \<Longrightarrow>
+         \<forall>SV1. (\<forall>SY42.
+                   bnd_sK2 (bnd_sK4 SV1 SY42) =
+                   bnd_sK5 (bnd_sK2 SV1) (bnd_sK2 SY42)) =
+               True"
+by (tactic {*nonfull_extcnf_combined_tac @{context} [Universal]*})
+
+     (* (Annotated_step ("14", "extcnf_forall_pos"), *)
+lemma "(\<forall>SY35 SY36.
+             bnd_sK1 (bnd_sK3 SY35 SY36) =
+             bnd_sK4 (bnd_sK1 SY35) (bnd_sK1 SY36)) =
+         True \<Longrightarrow>
+         \<forall>SV2. (\<forall>SY43.
+                   bnd_sK1 (bnd_sK3 SV2 SY43) =
+                   bnd_sK4 (bnd_sK1 SV2) (bnd_sK1 SY43)) =
+               True"
+by (tactic {*nonfull_extcnf_combined_tac @{context} [Universal]*})
+
+
+(*from SYO198^5.p.out*)
+   (* [[(Annotated_step ("11", "extcnf_forall_special_pos"), *)
+lemma "(\<forall>SX0\<Colon>bool \<Rightarrow> bool.
+             \<not> \<not> (\<not> SX0 bnd_sK1_Xx \<or> \<not> SX0 bnd_sK2_Xy)) =
+         True \<Longrightarrow>
+         (\<not> \<not> (\<not> True \<or> \<not> True)) = True"
+apply (tactic {*extcnf_forall_special_pos_tac 1*})
+done
+
+     (* (Annotated_step ("13", "extcnf_forall_special_pos"), *)
+lemma "(\<forall>SX0\<Colon>bool \<Rightarrow> bool.
+             \<not> \<not> (\<not> SX0 bnd_sK1_Xx \<or> \<not> SX0 bnd_sK2_Xy)) =
+         True \<Longrightarrow>
+         (\<not> \<not> (\<not> bnd_sK1_Xx \<or> \<not> bnd_sK2_Xy)) = True"
+apply (tactic {*extcnf_forall_special_pos_tac 1*})
+done
+
+   (* [[(Annotated_step ("8", "polarity_switch"), *)
+lemma "(\<forall>(Xx\<Colon>bool) (Xy\<Colon>bool) Xz\<Colon>bool. True \<and> True \<longrightarrow> True) =
+         False \<Longrightarrow>
+         (\<not> (\<forall>(Xx\<Colon>bool) (Xy\<Colon>bool) Xz\<Colon>bool.
+                True \<and> True \<longrightarrow> True)) =
+         True"
+apply (tactic {*nonfull_extcnf_combined_tac @{context} [Polarity_switch]*})
+done
+
+lemma "((\<forall>SY22 SY23 SY24.
+       bnd_sK1_RRR SY22 SY23 \<and> bnd_sK1_RRR SY23 SY24 \<longrightarrow>
+       bnd_sK1_RRR SY22 SY24) \<and>
+   (\<forall>SY25.
+       (\<forall>SY26. SY25 SY26 \<longrightarrow> bnd_sK1_RRR SY26 (bnd_sK2_U SY25)) \<and>
+       (\<forall>SY27.
+           (\<forall>SY28. SY25 SY28 \<longrightarrow> bnd_sK1_RRR SY28 SY27) \<longrightarrow>
+           bnd_sK1_RRR (bnd_sK2_U SY25) SY27)) \<longrightarrow>
+   (\<forall>SY29. \<exists>SY30. \<forall>SY31. SY29 SY31 \<longrightarrow> bnd_sK1_RRR SY31 SY30)) =
+  False \<Longrightarrow>
+  (\<forall>SY25.
+      (\<forall>SY26. SY25 SY26 \<longrightarrow> bnd_sK1_RRR SY26 (bnd_sK2_U SY25)) \<and>
+      (\<forall>SY27.
+          (\<forall>SY28. SY25 SY28 \<longrightarrow> bnd_sK1_RRR SY28 SY27) \<longrightarrow>
+          bnd_sK1_RRR (bnd_sK2_U SY25) SY27)) =
+  True"
+apply (tactic {*standard_cnf_tac @{context} 1*})
+done
+
+lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow>
+   (\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) \<longrightarrow>
+   (\<forall>A B. A = B \<longrightarrow>
+          (\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) \<longrightarrow>
+   (\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow>
+          bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset)) =
+  False \<Longrightarrow>
+ (\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) =
+  True"
+apply (tactic {*standard_cnf_tac @{context} 1*})
+done
+
+lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow>
+   (\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) \<longrightarrow>
+   (\<forall>A B. A = B \<longrightarrow>
+          (\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) \<longrightarrow>
+   (\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow>
+          bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset)) =
+  False \<Longrightarrow>
+  (\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) =
+  True"
+apply (tactic {*standard_cnf_tac @{context} 1*})
+done
+
+lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow>
+   (\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) \<longrightarrow>
+   (\<forall>A B. A = B \<longrightarrow>
+          (\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) \<longrightarrow>
+   (\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow>
+          bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset)) =
+  False \<Longrightarrow>
+  (\<forall>A B. A = B \<longrightarrow>
+         (\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) =
+  True"
+apply (tactic {*standard_cnf_tac @{context} 1*})
+done
+
+lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow>
+   (\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) \<longrightarrow>
+   (\<forall>A B. A = B \<longrightarrow>
+          (\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) \<longrightarrow>
+   (\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow>
+          bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset)) =
+  False \<Longrightarrow>
+  (\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow>
+         bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset) =
+  False"
+apply (tactic {*standard_cnf_tac @{context} 1*})
+done
+
+(*nested conjunctions*)
+lemma "((((\<forall>Xx. bnd_cP bnd_e Xx = Xx) \<and>
+     (\<forall>Xy. bnd_cP Xy bnd_e = Xy)) \<and>
+    (\<forall>Xz. bnd_cP Xz Xz = bnd_e)) \<and>
+   (\<forall>Xx Xy Xz.
+       bnd_cP (bnd_cP Xx Xy) Xz = bnd_cP Xx (bnd_cP Xy Xz)) \<longrightarrow>
+   (\<forall>Xa Xb. bnd_cP Xa Xb = bnd_cP Xb Xa)) =
+  False \<Longrightarrow>
+  (\<forall>Xx. bnd_cP bnd_e Xx = Xx) =
+  True"
+apply (tactic {*standard_cnf_tac @{context} 1*})
+done
+
+end
\ No newline at end of file