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