src/HOL/Tools/ATP/atp_problem_generate.ML
author blanchet
Tue, 24 Apr 2012 09:47:40 +0200
changeset 48586 04400144c6fc
parent 48584 bd0683000a0f
child 48589 39229c760636
permissions -rw-r--r--
handle TPTP definitions as definitions in Nitpick rather than as axioms
     1 (*  Title:      HOL/Tools/ATP/atp_problem_generate.ML
     2     Author:     Fabian Immler, TU Muenchen
     3     Author:     Makarius
     4     Author:     Jasmin Blanchette, TU Muenchen
     5 
     6 Translation of HOL to FOL for Metis and Sledgehammer.
     7 *)
     8 
     9 signature ATP_PROBLEM_GENERATE =
    10 sig
    11   type ('a, 'b) ho_term = ('a, 'b) ATP_Problem.ho_term
    12   type connective = ATP_Problem.connective
    13   type ('a, 'b, 'c) formula = ('a, 'b, 'c) ATP_Problem.formula
    14   type atp_format = ATP_Problem.atp_format
    15   type formula_kind = ATP_Problem.formula_kind
    16   type 'a problem = 'a ATP_Problem.problem
    17 
    18   datatype scope = Global | Local | Assum | Chained
    19   datatype status =
    20     General | Induction | Intro | Inductive | Elim | Simp | Def
    21   type stature = scope * status
    22 
    23   datatype polymorphism = Polymorphic | Raw_Monomorphic | Mangled_Monomorphic
    24   datatype strictness = Strict | Non_Strict
    25   datatype granularity = All_Vars | Positively_Naked_Vars | Ghost_Type_Arg_Vars
    26   datatype type_level =
    27     All_Types |
    28     Noninf_Nonmono_Types of strictness * granularity |
    29     Fin_Nonmono_Types of granularity |
    30     Const_Arg_Types |
    31     No_Types
    32   type type_enc
    33 
    34   val no_lamsN : string
    35   val hide_lamsN : string
    36   val liftingN : string
    37   val combsN : string
    38   val combs_and_liftingN : string
    39   val combs_or_liftingN : string
    40   val lam_liftingN : string
    41   val keep_lamsN : string
    42   val schematic_var_prefix : string
    43   val fixed_var_prefix : string
    44   val tvar_prefix : string
    45   val tfree_prefix : string
    46   val const_prefix : string
    47   val type_const_prefix : string
    48   val class_prefix : string
    49   val lam_lifted_prefix : string
    50   val lam_lifted_mono_prefix : string
    51   val lam_lifted_poly_prefix : string
    52   val skolem_const_prefix : string
    53   val old_skolem_const_prefix : string
    54   val new_skolem_const_prefix : string
    55   val combinator_prefix : string
    56   val type_decl_prefix : string
    57   val sym_decl_prefix : string
    58   val guards_sym_formula_prefix : string
    59   val tags_sym_formula_prefix : string
    60   val fact_prefix : string
    61   val conjecture_prefix : string
    62   val helper_prefix : string
    63   val class_rel_clause_prefix : string
    64   val arity_clause_prefix : string
    65   val tfree_clause_prefix : string
    66   val lam_fact_prefix : string
    67   val typed_helper_suffix : string
    68   val untyped_helper_suffix : string
    69   val predicator_name : string
    70   val app_op_name : string
    71   val type_guard_name : string
    72   val type_tag_name : string
    73   val native_type_prefix : string
    74   val prefixed_predicator_name : string
    75   val prefixed_app_op_name : string
    76   val prefixed_type_tag_name : string
    77   val ascii_of : string -> string
    78   val unascii_of : string -> string
    79   val unprefix_and_unascii : string -> string -> string option
    80   val proxy_table : (string * (string * (thm * (string * string)))) list
    81   val proxify_const : string -> (string * string) option
    82   val invert_const : string -> string
    83   val unproxify_const : string -> string
    84   val new_skolem_var_name_from_const : string -> string
    85   val atp_irrelevant_consts : string list
    86   val atp_schematic_consts_of : term -> typ list Symtab.table
    87   val is_type_enc_higher_order : type_enc -> bool
    88   val polymorphism_of_type_enc : type_enc -> polymorphism
    89   val level_of_type_enc : type_enc -> type_level
    90   val is_type_enc_quasi_sound : type_enc -> bool
    91   val is_type_enc_fairly_sound : type_enc -> bool
    92   val type_enc_from_string : strictness -> string -> type_enc
    93   val adjust_type_enc : atp_format -> type_enc -> type_enc
    94   val mk_aconns :
    95     connective -> ('a, 'b, 'c) formula list -> ('a, 'b, 'c) formula
    96   val unmangled_const : string -> string * (string, 'b) ho_term list
    97   val unmangled_const_name : string -> string list
    98   val helper_table : ((string * bool) * thm list) list
    99   val trans_lams_from_string :
   100     Proof.context -> type_enc -> string -> term list -> term list * term list
   101   val factsN : string
   102   val prepare_atp_problem :
   103     Proof.context -> atp_format -> formula_kind -> formula_kind -> type_enc
   104     -> bool -> string -> bool -> bool -> bool -> term list -> term
   105     -> ((string * stature) * term) list
   106     -> string problem * string Symtab.table * (string * stature) list vector
   107        * (string * term) list * int Symtab.table
   108   val atp_problem_selection_weights : string problem -> (string * real) list
   109   val atp_problem_term_order_info : string problem -> (string * int) list
   110 end;
   111 
   112 structure ATP_Problem_Generate : ATP_PROBLEM_GENERATE =
   113 struct
   114 
   115 open ATP_Util
   116 open ATP_Problem
   117 
   118 type name = string * string
   119 
   120 val no_lamsN = "no_lams" (* used internally; undocumented *)
   121 val hide_lamsN = "hide_lams"
   122 val liftingN = "lifting"
   123 val combsN = "combs"
   124 val combs_and_liftingN = "combs_and_lifting"
   125 val combs_or_liftingN = "combs_or_lifting"
   126 val keep_lamsN = "keep_lams"
   127 val lam_liftingN = "lam_lifting" (* legacy *)
   128 
   129 (* It's still unclear whether all TFF1 implementations will support type
   130    signatures such as "!>[A : $tType] : $o", with ghost type variables. *)
   131 val avoid_first_order_ghost_type_vars = false
   132 
   133 val bound_var_prefix = "B_"
   134 val all_bound_var_prefix = "A_"
   135 val exist_bound_var_prefix = "E_"
   136 val schematic_var_prefix = "V_"
   137 val fixed_var_prefix = "v_"
   138 val tvar_prefix = "T_"
   139 val tfree_prefix = "t_"
   140 val const_prefix = "c_"
   141 val type_const_prefix = "tc_"
   142 val native_type_prefix = "n_"
   143 val class_prefix = "cl_"
   144 
   145 (* Freshness almost guaranteed! *)
   146 val atp_prefix = "ATP" ^ Long_Name.separator
   147 val atp_weak_prefix = "ATP:"
   148 
   149 val lam_lifted_prefix = atp_weak_prefix ^ "Lam"
   150 val lam_lifted_mono_prefix = lam_lifted_prefix ^ "m"
   151 val lam_lifted_poly_prefix = lam_lifted_prefix ^ "p"
   152 
   153 val skolem_const_prefix = atp_prefix ^ "Sko"
   154 val old_skolem_const_prefix = skolem_const_prefix ^ "o"
   155 val new_skolem_const_prefix = skolem_const_prefix ^ "n"
   156 
   157 val combinator_prefix = "COMB"
   158 
   159 val type_decl_prefix = "ty_"
   160 val sym_decl_prefix = "sy_"
   161 val guards_sym_formula_prefix = "gsy_"
   162 val tags_sym_formula_prefix = "tsy_"
   163 val uncurried_alias_eq_prefix = "unc_"
   164 val fact_prefix = "fact_"
   165 val conjecture_prefix = "conj_"
   166 val helper_prefix = "help_"
   167 val class_rel_clause_prefix = "clar_"
   168 val arity_clause_prefix = "arity_"
   169 val tfree_clause_prefix = "tfree_"
   170 
   171 val lam_fact_prefix = "ATP.lambda_"
   172 val typed_helper_suffix = "_T"
   173 val untyped_helper_suffix = "_U"
   174 
   175 val predicator_name = "pp"
   176 val app_op_name = "aa"
   177 val type_guard_name = "gg"
   178 val type_tag_name = "tt"
   179 
   180 val prefixed_predicator_name = const_prefix ^ predicator_name
   181 val prefixed_app_op_name = const_prefix ^ app_op_name
   182 val prefixed_type_tag_name = const_prefix ^ type_tag_name
   183 
   184 (*Escaping of special characters.
   185   Alphanumeric characters are left unchanged.
   186   The character _ goes to __
   187   Characters in the range ASCII space to / go to _A to _P, respectively.
   188   Other characters go to _nnn where nnn is the decimal ASCII code.*)
   189 val upper_a_minus_space = Char.ord #"A" - Char.ord #" "
   190 
   191 fun stringN_of_int 0 _ = ""
   192   | stringN_of_int k n =
   193     stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10)
   194 
   195 fun ascii_of_char c =
   196   if Char.isAlphaNum c then
   197     String.str c
   198   else if c = #"_" then
   199     "__"
   200   else if #" " <= c andalso c <= #"/" then
   201     "_" ^ String.str (Char.chr (Char.ord c + upper_a_minus_space))
   202   else
   203     (* fixed width, in case more digits follow *)
   204     "_" ^ stringN_of_int 3 (Char.ord c)
   205 
   206 val ascii_of = String.translate ascii_of_char
   207 
   208 (** Remove ASCII armoring from names in proof files **)
   209 
   210 (* We don't raise error exceptions because this code can run inside a worker
   211    thread. Also, the errors are impossible. *)
   212 val unascii_of =
   213   let
   214     fun un rcs [] = String.implode(rev rcs)
   215       | un rcs [#"_"] = un (#"_" :: rcs) [] (* ERROR *)
   216         (* Three types of _ escapes: __, _A to _P, _nnn *)
   217       | un rcs (#"_" :: #"_" :: cs) = un (#"_" :: rcs) cs
   218       | un rcs (#"_" :: c :: cs) =
   219         if #"A" <= c andalso c<= #"P" then
   220           (* translation of #" " to #"/" *)
   221           un (Char.chr (Char.ord c - upper_a_minus_space) :: rcs) cs
   222         else
   223           let val digits = List.take (c :: cs, 3) handle General.Subscript => [] in
   224             case Int.fromString (String.implode digits) of
   225               SOME n => un (Char.chr n :: rcs) (List.drop (cs, 2))
   226             | NONE => un (c :: #"_" :: rcs) cs (* ERROR *)
   227           end
   228       | un rcs (c :: cs) = un (c :: rcs) cs
   229   in un [] o String.explode end
   230 
   231 (* If string s has the prefix s1, return the result of deleting it,
   232    un-ASCII'd. *)
   233 fun unprefix_and_unascii s1 s =
   234   if String.isPrefix s1 s then
   235     SOME (unascii_of (String.extract (s, size s1, NONE)))
   236   else
   237     NONE
   238 
   239 val proxy_table =
   240   [("c_False", (@{const_name False}, (@{thm fFalse_def},
   241        ("fFalse", @{const_name ATP.fFalse})))),
   242    ("c_True", (@{const_name True}, (@{thm fTrue_def},
   243        ("fTrue", @{const_name ATP.fTrue})))),
   244    ("c_Not", (@{const_name Not}, (@{thm fNot_def},
   245        ("fNot", @{const_name ATP.fNot})))),
   246    ("c_conj", (@{const_name conj}, (@{thm fconj_def},
   247        ("fconj", @{const_name ATP.fconj})))),
   248    ("c_disj", (@{const_name disj}, (@{thm fdisj_def},
   249        ("fdisj", @{const_name ATP.fdisj})))),
   250    ("c_implies", (@{const_name implies}, (@{thm fimplies_def},
   251        ("fimplies", @{const_name ATP.fimplies})))),
   252    ("equal", (@{const_name HOL.eq}, (@{thm fequal_def},
   253        ("fequal", @{const_name ATP.fequal})))),
   254    ("c_All", (@{const_name All}, (@{thm fAll_def},
   255        ("fAll", @{const_name ATP.fAll})))),
   256    ("c_Ex", (@{const_name Ex}, (@{thm fEx_def},
   257        ("fEx", @{const_name ATP.fEx}))))]
   258 
   259 val proxify_const = AList.lookup (op =) proxy_table #> Option.map (snd o snd)
   260 
   261 (* Readable names for the more common symbolic functions. Do not mess with the
   262    table unless you know what you are doing. *)
   263 val const_trans_table =
   264   [(@{type_name Product_Type.prod}, "prod"),
   265    (@{type_name Sum_Type.sum}, "sum"),
   266    (@{const_name False}, "False"),
   267    (@{const_name True}, "True"),
   268    (@{const_name Not}, "Not"),
   269    (@{const_name conj}, "conj"),
   270    (@{const_name disj}, "disj"),
   271    (@{const_name implies}, "implies"),
   272    (@{const_name HOL.eq}, "equal"),
   273    (@{const_name All}, "All"),
   274    (@{const_name Ex}, "Ex"),
   275    (@{const_name If}, "If"),
   276    (@{const_name Set.member}, "member"),
   277    (@{const_name Meson.COMBI}, combinator_prefix ^ "I"),
   278    (@{const_name Meson.COMBK}, combinator_prefix ^ "K"),
   279    (@{const_name Meson.COMBB}, combinator_prefix ^ "B"),
   280    (@{const_name Meson.COMBC}, combinator_prefix ^ "C"),
   281    (@{const_name Meson.COMBS}, combinator_prefix ^ "S")]
   282   |> Symtab.make
   283   |> fold (Symtab.update o swap o snd o snd o snd) proxy_table
   284 
   285 (* Invert the table of translations between Isabelle and ATPs. *)
   286 val const_trans_table_inv =
   287   const_trans_table |> Symtab.dest |> map swap |> Symtab.make
   288 val const_trans_table_unprox =
   289   Symtab.empty
   290   |> fold (fn (_, (isa, (_, (_, atp)))) => Symtab.update (atp, isa)) proxy_table
   291 
   292 val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
   293 val unproxify_const = perhaps (Symtab.lookup const_trans_table_unprox)
   294 
   295 fun lookup_const c =
   296   case Symtab.lookup const_trans_table c of
   297     SOME c' => c'
   298   | NONE => ascii_of c
   299 
   300 fun ascii_of_indexname (v, 0) = ascii_of v
   301   | ascii_of_indexname (v, i) = ascii_of v ^ "_" ^ string_of_int i
   302 
   303 fun make_bound_var x = bound_var_prefix ^ ascii_of x
   304 fun make_all_bound_var x = all_bound_var_prefix ^ ascii_of x
   305 fun make_exist_bound_var x = exist_bound_var_prefix ^ ascii_of x
   306 fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v
   307 fun make_fixed_var x = fixed_var_prefix ^ ascii_of x
   308 
   309 fun make_schematic_type_var (x, i) =
   310   tvar_prefix ^ (ascii_of_indexname (unprefix "'" x, i))
   311 fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (unprefix "'" x))
   312 
   313 (* "HOL.eq" and choice are mapped to the ATP's equivalents *)
   314 local
   315   val choice_const = (fst o dest_Const o HOLogic.choice_const) Term.dummyT
   316   fun default c = const_prefix ^ lookup_const c
   317 in
   318   fun make_fixed_const _ @{const_name HOL.eq} = tptp_old_equal
   319     | make_fixed_const (SOME (THF (_, _, THF_With_Choice))) c =
   320       if c = choice_const then tptp_choice else default c
   321     | make_fixed_const _ c = default c
   322 end
   323 
   324 fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
   325 
   326 fun make_type_class clas = class_prefix ^ ascii_of clas
   327 
   328 fun new_skolem_var_name_from_const s =
   329   let val ss = s |> space_explode Long_Name.separator in
   330     nth ss (length ss - 2)
   331   end
   332 
   333 (* These are either simplified away by "Meson.presimplify" (most of the time) or
   334    handled specially via "fFalse", "fTrue", ..., "fequal". *)
   335 val atp_irrelevant_consts =
   336   [@{const_name False}, @{const_name True}, @{const_name Not},
   337    @{const_name conj}, @{const_name disj}, @{const_name implies},
   338    @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
   339 
   340 val atp_monomorph_bad_consts =
   341   atp_irrelevant_consts @
   342   (* These are ignored anyway by the relevance filter (unless they appear in
   343      higher-order places) but not by the monomorphizer. *)
   344   [@{const_name all}, @{const_name "==>"}, @{const_name "=="},
   345    @{const_name Trueprop}, @{const_name All}, @{const_name Ex},
   346    @{const_name Ex1}, @{const_name Ball}, @{const_name Bex}]
   347 
   348 fun add_schematic_const (x as (_, T)) =
   349   Monomorph.typ_has_tvars T ? Symtab.insert_list (op =) x
   350 val add_schematic_consts_of =
   351   Term.fold_aterms (fn Const (x as (s, _)) =>
   352                        not (member (op =) atp_monomorph_bad_consts s)
   353                        ? add_schematic_const x
   354                       | _ => I)
   355 fun atp_schematic_consts_of t = add_schematic_consts_of t Symtab.empty
   356 
   357 (** Definitions and functions for FOL clauses and formulas for TPTP **)
   358 
   359 (** Isabelle arities **)
   360 
   361 type arity_atom = name * name * name list
   362 
   363 val type_class = the_single @{sort type}
   364 
   365 type arity_clause =
   366   {name : string,
   367    prem_atoms : arity_atom list,
   368    concl_atom : arity_atom}
   369 
   370 fun add_prem_atom tvar =
   371   fold (fn s => s <> type_class ? cons (`make_type_class s, `I tvar, []))
   372 
   373 (* Arity of type constructor "tcon :: (arg1, ..., argN) res" *)
   374 fun make_axiom_arity_clause (tcons, name, (cls, args)) =
   375   let
   376     val tvars = map (prefix tvar_prefix o string_of_int) (1 upto length args)
   377     val tvars_srts = ListPair.zip (tvars, args)
   378   in
   379     {name = name,
   380      prem_atoms = [] |> fold (uncurry add_prem_atom) tvars_srts,
   381      concl_atom = (`make_type_class cls, `make_fixed_type_const tcons,
   382                    tvars ~~ tvars)}
   383   end
   384 
   385 fun arity_clause _ _ (_, []) = []
   386   | arity_clause seen n (tcons, ("HOL.type", _) :: ars) =  (* ignore *)
   387     arity_clause seen n (tcons, ars)
   388   | arity_clause seen n (tcons, (ar as (class, _)) :: ars) =
   389     if member (op =) seen class then
   390       (* multiple arities for the same (tycon, class) pair *)
   391       make_axiom_arity_clause (tcons,
   392           lookup_const tcons ^ "___" ^ ascii_of class ^ "_" ^ string_of_int n,
   393           ar) ::
   394       arity_clause seen (n + 1) (tcons, ars)
   395     else
   396       make_axiom_arity_clause (tcons, lookup_const tcons ^ "___" ^
   397                                ascii_of class, ar) ::
   398       arity_clause (class :: seen) n (tcons, ars)
   399 
   400 fun multi_arity_clause [] = []
   401   | multi_arity_clause ((tcons, ars) :: tc_arlists) =
   402     arity_clause [] 1 (tcons, ars) @ multi_arity_clause tc_arlists
   403 
   404 (* Generate all pairs (tycon, class, sorts) such that tycon belongs to class in
   405    theory thy provided its arguments have the corresponding sorts. *)
   406 fun type_class_pairs thy tycons classes =
   407   let
   408     val alg = Sign.classes_of thy
   409     fun domain_sorts tycon = Sorts.mg_domain alg tycon o single
   410     fun add_class tycon class =
   411       cons (class, domain_sorts tycon class)
   412       handle Sorts.CLASS_ERROR _ => I
   413     fun try_classes tycon = (tycon, fold (add_class tycon) classes [])
   414   in map try_classes tycons end
   415 
   416 (*Proving one (tycon, class) membership may require proving others, so iterate.*)
   417 fun iter_type_class_pairs _ _ [] = ([], [])
   418   | iter_type_class_pairs thy tycons classes =
   419       let
   420         fun maybe_insert_class s =
   421           (s <> type_class andalso not (member (op =) classes s))
   422           ? insert (op =) s
   423         val cpairs = type_class_pairs thy tycons classes
   424         val newclasses =
   425           [] |> fold (fold (fold (fold maybe_insert_class) o snd) o snd) cpairs
   426         val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
   427       in (classes' @ classes, union (op =) cpairs' cpairs) end
   428 
   429 fun make_arity_clauses thy tycons =
   430   iter_type_class_pairs thy tycons ##> multi_arity_clause
   431 
   432 
   433 (** Isabelle class relations **)
   434 
   435 type class_rel_clause =
   436   {name : string,
   437    subclass : name,
   438    superclass : name}
   439 
   440 (* Generate all pairs (sub, super) such that sub is a proper subclass of super
   441    in theory "thy". *)
   442 fun class_pairs _ [] _ = []
   443   | class_pairs thy subs supers =
   444       let
   445         val class_less = Sorts.class_less (Sign.classes_of thy)
   446         fun add_super sub super = class_less (sub, super) ? cons (sub, super)
   447         fun add_supers sub = fold (add_super sub) supers
   448       in fold add_supers subs [] end
   449 
   450 fun make_class_rel_clause (sub, super) =
   451   {name = sub ^ "_" ^ super, subclass = `make_type_class sub,
   452    superclass = `make_type_class super}
   453 
   454 fun make_class_rel_clauses thy subs supers =
   455   map make_class_rel_clause (class_pairs thy subs supers)
   456 
   457 (* intermediate terms *)
   458 datatype iterm =
   459   IConst of name * typ * typ list |
   460   IVar of name * typ |
   461   IApp of iterm * iterm |
   462   IAbs of (name * typ) * iterm
   463 
   464 fun ityp_of (IConst (_, T, _)) = T
   465   | ityp_of (IVar (_, T)) = T
   466   | ityp_of (IApp (t1, _)) = snd (dest_funT (ityp_of t1))
   467   | ityp_of (IAbs ((_, T), tm)) = T --> ityp_of tm
   468 
   469 (*gets the head of a combinator application, along with the list of arguments*)
   470 fun strip_iterm_comb u =
   471   let
   472     fun stripc (IApp (t, u), ts) = stripc (t, u :: ts)
   473       | stripc x = x
   474   in stripc (u, []) end
   475 
   476 fun atomic_types_of T = fold_atyps (insert (op =)) T []
   477 
   478 val tvar_a_str = "'a"
   479 val tvar_a = TVar ((tvar_a_str, 0), HOLogic.typeS)
   480 val tvar_a_name = (make_schematic_type_var (tvar_a_str, 0), tvar_a_str)
   481 val itself_name = `make_fixed_type_const @{type_name itself}
   482 val TYPE_name = `(make_fixed_const NONE) @{const_name TYPE}
   483 val tvar_a_atype = AType (tvar_a_name, [])
   484 val a_itself_atype = AType (itself_name, [tvar_a_atype])
   485 
   486 fun new_skolem_const_name s num_T_args =
   487   [new_skolem_const_prefix, s, string_of_int num_T_args]
   488   |> Long_Name.implode
   489 
   490 fun robust_const_type thy s =
   491   if s = app_op_name then
   492     Logic.varifyT_global @{typ "('a => 'b) => 'a => 'b"}
   493   else if String.isPrefix lam_lifted_prefix s then
   494     Logic.varifyT_global @{typ "'a => 'b"}
   495   else
   496     (* Old Skolems throw a "TYPE" exception here, which will be caught. *)
   497     s |> Sign.the_const_type thy
   498 
   499 val robust_const_ary =
   500   let
   501     fun ary (Type (@{type_name fun}, [_, T])) = 1 + ary T
   502       | ary _ = 0
   503   in ary oo robust_const_type end
   504 
   505 (* This function only makes sense if "T" is as general as possible. *)
   506 fun robust_const_typargs thy (s, T) =
   507   if s = app_op_name then
   508     let val (T1, T2) = T |> domain_type |> dest_funT in [T1, T2] end
   509   else if String.isPrefix old_skolem_const_prefix s then
   510     [] |> Term.add_tvarsT T |> rev |> map TVar
   511   else if String.isPrefix lam_lifted_prefix s then
   512     if String.isPrefix lam_lifted_poly_prefix s then
   513       let val (T1, T2) = T |> dest_funT in [T1, T2] end
   514     else
   515       []
   516   else
   517     (s, T) |> Sign.const_typargs thy
   518 
   519 (* Converts an Isabelle/HOL term (with combinators) into an intermediate term.
   520    Also accumulates sort infomation. *)
   521 fun iterm_from_term thy format bs (P $ Q) =
   522     let
   523       val (P', P_atomics_Ts) = iterm_from_term thy format bs P
   524       val (Q', Q_atomics_Ts) = iterm_from_term thy format bs Q
   525     in (IApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end
   526   | iterm_from_term thy format _ (Const (c, T)) =
   527     (IConst (`(make_fixed_const (SOME format)) c, T,
   528              robust_const_typargs thy (c, T)),
   529      atomic_types_of T)
   530   | iterm_from_term _ _ _ (Free (s, T)) =
   531     (IConst (`make_fixed_var s, T, []), atomic_types_of T)
   532   | iterm_from_term _ format _ (Var (v as (s, _), T)) =
   533     (if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then
   534        let
   535          val Ts = T |> strip_type |> swap |> op ::
   536          val s' = new_skolem_const_name s (length Ts)
   537        in IConst (`(make_fixed_const (SOME format)) s', T, Ts) end
   538      else
   539        IVar ((make_schematic_var v, s), T), atomic_types_of T)
   540   | iterm_from_term _ _ bs (Bound j) =
   541     nth bs j |> (fn (_, (name, T)) => (IConst (name, T, []), atomic_types_of T))
   542   | iterm_from_term thy format bs (Abs (s, T, t)) =
   543     let
   544       fun vary s = s |> AList.defined (op =) bs s ? vary o Symbol.bump_string
   545       val s = vary s
   546       val name = `make_bound_var s
   547       val (tm, atomic_Ts) = iterm_from_term thy format ((s, (name, T)) :: bs) t
   548     in (IAbs ((name, T), tm), union (op =) atomic_Ts (atomic_types_of T)) end
   549 
   550 datatype scope = Global | Local | Assum | Chained
   551 datatype status = General | Induction | Intro | Inductive | Elim | Simp | Def
   552 type stature = scope * status
   553 
   554 datatype order = First_Order | Higher_Order
   555 datatype polymorphism = Polymorphic | Raw_Monomorphic | Mangled_Monomorphic
   556 datatype strictness = Strict | Non_Strict
   557 datatype granularity = All_Vars | Positively_Naked_Vars | Ghost_Type_Arg_Vars
   558 datatype type_level =
   559   All_Types |
   560   Noninf_Nonmono_Types of strictness * granularity |
   561   Fin_Nonmono_Types of granularity |
   562   Const_Arg_Types |
   563   No_Types
   564 
   565 datatype type_enc =
   566   Simple_Types of order * polymorphism * type_level |
   567   Guards of polymorphism * type_level |
   568   Tags of polymorphism * type_level
   569 
   570 fun is_type_enc_higher_order (Simple_Types (Higher_Order, _, _)) = true
   571   | is_type_enc_higher_order _ = false
   572 
   573 fun polymorphism_of_type_enc (Simple_Types (_, poly, _)) = poly
   574   | polymorphism_of_type_enc (Guards (poly, _)) = poly
   575   | polymorphism_of_type_enc (Tags (poly, _)) = poly
   576 
   577 fun level_of_type_enc (Simple_Types (_, _, level)) = level
   578   | level_of_type_enc (Guards (_, level)) = level
   579   | level_of_type_enc (Tags (_, level)) = level
   580 
   581 fun granularity_of_type_level (Noninf_Nonmono_Types (_, grain)) = grain
   582   | granularity_of_type_level (Fin_Nonmono_Types grain) = grain
   583   | granularity_of_type_level _ = All_Vars
   584 
   585 fun is_type_level_quasi_sound All_Types = true
   586   | is_type_level_quasi_sound (Noninf_Nonmono_Types _) = true
   587   | is_type_level_quasi_sound _ = false
   588 val is_type_enc_quasi_sound = is_type_level_quasi_sound o level_of_type_enc
   589 
   590 fun is_type_level_fairly_sound (Fin_Nonmono_Types _) = true
   591   | is_type_level_fairly_sound level = is_type_level_quasi_sound level
   592 val is_type_enc_fairly_sound = is_type_level_fairly_sound o level_of_type_enc
   593 
   594 fun is_type_level_monotonicity_based (Noninf_Nonmono_Types _) = true
   595   | is_type_level_monotonicity_based (Fin_Nonmono_Types _) = true
   596   | is_type_level_monotonicity_based _ = false
   597 
   598 (* "_query", "_bang", and "_at" are for the ASCII-challenged Metis and
   599    Mirabelle. *)
   600 val queries = ["?", "_query"]
   601 val bangs = ["!", "_bang"]
   602 val ats = ["@", "_at"]
   603 
   604 fun try_unsuffixes ss s =
   605   fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE
   606 
   607 fun try_nonmono constr suffixes fallback s =
   608   case try_unsuffixes suffixes s of
   609     SOME s =>
   610     (case try_unsuffixes suffixes s of
   611        SOME s => (constr Positively_Naked_Vars, s)
   612      | NONE =>
   613        case try_unsuffixes ats s of
   614          SOME s => (constr Ghost_Type_Arg_Vars, s)
   615        | NONE => (constr All_Vars, s))
   616   | NONE => fallback s
   617 
   618 fun type_enc_from_string strictness s =
   619   (case try (unprefix "poly_") s of
   620      SOME s => (SOME Polymorphic, s)
   621    | NONE =>
   622      case try (unprefix "raw_mono_") s of
   623        SOME s => (SOME Raw_Monomorphic, s)
   624      | NONE =>
   625        case try (unprefix "mono_") s of
   626          SOME s => (SOME Mangled_Monomorphic, s)
   627        | NONE => (NONE, s))
   628   ||> (pair All_Types
   629        |> try_nonmono Fin_Nonmono_Types bangs
   630        |> try_nonmono (curry Noninf_Nonmono_Types strictness) queries)
   631   |> (fn (poly, (level, core)) =>
   632          case (core, (poly, level)) of
   633            ("native", (SOME poly, _)) =>
   634            (case (poly, level) of
   635               (Polymorphic, All_Types) =>
   636               Simple_Types (First_Order, Polymorphic, All_Types)
   637             | (Mangled_Monomorphic, _) =>
   638               if granularity_of_type_level level = All_Vars then
   639                 Simple_Types (First_Order, Mangled_Monomorphic, level)
   640               else
   641                 raise Same.SAME
   642             | _ => raise Same.SAME)
   643          | ("native_higher", (SOME poly, _)) =>
   644            (case (poly, level) of
   645               (Polymorphic, All_Types) =>
   646               Simple_Types (Higher_Order, Polymorphic, All_Types)
   647             | (_, Noninf_Nonmono_Types _) => raise Same.SAME
   648             | (Mangled_Monomorphic, _) =>
   649               if granularity_of_type_level level = All_Vars then
   650                 Simple_Types (Higher_Order, Mangled_Monomorphic, level)
   651               else
   652                 raise Same.SAME
   653             | _ => raise Same.SAME)
   654          | ("guards", (SOME poly, _)) =>
   655            if poly = Mangled_Monomorphic andalso
   656               granularity_of_type_level level = Ghost_Type_Arg_Vars then
   657              raise Same.SAME
   658            else
   659              Guards (poly, level)
   660          | ("tags", (SOME poly, _)) =>
   661            if granularity_of_type_level level = Ghost_Type_Arg_Vars then
   662              raise Same.SAME
   663            else
   664              Tags (poly, level)
   665          | ("args", (SOME poly, All_Types (* naja *))) =>
   666            Guards (poly, Const_Arg_Types)
   667          | ("erased", (NONE, All_Types (* naja *))) =>
   668            Guards (Polymorphic, No_Types)
   669          | _ => raise Same.SAME)
   670   handle Same.SAME => error ("Unknown type encoding: " ^ quote s ^ ".")
   671 
   672 fun adjust_type_enc (THF (TPTP_Monomorphic, _, _))
   673                     (Simple_Types (order, _, level)) =
   674     Simple_Types (order, Mangled_Monomorphic, level)
   675   | adjust_type_enc (THF _) type_enc = type_enc
   676   | adjust_type_enc (TFF (TPTP_Monomorphic, _)) (Simple_Types (_, _, level)) =
   677     Simple_Types (First_Order, Mangled_Monomorphic, level)
   678   | adjust_type_enc (DFG DFG_Sorted) (Simple_Types (_, _, level)) =
   679     Simple_Types (First_Order, Mangled_Monomorphic, level)
   680   | adjust_type_enc (TFF _) (Simple_Types (_, poly, level)) =
   681     Simple_Types (First_Order, poly, level)
   682   | adjust_type_enc format (Simple_Types (_, poly, level)) =
   683     adjust_type_enc format (Guards (poly, level))
   684   | adjust_type_enc CNF_UEQ (type_enc as Guards stuff) =
   685     (if is_type_enc_fairly_sound type_enc then Tags else Guards) stuff
   686   | adjust_type_enc _ type_enc = type_enc
   687 
   688 fun is_fol_term t =
   689   case t of
   690     @{const Not} $ t1 => is_fol_term t1
   691   | Const (@{const_name All}, _) $ Abs (_, _, t') => is_fol_term t'
   692   | Const (@{const_name All}, _) $ t1 => is_fol_term t1
   693   | Const (@{const_name Ex}, _) $ Abs (_, _, t') => is_fol_term t'
   694   | Const (@{const_name Ex}, _) $ t1 => is_fol_term t1
   695   | @{const HOL.conj} $ t1 $ t2 => is_fol_term t1 andalso is_fol_term t2
   696   | @{const HOL.disj} $ t1 $ t2 => is_fol_term t1 andalso is_fol_term t2
   697   | @{const HOL.implies} $ t1 $ t2 => is_fol_term t1 andalso is_fol_term t2
   698   | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
   699     is_fol_term t1 andalso is_fol_term t2
   700   | _ => not (exists_subterm (fn Abs _ => true | _ => false) t)
   701 
   702 fun simple_translate_lambdas do_lambdas ctxt t =
   703   if is_fol_term t then
   704     t
   705   else
   706     let
   707       fun trans Ts t =
   708         case t of
   709           @{const Not} $ t1 => @{const Not} $ trans Ts t1
   710         | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
   711           t0 $ Abs (s, T, trans (T :: Ts) t')
   712         | (t0 as Const (@{const_name All}, _)) $ t1 =>
   713           trans Ts (t0 $ eta_expand Ts t1 1)
   714         | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
   715           t0 $ Abs (s, T, trans (T :: Ts) t')
   716         | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
   717           trans Ts (t0 $ eta_expand Ts t1 1)
   718         | (t0 as @{const HOL.conj}) $ t1 $ t2 =>
   719           t0 $ trans Ts t1 $ trans Ts t2
   720         | (t0 as @{const HOL.disj}) $ t1 $ t2 =>
   721           t0 $ trans Ts t1 $ trans Ts t2
   722         | (t0 as @{const HOL.implies}) $ t1 $ t2 =>
   723           t0 $ trans Ts t1 $ trans Ts t2
   724         | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
   725             $ t1 $ t2 =>
   726           t0 $ trans Ts t1 $ trans Ts t2
   727         | _ =>
   728           if not (exists_subterm (fn Abs _ => true | _ => false) t) then t
   729           else t |> Envir.eta_contract |> do_lambdas ctxt Ts
   730       val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
   731     in t |> trans [] |> singleton (Variable.export_terms ctxt' ctxt) end
   732 
   733 fun do_cheaply_conceal_lambdas Ts (t1 $ t2) =
   734     do_cheaply_conceal_lambdas Ts t1
   735     $ do_cheaply_conceal_lambdas Ts t2
   736   | do_cheaply_conceal_lambdas Ts (Abs (_, T, t)) =
   737     Const (lam_lifted_poly_prefix ^ serial_string (),
   738            T --> fastype_of1 (T :: Ts, t))
   739   | do_cheaply_conceal_lambdas _ t = t
   740 
   741 fun concealed_bound_name j = atp_weak_prefix ^ string_of_int j
   742 fun conceal_bounds Ts t =
   743   subst_bounds (map (Free o apfst concealed_bound_name)
   744                     (0 upto length Ts - 1 ~~ Ts), t)
   745 fun reveal_bounds Ts =
   746   subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
   747                     (0 upto length Ts - 1 ~~ Ts))
   748 
   749 fun do_introduce_combinators ctxt Ts t =
   750   let val thy = Proof_Context.theory_of ctxt in
   751     t |> conceal_bounds Ts
   752       |> cterm_of thy
   753       |> Meson_Clausify.introduce_combinators_in_cterm
   754       |> prop_of |> Logic.dest_equals |> snd
   755       |> reveal_bounds Ts
   756   end
   757   (* A type variable of sort "{}" will make abstraction fail. *)
   758   handle THM _ => t |> do_cheaply_conceal_lambdas Ts
   759 val introduce_combinators = simple_translate_lambdas do_introduce_combinators
   760 
   761 fun constify_lifted (t $ u) = constify_lifted t $ constify_lifted u
   762   | constify_lifted (Abs (s, T, t)) = Abs (s, T, constify_lifted t)
   763   | constify_lifted (Free (x as (s, _))) =
   764     (if String.isPrefix lam_lifted_prefix s then Const else Free) x
   765   | constify_lifted t = t
   766 
   767 (* Requires bound variables not to clash with any schematic variables (as should
   768    be the case right after lambda-lifting). *)
   769 fun open_form unprefix (t as Const (@{const_name All}, _) $ Abs (s, T, t')) =
   770     (case try unprefix s of
   771        SOME s =>
   772        let
   773          val names = Name.make_context (map fst (Term.add_var_names t' []))
   774          val (s, _) = Name.variant s names
   775        in open_form unprefix (subst_bound (Var ((s, 0), T), t')) end
   776      | NONE => t)
   777   | open_form _ t = t
   778 
   779 fun lift_lams_part_1 ctxt type_enc =
   780   map close_form #> rpair ctxt
   781   #-> Lambda_Lifting.lift_lambdas
   782           (SOME ((if polymorphism_of_type_enc type_enc = Polymorphic then
   783                     lam_lifted_poly_prefix
   784                   else
   785                     lam_lifted_mono_prefix) ^ "_a"))
   786           Lambda_Lifting.is_quantifier
   787   #> fst
   788 
   789 fun lift_lams_part_2 ctxt (facts, lifted) =
   790   (facts, lifted)
   791   (* Lambda-lifting sometimes leaves some lambdas around; we need some way to get rid
   792      of them *)
   793   |> pairself (map (introduce_combinators ctxt))
   794   |> pairself (map constify_lifted)
   795   |>> map (open_form (unprefix close_form_prefix))
   796   ||> map (open_form I)
   797 
   798 fun lift_lams ctxt = lift_lams_part_2 ctxt oo lift_lams_part_1 ctxt
   799 
   800 fun intentionalize_def (Const (@{const_name All}, _) $ Abs (_, _, t)) =
   801     intentionalize_def t
   802   | intentionalize_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
   803     let
   804       fun lam T t = Abs (Name.uu, T, t)
   805       val (head, args) = strip_comb t ||> rev
   806       val head_T = fastype_of head
   807       val n = length args
   808       val arg_Ts = head_T |> binder_types |> take n |> rev
   809       val u = u |> subst_atomic (args ~~ map Bound (0 upto n - 1))
   810     in HOLogic.eq_const head_T $ head $ fold lam arg_Ts u end
   811   | intentionalize_def t = t
   812 
   813 type translated_formula =
   814   {name : string,
   815    stature : stature,
   816    kind : formula_kind,
   817    iformula : (name, typ, iterm) formula,
   818    atomic_types : typ list}
   819 
   820 fun update_iformula f ({name, stature, kind, iformula, atomic_types}
   821                        : translated_formula) =
   822   {name = name, stature = stature, kind = kind, iformula = f iformula,
   823    atomic_types = atomic_types} : translated_formula
   824 
   825 fun fact_lift f ({iformula, ...} : translated_formula) = f iformula
   826 
   827 fun insert_type thy get_T x xs =
   828   let val T = get_T x in
   829     if exists (type_instance thy T o get_T) xs then xs
   830     else x :: filter_out (type_generalization thy T o get_T) xs
   831   end
   832 
   833 (* The Booleans indicate whether all type arguments should be kept. *)
   834 datatype type_arg_policy =
   835   Explicit_Type_Args of bool (* infer_from_term_args *) |
   836   Mangled_Type_Args |
   837   No_Type_Args
   838 
   839 fun type_arg_policy monom_constrs type_enc s =
   840   let val poly = polymorphism_of_type_enc type_enc in
   841     if s = type_tag_name then
   842       if poly = Mangled_Monomorphic then Mangled_Type_Args
   843       else Explicit_Type_Args false
   844     else case type_enc of
   845       Simple_Types (_, Polymorphic, _) => Explicit_Type_Args false
   846     | Tags (_, All_Types) => No_Type_Args
   847     | _ =>
   848       let val level = level_of_type_enc type_enc in
   849         if level = No_Types orelse s = @{const_name HOL.eq} orelse
   850            (s = app_op_name andalso level = Const_Arg_Types) then
   851           No_Type_Args
   852         else if poly = Mangled_Monomorphic then
   853           Mangled_Type_Args
   854         else if member (op =) monom_constrs s andalso
   855                 granularity_of_type_level level = Positively_Naked_Vars then
   856           No_Type_Args
   857         else
   858           Explicit_Type_Args
   859               (level = All_Types orelse
   860                granularity_of_type_level level = Ghost_Type_Arg_Vars)
   861       end
   862   end
   863 
   864 (* Make atoms for sorted type variables. *)
   865 fun generic_add_sorts_on_type (_, []) = I
   866   | generic_add_sorts_on_type ((x, i), s :: ss) =
   867     generic_add_sorts_on_type ((x, i), ss)
   868     #> (if s = the_single @{sort HOL.type} then
   869           I
   870         else if i = ~1 then
   871           insert (op =) (`make_type_class s, `make_fixed_type_var x)
   872         else
   873           insert (op =) (`make_type_class s,
   874                          (make_schematic_type_var (x, i), x)))
   875 fun add_sorts_on_tfree (TFree (s, S)) = generic_add_sorts_on_type ((s, ~1), S)
   876   | add_sorts_on_tfree _ = I
   877 fun add_sorts_on_tvar (TVar z) = generic_add_sorts_on_type z
   878   | add_sorts_on_tvar _ = I
   879 
   880 fun type_class_formula type_enc class arg =
   881   AAtom (ATerm (class, arg ::
   882       (case type_enc of
   883          Simple_Types (First_Order, Polymorphic, _) =>
   884          if avoid_first_order_ghost_type_vars then [ATerm (TYPE_name, [arg])]
   885          else []
   886        | _ => [])))
   887 fun formulas_for_types type_enc add_sorts_on_typ Ts =
   888   [] |> level_of_type_enc type_enc <> No_Types ? fold add_sorts_on_typ Ts
   889      |> map (fn (class, name) =>
   890                 type_class_formula type_enc class (ATerm (name, [])))
   891 
   892 fun mk_aconns c phis =
   893   let val (phis', phi') = split_last phis in
   894     fold_rev (mk_aconn c) phis' phi'
   895   end
   896 fun mk_ahorn [] phi = phi
   897   | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])
   898 fun mk_aquant _ [] phi = phi
   899   | mk_aquant q xs (phi as AQuant (q', xs', phi')) =
   900     if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)
   901   | mk_aquant q xs phi = AQuant (q, xs, phi)
   902 
   903 fun close_universally add_term_vars phi =
   904   let
   905     fun add_formula_vars bounds (AQuant (_, xs, phi)) =
   906         add_formula_vars (map fst xs @ bounds) phi
   907       | add_formula_vars bounds (AConn (_, phis)) =
   908         fold (add_formula_vars bounds) phis
   909       | add_formula_vars bounds (AAtom tm) = add_term_vars bounds tm
   910   in mk_aquant AForall (add_formula_vars [] phi []) phi end
   911 
   912 fun add_term_vars bounds (ATerm (name as (s, _), tms)) =
   913     (if is_tptp_variable s andalso
   914         not (String.isPrefix tvar_prefix s) andalso
   915         not (member (op =) bounds name) then
   916        insert (op =) (name, NONE)
   917      else
   918        I)
   919     #> fold (add_term_vars bounds) tms
   920   | add_term_vars bounds (AAbs ((name, _), tm)) =
   921     add_term_vars (name :: bounds) tm
   922 fun close_formula_universally phi = close_universally add_term_vars phi
   923 
   924 fun add_iterm_vars bounds (IApp (tm1, tm2)) =
   925     fold (add_iterm_vars bounds) [tm1, tm2]
   926   | add_iterm_vars _ (IConst _) = I
   927   | add_iterm_vars bounds (IVar (name, T)) =
   928     not (member (op =) bounds name) ? insert (op =) (name, SOME T)
   929   | add_iterm_vars bounds (IAbs (_, tm)) = add_iterm_vars bounds tm
   930 fun close_iformula_universally phi = close_universally add_iterm_vars phi
   931 
   932 val fused_infinite_type_name = "ATP.fused_inf" (* shouldn't clash *)
   933 val fused_infinite_type = Type (fused_infinite_type_name, [])
   934 
   935 fun tvar_name (x as (s, _)) = (make_schematic_type_var x, s)
   936 
   937 fun ho_term_from_typ format type_enc =
   938   let
   939     fun term (Type (s, Ts)) =
   940       ATerm (case (is_type_enc_higher_order type_enc, s) of
   941                (true, @{type_name bool}) => `I tptp_bool_type
   942              | (true, @{type_name fun}) => `I tptp_fun_type
   943              | _ => if s = fused_infinite_type_name andalso
   944                        is_format_typed format then
   945                       `I tptp_individual_type
   946                     else
   947                       `make_fixed_type_const s,
   948              map term Ts)
   949     | term (TFree (s, _)) = ATerm (`make_fixed_type_var s, [])
   950     | term (TVar (x, _)) = ATerm (tvar_name x, [])
   951   in term end
   952 
   953 fun ho_term_for_type_arg format type_enc T =
   954   if T = dummyT then NONE else SOME (ho_term_from_typ format type_enc T)
   955 
   956 (* This shouldn't clash with anything else. *)
   957 val uncurried_alias_sep = "\000"
   958 val mangled_type_sep = "\001"
   959 
   960 val ascii_of_uncurried_alias_sep = ascii_of uncurried_alias_sep
   961 
   962 fun generic_mangled_type_name f (ATerm (name, [])) = f name
   963   | generic_mangled_type_name f (ATerm (name, tys)) =
   964     f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys)
   965     ^ ")"
   966   | generic_mangled_type_name _ _ = raise Fail "unexpected type abstraction"
   967 
   968 fun mangled_type format type_enc =
   969   generic_mangled_type_name fst o ho_term_from_typ format type_enc
   970 
   971 fun make_native_type s =
   972   if s = tptp_bool_type orelse s = tptp_fun_type orelse
   973      s = tptp_individual_type then
   974     s
   975   else
   976     native_type_prefix ^ ascii_of s
   977 
   978 fun ho_type_from_ho_term type_enc pred_sym ary =
   979   let
   980     fun to_mangled_atype ty =
   981       AType ((make_native_type (generic_mangled_type_name fst ty),
   982               generic_mangled_type_name snd ty), [])
   983     fun to_poly_atype (ATerm (name, tys)) = AType (name, map to_poly_atype tys)
   984       | to_poly_atype _ = raise Fail "unexpected type abstraction"
   985     val to_atype =
   986       if polymorphism_of_type_enc type_enc = Polymorphic then to_poly_atype
   987       else to_mangled_atype
   988     fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1))
   989     fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty
   990       | to_fo ary (ATerm (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys
   991       | to_fo _ _ = raise Fail "unexpected type abstraction"
   992     fun to_ho (ty as ATerm ((s, _), tys)) =
   993         if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty
   994       | to_ho _ = raise Fail "unexpected type abstraction"
   995   in if is_type_enc_higher_order type_enc then to_ho else to_fo ary end
   996 
   997 fun ho_type_from_typ format type_enc pred_sym ary =
   998   ho_type_from_ho_term type_enc pred_sym ary
   999   o ho_term_from_typ format type_enc
  1000 
  1001 fun aliased_uncurried ary (s, s') =
  1002   (s ^ ascii_of_uncurried_alias_sep ^ string_of_int ary, s' ^ string_of_int ary)
  1003 fun unaliased_uncurried (s, s') =
  1004   case space_explode uncurried_alias_sep s of
  1005     [_] => (s, s')
  1006   | [s1, s2] => (s1, unsuffix s2 s')
  1007   | _ => raise Fail "ill-formed explicit application alias"
  1008 
  1009 fun raw_mangled_const_name type_name ty_args (s, s') =
  1010   let
  1011     fun type_suffix f g =
  1012       fold_rev (curry (op ^) o g o prefix mangled_type_sep o type_name f)
  1013                ty_args ""
  1014   in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end
  1015 fun mangled_const_name format type_enc =
  1016   map_filter (ho_term_for_type_arg format type_enc)
  1017   #> raw_mangled_const_name generic_mangled_type_name
  1018 
  1019 val parse_mangled_ident =
  1020   Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
  1021 
  1022 fun parse_mangled_type x =
  1023   (parse_mangled_ident
  1024    -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
  1025                     [] >> ATerm) x
  1026 and parse_mangled_types x =
  1027   (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x
  1028 
  1029 fun unmangled_type s =
  1030   s |> suffix ")" |> raw_explode
  1031     |> Scan.finite Symbol.stopper
  1032            (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^
  1033                                                 quote s)) parse_mangled_type))
  1034     |> fst
  1035 
  1036 fun unmangled_const_name s =
  1037   (s, s) |> unaliased_uncurried |> fst |> space_explode mangled_type_sep
  1038 fun unmangled_const s =
  1039   let val ss = unmangled_const_name s in
  1040     (hd ss, map unmangled_type (tl ss))
  1041   end
  1042 
  1043 fun introduce_proxies_in_iterm type_enc =
  1044   let
  1045     fun tweak_ho_quant ho_quant T [IAbs _] = IConst (`I ho_quant, T, [])
  1046       | tweak_ho_quant ho_quant (T as Type (_, [p_T as Type (_, [x_T, _]), _]))
  1047                        _ =
  1048         (* Eta-expand "!!" and "??", to work around LEO-II 1.2.8 parser
  1049            limitation. This works in conjuction with special code in
  1050            "ATP_Problem" that uses the syntactic sugar "!" and "?" whenever
  1051            possible. *)
  1052         IAbs ((`I "P", p_T),
  1053               IApp (IConst (`I ho_quant, T, []),
  1054                     IAbs ((`I "X", x_T),
  1055                           IApp (IConst (`I "P", p_T, []),
  1056                                 IConst (`I "X", x_T, [])))))
  1057       | tweak_ho_quant _ _ _ = raise Fail "unexpected type for quantifier"
  1058     fun intro top_level args (IApp (tm1, tm2)) =
  1059         IApp (intro top_level (tm2 :: args) tm1, intro false [] tm2)
  1060       | intro top_level args (IConst (name as (s, _), T, T_args)) =
  1061         (case proxify_const s of
  1062            SOME proxy_base =>
  1063            if top_level orelse is_type_enc_higher_order type_enc then
  1064              case (top_level, s) of
  1065                (_, "c_False") => IConst (`I tptp_false, T, [])
  1066              | (_, "c_True") => IConst (`I tptp_true, T, [])
  1067              | (false, "c_Not") => IConst (`I tptp_not, T, [])
  1068              | (false, "c_conj") => IConst (`I tptp_and, T, [])
  1069              | (false, "c_disj") => IConst (`I tptp_or, T, [])
  1070              | (false, "c_implies") => IConst (`I tptp_implies, T, [])
  1071              | (false, "c_All") => tweak_ho_quant tptp_ho_forall T args
  1072              | (false, "c_Ex") => tweak_ho_quant tptp_ho_exists T args
  1073              | (false, s) =>
  1074                if is_tptp_equal s andalso length args = 2 then
  1075                  IConst (`I tptp_equal, T, [])
  1076                else
  1077                  (* Use a proxy even for partially applied THF0 equality,
  1078                     because the LEO-II and Satallax parsers complain about not
  1079                     being able to infer the type of "=". *)
  1080                  IConst (proxy_base |>> prefix const_prefix, T, T_args)
  1081              | _ => IConst (name, T, [])
  1082            else
  1083              IConst (proxy_base |>> prefix const_prefix, T, T_args)
  1084           | NONE => if s = tptp_choice then tweak_ho_quant tptp_choice T args
  1085                     else IConst (name, T, T_args))
  1086       | intro _ _ (IAbs (bound, tm)) = IAbs (bound, intro false [] tm)
  1087       | intro _ _ tm = tm
  1088   in intro true [] end
  1089 
  1090 fun mangle_type_args_in_const format type_enc (name as (s, _)) T_args =
  1091   case unprefix_and_unascii const_prefix s of
  1092     NONE => (name, T_args)
  1093   | SOME s'' =>
  1094     case type_arg_policy [] type_enc (invert_const s'') of
  1095       Mangled_Type_Args => (mangled_const_name format type_enc T_args name, [])
  1096     | _ => (name, T_args)
  1097 fun mangle_type_args_in_iterm format type_enc =
  1098   if polymorphism_of_type_enc type_enc = Mangled_Monomorphic then
  1099     let
  1100       fun mangle (IApp (tm1, tm2)) = IApp (mangle tm1, mangle tm2)
  1101         | mangle (tm as IConst (_, _, [])) = tm
  1102         | mangle (IConst (name, T, T_args)) =
  1103           mangle_type_args_in_const format type_enc name T_args
  1104           |> (fn (name, T_args) => IConst (name, T, T_args))
  1105         | mangle (IAbs (bound, tm)) = IAbs (bound, mangle tm)
  1106         | mangle tm = tm
  1107     in mangle end
  1108   else
  1109     I
  1110 
  1111 fun chop_fun 0 T = ([], T)
  1112   | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
  1113     chop_fun (n - 1) ran_T |>> cons dom_T
  1114   | chop_fun _ T = ([], T)
  1115 
  1116 fun filter_const_type_args _ _ _ [] = []
  1117   | filter_const_type_args thy s ary T_args =
  1118     let
  1119       val U = robust_const_type thy s
  1120       val arg_U_vars = fold Term.add_tvarsT (U |> chop_fun ary |> fst) []
  1121       val U_args = (s, U) |> robust_const_typargs thy
  1122     in
  1123       U_args ~~ T_args
  1124       |> map (fn (U, T) =>
  1125                  if member (op =) arg_U_vars (dest_TVar U) then dummyT else T)
  1126     end
  1127     handle TYPE _ => T_args
  1128 
  1129 fun filter_type_args_in_const _ _ _ _ _ [] = []
  1130   | filter_type_args_in_const thy monom_constrs type_enc ary s T_args =
  1131     case unprefix_and_unascii const_prefix s of
  1132       NONE =>
  1133       if level_of_type_enc type_enc = No_Types orelse s = tptp_choice then []
  1134       else T_args
  1135     | SOME s'' =>
  1136       let
  1137         val s'' = invert_const s''
  1138         fun filter_T_args false = T_args
  1139           | filter_T_args true = filter_const_type_args thy s'' ary T_args
  1140       in
  1141         case type_arg_policy monom_constrs type_enc s'' of
  1142           Explicit_Type_Args infer_from_term_args =>
  1143           filter_T_args infer_from_term_args
  1144         | No_Type_Args => []
  1145         | Mangled_Type_Args => raise Fail "unexpected (un)mangled symbol"
  1146       end
  1147 fun filter_type_args_in_iterm thy monom_constrs type_enc =
  1148   let
  1149     fun filt ary (IApp (tm1, tm2)) = IApp (filt (ary + 1) tm1, filt 0 tm2)
  1150       | filt ary (IConst (name as (s, _), T, T_args)) =
  1151         filter_type_args_in_const thy monom_constrs type_enc ary s T_args
  1152         |> (fn T_args => IConst (name, T, T_args))
  1153       | filt _ (IAbs (bound, tm)) = IAbs (bound, filt 0 tm)
  1154       | filt _ tm = tm
  1155   in filt 0 end
  1156 
  1157 fun iformula_from_prop ctxt format type_enc eq_as_iff =
  1158   let
  1159     val thy = Proof_Context.theory_of ctxt
  1160     fun do_term bs t atomic_Ts =
  1161       iterm_from_term thy format bs (Envir.eta_contract t)
  1162       |>> (introduce_proxies_in_iterm type_enc
  1163            #> mangle_type_args_in_iterm format type_enc #> AAtom)
  1164       ||> union (op =) atomic_Ts
  1165     fun do_quant bs q pos s T t' =
  1166       let
  1167         val s = singleton (Name.variant_list (map fst bs)) s
  1168         val universal = Option.map (q = AExists ? not) pos
  1169         val name =
  1170           s |> `(case universal of
  1171                    SOME true => make_all_bound_var
  1172                  | SOME false => make_exist_bound_var
  1173                  | NONE => make_bound_var)
  1174       in
  1175         do_formula ((s, (name, T)) :: bs) pos t'
  1176         #>> mk_aquant q [(name, SOME T)]
  1177         ##> union (op =) (atomic_types_of T)
  1178       end
  1179     and do_conn bs c pos1 t1 pos2 t2 =
  1180       do_formula bs pos1 t1 ##>> do_formula bs pos2 t2 #>> uncurry (mk_aconn c)
  1181     and do_formula bs pos t =
  1182       case t of
  1183         @{const Trueprop} $ t1 => do_formula bs pos t1
  1184       | @{const Not} $ t1 => do_formula bs (Option.map not pos) t1 #>> mk_anot
  1185       | Const (@{const_name All}, _) $ Abs (s, T, t') =>
  1186         do_quant bs AForall pos s T t'
  1187       | (t0 as Const (@{const_name All}, _)) $ t1 =>
  1188         do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)
  1189       | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
  1190         do_quant bs AExists pos s T t'
  1191       | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
  1192         do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)
  1193       | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd pos t1 pos t2
  1194       | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr pos t1 pos t2
  1195       | @{const HOL.implies} $ t1 $ t2 =>
  1196         do_conn bs AImplies (Option.map not pos) t1 pos t2
  1197       | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
  1198         if eq_as_iff then do_conn bs AIff NONE t1 NONE t2 else do_term bs t
  1199       | _ => do_term bs t
  1200   in do_formula [] end
  1201 
  1202 fun presimplify_term thy t =
  1203   if exists_Const (member (op =) Meson.presimplified_consts o fst) t then
  1204     t |> Skip_Proof.make_thm thy
  1205       |> Meson.presimplify
  1206       |> prop_of
  1207   else
  1208     t
  1209 
  1210 fun preprocess_abstractions_in_terms trans_lams facts =
  1211   let
  1212     val (facts, lambda_ts) =
  1213       facts |> map (snd o snd) |> trans_lams
  1214             |>> map2 (fn (name, (kind, _)) => fn t => (name, (kind, t))) facts
  1215     val lam_facts =
  1216       map2 (fn t => fn j =>
  1217                ((lam_fact_prefix ^ Int.toString j, (Global, Def)), (Axiom, t)))
  1218            lambda_ts (1 upto length lambda_ts)
  1219   in (facts, lam_facts) end
  1220 
  1221 (* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
  1222    same in Sledgehammer to prevent the discovery of unreplayable proofs. *)
  1223 fun freeze_term t =
  1224   let
  1225     fun freeze (t $ u) = freeze t $ freeze u
  1226       | freeze (Abs (s, T, t)) = Abs (s, T, freeze t)
  1227       | freeze (Var ((s, i), T)) =
  1228         Free (atp_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
  1229       | freeze t = t
  1230   in t |> exists_subterm is_Var t ? freeze end
  1231 
  1232 fun presimp_prop ctxt t =
  1233   let
  1234     val thy = Proof_Context.theory_of ctxt
  1235     val t = t |> Envir.beta_eta_contract
  1236               |> transform_elim_prop
  1237               |> Object_Logic.atomize_term thy
  1238     val need_trueprop = (fastype_of t = @{typ bool})
  1239   in
  1240     t |> need_trueprop ? HOLogic.mk_Trueprop
  1241       |> extensionalize_term ctxt
  1242       |> presimplify_term thy
  1243       |> HOLogic.dest_Trueprop
  1244   end
  1245   handle TERM _ => @{const True}
  1246 
  1247 fun make_formula ctxt format type_enc eq_as_iff name stature kind t =
  1248   let
  1249     val (iformula, atomic_Ts) =
  1250       iformula_from_prop ctxt format type_enc eq_as_iff
  1251                          (SOME (kind <> Conjecture)) t []
  1252       |>> close_iformula_universally
  1253   in
  1254     {name = name, stature = stature, kind = kind, iformula = iformula,
  1255      atomic_types = atomic_Ts}
  1256   end
  1257 
  1258 fun make_fact ctxt format type_enc eq_as_iff ((name, stature), t) =
  1259   case t |> make_formula ctxt format type_enc (eq_as_iff andalso format <> CNF)
  1260                          name stature Axiom of
  1261     formula as {iformula = AAtom (IConst ((s, _), _, _)), ...} =>
  1262     if s = tptp_true then NONE else SOME formula
  1263   | formula => SOME formula
  1264 
  1265 fun s_not_prop (@{const Trueprop} $ t) = @{const Trueprop} $ s_not t
  1266   | s_not_prop _ = @{prop True} (* "t" is too meta for "metis" *)
  1267 (*
  1268   | s_not_prop (@{const "==>"} $ t $ @{prop False}) = t
  1269   | s_not_prop t = @{const "==>"} $ t $ @{prop False}
  1270 *)
  1271 
  1272 fun make_conjecture ctxt format type_enc =
  1273   map (fn ((name, stature), (kind, t)) =>
  1274           t |> kind = Conjecture ? s_not
  1275             |> make_formula ctxt format type_enc (format <> CNF) name stature
  1276                             kind)
  1277 
  1278 (** Finite and infinite type inference **)
  1279 
  1280 fun tvar_footprint thy s ary =
  1281   (case unprefix_and_unascii const_prefix s of
  1282      SOME s =>
  1283      s |> invert_const |> robust_const_type thy |> chop_fun ary |> fst
  1284        |> map (fn T => Term.add_tvarsT T [] |> map fst)
  1285    | NONE => [])
  1286   handle TYPE _ => []
  1287 
  1288 fun ghost_type_args thy s ary =
  1289   if is_tptp_equal s then
  1290     0 upto ary - 1
  1291   else
  1292     let
  1293       val footprint = tvar_footprint thy s ary
  1294       val eq = (s = @{const_name HOL.eq})
  1295       fun ghosts _ [] = []
  1296         | ghosts seen ((i, tvars) :: args) =
  1297           ghosts (union (op =) seen tvars) args
  1298           |> (eq orelse exists (fn tvar => not (member (op =) seen tvar)) tvars)
  1299              ? cons i
  1300     in
  1301       if forall null footprint then
  1302         []
  1303       else
  1304         0 upto length footprint - 1 ~~ footprint
  1305         |> sort (rev_order o list_ord Term_Ord.indexname_ord o pairself snd)
  1306         |> ghosts []
  1307     end
  1308 
  1309 type monotonicity_info =
  1310   {maybe_finite_Ts : typ list,
  1311    surely_finite_Ts : typ list,
  1312    maybe_infinite_Ts : typ list,
  1313    surely_infinite_Ts : typ list,
  1314    maybe_nonmono_Ts : typ list}
  1315 
  1316 (* These types witness that the type classes they belong to allow infinite
  1317    models and hence that any types with these type classes is monotonic. *)
  1318 val known_infinite_types =
  1319   [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}]
  1320 
  1321 fun is_type_kind_of_surely_infinite ctxt strictness cached_Ts T =
  1322   strictness <> Strict andalso is_type_surely_infinite ctxt true cached_Ts T
  1323 
  1324 (* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are
  1325    dangerous because their "exhaust" properties can easily lead to unsound ATP
  1326    proofs. On the other hand, all HOL infinite types can be given the same
  1327    models in first-order logic (via Löwenheim-Skolem). *)
  1328 
  1329 fun should_encode_type _ (_ : monotonicity_info) All_Types _ = true
  1330   | should_encode_type ctxt {maybe_finite_Ts, surely_infinite_Ts,
  1331                              maybe_nonmono_Ts, ...}
  1332                        (Noninf_Nonmono_Types (strictness, grain)) T =
  1333     let val thy = Proof_Context.theory_of ctxt in
  1334       grain = Ghost_Type_Arg_Vars orelse
  1335       (exists (type_intersect thy T) maybe_nonmono_Ts andalso
  1336        not (exists (type_instance thy T) surely_infinite_Ts orelse
  1337             (not (member (type_equiv thy) maybe_finite_Ts T) andalso
  1338              is_type_kind_of_surely_infinite ctxt strictness surely_infinite_Ts
  1339                                              T)))
  1340     end
  1341   | should_encode_type ctxt {surely_finite_Ts, maybe_infinite_Ts,
  1342                              maybe_nonmono_Ts, ...}
  1343                        (Fin_Nonmono_Types grain) T =
  1344     let val thy = Proof_Context.theory_of ctxt in
  1345       grain = Ghost_Type_Arg_Vars orelse
  1346       (exists (type_intersect thy T) maybe_nonmono_Ts andalso
  1347        (exists (type_generalization thy T) surely_finite_Ts orelse
  1348         (not (member (type_equiv thy) maybe_infinite_Ts T) andalso
  1349          is_type_surely_finite ctxt T)))
  1350     end
  1351   | should_encode_type _ _ _ _ = false
  1352 
  1353 fun should_guard_type ctxt mono (Guards (_, level)) should_guard_var T =
  1354     should_guard_var () andalso should_encode_type ctxt mono level T
  1355   | should_guard_type _ _ _ _ _ = false
  1356 
  1357 fun is_maybe_universal_var (IConst ((s, _), _, _)) =
  1358     String.isPrefix bound_var_prefix s orelse
  1359     String.isPrefix all_bound_var_prefix s
  1360   | is_maybe_universal_var (IVar _) = true
  1361   | is_maybe_universal_var _ = false
  1362 
  1363 datatype site =
  1364   Top_Level of bool option |
  1365   Eq_Arg of bool option |
  1366   Elsewhere
  1367 
  1368 fun should_tag_with_type _ _ _ (Top_Level _) _ _ = false
  1369   | should_tag_with_type ctxt mono (Tags (_, level)) site u T =
  1370     if granularity_of_type_level level = All_Vars then
  1371       should_encode_type ctxt mono level T
  1372     else
  1373       (case (site, is_maybe_universal_var u) of
  1374          (Eq_Arg _, true) => should_encode_type ctxt mono level T
  1375        | _ => false)
  1376   | should_tag_with_type _ _ _ _ _ _ = false
  1377 
  1378 fun fused_type ctxt mono level =
  1379   let
  1380     val should_encode = should_encode_type ctxt mono level
  1381     fun fuse 0 T = if should_encode T then T else fused_infinite_type
  1382       | fuse ary (Type (@{type_name fun}, [T1, T2])) =
  1383         fuse 0 T1 --> fuse (ary - 1) T2
  1384       | fuse _ _ = raise Fail "expected function type"
  1385   in fuse end
  1386 
  1387 (** predicators and application operators **)
  1388 
  1389 type sym_info =
  1390   {pred_sym : bool, min_ary : int, max_ary : int, types : typ list,
  1391    in_conj : bool}
  1392 
  1393 fun default_sym_tab_entries type_enc =
  1394   (make_fixed_const NONE @{const_name undefined},
  1395        {pred_sym = false, min_ary = 0, max_ary = 0, types = [],
  1396         in_conj = false}) ::
  1397   ([tptp_false, tptp_true]
  1398    |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, types = [],
  1399                   in_conj = false})) @
  1400   ([tptp_equal, tptp_old_equal]
  1401    |> map (rpair {pred_sym = true, min_ary = 2, max_ary = 2, types = [],
  1402                   in_conj = false}))
  1403   |> not (is_type_enc_higher_order type_enc)
  1404      ? cons (prefixed_predicator_name,
  1405              {pred_sym = true, min_ary = 1, max_ary = 1, types = [],
  1406               in_conj = false})
  1407 
  1408 datatype app_op_level =
  1409   Min_App_Op |
  1410   Sufficient_App_Op |
  1411   Sufficient_App_Op_And_Predicator |
  1412   Full_App_Op_And_Predicator
  1413 
  1414 fun sym_table_for_facts ctxt type_enc app_op_level conjs facts =
  1415   let
  1416     val thy = Proof_Context.theory_of ctxt
  1417     fun consider_var_ary const_T var_T max_ary =
  1418       let
  1419         fun iter ary T =
  1420           if ary = max_ary orelse type_instance thy var_T T orelse
  1421              type_instance thy T var_T then
  1422             ary
  1423           else
  1424             iter (ary + 1) (range_type T)
  1425       in iter 0 const_T end
  1426     fun add_universal_var T (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
  1427       if (app_op_level = Sufficient_App_Op andalso can dest_funT T) orelse
  1428          (app_op_level = Sufficient_App_Op_And_Predicator andalso
  1429           (can dest_funT T orelse T = @{typ bool})) then
  1430         let
  1431           val bool_vars' =
  1432             bool_vars orelse
  1433             (app_op_level = Sufficient_App_Op_And_Predicator andalso
  1434              body_type T = @{typ bool})
  1435           fun repair_min_ary {pred_sym, min_ary, max_ary, types, in_conj} =
  1436             {pred_sym = pred_sym andalso not bool_vars',
  1437              min_ary = fold (fn T' => consider_var_ary T' T) types min_ary,
  1438              max_ary = max_ary, types = types, in_conj = in_conj}
  1439           val fun_var_Ts' =
  1440             fun_var_Ts |> can dest_funT T ? insert_type thy I T
  1441         in
  1442           if bool_vars' = bool_vars andalso
  1443              pointer_eq (fun_var_Ts', fun_var_Ts) then
  1444             accum
  1445           else
  1446             ((bool_vars', fun_var_Ts'), Symtab.map (K repair_min_ary) sym_tab)
  1447         end
  1448       else
  1449         accum
  1450       fun add_iterm_syms conj_fact top_level tm
  1451                          (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
  1452         let val (head, args) = strip_iterm_comb tm in
  1453           (case head of
  1454              IConst ((s, _), T, _) =>
  1455              if String.isPrefix bound_var_prefix s orelse
  1456                 String.isPrefix all_bound_var_prefix s then
  1457                add_universal_var T accum
  1458              else if String.isPrefix exist_bound_var_prefix s then
  1459                accum
  1460              else
  1461                let val ary = length args in
  1462                  ((bool_vars, fun_var_Ts),
  1463                   case Symtab.lookup sym_tab s of
  1464                     SOME {pred_sym, min_ary, max_ary, types, in_conj} =>
  1465                     let
  1466                       val pred_sym =
  1467                         pred_sym andalso top_level andalso not bool_vars
  1468                       val types' = types |> insert_type thy I T
  1469                       val in_conj = in_conj orelse conj_fact
  1470                       val min_ary =
  1471                         if (app_op_level = Sufficient_App_Op orelse
  1472                             app_op_level = Sufficient_App_Op_And_Predicator)
  1473                            andalso not (pointer_eq (types', types)) then
  1474                           fold (consider_var_ary T) fun_var_Ts min_ary
  1475                         else
  1476                           min_ary
  1477                     in
  1478                       Symtab.update (s, {pred_sym = pred_sym,
  1479                                          min_ary = Int.min (ary, min_ary),
  1480                                          max_ary = Int.max (ary, max_ary),
  1481                                          types = types', in_conj = in_conj})
  1482                                     sym_tab
  1483                     end
  1484                   | NONE =>
  1485                     let
  1486                       val pred_sym = top_level andalso not bool_vars
  1487                       val ary =
  1488                         case unprefix_and_unascii const_prefix s of
  1489                           SOME s =>
  1490                           (if String.isSubstring uncurried_alias_sep s then
  1491                              ary
  1492                            else case try (robust_const_ary thy
  1493                                           o invert_const o hd
  1494                                           o unmangled_const_name) s of
  1495                              SOME ary0 => Int.min (ary0, ary)
  1496                            | NONE => ary)
  1497                         | NONE => ary
  1498                       val min_ary =
  1499                         case app_op_level of
  1500                           Min_App_Op => ary
  1501                         | Full_App_Op_And_Predicator => 0
  1502                         | _ => fold (consider_var_ary T) fun_var_Ts ary
  1503                     in
  1504                       Symtab.update_new (s,
  1505                           {pred_sym = pred_sym, min_ary = min_ary,
  1506                            max_ary = ary, types = [T], in_conj = conj_fact})
  1507                           sym_tab
  1508                     end)
  1509                end
  1510            | IVar (_, T) => add_universal_var T accum
  1511            | IAbs ((_, T), tm) =>
  1512              accum |> add_universal_var T |> add_iterm_syms conj_fact false tm
  1513            | _ => accum)
  1514           |> fold (add_iterm_syms conj_fact false) args
  1515         end
  1516     fun add_fact_syms conj_fact =
  1517       K (add_iterm_syms conj_fact true) |> formula_fold NONE |> fact_lift
  1518   in
  1519     ((false, []), Symtab.empty)
  1520     |> fold (add_fact_syms true) conjs
  1521     |> fold (add_fact_syms false) facts
  1522     |> snd
  1523     |> fold Symtab.update (default_sym_tab_entries type_enc)
  1524   end
  1525 
  1526 fun min_ary_of sym_tab s =
  1527   case Symtab.lookup sym_tab s of
  1528     SOME ({min_ary, ...} : sym_info) => min_ary
  1529   | NONE =>
  1530     case unprefix_and_unascii const_prefix s of
  1531       SOME s =>
  1532       let val s = s |> unmangled_const_name |> hd |> invert_const in
  1533         if s = predicator_name then 1
  1534         else if s = app_op_name then 2
  1535         else if s = type_guard_name then 1
  1536         else 0
  1537       end
  1538     | NONE => 0
  1539 
  1540 (* True if the constant ever appears outside of the top-level position in
  1541    literals, or if it appears with different arities (e.g., because of different
  1542    type instantiations). If false, the constant always receives all of its
  1543    arguments and is used as a predicate. *)
  1544 fun is_pred_sym sym_tab s =
  1545   case Symtab.lookup sym_tab s of
  1546     SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
  1547     pred_sym andalso min_ary = max_ary
  1548   | NONE => false
  1549 
  1550 val app_op = `(make_fixed_const NONE) app_op_name
  1551 val predicator_combconst =
  1552   IConst (`(make_fixed_const NONE) predicator_name, @{typ "bool => bool"}, [])
  1553 
  1554 fun list_app head args = fold (curry (IApp o swap)) args head
  1555 fun predicator tm = IApp (predicator_combconst, tm)
  1556 
  1557 fun mk_app_op format type_enc head arg =
  1558   let
  1559     val head_T = ityp_of head
  1560     val (arg_T, res_T) = dest_funT head_T
  1561     val app =
  1562       IConst (app_op, head_T --> head_T, [arg_T, res_T])
  1563       |> mangle_type_args_in_iterm format type_enc
  1564   in list_app app [head, arg] end
  1565 
  1566 fun firstorderize_fact thy monom_constrs format type_enc sym_tab
  1567                        uncurried_aliases =
  1568   let
  1569     fun do_app arg head = mk_app_op format type_enc head arg
  1570     fun list_app_ops head args = fold do_app args head
  1571     fun introduce_app_ops tm =
  1572       let val (head, args) = tm |> strip_iterm_comb ||> map introduce_app_ops in
  1573         case head of
  1574           IConst (name as (s, _), T, T_args) =>
  1575           if uncurried_aliases andalso String.isPrefix const_prefix s then
  1576             let
  1577               val ary = length args
  1578               val name =
  1579                 name |> ary > min_ary_of sym_tab s ? aliased_uncurried ary
  1580             in list_app (IConst (name, T, T_args)) args end
  1581           else
  1582             args |> chop (min_ary_of sym_tab s)
  1583                  |>> list_app head |-> list_app_ops
  1584         | _ => list_app_ops head args
  1585       end
  1586     fun introduce_predicators tm =
  1587       case strip_iterm_comb tm of
  1588         (IConst ((s, _), _, _), _) =>
  1589         if is_pred_sym sym_tab s then tm else predicator tm
  1590       | _ => predicator tm
  1591     val do_iterm =
  1592       not (is_type_enc_higher_order type_enc)
  1593       ? (introduce_app_ops #> introduce_predicators)
  1594       #> filter_type_args_in_iterm thy monom_constrs type_enc
  1595   in update_iformula (formula_map do_iterm) end
  1596 
  1597 (** Helper facts **)
  1598 
  1599 val not_ffalse = @{lemma "~ fFalse" by (unfold fFalse_def) fast}
  1600 val ftrue = @{lemma "fTrue" by (unfold fTrue_def) fast}
  1601 
  1602 (* The Boolean indicates that a fairly sound type encoding is needed. *)
  1603 val helper_table =
  1604   [(("COMBI", false), @{thms Meson.COMBI_def}),
  1605    (("COMBK", false), @{thms Meson.COMBK_def}),
  1606    (("COMBB", false), @{thms Meson.COMBB_def}),
  1607    (("COMBC", false), @{thms Meson.COMBC_def}),
  1608    (("COMBS", false), @{thms Meson.COMBS_def}),
  1609    ((predicator_name, false), [not_ffalse, ftrue]),
  1610    (("fFalse", false), [not_ffalse]),
  1611    (("fFalse", true), @{thms True_or_False}),
  1612    (("fTrue", false), [ftrue]),
  1613    (("fTrue", true), @{thms True_or_False}),
  1614    (("fNot", false),
  1615     @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1]
  1616            fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),
  1617    (("fconj", false),
  1618     @{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q"
  1619         by (unfold fconj_def) fast+}),
  1620    (("fdisj", false),
  1621     @{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q"
  1622         by (unfold fdisj_def) fast+}),
  1623    (("fimplies", false),
  1624     @{lemma "P | fimplies P Q" "~ Q | fimplies P Q" "~ fimplies P Q | ~ P | Q"
  1625         by (unfold fimplies_def) fast+}),
  1626    (("fequal", true),
  1627     (* This is a lie: Higher-order equality doesn't need a sound type encoding.
  1628        However, this is done so for backward compatibility: Including the
  1629        equality helpers by default in Metis breaks a few existing proofs. *)
  1630     @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1]
  1631            fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),
  1632    (* Partial characterization of "fAll" and "fEx". A complete characterization
  1633       would require the axiom of choice for replay with Metis. *)
  1634    (("fAll", false), [@{lemma "~ fAll P | P x" by (auto simp: fAll_def)}]),
  1635    (("fEx", false), [@{lemma "~ P x | fEx P" by (auto simp: fEx_def)}]),
  1636    (("If", true), @{thms if_True if_False True_or_False})]
  1637   |> map (apsnd (map zero_var_indexes))
  1638 
  1639 fun atype_of_type_vars (Simple_Types (_, Polymorphic, _)) = SOME atype_of_types
  1640   | atype_of_type_vars _ = NONE
  1641 
  1642 fun bound_tvars type_enc sorts Ts =
  1643   (sorts ? mk_ahorn (formulas_for_types type_enc add_sorts_on_tvar Ts))
  1644   #> mk_aquant AForall
  1645         (map_filter (fn TVar (x as (s, _), _) =>
  1646                         SOME ((make_schematic_type_var x, s),
  1647                               atype_of_type_vars type_enc)
  1648                       | _ => NONE) Ts)
  1649 
  1650 fun eq_formula type_enc atomic_Ts bounds pred_sym tm1 tm2 =
  1651   (if pred_sym then AConn (AIff, [AAtom tm1, AAtom tm2])
  1652    else AAtom (ATerm (`I tptp_equal, [tm1, tm2])))
  1653   |> mk_aquant AForall bounds
  1654   |> close_formula_universally
  1655   |> bound_tvars type_enc true atomic_Ts
  1656 
  1657 val helper_rank = default_rank
  1658 val min_rank = 9 * helper_rank div 10
  1659 val max_rank = 4 * min_rank
  1660 
  1661 fun rank_of_fact_num n j = min_rank + (max_rank - min_rank) * j div n
  1662 
  1663 val type_tag = `(make_fixed_const NONE) type_tag_name
  1664 
  1665 fun should_specialize_helper type_enc t =
  1666   polymorphism_of_type_enc type_enc <> Polymorphic andalso
  1667   level_of_type_enc type_enc <> No_Types andalso
  1668   not (null (Term.hidden_polymorphism t))
  1669 
  1670 fun helper_facts_for_sym ctxt format type_enc (s, {types, ...} : sym_info) =
  1671   case unprefix_and_unascii const_prefix s of
  1672     SOME mangled_s =>
  1673     let
  1674       val thy = Proof_Context.theory_of ctxt
  1675       val unmangled_s = mangled_s |> unmangled_const_name |> hd
  1676       fun dub needs_fairly_sound j k =
  1677         unmangled_s ^ "_" ^ string_of_int j ^ "_" ^ string_of_int k ^
  1678         (if mangled_s = unmangled_s then "" else "_" ^ ascii_of mangled_s) ^
  1679         (if needs_fairly_sound then typed_helper_suffix
  1680          else untyped_helper_suffix)
  1681       fun dub_and_inst needs_fairly_sound (th, j) =
  1682         let val t = prop_of th in
  1683           if should_specialize_helper type_enc t then
  1684             map (fn T => specialize_type thy (invert_const unmangled_s, T) t)
  1685                 types
  1686           else
  1687             [t]
  1688         end
  1689         |> tag_list 1
  1690         |> map (fn (k, t) => ((dub needs_fairly_sound j k, (Global, Def)), t))
  1691       val make_facts = map_filter (make_fact ctxt format type_enc false)
  1692       val fairly_sound = is_type_enc_fairly_sound type_enc
  1693     in
  1694       helper_table
  1695       |> maps (fn ((helper_s, needs_fairly_sound), ths) =>
  1696                   if helper_s <> unmangled_s orelse
  1697                      (needs_fairly_sound andalso not fairly_sound) then
  1698                     []
  1699                   else
  1700                     ths ~~ (1 upto length ths)
  1701                     |> maps (dub_and_inst needs_fairly_sound)
  1702                     |> make_facts)
  1703     end
  1704   | NONE => []
  1705 fun helper_facts_for_sym_table ctxt format type_enc sym_tab =
  1706   Symtab.fold_rev (append o helper_facts_for_sym ctxt format type_enc) sym_tab
  1707                   []
  1708 
  1709 (***************************************************************)
  1710 (* Type Classes Present in the Axiom or Conjecture Clauses     *)
  1711 (***************************************************************)
  1712 
  1713 fun set_insert (x, s) = Symtab.update (x, ()) s
  1714 
  1715 fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
  1716 
  1717 (* Remove this trivial type class (FIXME: similar code elsewhere) *)
  1718 fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset
  1719 
  1720 fun classes_of_terms get_Ts =
  1721   map (map snd o get_Ts)
  1722   #> List.foldl add_classes Symtab.empty
  1723   #> delete_type #> Symtab.keys
  1724 
  1725 val tfree_classes_of_terms = classes_of_terms Misc_Legacy.term_tfrees
  1726 val tvar_classes_of_terms = classes_of_terms Misc_Legacy.term_tvars
  1727 
  1728 fun fold_type_constrs f (Type (s, Ts)) x =
  1729     fold (fold_type_constrs f) Ts (f (s, x))
  1730   | fold_type_constrs _ _ x = x
  1731 
  1732 (* Type constructors used to instantiate overloaded constants are the only ones
  1733    needed. *)
  1734 fun add_type_constrs_in_term thy =
  1735   let
  1736     fun add (Const (@{const_name Meson.skolem}, _) $ _) = I
  1737       | add (t $ u) = add t #> add u
  1738       | add (Const x) =
  1739         x |> robust_const_typargs thy |> fold (fold_type_constrs set_insert)
  1740       | add (Abs (_, _, u)) = add u
  1741       | add _ = I
  1742   in add end
  1743 
  1744 fun type_constrs_of_terms thy ts =
  1745   Symtab.keys (fold (add_type_constrs_in_term thy) ts Symtab.empty)
  1746 
  1747 fun extract_lambda_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
  1748     let val (head, args) = strip_comb t in
  1749       (head |> dest_Const |> fst,
  1750        fold_rev (fn t as Var ((s, _), T) =>
  1751                     (fn u => Abs (s, T, abstract_over (t, u)))
  1752                   | _ => raise Fail "expected Var") args u)
  1753     end
  1754   | extract_lambda_def _ = raise Fail "malformed lifted lambda"
  1755 
  1756 fun trans_lams_from_string ctxt type_enc lam_trans =
  1757   if lam_trans = no_lamsN then
  1758     rpair []
  1759   else if lam_trans = hide_lamsN then
  1760     lift_lams ctxt type_enc ##> K []
  1761   else if lam_trans = liftingN orelse lam_trans = lam_liftingN then
  1762     lift_lams ctxt type_enc
  1763   else if lam_trans = combsN then
  1764     map (introduce_combinators ctxt) #> rpair []
  1765   else if lam_trans = combs_and_liftingN then
  1766     lift_lams_part_1 ctxt type_enc
  1767     ##> maps (fn t => [t, introduce_combinators ctxt (intentionalize_def t)])
  1768     #> lift_lams_part_2 ctxt
  1769   else if lam_trans = combs_or_liftingN then
  1770     lift_lams_part_1 ctxt type_enc
  1771     ##> map (fn t => case head_of (strip_qnt_body @{const_name All} t) of
  1772                        @{term "op =::bool => bool => bool"} => t
  1773                      | _ => introduce_combinators ctxt (intentionalize_def t))
  1774     #> lift_lams_part_2 ctxt
  1775   else if lam_trans = keep_lamsN then
  1776     map (Envir.eta_contract) #> rpair []
  1777   else
  1778     error ("Unknown lambda translation scheme: " ^ quote lam_trans ^ ".")
  1779 
  1780 fun translate_formulas ctxt format prem_kind type_enc lam_trans presimp hyp_ts
  1781                        concl_t facts =
  1782   let
  1783     val thy = Proof_Context.theory_of ctxt
  1784     val trans_lams = trans_lams_from_string ctxt type_enc lam_trans
  1785     val fact_ts = facts |> map snd
  1786     (* Remove existing facts from the conjecture, as this can dramatically
  1787        boost an ATP's performance (for some reason). *)
  1788     val hyp_ts =
  1789       hyp_ts
  1790       |> map (fn t => if member (op aconv) fact_ts t then @{prop True} else t)
  1791     val facts = facts |> map (apsnd (pair Axiom))
  1792     val conjs =
  1793       map (pair prem_kind) hyp_ts @ [(Conjecture, s_not_prop concl_t)]
  1794       |> map (apsnd freeze_term)
  1795       |> map2 (pair o rpair (Local, General) o string_of_int)
  1796               (0 upto length hyp_ts)
  1797     val ((conjs, facts), lam_facts) =
  1798       (conjs, facts)
  1799       |> presimp ? pairself (map (apsnd (apsnd (presimp_prop ctxt))))
  1800       |> (if lam_trans = no_lamsN then
  1801             rpair []
  1802           else
  1803             op @
  1804             #> preprocess_abstractions_in_terms trans_lams
  1805             #>> chop (length conjs))
  1806     val conjs = conjs |> make_conjecture ctxt format type_enc
  1807     val (fact_names, facts) =
  1808       facts
  1809       |> map_filter (fn (name, (_, t)) =>
  1810                         make_fact ctxt format type_enc true (name, t)
  1811                         |> Option.map (pair name))
  1812       |> ListPair.unzip
  1813     val lifted = lam_facts |> map (extract_lambda_def o snd o snd)
  1814     val lam_facts =
  1815       lam_facts |> map_filter (make_fact ctxt format type_enc true o apsnd snd)
  1816     val all_ts = concl_t :: hyp_ts @ fact_ts
  1817     val subs = tfree_classes_of_terms all_ts
  1818     val supers = tvar_classes_of_terms all_ts
  1819     val tycons = type_constrs_of_terms thy all_ts
  1820     val (supers, arity_clauses) =
  1821       if level_of_type_enc type_enc = No_Types then ([], [])
  1822       else make_arity_clauses thy tycons supers
  1823     val class_rel_clauses = make_class_rel_clauses thy subs supers
  1824   in
  1825     (fact_names |> map single, union (op =) subs supers, conjs,
  1826      facts @ lam_facts, class_rel_clauses, arity_clauses, lifted)
  1827   end
  1828 
  1829 val type_guard = `(make_fixed_const NONE) type_guard_name
  1830 
  1831 fun type_guard_iterm format type_enc T tm =
  1832   IApp (IConst (type_guard, T --> @{typ bool}, [T])
  1833         |> mangle_type_args_in_iterm format type_enc, tm)
  1834 
  1835 fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum
  1836   | is_var_positively_naked_in_term name _ (ATerm ((s, _), tms)) accum =
  1837     accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm (name, [])))
  1838   | is_var_positively_naked_in_term _ _ _ _ = true
  1839 
  1840 fun is_var_ghost_type_arg_in_term thy polym_constrs name pos tm accum =
  1841   is_var_positively_naked_in_term name pos tm accum orelse
  1842   let
  1843     val var = ATerm (name, [])
  1844     fun is_nasty_in_term (ATerm (_, [])) = false
  1845       | is_nasty_in_term (ATerm ((s, _), tms)) =
  1846         let
  1847           val ary = length tms
  1848           val polym_constr = member (op =) polym_constrs s
  1849           val ghosts = ghost_type_args thy s ary
  1850         in
  1851           exists (fn (j, tm) =>
  1852                      if polym_constr then
  1853                        member (op =) ghosts j andalso
  1854                        (tm = var orelse is_nasty_in_term tm)
  1855                      else
  1856                        tm = var andalso member (op =) ghosts j)
  1857                  (0 upto ary - 1 ~~ tms)
  1858           orelse (not polym_constr andalso exists is_nasty_in_term tms)
  1859         end
  1860       | is_nasty_in_term _ = true
  1861   in is_nasty_in_term tm end
  1862 
  1863 fun should_guard_var_in_formula thy polym_constrs level pos phi (SOME true)
  1864                                 name =
  1865     (case granularity_of_type_level level of
  1866        All_Vars => true
  1867      | Positively_Naked_Vars =>
  1868        formula_fold pos (is_var_positively_naked_in_term name) phi false
  1869      | Ghost_Type_Arg_Vars =>
  1870        formula_fold pos (is_var_ghost_type_arg_in_term thy polym_constrs name) phi
  1871                     false)
  1872   | should_guard_var_in_formula _ _ _ _ _ _ _ = true
  1873 
  1874 fun always_guard_var_in_formula _ _ _ _ _ _ _ = true
  1875 
  1876 fun should_generate_tag_bound_decl _ _ _ (SOME true) _ = false
  1877   | should_generate_tag_bound_decl ctxt mono (Tags (_, level)) _ T =
  1878     granularity_of_type_level level <> All_Vars andalso
  1879     should_encode_type ctxt mono level T
  1880   | should_generate_tag_bound_decl _ _ _ _ _ = false
  1881 
  1882 fun mk_aterm format type_enc name T_args args =
  1883   ATerm (name, map_filter (ho_term_for_type_arg format type_enc) T_args @ args)
  1884 
  1885 fun do_bound_type ctxt format mono type_enc =
  1886   case type_enc of
  1887     Simple_Types (_, _, level) =>
  1888     fused_type ctxt mono level 0
  1889     #> ho_type_from_typ format type_enc false 0 #> SOME
  1890   | _ => K NONE
  1891 
  1892 fun tag_with_type ctxt format mono type_enc pos T tm =
  1893   IConst (type_tag, T --> T, [T])
  1894   |> mangle_type_args_in_iterm format type_enc
  1895   |> ho_term_from_iterm ctxt format mono type_enc pos
  1896   |> (fn ATerm (s, tms) => ATerm (s, tms @ [tm])
  1897        | _ => raise Fail "unexpected lambda-abstraction")
  1898 and ho_term_from_iterm ctxt format mono type_enc pos =
  1899   let
  1900     fun beta_red bs (IApp (IAbs ((name, _), tm), tm')) =
  1901         beta_red ((name, beta_red bs tm') :: bs) tm
  1902       | beta_red bs (IApp tmp) = IApp (pairself (beta_red bs) tmp)
  1903       | beta_red bs (tm as IConst (name, _, _)) =
  1904         (case AList.lookup (op =) bs name of
  1905            SOME tm' => tm'
  1906          | NONE => tm)
  1907       | beta_red bs (IAbs ((name, T), tm)) =
  1908         IAbs ((name, T), beta_red (AList.delete (op =) name bs) tm)
  1909       | beta_red _ tm = tm
  1910     fun term site u =
  1911       let
  1912         val (head, args) = strip_iterm_comb u
  1913         val pos =
  1914           case site of
  1915             Top_Level pos => pos
  1916           | Eq_Arg pos => pos
  1917           | _ => NONE
  1918         val t =
  1919           case head of
  1920             IConst (name as (s, _), _, T_args) =>
  1921             let
  1922               val arg_site = if is_tptp_equal s then Eq_Arg pos else Elsewhere
  1923             in
  1924               map (term arg_site) args |> mk_aterm format type_enc name T_args
  1925             end
  1926           | IVar (name, _) =>
  1927             map (term Elsewhere) args |> mk_aterm format type_enc name []
  1928           | IAbs ((name, T), tm) =>
  1929             if is_type_enc_higher_order type_enc then
  1930               AAbs ((name, ho_type_from_typ format type_enc true 0 T),
  1931                     term Elsewhere tm)
  1932             else
  1933               raise Fail "unexpected lambda-abstraction"
  1934           | IApp _ => raise Fail "impossible \"IApp\""
  1935         val T = ityp_of u
  1936       in
  1937         if should_tag_with_type ctxt mono type_enc site u T then
  1938           tag_with_type ctxt format mono type_enc pos T t
  1939         else
  1940           t
  1941       end
  1942   in term (Top_Level pos) o beta_red [] end
  1943 and formula_from_iformula ctxt polym_constrs format mono type_enc
  1944                           should_guard_var =
  1945   let
  1946     val thy = Proof_Context.theory_of ctxt
  1947     val level = level_of_type_enc type_enc
  1948     val do_term = ho_term_from_iterm ctxt format mono type_enc
  1949     fun do_out_of_bound_type pos phi universal (name, T) =
  1950       if should_guard_type ctxt mono type_enc
  1951              (fn () => should_guard_var thy polym_constrs level pos phi
  1952                                         universal name) T then
  1953         IVar (name, T)
  1954         |> type_guard_iterm format type_enc T
  1955         |> do_term pos |> AAtom |> SOME
  1956       else if should_generate_tag_bound_decl ctxt mono type_enc universal T then
  1957         let
  1958           val var = ATerm (name, [])
  1959           val tagged_var = tag_with_type ctxt format mono type_enc pos T var
  1960         in SOME (AAtom (ATerm (`I tptp_equal, [tagged_var, var]))) end
  1961       else
  1962         NONE
  1963     fun do_formula pos (AQuant (q, xs, phi)) =
  1964         let
  1965           val phi = phi |> do_formula pos
  1966           val universal = Option.map (q = AExists ? not) pos
  1967           val do_bound_type = do_bound_type ctxt format mono type_enc
  1968         in
  1969           AQuant (q, xs |> map (apsnd (fn NONE => NONE
  1970                                         | SOME T => do_bound_type T)),
  1971                   (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))
  1972                       (map_filter
  1973                            (fn (_, NONE) => NONE
  1974                              | (s, SOME T) =>
  1975                                do_out_of_bound_type pos phi universal (s, T))
  1976                            xs)
  1977                       phi)
  1978         end
  1979       | do_formula pos (AConn conn) = aconn_map pos do_formula conn
  1980       | do_formula pos (AAtom tm) = AAtom (do_term pos tm)
  1981   in do_formula end
  1982 
  1983 (* Each fact is given a unique fact number to avoid name clashes (e.g., because
  1984    of monomorphization). The TPTP explicitly forbids name clashes, and some of
  1985    the remote provers might care. *)
  1986 fun formula_line_for_fact ctxt polym_constrs format prefix encode freshen pos
  1987         mono type_enc rank (j, {name, stature, kind, iformula, atomic_types}) =
  1988   (prefix ^ (if freshen then string_of_int j ^ "_" else "") ^ encode name, kind,
  1989    iformula
  1990    |> formula_from_iformula ctxt polym_constrs format mono type_enc
  1991           should_guard_var_in_formula (if pos then SOME true else NONE)
  1992    |> close_formula_universally
  1993    |> bound_tvars type_enc true atomic_types,
  1994    NONE,
  1995    let val rank = rank j in
  1996      case snd stature of
  1997        Intro => isabelle_info introN rank
  1998      | Inductive => isabelle_info inductiveN rank
  1999      | Elim => isabelle_info elimN rank
  2000      | Simp => isabelle_info simpN rank
  2001      | Def => isabelle_info defN rank
  2002      | _ => isabelle_info "" rank
  2003    end)
  2004   |> Formula
  2005 
  2006 fun formula_line_for_class_rel_clause type_enc
  2007         ({name, subclass, superclass, ...} : class_rel_clause) =
  2008   let val ty_arg = ATerm (tvar_a_name, []) in
  2009     Formula (class_rel_clause_prefix ^ ascii_of name, Axiom,
  2010              AConn (AImplies,
  2011                     [type_class_formula type_enc subclass ty_arg,
  2012                      type_class_formula type_enc superclass ty_arg])
  2013              |> mk_aquant AForall
  2014                           [(tvar_a_name, atype_of_type_vars type_enc)],
  2015              NONE, isabelle_info inductiveN helper_rank)
  2016   end
  2017 
  2018 fun formula_from_arity_atom type_enc (class, t, args) =
  2019   ATerm (t, map (fn arg => ATerm (arg, [])) args)
  2020   |> type_class_formula type_enc class
  2021 
  2022 fun formula_line_for_arity_clause type_enc
  2023         ({name, prem_atoms, concl_atom} : arity_clause) =
  2024   Formula (arity_clause_prefix ^ name, Axiom,
  2025            mk_ahorn (map (formula_from_arity_atom type_enc) prem_atoms)
  2026                     (formula_from_arity_atom type_enc concl_atom)
  2027            |> mk_aquant AForall
  2028                   (map (rpair (atype_of_type_vars type_enc)) (#3 concl_atom)),
  2029            NONE, isabelle_info inductiveN helper_rank)
  2030 
  2031 fun formula_line_for_conjecture ctxt polym_constrs format mono type_enc
  2032         ({name, kind, iformula, atomic_types, ...} : translated_formula) =
  2033   Formula (conjecture_prefix ^ name, kind,
  2034            iformula
  2035            |> formula_from_iformula ctxt polym_constrs format mono type_enc
  2036                   should_guard_var_in_formula (SOME false)
  2037            |> close_formula_universally
  2038            |> bound_tvars type_enc true atomic_types, NONE, [])
  2039 
  2040 fun type_enc_needs_free_types (Simple_Types (_, Polymorphic, _)) = true
  2041   | type_enc_needs_free_types (Simple_Types _) = false
  2042   | type_enc_needs_free_types _ = true
  2043 
  2044 fun formula_line_for_free_type j phi =
  2045   Formula (tfree_clause_prefix ^ string_of_int j, Hypothesis, phi, NONE, [])
  2046 fun formula_lines_for_free_types type_enc (facts : translated_formula list) =
  2047   if type_enc_needs_free_types type_enc then
  2048     let
  2049       val phis =
  2050         fold (union (op =)) (map #atomic_types facts) []
  2051         |> formulas_for_types type_enc add_sorts_on_tfree
  2052     in map2 formula_line_for_free_type (0 upto length phis - 1) phis end
  2053   else
  2054     []
  2055 
  2056 (** Symbol declarations **)
  2057 
  2058 fun decl_line_for_class order s =
  2059   let val name as (s, _) = `make_type_class s in
  2060     Decl (sym_decl_prefix ^ s, name,
  2061           if order = First_Order then
  2062             ATyAbs ([tvar_a_name],
  2063                     if avoid_first_order_ghost_type_vars then
  2064                       AFun (a_itself_atype, bool_atype)
  2065                     else
  2066                       bool_atype)
  2067           else
  2068             AFun (atype_of_types, bool_atype))
  2069   end
  2070 
  2071 fun decl_lines_for_classes type_enc classes =
  2072   case type_enc of
  2073     Simple_Types (order, Polymorphic, _) =>
  2074     map (decl_line_for_class order) classes
  2075   | _ => []
  2076 
  2077 fun sym_decl_table_for_facts thy format type_enc sym_tab
  2078                              (conjs, facts, extra_tms) =
  2079   let
  2080     fun add_iterm_syms tm =
  2081       let val (head, args) = strip_iterm_comb tm in
  2082         (case head of
  2083            IConst ((s, s'), T, T_args) =>
  2084            let
  2085              val (pred_sym, in_conj) =
  2086                case Symtab.lookup sym_tab s of
  2087                  SOME ({pred_sym, in_conj, ...} : sym_info) =>
  2088                  (pred_sym, in_conj)
  2089                | NONE => (false, false)
  2090              val decl_sym =
  2091                (case type_enc of
  2092                   Guards _ => not pred_sym
  2093                 | _ => true) andalso
  2094                is_tptp_user_symbol s
  2095            in
  2096              if decl_sym then
  2097                Symtab.map_default (s, [])
  2098                    (insert_type thy #3 (s', T_args, T, pred_sym, length args,
  2099                                         in_conj))
  2100              else
  2101                I
  2102            end
  2103          | IAbs (_, tm) => add_iterm_syms tm
  2104          | _ => I)
  2105         #> fold add_iterm_syms args
  2106       end
  2107     val add_fact_syms = K add_iterm_syms |> formula_fold NONE |> fact_lift
  2108     fun add_formula_var_types (AQuant (_, xs, phi)) =
  2109         fold (fn (_, SOME T) => insert_type thy I T | _ => I) xs
  2110         #> add_formula_var_types phi
  2111       | add_formula_var_types (AConn (_, phis)) =
  2112         fold add_formula_var_types phis
  2113       | add_formula_var_types _ = I
  2114     fun var_types () =
  2115       if polymorphism_of_type_enc type_enc = Polymorphic then [tvar_a]
  2116       else fold (fact_lift add_formula_var_types) (conjs @ facts) []
  2117     fun add_undefined_const T =
  2118       let
  2119         val (s, s') =
  2120           `(make_fixed_const NONE) @{const_name undefined}
  2121           |> (case type_arg_policy [] type_enc @{const_name undefined} of
  2122                 Mangled_Type_Args => mangled_const_name format type_enc [T]
  2123               | _ => I)
  2124       in
  2125         Symtab.map_default (s, [])
  2126                            (insert_type thy #3 (s', [T], T, false, 0, false))
  2127       end
  2128     fun add_TYPE_const () =
  2129       let val (s, s') = TYPE_name in
  2130         Symtab.map_default (s, [])
  2131             (insert_type thy #3
  2132                          (s', [tvar_a], @{typ "'a itself"}, false, 0, false))
  2133       end
  2134   in
  2135     Symtab.empty
  2136     |> is_type_enc_fairly_sound type_enc
  2137        ? (fold (fold add_fact_syms) [conjs, facts]
  2138           #> fold add_iterm_syms extra_tms
  2139           #> (case type_enc of
  2140                 Simple_Types (First_Order, Polymorphic, _) =>
  2141                 if avoid_first_order_ghost_type_vars then add_TYPE_const ()
  2142                 else I
  2143               | Simple_Types _ => I
  2144               | _ => fold add_undefined_const (var_types ())))
  2145   end
  2146 
  2147 (* We add "bool" in case the helper "True_or_False" is included later. *)
  2148 fun default_mono level =
  2149   {maybe_finite_Ts = [@{typ bool}],
  2150    surely_finite_Ts = [@{typ bool}],
  2151    maybe_infinite_Ts = known_infinite_types,
  2152    surely_infinite_Ts =
  2153      case level of
  2154        Noninf_Nonmono_Types (Strict, _) => []
  2155      | _ => known_infinite_types,
  2156    maybe_nonmono_Ts = [@{typ bool}]}
  2157 
  2158 (* This inference is described in section 2.3 of Claessen et al.'s "Sorting it
  2159    out with monotonicity" paper presented at CADE 2011. *)
  2160 fun add_iterm_mononotonicity_info _ _ (SOME false) _ mono = mono
  2161   | add_iterm_mononotonicity_info ctxt level _
  2162         (IApp (IApp (IConst ((s, _), Type (_, [T, _]), _), tm1), tm2))
  2163         (mono as {maybe_finite_Ts, surely_finite_Ts, maybe_infinite_Ts,
  2164                   surely_infinite_Ts, maybe_nonmono_Ts}) =
  2165     let val thy = Proof_Context.theory_of ctxt in
  2166       if is_tptp_equal s andalso exists is_maybe_universal_var [tm1, tm2] then
  2167         case level of
  2168           Noninf_Nonmono_Types (strictness, _) =>
  2169           if exists (type_instance thy T) surely_infinite_Ts orelse
  2170              member (type_equiv thy) maybe_finite_Ts T then
  2171             mono
  2172           else if is_type_kind_of_surely_infinite ctxt strictness
  2173                                                   surely_infinite_Ts T then
  2174             {maybe_finite_Ts = maybe_finite_Ts,
  2175              surely_finite_Ts = surely_finite_Ts,
  2176              maybe_infinite_Ts = maybe_infinite_Ts,
  2177              surely_infinite_Ts = surely_infinite_Ts |> insert_type thy I T,
  2178              maybe_nonmono_Ts = maybe_nonmono_Ts}
  2179           else
  2180             {maybe_finite_Ts = maybe_finite_Ts |> insert (type_equiv thy) T,
  2181              surely_finite_Ts = surely_finite_Ts,
  2182              maybe_infinite_Ts = maybe_infinite_Ts,
  2183              surely_infinite_Ts = surely_infinite_Ts,
  2184              maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type thy I T}
  2185         | Fin_Nonmono_Types _ =>
  2186           if exists (type_instance thy T) surely_finite_Ts orelse
  2187              member (type_equiv thy) maybe_infinite_Ts T then
  2188             mono
  2189           else if is_type_surely_finite ctxt T then
  2190             {maybe_finite_Ts = maybe_finite_Ts,
  2191              surely_finite_Ts = surely_finite_Ts |> insert_type thy I T,
  2192              maybe_infinite_Ts = maybe_infinite_Ts,
  2193              surely_infinite_Ts = surely_infinite_Ts,
  2194              maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type thy I T}
  2195           else
  2196             {maybe_finite_Ts = maybe_finite_Ts,
  2197              surely_finite_Ts = surely_finite_Ts,
  2198              maybe_infinite_Ts = maybe_infinite_Ts |> insert (type_equiv thy) T,
  2199              surely_infinite_Ts = surely_infinite_Ts,
  2200              maybe_nonmono_Ts = maybe_nonmono_Ts}
  2201         | _ => mono
  2202       else
  2203         mono
  2204     end
  2205   | add_iterm_mononotonicity_info _ _ _ _ mono = mono
  2206 fun add_fact_mononotonicity_info ctxt level
  2207         ({kind, iformula, ...} : translated_formula) =
  2208   formula_fold (SOME (kind <> Conjecture))
  2209                (add_iterm_mononotonicity_info ctxt level) iformula
  2210 fun mononotonicity_info_for_facts ctxt type_enc facts =
  2211   let val level = level_of_type_enc type_enc in
  2212     default_mono level
  2213     |> is_type_level_monotonicity_based level
  2214        ? fold (add_fact_mononotonicity_info ctxt level) facts
  2215   end
  2216 
  2217 fun add_iformula_monotonic_types ctxt mono type_enc =
  2218   let
  2219     val thy = Proof_Context.theory_of ctxt
  2220     val level = level_of_type_enc type_enc
  2221     val should_encode = should_encode_type ctxt mono level
  2222     fun add_type T = not (should_encode T) ? insert_type thy I T
  2223     fun add_args (IApp (tm1, tm2)) = add_args tm1 #> add_term tm2
  2224       | add_args _ = I
  2225     and add_term tm = add_type (ityp_of tm) #> add_args tm
  2226   in formula_fold NONE (K add_term) end
  2227 fun add_fact_monotonic_types ctxt mono type_enc =
  2228   add_iformula_monotonic_types ctxt mono type_enc |> fact_lift
  2229 fun monotonic_types_for_facts ctxt mono type_enc facts =
  2230   let val level = level_of_type_enc type_enc in
  2231     [] |> (polymorphism_of_type_enc type_enc = Polymorphic andalso
  2232            is_type_level_monotonicity_based level andalso
  2233            granularity_of_type_level level <> Ghost_Type_Arg_Vars)
  2234           ? fold (add_fact_monotonic_types ctxt mono type_enc) facts
  2235   end
  2236 
  2237 fun formula_line_for_guards_mono_type ctxt format mono type_enc T =
  2238   Formula (guards_sym_formula_prefix ^
  2239            ascii_of (mangled_type format type_enc T),
  2240            Axiom,
  2241            IConst (`make_bound_var "X", T, [])
  2242            |> type_guard_iterm format type_enc T
  2243            |> AAtom
  2244            |> formula_from_iformula ctxt [] format mono type_enc
  2245                                     always_guard_var_in_formula (SOME true)
  2246            |> close_formula_universally
  2247            |> bound_tvars type_enc true (atomic_types_of T),
  2248            NONE, isabelle_info inductiveN helper_rank)
  2249 
  2250 fun formula_line_for_tags_mono_type ctxt format mono type_enc T =
  2251   let val x_var = ATerm (`make_bound_var "X", []) in
  2252     Formula (tags_sym_formula_prefix ^
  2253              ascii_of (mangled_type format type_enc T),
  2254              Axiom,
  2255              eq_formula type_enc (atomic_types_of T) [] false
  2256                   (tag_with_type ctxt format mono type_enc NONE T x_var) x_var,
  2257              NONE, isabelle_info defN helper_rank)
  2258   end
  2259 
  2260 fun problem_lines_for_mono_types ctxt format mono type_enc Ts =
  2261   case type_enc of
  2262     Simple_Types _ => []
  2263   | Guards _ =>
  2264     map (formula_line_for_guards_mono_type ctxt format mono type_enc) Ts
  2265   | Tags _ => map (formula_line_for_tags_mono_type ctxt format mono type_enc) Ts
  2266 
  2267 fun decl_line_for_sym ctxt format mono type_enc s
  2268                       (s', T_args, T, pred_sym, ary, _) =
  2269   let
  2270     val thy = Proof_Context.theory_of ctxt
  2271     val (T, T_args) =
  2272       if null T_args then
  2273         (T, [])
  2274       else case unprefix_and_unascii const_prefix s of
  2275         SOME s' =>
  2276         let
  2277           val s' = s' |> invert_const
  2278           val T = s' |> robust_const_type thy
  2279         in (T, robust_const_typargs thy (s', T)) end
  2280       | NONE => raise Fail "unexpected type arguments"
  2281   in
  2282     Decl (sym_decl_prefix ^ s, (s, s'),
  2283           T |> fused_type ctxt mono (level_of_type_enc type_enc) ary
  2284             |> ho_type_from_typ format type_enc pred_sym ary
  2285             |> not (null T_args)
  2286                ? curry ATyAbs (map (tvar_name o fst o dest_TVar) T_args))
  2287   end
  2288 
  2289 fun honor_conj_sym_kind in_conj conj_sym_kind =
  2290   if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
  2291   else (Axiom, I)
  2292 
  2293 fun formula_line_for_guards_sym_decl ctxt format conj_sym_kind mono type_enc n s
  2294                                      j (s', T_args, T, _, ary, in_conj) =
  2295   let
  2296     val thy = Proof_Context.theory_of ctxt
  2297     val (kind, maybe_negate) = honor_conj_sym_kind in_conj conj_sym_kind
  2298     val (arg_Ts, res_T) = chop_fun ary T
  2299     val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)
  2300     val bounds =
  2301       bound_names ~~ arg_Ts |> map (fn (name, T) => IConst (name, T, []))
  2302     val bound_Ts =
  2303       if exists (curry (op =) dummyT) T_args then
  2304         case level_of_type_enc type_enc of
  2305           All_Types => map SOME arg_Ts
  2306         | level =>
  2307           if granularity_of_type_level level = Ghost_Type_Arg_Vars then
  2308             let val ghosts = ghost_type_args thy s ary in
  2309               map2 (fn j => if member (op =) ghosts j then SOME else K NONE)
  2310                    (0 upto ary - 1) arg_Ts
  2311             end
  2312           else
  2313             replicate ary NONE
  2314       else
  2315         replicate ary NONE
  2316   in
  2317     Formula (guards_sym_formula_prefix ^ s ^
  2318              (if n > 1 then "_" ^ string_of_int j else ""), kind,
  2319              IConst ((s, s'), T, T_args)
  2320              |> fold (curry (IApp o swap)) bounds
  2321              |> type_guard_iterm format type_enc res_T
  2322              |> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)
  2323              |> formula_from_iformula ctxt [] format mono type_enc
  2324                                       always_guard_var_in_formula (SOME true)
  2325              |> close_formula_universally
  2326              |> bound_tvars type_enc (n > 1) (atomic_types_of T)
  2327              |> maybe_negate,
  2328              NONE, isabelle_info inductiveN helper_rank)
  2329   end
  2330 
  2331 fun formula_lines_for_tags_sym_decl ctxt format conj_sym_kind mono type_enc n s
  2332         (j, (s', T_args, T, pred_sym, ary, in_conj)) =
  2333   let
  2334     val thy = Proof_Context.theory_of ctxt
  2335     val level = level_of_type_enc type_enc
  2336     val grain = granularity_of_type_level level
  2337     val ident_base =
  2338       tags_sym_formula_prefix ^ s ^
  2339       (if n > 1 then "_" ^ string_of_int j else "")
  2340     val (kind, maybe_negate) = honor_conj_sym_kind in_conj conj_sym_kind
  2341     val (arg_Ts, res_T) = chop_fun ary T
  2342     val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)
  2343     val bounds = bound_names |> map (fn name => ATerm (name, []))
  2344     val cst = mk_aterm format type_enc (s, s') T_args
  2345     val eq = maybe_negate oo eq_formula type_enc (atomic_types_of T) [] pred_sym
  2346     val should_encode = should_encode_type ctxt mono level
  2347     val tag_with = tag_with_type ctxt format mono type_enc NONE
  2348     val add_formula_for_res =
  2349       if should_encode res_T then
  2350         let
  2351           val tagged_bounds =
  2352             if grain = Ghost_Type_Arg_Vars then
  2353               let val ghosts = ghost_type_args thy s ary in
  2354                 map2 (fn (j, arg_T) => member (op =) ghosts j ? tag_with arg_T)
  2355                      (0 upto ary - 1 ~~ arg_Ts) bounds
  2356               end
  2357             else
  2358               bounds
  2359         in
  2360           cons (Formula (ident_base ^ "_res", kind,
  2361                          eq (tag_with res_T (cst bounds)) (cst tagged_bounds),
  2362                          NONE, isabelle_info defN helper_rank))
  2363         end
  2364       else
  2365         I
  2366   in [] |> not pred_sym ? add_formula_for_res end
  2367 
  2368 fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd
  2369 
  2370 fun rationalize_decls thy (decls as decl :: (decls' as _ :: _)) =
  2371     let
  2372       val T = result_type_of_decl decl
  2373               |> map_type_tvar (fn (z, _) => TVar (z, HOLogic.typeS))
  2374     in
  2375       if forall (type_generalization thy T o result_type_of_decl) decls' then
  2376         [decl]
  2377       else
  2378         decls
  2379     end
  2380   | rationalize_decls _ decls = decls
  2381 
  2382 fun problem_lines_for_sym_decls ctxt format conj_sym_kind mono type_enc
  2383                                 (s, decls) =
  2384   case type_enc of
  2385     Simple_Types _ => [decl_line_for_sym ctxt format mono type_enc s (hd decls)]
  2386   | Guards (_, level) =>
  2387     let
  2388       val thy = Proof_Context.theory_of ctxt
  2389       val decls = decls |> rationalize_decls thy
  2390       val n = length decls
  2391       val decls =
  2392         decls |> filter (should_encode_type ctxt mono level
  2393                          o result_type_of_decl)
  2394     in
  2395       (0 upto length decls - 1, decls)
  2396       |-> map2 (formula_line_for_guards_sym_decl ctxt format conj_sym_kind mono
  2397                                                  type_enc n s)
  2398     end
  2399   | Tags (_, level) =>
  2400     if granularity_of_type_level level = All_Vars then
  2401       []
  2402     else
  2403       let val n = length decls in
  2404         (0 upto n - 1 ~~ decls)
  2405         |> maps (formula_lines_for_tags_sym_decl ctxt format conj_sym_kind mono
  2406                                                  type_enc n s)
  2407       end
  2408 
  2409 fun problem_lines_for_sym_decl_table ctxt format conj_sym_kind mono type_enc
  2410                                      mono_Ts sym_decl_tab =
  2411   let
  2412     val syms = sym_decl_tab |> Symtab.dest |> sort_wrt fst
  2413     val mono_lines =
  2414       problem_lines_for_mono_types ctxt format mono type_enc mono_Ts
  2415     val decl_lines =
  2416       fold_rev (append o problem_lines_for_sym_decls ctxt format conj_sym_kind
  2417                              mono type_enc)
  2418                syms []
  2419   in mono_lines @ decl_lines end
  2420 
  2421 fun pair_append (xs1, xs2) (ys1, ys2) = (xs1 @ ys1, xs2 @ ys2)
  2422 
  2423 fun do_uncurried_alias_lines_for_sym ctxt monom_constrs format conj_sym_kind
  2424         mono type_enc sym_tab0 sym_tab base_s0 types in_conj =
  2425   let
  2426     fun do_alias ary =
  2427       let
  2428         val thy = Proof_Context.theory_of ctxt
  2429         val (kind, maybe_negate) = honor_conj_sym_kind in_conj conj_sym_kind
  2430         val base_name = base_s0 |> `(make_fixed_const (SOME format))
  2431         val T = case types of [T] => T | _ => robust_const_type thy base_s0
  2432         val T_args = robust_const_typargs thy (base_s0, T)
  2433         val (base_name as (base_s, _), T_args) =
  2434           mangle_type_args_in_const format type_enc base_name T_args
  2435         val base_ary = min_ary_of sym_tab0 base_s
  2436         fun do_const name = IConst (name, T, T_args)
  2437         val filter_ty_args =
  2438           filter_type_args_in_iterm thy monom_constrs type_enc
  2439         val ho_term_of =
  2440           ho_term_from_iterm ctxt format mono type_enc (SOME true)
  2441         val name1 as (s1, _) =
  2442           base_name |> ary - 1 > base_ary ? aliased_uncurried (ary - 1)
  2443         val name2 as (s2, _) = base_name |> aliased_uncurried ary
  2444         val (arg_Ts, _) = chop_fun ary T
  2445         val bound_names =
  2446           1 upto ary |> map (`I o make_bound_var o string_of_int)
  2447         val bounds = bound_names ~~ arg_Ts
  2448         val (first_bounds, last_bound) =
  2449           bounds |> map (fn (name, T) => IConst (name, T, [])) |> split_last
  2450         val tm1 =
  2451           mk_app_op format type_enc (list_app (do_const name1) first_bounds)
  2452                     last_bound
  2453           |> filter_ty_args
  2454         val tm2 =
  2455           list_app (do_const name2) (first_bounds @ [last_bound])
  2456           |> filter_ty_args
  2457         val do_bound_type = do_bound_type ctxt format mono type_enc
  2458         val eq =
  2459           eq_formula type_enc (atomic_types_of T)
  2460                      (map (apsnd do_bound_type) bounds) false
  2461                      (ho_term_of tm1) (ho_term_of tm2)
  2462       in
  2463         ([tm1, tm2],
  2464          [Formula (uncurried_alias_eq_prefix ^ s2, kind, eq |> maybe_negate,
  2465                    NONE, isabelle_info defN helper_rank)])
  2466         |> (if ary - 1 = base_ary orelse Symtab.defined sym_tab s1 then I
  2467             else pair_append (do_alias (ary - 1)))
  2468       end
  2469   in do_alias end
  2470 fun uncurried_alias_lines_for_sym ctxt monom_constrs format conj_sym_kind mono
  2471         type_enc sym_tab0 sym_tab
  2472         (s, {min_ary, types, in_conj, ...} : sym_info) =
  2473   case unprefix_and_unascii const_prefix s of
  2474     SOME mangled_s =>
  2475     if String.isSubstring uncurried_alias_sep mangled_s then
  2476       let
  2477         val base_s0 = mangled_s |> unmangled_const_name |> hd |> invert_const
  2478       in
  2479         do_uncurried_alias_lines_for_sym ctxt monom_constrs format conj_sym_kind
  2480             mono type_enc sym_tab0 sym_tab base_s0 types in_conj min_ary
  2481       end
  2482     else
  2483       ([], [])
  2484   | NONE => ([], [])
  2485 fun uncurried_alias_lines_for_sym_table ctxt monom_constrs format conj_sym_kind
  2486         mono type_enc uncurried_aliases sym_tab0 sym_tab =
  2487   ([], [])
  2488   |> uncurried_aliases
  2489      ? Symtab.fold_rev
  2490            (pair_append
  2491             o uncurried_alias_lines_for_sym ctxt monom_constrs format
  2492                   conj_sym_kind mono type_enc sym_tab0 sym_tab) sym_tab
  2493 
  2494 val implicit_declsN = "Should-be-implicit typings"
  2495 val explicit_declsN = "Explicit typings"
  2496 val uncurried_alias_eqsN = "Uncurried aliases"
  2497 val factsN = "Relevant facts"
  2498 val class_relsN = "Class relationships"
  2499 val aritiesN = "Arities"
  2500 val helpersN = "Helper facts"
  2501 val conjsN = "Conjectures"
  2502 val free_typesN = "Type variables"
  2503 
  2504 (* TFF allows implicit declarations of types, function symbols, and predicate
  2505    symbols (with "$i" as the type of individuals), but some provers (e.g.,
  2506    SNARK) require explicit declarations. The situation is similar for THF. *)
  2507 
  2508 fun default_type type_enc pred_sym s =
  2509   let
  2510     val ind =
  2511       case type_enc of
  2512         Simple_Types _ =>
  2513         if String.isPrefix type_const_prefix s orelse
  2514            String.isPrefix tfree_prefix s then
  2515           atype_of_types
  2516         else
  2517           individual_atype
  2518       | _ => individual_atype
  2519     fun typ 0 = if pred_sym then bool_atype else ind
  2520       | typ ary = AFun (ind, typ (ary - 1))
  2521   in typ end
  2522 
  2523 fun nary_type_constr_type n =
  2524   funpow n (curry AFun atype_of_types) atype_of_types
  2525 
  2526 fun undeclared_syms_in_problem type_enc problem =
  2527   let
  2528     fun do_sym (name as (s, _)) ty =
  2529       if is_tptp_user_symbol s then
  2530         Symtab.default (s, (name, ty))
  2531       else
  2532         I
  2533     fun do_type (AType (name, tys)) =
  2534         do_sym name (fn () => nary_type_constr_type (length tys))
  2535         #> fold do_type tys
  2536       | do_type (AFun (ty1, ty2)) = do_type ty1 #> do_type ty2
  2537       | do_type (ATyAbs (_, ty)) = do_type ty
  2538     fun do_term pred_sym (ATerm (name as (s, _), tms)) =
  2539         do_sym name (fn _ => default_type type_enc pred_sym s (length tms))
  2540         #> fold (do_term false) tms
  2541       | do_term _ (AAbs ((_, ty), tm)) = do_type ty #> do_term false tm
  2542     fun do_formula (AQuant (_, xs, phi)) =
  2543         fold do_type (map_filter snd xs) #> do_formula phi
  2544       | do_formula (AConn (_, phis)) = fold do_formula phis
  2545       | do_formula (AAtom tm) = do_term true tm
  2546     fun do_problem_line (Decl (_, _, ty)) = do_type ty
  2547       | do_problem_line (Formula (_, _, phi, _, _)) = do_formula phi
  2548   in
  2549     Symtab.empty
  2550     |> fold (fn (s, _) => Symtab.default (s, (("", ""), K tvar_a_atype)))
  2551             (declared_syms_in_problem problem)
  2552     |> fold (fold do_problem_line o snd) problem
  2553   end
  2554 
  2555 fun declare_undeclared_syms_in_atp_problem type_enc problem =
  2556   let
  2557     val decls =
  2558       Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *)
  2559                     | (s, (sym, ty)) =>
  2560                       cons (Decl (type_decl_prefix ^ s, sym, ty ())))
  2561                   (undeclared_syms_in_problem type_enc problem) []
  2562   in (implicit_declsN, decls) :: problem end
  2563 
  2564 fun exists_subdtype P =
  2565   let
  2566     fun ex U = P U orelse
  2567       (case U of Datatype.DtType (_, Us) => exists ex Us | _ => false)
  2568   in ex end
  2569 
  2570 fun is_poly_constr (_, Us) =
  2571   exists (exists_subdtype (fn Datatype.DtTFree _ => true | _ => false)) Us
  2572 
  2573 fun all_constrs_of_polymorphic_datatypes thy =
  2574   Symtab.fold (snd
  2575                #> #descr
  2576                #> maps (snd #> #3)
  2577                #> (fn cs => exists is_poly_constr cs ? append cs))
  2578               (Datatype.get_all thy) []
  2579   |> List.partition is_poly_constr
  2580   |> pairself (map fst)
  2581 
  2582 val app_op_and_predicator_threshold = 50
  2583 
  2584 fun prepare_atp_problem ctxt format conj_sym_kind prem_kind type_enc exporter
  2585                         lam_trans uncurried_aliases readable_names preproc
  2586                         hyp_ts concl_t facts =
  2587   let
  2588     val thy = Proof_Context.theory_of ctxt
  2589     val type_enc = type_enc |> adjust_type_enc format
  2590     (* Forcing explicit applications is expensive for polymorphic encodings,
  2591        because it takes only one existential variable ranging over "'a => 'b" to
  2592        ruin everything. Hence we do it only if there are few facts (which is
  2593        normally the case for "metis" and the minimizer). *)
  2594     val app_op_level =
  2595       if length facts > app_op_and_predicator_threshold then
  2596         if polymorphism_of_type_enc type_enc = Polymorphic then
  2597           Min_App_Op
  2598         else
  2599           Sufficient_App_Op
  2600       else
  2601         Sufficient_App_Op_And_Predicator
  2602     val lam_trans =
  2603       if lam_trans = keep_lamsN andalso
  2604          not (is_type_enc_higher_order type_enc) then
  2605         error ("Lambda translation scheme incompatible with first-order \
  2606                \encoding.")
  2607       else
  2608         lam_trans
  2609     val (fact_names, classes, conjs, facts, class_rel_clauses, arity_clauses,
  2610          lifted) =
  2611       translate_formulas ctxt format prem_kind type_enc lam_trans preproc hyp_ts
  2612                          concl_t facts
  2613     val sym_tab0 = sym_table_for_facts ctxt type_enc app_op_level conjs facts
  2614     val mono = conjs @ facts |> mononotonicity_info_for_facts ctxt type_enc
  2615     val (polym_constrs, monom_constrs) =
  2616       all_constrs_of_polymorphic_datatypes thy
  2617       |>> map (make_fixed_const (SOME format))
  2618     fun firstorderize in_helper =
  2619       firstorderize_fact thy monom_constrs format type_enc sym_tab0
  2620                          (uncurried_aliases andalso not in_helper)
  2621     val (conjs, facts) = (conjs, facts) |> pairself (map (firstorderize false))
  2622     val sym_tab = sym_table_for_facts ctxt type_enc Min_App_Op conjs facts
  2623     val helpers =
  2624       sym_tab |> helper_facts_for_sym_table ctxt format type_enc
  2625               |> map (firstorderize true)
  2626     val mono_Ts =
  2627       helpers @ conjs @ facts |> monotonic_types_for_facts ctxt mono type_enc
  2628     val class_decl_lines = decl_lines_for_classes type_enc classes
  2629     val (uncurried_alias_eq_tms, uncurried_alias_eq_lines) =
  2630       uncurried_alias_lines_for_sym_table ctxt monom_constrs format
  2631           conj_sym_kind mono type_enc uncurried_aliases sym_tab0 sym_tab
  2632     val sym_decl_lines =
  2633       (conjs, helpers @ facts, uncurried_alias_eq_tms)
  2634       |> sym_decl_table_for_facts thy format type_enc sym_tab
  2635       |> problem_lines_for_sym_decl_table ctxt format conj_sym_kind mono
  2636                                                type_enc mono_Ts
  2637     val num_facts = length facts
  2638     val fact_lines =
  2639       map (formula_line_for_fact ctxt polym_constrs format fact_prefix
  2640                ascii_of (not exporter) (not exporter) mono type_enc
  2641                (rank_of_fact_num num_facts))
  2642           (0 upto num_facts - 1 ~~ facts)
  2643     val helper_lines =
  2644       0 upto length helpers - 1 ~~ helpers
  2645       |> map (formula_line_for_fact ctxt polym_constrs format helper_prefix I
  2646                                     false true mono type_enc (K default_rank))
  2647     (* Reordering these might confuse the proof reconstruction code or the SPASS
  2648        FLOTTER hack. *)
  2649     val problem =
  2650       [(explicit_declsN, class_decl_lines @ sym_decl_lines),
  2651        (uncurried_alias_eqsN, uncurried_alias_eq_lines),
  2652        (factsN, fact_lines),
  2653        (class_relsN,
  2654         map (formula_line_for_class_rel_clause type_enc) class_rel_clauses),
  2655        (aritiesN, map (formula_line_for_arity_clause type_enc) arity_clauses),
  2656        (helpersN, helper_lines),
  2657        (free_typesN, formula_lines_for_free_types type_enc (facts @ conjs)),
  2658        (conjsN,
  2659         map (formula_line_for_conjecture ctxt polym_constrs format mono
  2660                                          type_enc) conjs)]
  2661     val problem =
  2662       problem
  2663       |> (case format of
  2664             CNF => ensure_cnf_problem
  2665           | CNF_UEQ => filter_cnf_ueq_problem
  2666           | FOF => I
  2667           | TFF (_, TPTP_Implicit) => I
  2668           | THF (_, TPTP_Implicit, _) => I
  2669           | _ => declare_undeclared_syms_in_atp_problem type_enc)
  2670     val (problem, pool) = problem |> nice_atp_problem readable_names format
  2671     fun add_sym_ary (s, {min_ary, ...} : sym_info) =
  2672       min_ary > 0 ? Symtab.insert (op =) (s, min_ary)
  2673   in
  2674     (problem,
  2675      case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
  2676      fact_names |> Vector.fromList,
  2677      lifted,
  2678      Symtab.empty |> Symtab.fold add_sym_ary sym_tab)
  2679   end
  2680 
  2681 (* FUDGE *)
  2682 val conj_weight = 0.0
  2683 val hyp_weight = 0.1
  2684 val fact_min_weight = 0.2
  2685 val fact_max_weight = 1.0
  2686 val type_info_default_weight = 0.8
  2687 
  2688 (* Weights are from 0.0 (most important) to 1.0 (least important). *)
  2689 fun atp_problem_selection_weights problem =
  2690   let
  2691     fun add_term_weights weight (ATerm (s, tms)) =
  2692         is_tptp_user_symbol s ? Symtab.default (s, weight)
  2693         #> fold (add_term_weights weight) tms
  2694       | add_term_weights weight (AAbs (_, tm)) = add_term_weights weight tm
  2695     fun add_line_weights weight (Formula (_, _, phi, _, _)) =
  2696         formula_fold NONE (K (add_term_weights weight)) phi
  2697       | add_line_weights _ _ = I
  2698     fun add_conjectures_weights [] = I
  2699       | add_conjectures_weights conjs =
  2700         let val (hyps, conj) = split_last conjs in
  2701           add_line_weights conj_weight conj
  2702           #> fold (add_line_weights hyp_weight) hyps
  2703         end
  2704     fun add_facts_weights facts =
  2705       let
  2706         val num_facts = length facts
  2707         fun weight_of j =
  2708           fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j
  2709                             / Real.fromInt num_facts
  2710       in
  2711         map weight_of (0 upto num_facts - 1) ~~ facts
  2712         |> fold (uncurry add_line_weights)
  2713       end
  2714     val get = these o AList.lookup (op =) problem
  2715   in
  2716     Symtab.empty
  2717     |> add_conjectures_weights (get free_typesN @ get conjsN)
  2718     |> add_facts_weights (get factsN)
  2719     |> fold (fold (add_line_weights type_info_default_weight) o get)
  2720             [explicit_declsN, class_relsN, aritiesN]
  2721     |> Symtab.dest
  2722     |> sort (prod_ord Real.compare string_ord o pairself swap)
  2723   end
  2724 
  2725 (* Ugly hack: may make innocent victims (collateral damage) *)
  2726 fun may_be_app s args = String.isPrefix app_op_name s andalso length args = 2
  2727 fun may_be_predicator s =
  2728   member (op =) [predicator_name, prefixed_predicator_name] s
  2729 
  2730 fun strip_predicator (tm as ATerm (s, [tm'])) =
  2731     if may_be_predicator s then tm' else tm
  2732   | strip_predicator tm = tm
  2733 
  2734 fun make_head_roll (ATerm (s, tms)) =
  2735     if may_be_app s tms then make_head_roll (hd tms) ||> append (tl tms)
  2736     else (s, tms)
  2737   | make_head_roll _ = ("", [])
  2738 
  2739 fun strip_up_to_predicator (AQuant (_, _, phi)) = strip_up_to_predicator phi
  2740   | strip_up_to_predicator (AConn (_, phis)) = maps strip_up_to_predicator phis
  2741   | strip_up_to_predicator (AAtom tm) = [strip_predicator tm]
  2742 
  2743 fun strip_ahorn_etc (AQuant (_, _, phi)) = strip_ahorn_etc phi
  2744   | strip_ahorn_etc (AConn (AImplies, [phi1, phi2])) =
  2745     strip_ahorn_etc phi2 |>> append (strip_up_to_predicator phi1)
  2746   | strip_ahorn_etc phi = ([], hd (strip_up_to_predicator phi))
  2747 
  2748 fun strip_iff_etc (AQuant (_, _, phi)) = strip_iff_etc phi
  2749   | strip_iff_etc (AConn (AIff, [phi1, phi2])) =
  2750     pairself strip_up_to_predicator (phi1, phi2)
  2751   | strip_iff_etc _ = ([], [])
  2752 
  2753 val max_term_order_weight = 2147483647
  2754 
  2755 fun atp_problem_term_order_info problem =
  2756   let
  2757     fun add_edge s s' =
  2758       Graph.default_node (s, ())
  2759       #> Graph.default_node (s', ())
  2760       #> Graph.add_edge_acyclic (s, s')
  2761     fun add_term_deps head (ATerm (s, args)) =
  2762         if is_tptp_user_symbol head then
  2763           (if is_tptp_user_symbol s then perhaps (try (add_edge s head)) else I)
  2764           #> fold (add_term_deps head) args
  2765         else
  2766           I
  2767       | add_term_deps head (AAbs (_, tm)) = add_term_deps head tm
  2768     fun add_intro_deps pred (Formula (_, role, phi, _, info)) =
  2769         if pred (role, info) then
  2770           let val (hyps, concl) = strip_ahorn_etc phi in
  2771             case make_head_roll concl of
  2772               (head, args as _ :: _) => fold (add_term_deps head) (hyps @ args)
  2773             | _ => I
  2774           end
  2775         else
  2776           I
  2777       | add_intro_deps _ _ = I
  2778     fun add_atom_eq_deps (SOME true) (ATerm (s, [lhs as _, rhs])) =
  2779         if is_tptp_equal s then
  2780           case make_head_roll lhs of
  2781             (head, args as _ :: _) => fold (add_term_deps head) (rhs :: args)
  2782           | _ => I
  2783         else
  2784           I
  2785       | add_atom_eq_deps _ _ = I
  2786     fun add_eq_deps pred (Formula (_, role, phi, _, info)) =
  2787         if pred (role, info) then
  2788           case strip_iff_etc phi of
  2789             ([lhs], rhs) =>
  2790             (case make_head_roll lhs of
  2791                (head, args as _ :: _) => fold (add_term_deps head) (rhs @ args)
  2792              | _ => I)
  2793           | _ => formula_fold (SOME (role <> Conjecture)) add_atom_eq_deps phi
  2794         else
  2795           I
  2796       | add_eq_deps _ _ = I
  2797     fun has_status status (_, info) =
  2798       extract_isabelle_status info = SOME status
  2799     fun is_conj (role, _) = (role = Conjecture orelse role = Hypothesis)
  2800     val graph =
  2801       Graph.empty
  2802       |> fold (fold (add_eq_deps (has_status defN)) o snd) problem
  2803       |> fold (fold (add_eq_deps (has_status simpN orf is_conj)) o snd) problem
  2804       |> fold (fold (add_intro_deps (has_status inductiveN)) o snd) problem
  2805       |> fold (fold (add_intro_deps (has_status introN)) o snd) problem
  2806     fun next_weight w = if w + w <= max_term_order_weight then w + w else w + 1
  2807     fun add_weights _ [] = I
  2808       | add_weights weight syms =
  2809         fold (AList.update (op =) o rpair weight) syms
  2810         #> add_weights (next_weight weight)
  2811                (fold (union (op =) o Graph.immediate_succs graph) syms [])
  2812   in
  2813     (* Sorting is not just for aesthetics: It specifies the precedence order
  2814        for the term ordering (KBO or LPO), from smaller to larger values. *)
  2815     [] |> add_weights 1 (Graph.minimals graph) |> sort (int_ord o pairself snd)
  2816   end
  2817 
  2818 end;