src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML
author blanchet
Thu, 26 Aug 2010 00:49:38 +0200
changeset 38987 69fea359d3f8
parent 38937 d19c3a7ce38b
child 38991 6628adcae4a7
permissions -rw-r--r--
renaming
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_translate.ML
     2     Author:     Fabian Immler, TU Muenchen
     3     Author:     Makarius
     4     Author:     Jasmin Blanchette, TU Muenchen
     5 
     6 Translation of HOL to FOL.
     7 *)
     8 
     9 signature SLEDGEHAMMER_TRANSLATE =
    10 sig
    11   type 'a problem = 'a ATP_Problem.problem
    12 
    13   val axiom_prefix : string
    14   val conjecture_prefix : string
    15   val helper_prefix : string
    16   val class_rel_clause_prefix : string
    17   val arity_clause_prefix : string
    18   val tfrees_name : string
    19   val prepare_problem :
    20     Proof.context -> bool -> bool -> bool -> bool -> term list -> term
    21     -> ((string * bool) * thm) list
    22     -> string problem * string Symtab.table * int * (string * bool) vector
    23 end;
    24 
    25 structure Sledgehammer_Translate : SLEDGEHAMMER_TRANSLATE =
    26 struct
    27 
    28 open ATP_Problem
    29 open Metis_Clauses
    30 open Sledgehammer_Util
    31 
    32 val axiom_prefix = "ax_"
    33 val conjecture_prefix = "conj_"
    34 val helper_prefix = "help_"
    35 val class_rel_clause_prefix = "clrel_";
    36 val arity_clause_prefix = "arity_"
    37 val tfrees_name = "tfrees"
    38 
    39 (* Freshness almost guaranteed! *)
    40 val sledgehammer_weak_prefix = "Sledgehammer:"
    41 
    42 datatype fol_formula =
    43   FOLFormula of {name: string,
    44                  kind: kind,
    45                  combformula: (name, combterm) formula,
    46                  ctypes_sorts: typ list}
    47 
    48 fun mk_anot phi = AConn (ANot, [phi])
    49 fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2])
    50 fun mk_ahorn [] phi = phi
    51   | mk_ahorn (phi :: phis) psi =
    52     AConn (AImplies, [fold (mk_aconn AAnd) phis phi, psi])
    53 
    54 fun combformula_for_prop thy =
    55   let
    56     val do_term = combterm_from_term thy
    57     fun do_quant bs q s T t' =
    58       let val s = Name.variant (map fst bs) s in
    59         do_formula ((s, T) :: bs) t'
    60         #>> (fn phi => AQuant (q, [`make_bound_var s], phi))
    61       end
    62     and do_conn bs c t1 t2 =
    63       do_formula bs t1 ##>> do_formula bs t2
    64       #>> (fn (phi1, phi2) => AConn (c, [phi1, phi2]))
    65     and do_formula bs t =
    66       case t of
    67         @{const Not} $ t1 =>
    68         do_formula bs t1 #>> (fn phi => AConn (ANot, [phi]))
    69       | Const (@{const_name All}, _) $ Abs (s, T, t') =>
    70         do_quant bs AForall s T t'
    71       | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
    72         do_quant bs AExists s T t'
    73       | @{const "op &"} $ t1 $ t2 => do_conn bs AAnd t1 t2
    74       | @{const "op |"} $ t1 $ t2 => do_conn bs AOr t1 t2
    75       | @{const "op -->"} $ t1 $ t2 => do_conn bs AImplies t1 t2
    76       | Const (@{const_name "op ="}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
    77         do_conn bs AIff t1 t2
    78       | _ => (fn ts => do_term bs (Envir.eta_contract t)
    79                        |>> AAtom ||> union (op =) ts)
    80   in do_formula [] end
    81 
    82 val presimplify_term = prop_of o Meson.presimplify oo Skip_Proof.make_thm
    83 
    84 fun concealed_bound_name j = sledgehammer_weak_prefix ^ Int.toString j
    85 fun conceal_bounds Ts t =
    86   subst_bounds (map (Free o apfst concealed_bound_name)
    87                     (0 upto length Ts - 1 ~~ Ts), t)
    88 fun reveal_bounds Ts =
    89   subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
    90                     (0 upto length Ts - 1 ~~ Ts))
    91 
    92 (* Removes the lambdas from an equation of the form "t = (%x. u)".
    93    (Cf. "extensionalize_theorem" in "Clausifier".) *)
    94 fun extensionalize_term t =
    95   let
    96     fun aux j (@{const Trueprop} $ t') = @{const Trueprop} $ aux j t'
    97       | aux j (t as Const (s, Type (_, [Type (_, [_, T']),
    98                                         Type (_, [_, res_T])]))
    99                     $ t2 $ Abs (var_s, var_T, t')) =
   100         if s = @{const_name "op ="} orelse s = @{const_name "=="} then
   101           let val var_t = Var ((var_s, j), var_T) in
   102             Const (s, T' --> T' --> res_T)
   103               $ betapply (t2, var_t) $ subst_bound (var_t, t')
   104             |> aux (j + 1)
   105           end
   106         else
   107           t
   108       | aux _ t = t
   109   in aux (maxidx_of_term t + 1) t end
   110 
   111 fun introduce_combinators_in_term ctxt kind t =
   112   let val thy = ProofContext.theory_of ctxt in
   113     if Meson.is_fol_term thy t then
   114       t
   115     else
   116       let
   117         fun aux Ts t =
   118           case t of
   119             @{const Not} $ t1 => @{const Not} $ aux Ts t1
   120           | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
   121             t0 $ Abs (s, T, aux (T :: Ts) t')
   122           | (t0 as Const (@{const_name All}, _)) $ t1 =>
   123             aux Ts (t0 $ eta_expand Ts t1 1)
   124           | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
   125             t0 $ Abs (s, T, aux (T :: Ts) t')
   126           | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
   127             aux Ts (t0 $ eta_expand Ts t1 1)
   128           | (t0 as @{const "op &"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
   129           | (t0 as @{const "op |"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
   130           | (t0 as @{const "op -->"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
   131           | (t0 as Const (@{const_name "op ="}, Type (_, [@{typ bool}, _])))
   132               $ t1 $ t2 =>
   133             t0 $ aux Ts t1 $ aux Ts t2
   134           | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
   135                    t
   136                  else
   137                    t |> conceal_bounds Ts
   138                      |> Envir.eta_contract
   139                      |> cterm_of thy
   140                      |> Clausifier.introduce_combinators_in_cterm
   141                      |> prop_of |> Logic.dest_equals |> snd
   142                      |> reveal_bounds Ts
   143         val ([t], ctxt') = Variable.import_terms true [t] ctxt
   144       in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
   145       handle THM _ =>
   146              (* A type variable of sort "{}" will make abstraction fail. *)
   147              if kind = Conjecture then HOLogic.false_const
   148              else HOLogic.true_const
   149   end
   150 
   151 (* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
   152    same in Sledgehammer to prevent the discovery of unreplable proofs. *)
   153 fun freeze_term t =
   154   let
   155     fun aux (t $ u) = aux t $ aux u
   156       | aux (Abs (s, T, t)) = Abs (s, T, aux t)
   157       | aux (Var ((s, i), T)) =
   158         Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
   159       | aux t = t
   160   in t |> exists_subterm is_Var t ? aux end
   161 
   162 (* "Object_Logic.atomize_term" isn't as powerful as it could be; for example,
   163     it leaves metaequalities over "prop"s alone. *)
   164 val atomize_term =
   165   let
   166     fun aux (@{const Trueprop} $ t1) = t1
   167       | aux (Const (@{const_name all}, _) $ Abs (s, T, t')) =
   168         HOLogic.all_const T $ Abs (s, T, aux t')
   169       | aux (@{const "==>"} $ t1 $ t2) = HOLogic.mk_imp (pairself aux (t1, t2))
   170       | aux (Const (@{const_name "=="}, Type (_, [@{typ prop}, _])) $ t1 $ t2) =
   171         HOLogic.eq_const HOLogic.boolT $ aux t1 $ aux t2
   172       | aux (Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2) =
   173         HOLogic.eq_const T $ t1 $ t2
   174       | aux _ = raise Fail "aux"
   175   in perhaps (try aux) end
   176 
   177 (* making axiom and conjecture formulas *)
   178 fun make_formula ctxt presimp name kind t =
   179   let
   180     val thy = ProofContext.theory_of ctxt
   181     val t = t |> Envir.beta_eta_contract
   182               |> transform_elim_term
   183               |> atomize_term
   184     val need_trueprop = (fastype_of t = HOLogic.boolT)
   185     val t = t |> need_trueprop ? HOLogic.mk_Trueprop
   186               |> extensionalize_term
   187               |> presimp ? presimplify_term thy
   188               |> perhaps (try (HOLogic.dest_Trueprop))
   189               |> introduce_combinators_in_term ctxt kind
   190               |> kind <> Axiom ? freeze_term
   191     val (combformula, ctypes_sorts) = combformula_for_prop thy t []
   192   in
   193     FOLFormula {name = name, combformula = combformula, kind = kind,
   194                 ctypes_sorts = ctypes_sorts}
   195   end
   196 
   197 fun make_axiom ctxt presimp ((name, chained), th) =
   198   case make_formula ctxt presimp name Axiom (prop_of th) of
   199     FOLFormula {combformula = AAtom (CombConst (("c_True", _), _, _)), ...} =>
   200     NONE
   201   | formula => SOME ((name, chained), formula)
   202 fun make_conjecture ctxt ts =
   203   let val last = length ts - 1 in
   204     map2 (fn j => make_formula ctxt true (Int.toString j)
   205                                (if j = last then Conjecture else Hypothesis))
   206          (0 upto last) ts
   207   end
   208 
   209 (** Helper facts **)
   210 
   211 fun count_combterm (CombConst ((s, _), _, _)) =
   212     Symtab.map_entry s (Integer.add 1)
   213   | count_combterm (CombVar _) = I
   214   | count_combterm (CombApp (t1, t2)) = fold count_combterm [t1, t2]
   215 fun count_combformula (AQuant (_, _, phi)) = count_combformula phi
   216   | count_combformula (AConn (_, phis)) = fold count_combformula phis
   217   | count_combformula (AAtom tm) = count_combterm tm
   218 fun count_fol_formula (FOLFormula {combformula, ...}) =
   219   count_combformula combformula
   220 
   221 val optional_helpers =
   222   [(["c_COMBI", "c_COMBK"], @{thms COMBI_def COMBK_def}),
   223    (["c_COMBB", "c_COMBC"], @{thms COMBB_def COMBC_def}),
   224    (["c_COMBS"], @{thms COMBS_def})]
   225 val optional_typed_helpers =
   226   [(["c_True", "c_False", "c_If"], @{thms True_or_False}),
   227    (["c_If"], @{thms if_True if_False})]
   228 val mandatory_helpers = @{thms fequal_def}
   229 
   230 val init_counters =
   231   [optional_helpers, optional_typed_helpers] |> maps (maps fst)
   232   |> sort_distinct string_ord |> map (rpair 0) |> Symtab.make
   233 
   234 fun get_helper_facts ctxt is_FO full_types conjectures axioms =
   235   let
   236     val ct = fold (fold count_fol_formula) [conjectures, axioms] init_counters
   237     fun is_needed c = the (Symtab.lookup ct c) > 0
   238     fun baptize th = ((Thm.get_name_hint th, false), th)
   239   in
   240     (optional_helpers
   241      |> full_types ? append optional_typed_helpers
   242      |> maps (fn (ss, ths) =>
   243                  if exists is_needed ss then map baptize ths else [])) @
   244     (if is_FO then [] else map baptize mandatory_helpers)
   245     |> map_filter (Option.map snd o make_axiom ctxt false)
   246   end
   247 
   248 fun prepare_formulas ctxt full_types hyp_ts concl_t axiom_pairs =
   249   let
   250     val thy = ProofContext.theory_of ctxt
   251     val axiom_ts = map (prop_of o snd) axiom_pairs
   252     val hyp_ts =
   253       if null hyp_ts then
   254         []
   255       else
   256         (* Remove existing axioms from the conjecture, as this can dramatically
   257            boost an ATP's performance (for some reason). *)
   258         let
   259           val axiom_table = fold (Termtab.update o rpair ()) axiom_ts
   260                                  Termtab.empty
   261         in hyp_ts |> filter_out (Termtab.defined axiom_table) end
   262     val goal_t = Logic.list_implies (hyp_ts, concl_t)
   263     val is_FO = Meson.is_fol_term thy goal_t
   264     val subs = tfree_classes_of_terms [goal_t]
   265     val supers = tvar_classes_of_terms axiom_ts
   266     val tycons = type_consts_of_terms thy (goal_t :: axiom_ts)
   267     (* TFrees in the conjecture; TVars in the axioms *)
   268     val conjectures = make_conjecture ctxt (hyp_ts @ [concl_t])
   269     val (axiom_names, axioms) =
   270       ListPair.unzip (map_filter (make_axiom ctxt true) axiom_pairs)
   271     val helper_facts = get_helper_facts ctxt is_FO full_types conjectures axioms
   272     val (supers', arity_clauses) = make_arity_clauses thy tycons supers
   273     val class_rel_clauses = make_class_rel_clauses thy subs supers'
   274   in
   275     (Vector.fromList axiom_names,
   276      (conjectures, axioms, helper_facts, class_rel_clauses, arity_clauses))
   277   end
   278 
   279 fun wrap_type ty t = ATerm ((type_wrapper_name, type_wrapper_name), [ty, t])
   280 
   281 fun fo_term_for_combtyp (CombTVar name) = ATerm (name, [])
   282   | fo_term_for_combtyp (CombTFree name) = ATerm (name, [])
   283   | fo_term_for_combtyp (CombType (name, tys)) =
   284     ATerm (name, map fo_term_for_combtyp tys)
   285 
   286 fun fo_literal_for_type_literal (TyLitVar (class, name)) =
   287     (true, ATerm (class, [ATerm (name, [])]))
   288   | fo_literal_for_type_literal (TyLitFree (class, name)) =
   289     (true, ATerm (class, [ATerm (name, [])]))
   290 
   291 fun formula_for_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
   292 
   293 fun fo_term_for_combterm full_types =
   294   let
   295     fun aux top_level u =
   296       let
   297         val (head, args) = strip_combterm_comb u
   298         val (x, ty_args) =
   299           case head of
   300             CombConst (name as (s, s'), _, ty_args) =>
   301             let val ty_args = if full_types then [] else ty_args in
   302               if s = "equal" then
   303                 if top_level andalso length args = 2 then (name, [])
   304                 else (("c_fequal", @{const_name fequal}), ty_args)
   305               else if top_level then
   306                 case s of
   307                   "c_False" => (("$false", s'), [])
   308                 | "c_True" => (("$true", s'), [])
   309                 | _ => (name, ty_args)
   310               else
   311                 (name, ty_args)
   312             end
   313           | CombVar (name, _) => (name, [])
   314           | CombApp _ => raise Fail "impossible \"CombApp\""
   315         val t = ATerm (x, map fo_term_for_combtyp ty_args @
   316                           map (aux false) args)
   317     in
   318       if full_types then wrap_type (fo_term_for_combtyp (combtyp_of u)) t else t
   319     end
   320   in aux true end
   321 
   322 fun formula_for_combformula full_types =
   323   let
   324     fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi)
   325       | aux (AConn (c, phis)) = AConn (c, map aux phis)
   326       | aux (AAtom tm) = AAtom (fo_term_for_combterm full_types tm)
   327   in aux end
   328 
   329 fun formula_for_axiom full_types (FOLFormula {combformula, ctypes_sorts, ...}) =
   330   mk_ahorn (map (formula_for_fo_literal o fo_literal_for_type_literal)
   331                 (type_literals_for_types ctypes_sorts))
   332            (formula_for_combformula full_types combformula)
   333 
   334 fun problem_line_for_fact prefix full_types
   335                           (formula as FOLFormula {name, kind, ...}) =
   336   Fof (prefix ^ ascii_of name, kind, formula_for_axiom full_types formula)
   337 
   338 fun problem_line_for_class_rel_clause (ClassRelClause {name, subclass,
   339                                                        superclass, ...}) =
   340   let val ty_arg = ATerm (("T", "T"), []) in
   341     Fof (class_rel_clause_prefix ^ ascii_of name, Axiom,
   342          AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
   343                            AAtom (ATerm (superclass, [ty_arg]))]))
   344   end
   345 
   346 fun fo_literal_for_arity_literal (TConsLit (c, t, args)) =
   347     (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
   348   | fo_literal_for_arity_literal (TVarLit (c, sort)) =
   349     (false, ATerm (c, [ATerm (sort, [])]))
   350 
   351 fun problem_line_for_arity_clause (ArityClause {name, conclLit, premLits,
   352                                                 ...}) =
   353   Fof (arity_clause_prefix ^ ascii_of name, Axiom,
   354        mk_ahorn (map (formula_for_fo_literal o apfst not
   355                       o fo_literal_for_arity_literal) premLits)
   356                 (formula_for_fo_literal
   357                      (fo_literal_for_arity_literal conclLit)))
   358 
   359 fun problem_line_for_conjecture full_types
   360                                 (FOLFormula {name, kind, combformula, ...}) =
   361   Fof (conjecture_prefix ^ name, kind,
   362        formula_for_combformula full_types combformula)
   363 
   364 fun free_type_literals_for_conjecture (FOLFormula {ctypes_sorts, ...}) =
   365   map fo_literal_for_type_literal (type_literals_for_types ctypes_sorts)
   366 
   367 fun problem_line_for_free_type lit =
   368   Fof (tfrees_name, Hypothesis, formula_for_fo_literal lit)
   369 fun problem_lines_for_free_types conjectures =
   370   let
   371     val litss = map free_type_literals_for_conjecture conjectures
   372     val lits = fold (union (op =)) litss []
   373   in map problem_line_for_free_type lits end
   374 
   375 (** "hBOOL" and "hAPP" **)
   376 
   377 type const_info = {min_arity: int, max_arity: int, sub_level: bool}
   378 
   379 fun consider_term top_level (ATerm ((s, _), ts)) =
   380   (if is_tptp_variable s then
   381      I
   382    else
   383      let val n = length ts in
   384        Symtab.map_default
   385            (s, {min_arity = n, max_arity = 0, sub_level = false})
   386            (fn {min_arity, max_arity, sub_level} =>
   387                {min_arity = Int.min (n, min_arity),
   388                 max_arity = Int.max (n, max_arity),
   389                 sub_level = sub_level orelse not top_level})
   390      end)
   391   #> fold (consider_term (top_level andalso s = type_wrapper_name)) ts
   392 fun consider_formula (AQuant (_, _, phi)) = consider_formula phi
   393   | consider_formula (AConn (_, phis)) = fold consider_formula phis
   394   | consider_formula (AAtom tm) = consider_term true tm
   395 
   396 fun consider_problem_line (Fof (_, _, phi)) = consider_formula phi
   397 fun consider_problem problem = fold (fold consider_problem_line o snd) problem
   398 
   399 fun const_table_for_problem explicit_apply problem =
   400   if explicit_apply then NONE
   401   else SOME (Symtab.empty |> consider_problem problem)
   402 
   403 fun min_arity_of thy full_types NONE s =
   404     (if s = "equal" orelse s = type_wrapper_name orelse
   405         String.isPrefix type_const_prefix s orelse
   406         String.isPrefix class_prefix s then
   407        16383 (* large number *)
   408      else if full_types then
   409        0
   410      else case strip_prefix_and_unascii const_prefix s of
   411        SOME s' => num_type_args thy (invert_const s')
   412      | NONE => 0)
   413   | min_arity_of _ _ (SOME the_const_tab) s =
   414     case Symtab.lookup the_const_tab s of
   415       SOME ({min_arity, ...} : const_info) => min_arity
   416     | NONE => 0
   417 
   418 fun full_type_of (ATerm ((s, _), [ty, _])) =
   419     if s = type_wrapper_name then ty else raise Fail "expected type wrapper"
   420   | full_type_of _ = raise Fail "expected type wrapper"
   421 
   422 fun list_hAPP_rev _ t1 [] = t1
   423   | list_hAPP_rev NONE t1 (t2 :: ts2) =
   424     ATerm (`I "hAPP", [list_hAPP_rev NONE t1 ts2, t2])
   425   | list_hAPP_rev (SOME ty) t1 (t2 :: ts2) =
   426     let val ty' = ATerm (`make_fixed_type_const @{type_name fun},
   427                          [full_type_of t2, ty]) in
   428       ATerm (`I "hAPP", [wrap_type ty' (list_hAPP_rev (SOME ty') t1 ts2), t2])
   429     end
   430 
   431 fun repair_applications_in_term thy full_types const_tab =
   432   let
   433     fun aux opt_ty (ATerm (name as (s, _), ts)) =
   434       if s = type_wrapper_name then
   435         case ts of
   436           [t1, t2] => ATerm (name, [aux NONE t1, aux (SOME t1) t2])
   437         | _ => raise Fail "malformed type wrapper"
   438       else
   439         let
   440           val ts = map (aux NONE) ts
   441           val (ts1, ts2) = chop (min_arity_of thy full_types const_tab s) ts
   442         in list_hAPP_rev opt_ty (ATerm (name, ts1)) (rev ts2) end
   443   in aux NONE end
   444 
   445 fun boolify t = ATerm (`I "hBOOL", [t])
   446 
   447 (* True if the constant ever appears outside of the top-level position in
   448    literals, or if it appears with different arities (e.g., because of different
   449    type instantiations). If false, the constant always receives all of its
   450    arguments and is used as a predicate. *)
   451 fun is_predicate NONE s =
   452     s = "equal" orelse s = "$false" orelse s = "$true" orelse
   453     String.isPrefix type_const_prefix s orelse String.isPrefix class_prefix s
   454   | is_predicate (SOME the_const_tab) s =
   455     case Symtab.lookup the_const_tab s of
   456       SOME {min_arity, max_arity, sub_level} =>
   457       not sub_level andalso min_arity = max_arity
   458     | NONE => false
   459 
   460 fun repair_predicates_in_term const_tab (t as ATerm ((s, _), ts)) =
   461   if s = type_wrapper_name then
   462     case ts of
   463       [_, t' as ATerm ((s', _), _)] =>
   464       if is_predicate const_tab s' then t' else boolify t
   465     | _ => raise Fail "malformed type wrapper"
   466   else
   467     t |> not (is_predicate const_tab s) ? boolify
   468 
   469 fun close_universally phi =
   470   let
   471     fun term_vars bounds (ATerm (name as (s, _), tms)) =
   472         (is_tptp_variable s andalso not (member (op =) bounds name))
   473           ? insert (op =) name
   474         #> fold (term_vars bounds) tms
   475     fun formula_vars bounds (AQuant (_, xs, phi)) =
   476         formula_vars (xs @ bounds) phi
   477       | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis
   478       | formula_vars bounds (AAtom tm) = term_vars bounds tm
   479   in
   480     case formula_vars [] phi [] of [] => phi | xs => AQuant (AForall, xs, phi)
   481   end
   482 
   483 fun repair_formula thy explicit_forall full_types const_tab =
   484   let
   485     fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi)
   486       | aux (AConn (c, phis)) = AConn (c, map aux phis)
   487       | aux (AAtom tm) =
   488         AAtom (tm |> repair_applications_in_term thy full_types const_tab
   489                   |> repair_predicates_in_term const_tab)
   490   in aux #> explicit_forall ? close_universally end
   491 
   492 fun repair_problem_line thy explicit_forall full_types const_tab
   493                         (Fof (ident, kind, phi)) =
   494   Fof (ident, kind, repair_formula thy explicit_forall full_types const_tab phi)
   495 fun repair_problem_with_const_table thy =
   496   map o apsnd o map ooo repair_problem_line thy
   497 
   498 fun repair_problem thy explicit_forall full_types explicit_apply problem =
   499   repair_problem_with_const_table thy explicit_forall full_types
   500       (const_table_for_problem explicit_apply problem) problem
   501 
   502 fun prepare_problem ctxt readable_names explicit_forall full_types
   503                     explicit_apply hyp_ts concl_t axiom_pairs =
   504   let
   505     val thy = ProofContext.theory_of ctxt
   506     val (axiom_names, (conjectures, axioms, helper_facts, class_rel_clauses,
   507                        arity_clauses)) =
   508       prepare_formulas ctxt full_types hyp_ts concl_t axiom_pairs
   509     val axiom_lines = map (problem_line_for_fact axiom_prefix full_types) axioms
   510     val helper_lines =
   511       map (problem_line_for_fact helper_prefix full_types) helper_facts
   512     val conjecture_lines =
   513       map (problem_line_for_conjecture full_types) conjectures
   514     val tfree_lines = problem_lines_for_free_types conjectures
   515     val class_rel_lines =
   516       map problem_line_for_class_rel_clause class_rel_clauses
   517     val arity_lines = map problem_line_for_arity_clause arity_clauses
   518     (* Reordering these might or might not confuse the proof reconstruction
   519        code or the SPASS Flotter hack. *)
   520     val problem =
   521       [("Relevant facts", axiom_lines),
   522        ("Class relationships", class_rel_lines),
   523        ("Arity declarations", arity_lines),
   524        ("Helper facts", helper_lines),
   525        ("Conjectures", conjecture_lines),
   526        ("Type variables", tfree_lines)]
   527       |> repair_problem thy explicit_forall full_types explicit_apply
   528     val (problem, pool) = nice_tptp_problem readable_names problem
   529     val conjecture_offset =
   530       length axiom_lines + length class_rel_lines + length arity_lines
   531       + length helper_lines
   532   in
   533     (problem,
   534      case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
   535      conjecture_offset, axiom_names)
   536   end
   537 
   538 end;