src/HOL/Tools/ATP/atp_problem_generate.ML
author blanchet
Wed, 06 Jun 2012 10:35:05 +0200
changeset 49102 94835838ed2c
parent 49096 6435b2c73038
child 49103 c75f36d190df
permissions -rw-r--r--
removed micro-optimization whose justification I can't recall
     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_role = ATP_Problem.formula_role
    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_role -> 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_choice
   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 phantom type variables. *)
   187 val avoid_first_order_phantom_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, _, choice, _))
   690                     (Native (order, poly, level)) =
   691     Native (adjust_order choice order, poly, level)
   692   | adjust_type_enc (THF (TPTP_Monomorphic, _, choice, _))
   693                          (Native (order, _, level)) =
   694     Native (adjust_order choice 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 ifact =
   823   {name : string,
   824    stature : stature,
   825    role : formula_role,
   826    iformula : (name, typ, iterm) formula,
   827    atomic_types : typ list}
   828 
   829 fun update_iformula f ({name, stature, role, iformula, atomic_types} : ifact) =
   830   {name = name, stature = stature, role = role, iformula = f iformula,
   831    atomic_types = atomic_types} : ifact
   832 
   833 fun ifact_lift f ({iformula, ...} : ifact) = f iformula
   834 
   835 fun insert_type thy get_T x xs =
   836   let val T = get_T x in
   837     if exists (type_instance thy T o get_T) xs then xs
   838     else x :: filter_out (type_generalization thy T o get_T) xs
   839   end
   840 
   841 (* The Booleans indicate whether all type arguments should be kept. *)
   842 datatype type_arg_policy =
   843   Explicit_Type_Args of bool (* infer_from_term_args *) |
   844   Mangled_Type_Args |
   845   No_Type_Args
   846 
   847 fun type_arg_policy monom_constrs type_enc s =
   848   let val poly = polymorphism_of_type_enc type_enc in
   849     if s = type_tag_name then
   850       if poly = Mangled_Monomorphic then Mangled_Type_Args
   851       else Explicit_Type_Args false
   852     else case type_enc of
   853       Native (_, Polymorphic, _) => Explicit_Type_Args false
   854     | Tags (_, All_Types) => No_Type_Args
   855     | _ =>
   856       let val level = level_of_type_enc type_enc in
   857         if level = No_Types orelse s = @{const_name HOL.eq} orelse
   858            (s = app_op_name andalso level = Const_Arg_Types) then
   859           No_Type_Args
   860         else if poly = Mangled_Monomorphic then
   861           Mangled_Type_Args
   862         else if member (op =) monom_constrs s andalso
   863                 granularity_of_type_level level <> Ghost_Type_Arg_Vars then
   864           No_Type_Args
   865         else
   866           Explicit_Type_Args
   867               (level = All_Types orelse
   868                granularity_of_type_level level = Ghost_Type_Arg_Vars)
   869       end
   870   end
   871 
   872 (* Make atoms for sorted type variables. *)
   873 fun generic_add_sorts_on_type (_, []) = I
   874   | generic_add_sorts_on_type ((x, i), s :: ss) =
   875     generic_add_sorts_on_type ((x, i), ss)
   876     #> (if s = the_single @{sort HOL.type} then
   877           I
   878         else if i = ~1 then
   879           insert (op =) (`make_type_class s, `make_fixed_type_var x)
   880         else
   881           insert (op =) (`make_type_class s,
   882                          (make_schematic_type_var (x, i), x)))
   883 fun add_sorts_on_tfree (TFree (s, S)) = generic_add_sorts_on_type ((s, ~1), S)
   884   | add_sorts_on_tfree _ = I
   885 fun add_sorts_on_tvar (TVar z) = generic_add_sorts_on_type z
   886   | add_sorts_on_tvar _ = I
   887 
   888 fun type_class_formula type_enc class arg =
   889   AAtom (ATerm (class, arg ::
   890       (case type_enc of
   891          Native (First_Order, Polymorphic, _) =>
   892          if avoid_first_order_phantom_type_vars then [ATerm (TYPE_name, [arg])]
   893          else []
   894        | _ => [])))
   895 fun formulas_for_types type_enc add_sorts_on_typ Ts =
   896   [] |> level_of_type_enc type_enc <> No_Types ? fold add_sorts_on_typ Ts
   897      |> map (fn (class, name) =>
   898                 type_class_formula type_enc class (ATerm (name, [])))
   899 
   900 fun mk_aconns c phis =
   901   let val (phis', phi') = split_last phis in
   902     fold_rev (mk_aconn c) phis' phi'
   903   end
   904 fun mk_ahorn [] phi = phi
   905   | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])
   906 fun mk_aquant _ [] phi = phi
   907   | mk_aquant q xs (phi as AQuant (q', xs', phi')) =
   908     if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)
   909   | mk_aquant q xs phi = AQuant (q, xs, phi)
   910 
   911 fun close_universally add_term_vars phi =
   912   let
   913     fun add_formula_vars bounds (AQuant (_, xs, phi)) =
   914         add_formula_vars (map fst xs @ bounds) phi
   915       | add_formula_vars bounds (AConn (_, phis)) =
   916         fold (add_formula_vars bounds) phis
   917       | add_formula_vars bounds (AAtom tm) = add_term_vars bounds tm
   918   in mk_aquant AForall (add_formula_vars [] phi []) phi end
   919 
   920 fun add_term_vars bounds (ATerm (name as (s, _), tms)) =
   921     (if is_tptp_variable s andalso
   922         not (String.isPrefix tvar_prefix s) andalso
   923         not (member (op =) bounds name) then
   924        insert (op =) (name, NONE)
   925      else
   926        I)
   927     #> fold (add_term_vars bounds) tms
   928   | add_term_vars bounds (AAbs (((name, _), tm), args)) =
   929     add_term_vars (name :: bounds) tm #> fold (add_term_vars bounds) args
   930 fun close_formula_universally phi = close_universally add_term_vars phi
   931 
   932 fun add_iterm_vars bounds (IApp (tm1, tm2)) =
   933     fold (add_iterm_vars bounds) [tm1, tm2]
   934   | add_iterm_vars _ (IConst _) = I
   935   | add_iterm_vars bounds (IVar (name, T)) =
   936     not (member (op =) bounds name) ? insert (op =) (name, SOME T)
   937   | add_iterm_vars bounds (IAbs (_, tm)) = add_iterm_vars bounds tm
   938 
   939 fun close_iformula_universally phi = close_universally add_iterm_vars phi
   940 
   941 val fused_infinite_type_name = "ATP.fused_inf" (* shouldn't clash *)
   942 val fused_infinite_type = Type (fused_infinite_type_name, [])
   943 
   944 fun tvar_name (x as (s, _)) = (make_schematic_type_var x, s)
   945 
   946 fun ho_term_from_typ type_enc =
   947   let
   948     fun term (Type (s, Ts)) =
   949       ATerm (case (is_type_enc_higher_order type_enc, s) of
   950                (true, @{type_name bool}) => `I tptp_bool_type
   951              | (true, @{type_name fun}) => `I tptp_fun_type
   952              | _ => if s = fused_infinite_type_name andalso
   953                        is_type_enc_native type_enc then
   954                       `I tptp_individual_type
   955                     else
   956                       `make_fixed_type_const s,
   957              map term Ts)
   958     | term (TFree (s, _)) = ATerm (`make_fixed_type_var s, [])
   959     | term (TVar (x, _)) = ATerm (tvar_name x, [])
   960   in term end
   961 
   962 fun ho_term_for_type_arg type_enc T =
   963   if T = dummyT then NONE else SOME (ho_term_from_typ type_enc T)
   964 
   965 (* This shouldn't clash with anything else. *)
   966 val uncurried_alias_sep = "\000"
   967 val mangled_type_sep = "\001"
   968 
   969 val ascii_of_uncurried_alias_sep = ascii_of uncurried_alias_sep
   970 
   971 fun generic_mangled_type_name f (ATerm (name, [])) = f name
   972   | generic_mangled_type_name f (ATerm (name, tys)) =
   973     f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys)
   974     ^ ")"
   975   | generic_mangled_type_name _ _ = raise Fail "unexpected type abstraction"
   976 
   977 fun mangled_type type_enc =
   978   generic_mangled_type_name fst o ho_term_from_typ type_enc
   979 
   980 fun make_native_type s =
   981   if s = tptp_bool_type orelse s = tptp_fun_type orelse
   982      s = tptp_individual_type then
   983     s
   984   else
   985     native_type_prefix ^ ascii_of s
   986 
   987 fun ho_type_from_ho_term type_enc pred_sym ary =
   988   let
   989     fun to_mangled_atype ty =
   990       AType ((make_native_type (generic_mangled_type_name fst ty),
   991               generic_mangled_type_name snd ty), [])
   992     fun to_poly_atype (ATerm (name, tys)) = AType (name, map to_poly_atype tys)
   993       | to_poly_atype _ = raise Fail "unexpected type abstraction"
   994     val to_atype =
   995       if polymorphism_of_type_enc type_enc = Polymorphic then to_poly_atype
   996       else to_mangled_atype
   997     fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1))
   998     fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty
   999       | to_fo ary (ATerm (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys
  1000       | to_fo _ _ = raise Fail "unexpected type abstraction"
  1001     fun to_ho (ty as ATerm ((s, _), tys)) =
  1002         if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty
  1003       | to_ho _ = raise Fail "unexpected type abstraction"
  1004   in if is_type_enc_higher_order type_enc then to_ho else to_fo ary end
  1005 
  1006 fun ho_type_from_typ type_enc pred_sym ary =
  1007   ho_type_from_ho_term type_enc pred_sym ary
  1008   o ho_term_from_typ type_enc
  1009 
  1010 fun aliased_uncurried ary (s, s') =
  1011   (s ^ ascii_of_uncurried_alias_sep ^ string_of_int ary, s' ^ string_of_int ary)
  1012 fun unaliased_uncurried (s, s') =
  1013   case space_explode uncurried_alias_sep s of
  1014     [_] => (s, s')
  1015   | [s1, s2] => (s1, unsuffix s2 s')
  1016   | _ => raise Fail "ill-formed explicit application alias"
  1017 
  1018 fun raw_mangled_const_name type_name ty_args (s, s') =
  1019   let
  1020     fun type_suffix f g =
  1021       fold_rev (curry (op ^) o g o prefix mangled_type_sep o type_name f)
  1022                ty_args ""
  1023   in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end
  1024 fun mangled_const_name type_enc =
  1025   map_filter (ho_term_for_type_arg type_enc)
  1026   #> raw_mangled_const_name generic_mangled_type_name
  1027 
  1028 val parse_mangled_ident =
  1029   Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
  1030 
  1031 fun parse_mangled_type x =
  1032   (parse_mangled_ident
  1033    -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
  1034                     [] >> ATerm) x
  1035 and parse_mangled_types x =
  1036   (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x
  1037 
  1038 fun unmangled_type s =
  1039   s |> suffix ")" |> raw_explode
  1040     |> Scan.finite Symbol.stopper
  1041            (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^
  1042                                                 quote s)) parse_mangled_type))
  1043     |> fst
  1044 
  1045 fun unmangled_const_name s =
  1046   (s, s) |> unaliased_uncurried |> fst |> space_explode mangled_type_sep
  1047 fun unmangled_const s =
  1048   let val ss = unmangled_const_name s in
  1049     (hd ss, map unmangled_type (tl ss))
  1050   end
  1051 
  1052 fun introduce_proxies_in_iterm type_enc =
  1053   let
  1054     fun tweak_ho_quant ho_quant T [IAbs _] = IConst (`I ho_quant, T, [])
  1055       | tweak_ho_quant ho_quant (T as Type (_, [p_T as Type (_, [x_T, _]), _]))
  1056                        _ =
  1057         (* Eta-expand "!!" and "??", to work around LEO-II 1.2.8 parser
  1058            limitation. This works in conjuction with special code in
  1059            "ATP_Problem" that uses the syntactic sugar "!" and "?" whenever
  1060            possible. *)
  1061         IAbs ((`I "P", p_T),
  1062               IApp (IConst (`I ho_quant, T, []),
  1063                     IAbs ((`I "X", x_T),
  1064                           IApp (IConst (`I "P", p_T, []),
  1065                                 IConst (`I "X", x_T, [])))))
  1066       | tweak_ho_quant _ _ _ = raise Fail "unexpected type for quantifier"
  1067     fun intro top_level args (IApp (tm1, tm2)) =
  1068         IApp (intro top_level (tm2 :: args) tm1, intro false [] tm2)
  1069       | intro top_level args (IConst (name as (s, _), T, T_args)) =
  1070         (case proxify_const s of
  1071            SOME proxy_base =>
  1072            if top_level orelse is_type_enc_higher_order type_enc then
  1073              case (top_level, s) of
  1074                (_, "c_False") => IConst (`I tptp_false, T, [])
  1075              | (_, "c_True") => IConst (`I tptp_true, T, [])
  1076              | (false, "c_Not") => IConst (`I tptp_not, T, [])
  1077              | (false, "c_conj") => IConst (`I tptp_and, T, [])
  1078              | (false, "c_disj") => IConst (`I tptp_or, T, [])
  1079              | (false, "c_implies") => IConst (`I tptp_implies, T, [])
  1080              | (false, "c_All") => tweak_ho_quant tptp_ho_forall T args
  1081              | (false, "c_Ex") => tweak_ho_quant tptp_ho_exists T args
  1082              | (false, s) =>
  1083                if is_tptp_equal s then
  1084                  if length args = 2 then
  1085                    IConst (`I tptp_equal, T, [])
  1086                  else
  1087                    (* Eta-expand partially applied THF equality, because the
  1088                       LEO-II and Satallax parsers complain about not being able to
  1089                       infer the type of "=". *)
  1090                    let val i_T = domain_type T in
  1091                      IAbs ((`I "Y", i_T),
  1092                            IAbs ((`I "Z", i_T),
  1093                                  IApp (IApp (IConst (`I tptp_equal, T, []),
  1094                                              IConst (`I "Y", i_T, [])),
  1095                                        IConst (`I "Z", i_T, []))))
  1096                    end
  1097                else
  1098                  IConst (name, T, [])
  1099              | _ => IConst (name, T, [])
  1100            else
  1101              IConst (proxy_base |>> prefix const_prefix, T, T_args)
  1102           | NONE => if s = tptp_choice then tweak_ho_quant tptp_choice T args
  1103                     else IConst (name, T, T_args))
  1104       | intro _ _ (IAbs (bound, tm)) = IAbs (bound, intro false [] tm)
  1105       | intro _ _ tm = tm
  1106   in intro true [] end
  1107 
  1108 fun mangle_type_args_in_const type_enc (name as (s, _)) T_args =
  1109   case unprefix_and_unascii const_prefix s of
  1110     NONE => (name, T_args)
  1111   | SOME s'' =>
  1112     case type_arg_policy [] type_enc (invert_const s'') of
  1113       Mangled_Type_Args => (mangled_const_name type_enc T_args name, [])
  1114     | _ => (name, T_args)
  1115 fun mangle_type_args_in_iterm type_enc =
  1116   if polymorphism_of_type_enc type_enc = Mangled_Monomorphic then
  1117     let
  1118       fun mangle (IApp (tm1, tm2)) = IApp (mangle tm1, mangle tm2)
  1119         | mangle (tm as IConst (_, _, [])) = tm
  1120         | mangle (IConst (name, T, T_args)) =
  1121           mangle_type_args_in_const type_enc name T_args
  1122           |> (fn (name, T_args) => IConst (name, T, T_args))
  1123         | mangle (IAbs (bound, tm)) = IAbs (bound, mangle tm)
  1124         | mangle tm = tm
  1125     in mangle end
  1126   else
  1127     I
  1128 
  1129 fun chop_fun 0 T = ([], T)
  1130   | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
  1131     chop_fun (n - 1) ran_T |>> cons dom_T
  1132   | chop_fun _ T = ([], T)
  1133 
  1134 fun filter_const_type_args _ _ _ [] = []
  1135   | filter_const_type_args thy s ary T_args =
  1136     let
  1137       val U = robust_const_type thy s
  1138       val arg_U_vars = fold Term.add_tvarsT (U |> chop_fun ary |> fst) []
  1139       val U_args = (s, U) |> robust_const_typargs thy
  1140     in
  1141       U_args ~~ T_args
  1142       |> map (fn (U, T) =>
  1143                  if member (op =) arg_U_vars (dest_TVar U) then dummyT else T)
  1144     end
  1145     handle TYPE _ => T_args
  1146 
  1147 fun filter_type_args_in_const _ _ _ _ _ [] = []
  1148   | filter_type_args_in_const thy monom_constrs type_enc ary s T_args =
  1149     case unprefix_and_unascii const_prefix s of
  1150       NONE =>
  1151       if level_of_type_enc type_enc = No_Types orelse s = tptp_choice then []
  1152       else T_args
  1153     | SOME s'' =>
  1154       let
  1155         val s'' = invert_const s''
  1156         fun filter_T_args false = T_args
  1157           | filter_T_args true = filter_const_type_args thy s'' ary T_args
  1158       in
  1159         case type_arg_policy monom_constrs type_enc s'' of
  1160           Explicit_Type_Args infer_from_term_args =>
  1161           filter_T_args infer_from_term_args
  1162         | No_Type_Args => []
  1163         | Mangled_Type_Args => raise Fail "unexpected (un)mangled symbol"
  1164       end
  1165 fun filter_type_args_in_iterm thy monom_constrs type_enc =
  1166   let
  1167     fun filt ary (IApp (tm1, tm2)) = IApp (filt (ary + 1) tm1, filt 0 tm2)
  1168       | filt ary (IConst (name as (s, _), T, T_args)) =
  1169         filter_type_args_in_const thy monom_constrs type_enc ary s T_args
  1170         |> (fn T_args => IConst (name, T, T_args))
  1171       | filt _ (IAbs (bound, tm)) = IAbs (bound, filt 0 tm)
  1172       | filt _ tm = tm
  1173   in filt 0 end
  1174 
  1175 fun iformula_from_prop ctxt type_enc iff_for_eq =
  1176   let
  1177     val thy = Proof_Context.theory_of ctxt
  1178     fun do_term bs t atomic_Ts =
  1179       iterm_from_term thy type_enc bs (Envir.eta_contract t)
  1180       |>> (introduce_proxies_in_iterm type_enc
  1181            #> mangle_type_args_in_iterm type_enc #> AAtom)
  1182       ||> union (op =) atomic_Ts
  1183     fun do_quant bs q pos s T t' =
  1184       let
  1185         val s = singleton (Name.variant_list (map fst bs)) s
  1186         val universal = Option.map (q = AExists ? not) pos
  1187         val name =
  1188           s |> `(case universal of
  1189                    SOME true => make_all_bound_var
  1190                  | SOME false => make_exist_bound_var
  1191                  | NONE => make_bound_var)
  1192       in
  1193         do_formula ((s, (name, T)) :: bs) pos t'
  1194         #>> mk_aquant q [(name, SOME T)]
  1195         ##> union (op =) (atomic_types_of T)
  1196       end
  1197     and do_conn bs c pos1 t1 pos2 t2 =
  1198       do_formula bs pos1 t1 ##>> do_formula bs pos2 t2 #>> uncurry (mk_aconn c)
  1199     and do_formula bs pos t =
  1200       case t of
  1201         @{const Trueprop} $ t1 => do_formula bs pos t1
  1202       | @{const Not} $ t1 => do_formula bs (Option.map not pos) t1 #>> mk_anot
  1203       | Const (@{const_name All}, _) $ Abs (s, T, t') =>
  1204         do_quant bs AForall pos s T t'
  1205       | (t0 as Const (@{const_name All}, _)) $ t1 =>
  1206         do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)
  1207       | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
  1208         do_quant bs AExists pos s T t'
  1209       | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
  1210         do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)
  1211       | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd pos t1 pos t2
  1212       | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr pos t1 pos t2
  1213       | @{const HOL.implies} $ t1 $ t2 =>
  1214         do_conn bs AImplies (Option.map not pos) t1 pos t2
  1215       | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
  1216         if iff_for_eq then do_conn bs AIff NONE t1 NONE t2 else do_term bs t
  1217       | _ => do_term bs t
  1218   in do_formula [] end
  1219 
  1220 fun presimplify_term thy t =
  1221   if exists_Const (member (op =) Meson.presimplified_consts o fst) t then
  1222     t |> Skip_Proof.make_thm thy
  1223       |> Meson.presimplify
  1224       |> prop_of
  1225   else
  1226     t
  1227 
  1228 fun preprocess_abstractions_in_terms trans_lams facts =
  1229   let
  1230     val (facts, lambda_ts) =
  1231       facts |> map (snd o snd) |> trans_lams
  1232             |>> map2 (fn (name, (role, _)) => fn t => (name, (role, t))) facts
  1233     val lam_facts =
  1234       map2 (fn t => fn j =>
  1235                ((lam_fact_prefix ^ Int.toString j, (Global, Def)), (Axiom, t)))
  1236            lambda_ts (1 upto length lambda_ts)
  1237   in (facts, lam_facts) end
  1238 
  1239 (* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
  1240    same in Sledgehammer to prevent the discovery of unreplayable proofs. *)
  1241 fun freeze_term t =
  1242   let
  1243     fun freeze (t $ u) = freeze t $ freeze u
  1244       | freeze (Abs (s, T, t)) = Abs (s, T, freeze t)
  1245       | freeze (Var ((s, i), T)) =
  1246         Free (atp_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
  1247       | freeze t = t
  1248   in t |> exists_subterm is_Var t ? freeze end
  1249 
  1250 fun presimp_prop ctxt type_enc t =
  1251   let
  1252     val thy = Proof_Context.theory_of ctxt
  1253     val t = t |> Envir.beta_eta_contract
  1254               |> transform_elim_prop
  1255               |> Object_Logic.atomize_term thy
  1256     val need_trueprop = (fastype_of t = @{typ bool})
  1257     val is_ho = is_type_enc_higher_order type_enc
  1258   in
  1259     t |> need_trueprop ? HOLogic.mk_Trueprop
  1260       |> (if is_ho then unextensionalize_def
  1261           else cong_extensionalize_term thy #> abs_extensionalize_term ctxt)
  1262       |> presimplify_term thy
  1263       |> HOLogic.dest_Trueprop
  1264   end
  1265   handle TERM _ => @{const True}
  1266 
  1267 (* Satallax prefers "=" to "<=>" (for definitions) and Metis (CNF) requires "="
  1268    for obscure technical reasons. *)
  1269 fun should_use_iff_for_eq CNF _ = false
  1270   | should_use_iff_for_eq (THF _) format = not (is_type_enc_higher_order format)
  1271   | should_use_iff_for_eq _ _ = true
  1272 
  1273 fun make_formula ctxt format type_enc iff_for_eq name stature role t =
  1274   let
  1275     val iff_for_eq = iff_for_eq andalso should_use_iff_for_eq format type_enc
  1276     val (iformula, atomic_Ts) =
  1277       iformula_from_prop ctxt type_enc iff_for_eq (SOME (role <> Conjecture)) t
  1278                          []
  1279       |>> close_iformula_universally
  1280   in
  1281     {name = name, stature = stature, role = role, iformula = iformula,
  1282      atomic_types = atomic_Ts}
  1283   end
  1284 
  1285 fun is_format_with_defs (THF (_, _, _, THF_With_Defs)) = true
  1286   | is_format_with_defs _ = false
  1287 
  1288 fun make_fact ctxt format type_enc iff_for_eq
  1289               ((name, stature as (_, status)), t) =
  1290   let
  1291     val role =
  1292       if is_format_with_defs format andalso status = Def andalso
  1293          is_legitimate_tptp_def t then
  1294         Definition
  1295       else
  1296         Axiom
  1297   in
  1298     case t |> make_formula ctxt format type_enc iff_for_eq name stature role of
  1299       formula as {iformula = AAtom (IConst ((s, _), _, _)), ...} =>
  1300       if s = tptp_true then NONE else SOME formula
  1301     | formula => SOME formula
  1302   end
  1303 
  1304 fun s_not_prop (@{const Trueprop} $ t) = @{const Trueprop} $ s_not t
  1305   | s_not_prop (@{const "==>"} $ t $ @{prop False}) = t
  1306   | s_not_prop t = @{const "==>"} $ t $ @{prop False}
  1307 
  1308 fun make_conjecture ctxt format type_enc =
  1309   map (fn ((name, stature), (role, t)) =>
  1310           let
  1311             (* FIXME: The commented-out code is a hack to get decent performance
  1312                out of LEO-II on the TPTP THF benchmarks. *)
  1313             val role =
  1314               if (* is_format_with_defs format andalso *)
  1315                  role <> Conjecture andalso is_legitimate_tptp_def t then
  1316                 Definition
  1317               else
  1318                 role
  1319           in
  1320             t |> role = Conjecture ? s_not
  1321               |> make_formula ctxt format type_enc true name stature role
  1322           end)
  1323 
  1324 (** Finite and infinite type inference **)
  1325 
  1326 fun tvar_footprint thy s ary =
  1327   (case unprefix_and_unascii const_prefix s of
  1328      SOME s =>
  1329      s |> invert_const |> robust_const_type thy |> chop_fun ary |> fst
  1330        |> map (fn T => Term.add_tvarsT T [] |> map fst)
  1331    | NONE => [])
  1332   handle TYPE _ => []
  1333 
  1334 fun type_arg_cover thy s ary =
  1335   if is_tptp_equal s then
  1336     0 upto ary - 1
  1337   else
  1338     let
  1339       val footprint = tvar_footprint thy s ary
  1340       val eq = (s = @{const_name HOL.eq})
  1341       fun cover _ [] = []
  1342         | cover seen ((i, tvars) :: args) =
  1343           cover (union (op =) seen tvars) args
  1344           |> (eq orelse exists (fn tvar => not (member (op =) seen tvar)) tvars)
  1345              ? cons i
  1346     in
  1347       if forall null footprint then
  1348         []
  1349       else
  1350         0 upto length footprint - 1 ~~ footprint
  1351         |> sort (rev_order o list_ord Term_Ord.indexname_ord o pairself snd)
  1352         |> cover []
  1353     end
  1354 
  1355 type monotonicity_info =
  1356   {maybe_finite_Ts : typ list,
  1357    surely_finite_Ts : typ list,
  1358    maybe_infinite_Ts : typ list,
  1359    surely_infinite_Ts : typ list,
  1360    maybe_nonmono_Ts : typ list}
  1361 
  1362 (* These types witness that the type classes they belong to allow infinite
  1363    models and hence that any types with these type classes is monotonic. *)
  1364 val known_infinite_types =
  1365   [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}]
  1366 
  1367 fun is_type_kind_of_surely_infinite ctxt strictness cached_Ts T =
  1368   strictness <> Strict andalso is_type_surely_infinite ctxt true cached_Ts T
  1369 
  1370 (* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are
  1371    dangerous because their "exhaust" properties can easily lead to unsound ATP
  1372    proofs. On the other hand, all HOL infinite types can be given the same
  1373    models in first-order logic (via Löwenheim-Skolem). *)
  1374 
  1375 fun should_encode_type _ (_ : monotonicity_info) All_Types _ = true
  1376   | should_encode_type ctxt {maybe_finite_Ts, surely_infinite_Ts,
  1377                              maybe_nonmono_Ts, ...}
  1378                        (Noninf_Nonmono_Types (strictness, grain)) T =
  1379     let val thy = Proof_Context.theory_of ctxt in
  1380       grain = Ghost_Type_Arg_Vars orelse
  1381       (exists (type_intersect thy T) maybe_nonmono_Ts andalso
  1382        not (exists (type_instance thy T) surely_infinite_Ts orelse
  1383             (not (member (type_equiv thy) maybe_finite_Ts T) andalso
  1384              is_type_kind_of_surely_infinite ctxt strictness surely_infinite_Ts
  1385                                              T)))
  1386     end
  1387   | should_encode_type ctxt {surely_finite_Ts, maybe_infinite_Ts,
  1388                              maybe_nonmono_Ts, ...}
  1389                        (Fin_Nonmono_Types grain) T =
  1390     let val thy = Proof_Context.theory_of ctxt in
  1391       grain = Ghost_Type_Arg_Vars orelse
  1392       (exists (type_intersect thy T) maybe_nonmono_Ts andalso
  1393        (exists (type_generalization thy T) surely_finite_Ts orelse
  1394         (not (member (type_equiv thy) maybe_infinite_Ts T) andalso
  1395          is_type_surely_finite ctxt T)))
  1396     end
  1397   | should_encode_type _ _ _ _ = false
  1398 
  1399 fun should_guard_type ctxt mono (Guards (_, level)) should_guard_var T =
  1400     should_guard_var () andalso should_encode_type ctxt mono level T
  1401   | should_guard_type _ _ _ _ _ = false
  1402 
  1403 fun is_maybe_universal_var (IConst ((s, _), _, _)) =
  1404     String.isPrefix bound_var_prefix s orelse
  1405     String.isPrefix all_bound_var_prefix s
  1406   | is_maybe_universal_var (IVar _) = true
  1407   | is_maybe_universal_var _ = false
  1408 
  1409 datatype site =
  1410   Top_Level of bool option |
  1411   Eq_Arg of bool option |
  1412   Elsewhere
  1413 
  1414 fun should_tag_with_type _ _ _ (Top_Level _) _ _ = false
  1415   | should_tag_with_type ctxt mono (Tags (_, level)) site u T =
  1416     if granularity_of_type_level level = All_Vars then
  1417       should_encode_type ctxt mono level T
  1418     else
  1419       (case (site, is_maybe_universal_var u) of
  1420          (Eq_Arg _, true) => should_encode_type ctxt mono level T
  1421        | _ => false)
  1422   | should_tag_with_type _ _ _ _ _ _ = false
  1423 
  1424 fun fused_type ctxt mono level =
  1425   let
  1426     val should_encode = should_encode_type ctxt mono level
  1427     fun fuse 0 T = if should_encode T then T else fused_infinite_type
  1428       | fuse ary (Type (@{type_name fun}, [T1, T2])) =
  1429         fuse 0 T1 --> fuse (ary - 1) T2
  1430       | fuse _ _ = raise Fail "expected function type"
  1431   in fuse end
  1432 
  1433 (** predicators and application operators **)
  1434 
  1435 type sym_info =
  1436   {pred_sym : bool, min_ary : int, max_ary : int, types : typ list,
  1437    in_conj : bool}
  1438 
  1439 fun default_sym_tab_entries type_enc =
  1440   (make_fixed_const NONE @{const_name undefined},
  1441        {pred_sym = false, min_ary = 0, max_ary = 0, types = [],
  1442         in_conj = false}) ::
  1443   ([tptp_false, tptp_true]
  1444    |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, types = [],
  1445                   in_conj = false})) @
  1446   ([tptp_equal, tptp_old_equal]
  1447    |> map (rpair {pred_sym = true, min_ary = 2, max_ary = 2, types = [],
  1448                   in_conj = false}))
  1449   |> not (is_type_enc_higher_order type_enc)
  1450      ? cons (prefixed_predicator_name,
  1451              {pred_sym = true, min_ary = 1, max_ary = 1, types = [],
  1452               in_conj = false})
  1453 
  1454 datatype app_op_level =
  1455   Min_App_Op |
  1456   Sufficient_App_Op |
  1457   Sufficient_App_Op_And_Predicator |
  1458   Full_App_Op_And_Predicator
  1459 
  1460 fun add_iterm_syms_to_sym_table ctxt app_op_level conj_fact =
  1461   let
  1462     val thy = Proof_Context.theory_of ctxt
  1463     fun consider_var_ary const_T var_T max_ary =
  1464       let
  1465         fun iter ary T =
  1466           if ary = max_ary orelse type_instance thy var_T T orelse
  1467              type_instance thy T var_T then
  1468             ary
  1469           else
  1470             iter (ary + 1) (range_type T)
  1471       in iter 0 const_T end
  1472     fun add_universal_var T (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
  1473       if (app_op_level = Sufficient_App_Op andalso can dest_funT T) orelse
  1474          (app_op_level = Sufficient_App_Op_And_Predicator andalso
  1475           (can dest_funT T orelse T = @{typ bool})) then
  1476         let
  1477           val bool_vars' =
  1478             bool_vars orelse
  1479             (app_op_level = Sufficient_App_Op_And_Predicator andalso
  1480              body_type T = @{typ bool})
  1481           fun repair_min_ary {pred_sym, min_ary, max_ary, types, in_conj} =
  1482             {pred_sym = pred_sym andalso not bool_vars',
  1483              min_ary = fold (fn T' => consider_var_ary T' T) types min_ary,
  1484              max_ary = max_ary, types = types, in_conj = in_conj}
  1485           val fun_var_Ts' =
  1486             fun_var_Ts |> can dest_funT T ? insert_type thy I T
  1487         in
  1488           if bool_vars' = bool_vars andalso
  1489              pointer_eq (fun_var_Ts', fun_var_Ts) then
  1490             accum
  1491           else
  1492             ((bool_vars', fun_var_Ts'), Symtab.map (K repair_min_ary) sym_tab)
  1493         end
  1494       else
  1495         accum
  1496     fun add_iterm_syms top_level tm
  1497                        (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
  1498       let val (head, args) = strip_iterm_comb tm in
  1499         (case head of
  1500            IConst ((s, _), T, _) =>
  1501            if String.isPrefix bound_var_prefix s orelse
  1502               String.isPrefix all_bound_var_prefix s then
  1503              add_universal_var T accum
  1504            else if String.isPrefix exist_bound_var_prefix s then
  1505              accum
  1506            else
  1507              let val ary = length args in
  1508                ((bool_vars, fun_var_Ts),
  1509                 case Symtab.lookup sym_tab s of
  1510                   SOME {pred_sym, min_ary, max_ary, types, in_conj} =>
  1511                   let
  1512                     val pred_sym =
  1513                       pred_sym andalso top_level andalso not bool_vars
  1514                     val types' = types |> insert_type thy I T
  1515                     val in_conj = in_conj orelse conj_fact
  1516                     val min_ary =
  1517                       if (app_op_level = Sufficient_App_Op orelse
  1518                           app_op_level = Sufficient_App_Op_And_Predicator)
  1519                          andalso not (pointer_eq (types', types)) then
  1520                         fold (consider_var_ary T) fun_var_Ts min_ary
  1521                       else
  1522                         min_ary
  1523                   in
  1524                     Symtab.update (s, {pred_sym = pred_sym,
  1525                                        min_ary = Int.min (ary, min_ary),
  1526                                        max_ary = Int.max (ary, max_ary),
  1527                                        types = types', in_conj = in_conj})
  1528                                   sym_tab
  1529                   end
  1530                 | NONE =>
  1531                   let
  1532                     val pred_sym = top_level andalso not bool_vars
  1533                     val ary =
  1534                       case unprefix_and_unascii const_prefix s of
  1535                         SOME s =>
  1536                         (if String.isSubstring uncurried_alias_sep s then
  1537                            ary
  1538                          else case try (robust_const_ary thy
  1539                                         o invert_const o hd
  1540                                         o unmangled_const_name) s of
  1541                            SOME ary0 => Int.min (ary0, ary)
  1542                          | NONE => ary)
  1543                       | NONE => ary
  1544                     val min_ary =
  1545                       case app_op_level of
  1546                         Min_App_Op => ary
  1547                       | Full_App_Op_And_Predicator => 0
  1548                       | _ => fold (consider_var_ary T) fun_var_Ts ary
  1549                   in
  1550                     Symtab.update_new (s,
  1551                         {pred_sym = pred_sym, min_ary = min_ary,
  1552                          max_ary = ary, types = [T], in_conj = conj_fact})
  1553                         sym_tab
  1554                   end)
  1555              end
  1556          | IVar (_, T) => add_universal_var T accum
  1557          | IAbs ((_, T), tm) =>
  1558            accum |> add_universal_var T |> add_iterm_syms false tm
  1559          | _ => accum)
  1560         |> fold (add_iterm_syms false) args
  1561       end
  1562   in add_iterm_syms end
  1563 
  1564 fun sym_table_for_facts ctxt type_enc app_op_level conjs facts =
  1565   let
  1566     fun add_iterm_syms conj_fact =
  1567       add_iterm_syms_to_sym_table ctxt app_op_level conj_fact true
  1568     fun add_fact_syms conj_fact =
  1569       K (add_iterm_syms conj_fact) |> formula_fold NONE |> ifact_lift
  1570   in
  1571     ((false, []), Symtab.empty)
  1572     |> fold (add_fact_syms true) conjs
  1573     |> fold (add_fact_syms false) facts
  1574     ||> fold Symtab.update (default_sym_tab_entries type_enc)
  1575   end
  1576 
  1577 fun min_ary_of sym_tab s =
  1578   case Symtab.lookup sym_tab s of
  1579     SOME ({min_ary, ...} : sym_info) => min_ary
  1580   | NONE =>
  1581     case unprefix_and_unascii const_prefix s of
  1582       SOME s =>
  1583       let val s = s |> unmangled_const_name |> hd |> invert_const in
  1584         if s = predicator_name then 1
  1585         else if s = app_op_name then 2
  1586         else if s = type_guard_name then 1
  1587         else 0
  1588       end
  1589     | NONE => 0
  1590 
  1591 (* True if the constant ever appears outside of the top-level position in
  1592    literals, or if it appears with different arities (e.g., because of different
  1593    type instantiations). If false, the constant always receives all of its
  1594    arguments and is used as a predicate. *)
  1595 fun is_pred_sym sym_tab s =
  1596   case Symtab.lookup sym_tab s of
  1597     SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
  1598     pred_sym andalso min_ary = max_ary
  1599   | NONE => false
  1600 
  1601 val fTrue_iconst =
  1602   IConst ((const_prefix ^ "fTrue", @{const_name ATP.fTrue}), @{typ bool}, [])
  1603 val predicator_iconst =
  1604   IConst (`(make_fixed_const NONE) predicator_name, @{typ "bool => bool"}, [])
  1605 
  1606 fun predicatify aggressive tm =
  1607   if aggressive then
  1608     IApp (IApp (IConst (`I tptp_equal, @{typ "bool => bool => bool"}, []), tm),
  1609           fTrue_iconst)
  1610   else
  1611     IApp (predicator_iconst, tm)
  1612 
  1613 val app_op = `(make_fixed_const NONE) app_op_name
  1614 
  1615 fun list_app head args = fold (curry (IApp o swap)) args head
  1616 
  1617 fun mk_app_op type_enc head arg =
  1618   let
  1619     val head_T = ityp_of head
  1620     val (arg_T, res_T) = dest_funT head_T
  1621     val app =
  1622       IConst (app_op, head_T --> head_T, [arg_T, res_T])
  1623       |> mangle_type_args_in_iterm type_enc
  1624   in list_app app [head, arg] end
  1625 
  1626 fun firstorderize_fact thy monom_constrs type_enc sym_tab uncurried_aliases
  1627                        aggressive =
  1628   let
  1629     fun do_app arg head = mk_app_op type_enc head arg
  1630     fun list_app_ops head args = fold do_app args head
  1631     fun introduce_app_ops tm =
  1632       let val (head, args) = tm |> strip_iterm_comb ||> map introduce_app_ops in
  1633         case head of
  1634           IConst (name as (s, _), T, T_args) =>
  1635           if uncurried_aliases andalso String.isPrefix const_prefix s then
  1636             let
  1637               val ary = length args
  1638               val name =
  1639                 name |> ary > min_ary_of sym_tab s ? aliased_uncurried ary
  1640             in list_app (IConst (name, T, T_args)) args end
  1641           else
  1642             args |> chop (min_ary_of sym_tab s)
  1643                  |>> list_app head |-> list_app_ops
  1644         | _ => list_app_ops head args
  1645       end
  1646     fun introduce_predicators tm =
  1647       case strip_iterm_comb tm of
  1648         (IConst ((s, _), _, _), _) =>
  1649         if is_pred_sym sym_tab s then tm else predicatify aggressive tm
  1650       | _ => predicatify aggressive tm
  1651     val do_iterm =
  1652       not (is_type_enc_higher_order type_enc)
  1653       ? (introduce_app_ops #> introduce_predicators)
  1654       #> filter_type_args_in_iterm thy monom_constrs type_enc
  1655   in update_iformula (formula_map do_iterm) end
  1656 
  1657 (** Helper facts **)
  1658 
  1659 val not_ffalse = @{lemma "~ fFalse" by (unfold fFalse_def) fast}
  1660 val ftrue = @{lemma "fTrue" by (unfold fTrue_def) fast}
  1661 
  1662 (* The Boolean indicates that a fairly sound type encoding is needed. *)
  1663 val base_helper_table =
  1664   [(("COMBI", false), [(Def, @{thm Meson.COMBI_def})]),
  1665    (("COMBK", false), [(Def, @{thm Meson.COMBK_def})]),
  1666    (("COMBB", false), [(Def, @{thm Meson.COMBB_def})]),
  1667    (("COMBC", false), [(Def, @{thm Meson.COMBC_def})]),
  1668    (("COMBS", false), [(Def, @{thm Meson.COMBS_def})]),
  1669    ((predicator_name, false), [(General, not_ffalse), (General, ftrue)]),
  1670    (("fFalse", false), [(General, not_ffalse)]),
  1671    (("fFalse", true), [(General, @{thm True_or_False})]),
  1672    (("fTrue", false), [(General, ftrue)]),
  1673    (("fTrue", true), [(General, @{thm True_or_False})]),
  1674    (("If", true),
  1675     [(Def, @{thm if_True}), (Def, @{thm if_False}),
  1676      (General, @{thm True_or_False})])]
  1677 
  1678 val helper_table =
  1679   base_helper_table @
  1680   [(("fNot", false),
  1681     @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1]
  1682            fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}
  1683     |> map (pair Def)),
  1684    (("fconj", false),
  1685     @{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q"
  1686         by (unfold fconj_def) fast+}
  1687     |> map (pair General)),
  1688    (("fdisj", false),
  1689     @{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q"
  1690         by (unfold fdisj_def) fast+}
  1691     |> map (pair General)),
  1692    (("fimplies", false),
  1693     @{lemma "P | fimplies P Q" "~ Q | fimplies P Q" "~ fimplies P Q | ~ P | Q"
  1694         by (unfold fimplies_def) fast+}
  1695     |> map (pair General)),
  1696    (("fequal", true),
  1697     (* This is a lie: Higher-order equality doesn't need a sound type encoding.
  1698        However, this is done so for backward compatibility: Including the
  1699        equality helpers by default in Metis breaks a few existing proofs. *)
  1700     @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1]
  1701            fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]}
  1702     |> map (pair General)),
  1703    (* Partial characterization of "fAll" and "fEx". A complete characterization
  1704       would require the axiom of choice for replay with Metis. *)
  1705    (("fAll", false),
  1706     [(General, @{lemma "~ fAll P | P x" by (auto simp: fAll_def)})]),
  1707    (("fEx", false),
  1708     [(General, @{lemma "~ P x | fEx P" by (auto simp: fEx_def)})])]
  1709   |> map (apsnd (map (apsnd zero_var_indexes)))
  1710 
  1711 val aggressive_helper_table =
  1712   base_helper_table @
  1713   [((predicator_name, true),
  1714     @{thms True_or_False fTrue_ne_fFalse} |> map (pair General)),
  1715    ((app_op_name, true),
  1716     [(General, @{lemma "EX x. ~ f x = g x | f = g" by blast})]),
  1717    (("fconj", false),
  1718     @{thms fconj_table fconj_laws fdisj_laws} |> map (pair Def)),
  1719    (("fdisj", false),
  1720     @{thms fdisj_table fconj_laws fdisj_laws} |> map (pair Def)),
  1721    (("fimplies", false),
  1722     @{thms fimplies_table fconj_laws fdisj_laws fimplies_laws}
  1723     |> map (pair Def)),
  1724    (("fequal", false),
  1725     (@{thms fequal_table} |> map (pair Def)) @
  1726     (@{thms fequal_laws} |> map (pair General))),
  1727    (("fAll", false),
  1728     @{thms fAll_table fComp_law fAll_law fEx_law} |> map (pair Def)),
  1729    (("fEx", false),
  1730     @{thms fEx_table fComp_law fAll_law fEx_law} |> map (pair Def))]
  1731   |> map (apsnd (map (apsnd zero_var_indexes)))
  1732 
  1733 fun atype_of_type_vars (Native (_, Polymorphic, _)) = SOME atype_of_types
  1734   | atype_of_type_vars _ = NONE
  1735 
  1736 fun bound_tvars type_enc sorts Ts =
  1737   (sorts ? mk_ahorn (formulas_for_types type_enc add_sorts_on_tvar Ts))
  1738   #> mk_aquant AForall
  1739         (map_filter (fn TVar (x as (s, _), _) =>
  1740                         SOME ((make_schematic_type_var x, s),
  1741                               atype_of_type_vars type_enc)
  1742                       | _ => NONE) Ts)
  1743 
  1744 fun eq_formula type_enc atomic_Ts bounds pred_sym tm1 tm2 =
  1745   (if pred_sym then AConn (AIff, [AAtom tm1, AAtom tm2])
  1746    else AAtom (ATerm (`I tptp_equal, [tm1, tm2])))
  1747   |> mk_aquant AForall bounds
  1748   |> close_formula_universally
  1749   |> bound_tvars type_enc true atomic_Ts
  1750 
  1751 val helper_rank = default_rank
  1752 val min_rank = 9 * helper_rank div 10
  1753 val max_rank = 4 * min_rank
  1754 
  1755 fun rank_of_fact_num n j = min_rank + (max_rank - min_rank) * j div n
  1756 
  1757 val type_tag = `(make_fixed_const NONE) type_tag_name
  1758 
  1759 fun could_specialize_helpers type_enc =
  1760   polymorphism_of_type_enc type_enc <> Polymorphic andalso
  1761   level_of_type_enc type_enc <> No_Types
  1762 fun should_specialize_helper type_enc t =
  1763   could_specialize_helpers type_enc andalso
  1764   not (null (Term.hidden_polymorphism t))
  1765 
  1766 fun add_helper_facts_for_sym ctxt format type_enc aggressive
  1767                              (s, {types, ...} : sym_info) =
  1768   case unprefix_and_unascii const_prefix s of
  1769     SOME mangled_s =>
  1770     let
  1771       val thy = Proof_Context.theory_of ctxt
  1772       val unmangled_s = mangled_s |> unmangled_const_name |> hd
  1773       fun dub needs_fairly_sound j k =
  1774         ascii_of unmangled_s ^ "_" ^ string_of_int j ^ "_" ^ string_of_int k ^
  1775         (if mangled_s = unmangled_s then "" else "_" ^ ascii_of mangled_s) ^
  1776         (if needs_fairly_sound then typed_helper_suffix
  1777          else untyped_helper_suffix)
  1778       fun specialize_helper t T =
  1779         if unmangled_s = app_op_name then
  1780           let
  1781             val tyenv =
  1782               Sign.typ_match thy (alpha_to_beta, domain_type T) Vartab.empty
  1783           in monomorphic_term tyenv t end
  1784         else
  1785           specialize_type thy (invert_const unmangled_s, T) t
  1786       fun dub_and_inst needs_fairly_sound ((status, t), j) =
  1787         (if should_specialize_helper type_enc t then
  1788            map_filter (try (specialize_helper t)) types
  1789          else
  1790            [t])
  1791         |> tag_list 1
  1792         |> map (fn (k, t) =>
  1793                    ((dub needs_fairly_sound j k, (Global, status)), t))
  1794       val make_facts = map_filter (make_fact ctxt format type_enc false)
  1795       val fairly_sound = is_type_enc_fairly_sound type_enc
  1796       val could_specialize = could_specialize_helpers type_enc
  1797     in
  1798       fold (fn ((helper_s, needs_fairly_sound), ths) =>
  1799                if (needs_fairly_sound andalso not fairly_sound) orelse
  1800                   (helper_s <> unmangled_s andalso
  1801                    (not aggressive orelse could_specialize)) then
  1802                  I
  1803                else
  1804                  ths ~~ (1 upto length ths)
  1805                  |> maps (dub_and_inst needs_fairly_sound
  1806                           o apfst (apsnd prop_of))
  1807                  |> make_facts
  1808                  |> union (op = o pairself #iformula))
  1809            (if aggressive then aggressive_helper_table else helper_table)
  1810     end
  1811   | NONE => I
  1812 fun helper_facts_for_sym_table ctxt format type_enc aggressive sym_tab =
  1813   Symtab.fold_rev (add_helper_facts_for_sym ctxt format type_enc aggressive)
  1814                   sym_tab []
  1815 
  1816 (***************************************************************)
  1817 (* Type Classes Present in the Axiom or Conjecture Clauses     *)
  1818 (***************************************************************)
  1819 
  1820 fun set_insert (x, s) = Symtab.update (x, ()) s
  1821 
  1822 fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
  1823 
  1824 (* Remove this trivial type class (FIXME: similar code elsewhere) *)
  1825 fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset
  1826 
  1827 fun classes_of_terms get_Ts =
  1828   map (map snd o get_Ts)
  1829   #> List.foldl add_classes Symtab.empty
  1830   #> delete_type #> Symtab.keys
  1831 
  1832 val tfree_classes_of_terms = classes_of_terms Misc_Legacy.term_tfrees
  1833 val tvar_classes_of_terms = classes_of_terms Misc_Legacy.term_tvars
  1834 
  1835 fun fold_type_constrs f (Type (s, Ts)) x =
  1836     fold (fold_type_constrs f) Ts (f (s, x))
  1837   | fold_type_constrs _ _ x = x
  1838 
  1839 (* Type constructors used to instantiate overloaded constants are the only ones
  1840    needed. *)
  1841 fun add_type_constrs_in_term thy =
  1842   let
  1843     fun add (Const (@{const_name Meson.skolem}, _) $ _) = I
  1844       | add (t $ u) = add t #> add u
  1845       | add (Const x) =
  1846         x |> robust_const_typargs thy |> fold (fold_type_constrs set_insert)
  1847       | add (Abs (_, _, u)) = add u
  1848       | add _ = I
  1849   in add end
  1850 
  1851 fun type_constrs_of_terms thy ts =
  1852   Symtab.keys (fold (add_type_constrs_in_term thy) ts Symtab.empty)
  1853 
  1854 fun extract_lambda_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
  1855     let val (head, args) = strip_comb t in
  1856       (head |> dest_Const |> fst,
  1857        fold_rev (fn t as Var ((s, _), T) =>
  1858                     (fn u => Abs (s, T, abstract_over (t, u)))
  1859                   | _ => raise Fail "expected \"Var\"") args u)
  1860     end
  1861   | extract_lambda_def _ = raise Fail "malformed lifted lambda"
  1862 
  1863 fun trans_lams_from_string ctxt type_enc lam_trans =
  1864   if lam_trans = no_lamsN then
  1865     rpair []
  1866   else if lam_trans = hide_lamsN then
  1867     lift_lams ctxt type_enc ##> K []
  1868   else if lam_trans = liftingN orelse lam_trans = lam_liftingN then
  1869     lift_lams ctxt type_enc
  1870   else if lam_trans = combsN then
  1871     map (introduce_combinators ctxt) #> rpair []
  1872   else if lam_trans = combs_and_liftingN then
  1873     lift_lams_part_1 ctxt type_enc
  1874     ##> maps (fn t => [t, introduce_combinators ctxt (intentionalize_def t)])
  1875     #> lift_lams_part_2 ctxt
  1876   else if lam_trans = combs_or_liftingN then
  1877     lift_lams_part_1 ctxt type_enc
  1878     ##> map (fn t => case head_of (strip_qnt_body @{const_name All} t) of
  1879                        @{term "op =::bool => bool => bool"} => t
  1880                      | _ => introduce_combinators ctxt (intentionalize_def t))
  1881     #> lift_lams_part_2 ctxt
  1882   else if lam_trans = keep_lamsN then
  1883     map (Envir.eta_contract) #> rpair []
  1884   else
  1885     error ("Unknown lambda translation scheme: " ^ quote lam_trans ^ ".")
  1886 
  1887 val pull_and_reorder_definitions =
  1888   let
  1889     fun add_consts (IApp (t, u)) = fold add_consts [t, u]
  1890       | add_consts (IAbs (_, t)) = add_consts t
  1891       | add_consts (IConst (name, _, _)) = insert (op =) name
  1892       | add_consts (IVar _) = I
  1893     fun consts_of_hs l_or_r ({iformula, ...} : ifact) =
  1894       case iformula of
  1895         AAtom (IApp (IApp (IConst _, t), u)) => add_consts (l_or_r (t, u)) []
  1896       | _ => []
  1897     (* Quadratic, but usually OK. *)
  1898     fun reorder [] [] = []
  1899       | reorder (fact :: skipped) [] =
  1900         fact :: reorder [] skipped (* break cycle *)
  1901       | reorder skipped (fact :: facts) =
  1902         let val rhs_consts = consts_of_hs snd fact in
  1903           if exists (exists (member (op =) rhs_consts o the_single
  1904                      o consts_of_hs fst))
  1905                     [skipped, facts] then
  1906             reorder (fact :: skipped) facts
  1907           else
  1908             fact :: reorder [] (facts @ skipped)
  1909         end
  1910   in List.partition (curry (op =) Definition o #role) #>> reorder [] #> op @ end
  1911 
  1912 fun translate_formulas ctxt prem_role format type_enc lam_trans presimp hyp_ts
  1913                        concl_t facts =
  1914   let
  1915     val thy = Proof_Context.theory_of ctxt
  1916     val trans_lams = trans_lams_from_string ctxt type_enc lam_trans
  1917     val fact_ts = facts |> map snd
  1918     (* Remove existing facts from the conjecture, as this can dramatically
  1919        boost an ATP's performance (for some reason). *)
  1920     val hyp_ts =
  1921       hyp_ts
  1922       |> map (fn t => if member (op aconv) fact_ts t then @{prop True} else t)
  1923     val facts = facts |> map (apsnd (pair Axiom))
  1924     val conjs =
  1925       map (pair prem_role) hyp_ts @ [(Conjecture, s_not_prop concl_t)]
  1926       |> map (apsnd freeze_term)
  1927       |> map2 (pair o rpair (Local, General) o string_of_int)
  1928               (0 upto length hyp_ts)
  1929     val ((conjs, facts), lam_facts) =
  1930       (conjs, facts)
  1931       |> presimp ? pairself (map (apsnd (apsnd (presimp_prop ctxt type_enc))))
  1932       |> (if lam_trans = no_lamsN then
  1933             rpair []
  1934           else
  1935             op @
  1936             #> preprocess_abstractions_in_terms trans_lams
  1937             #>> chop (length conjs))
  1938     val conjs =
  1939       conjs |> make_conjecture ctxt format type_enc
  1940             |> pull_and_reorder_definitions
  1941     val facts =
  1942       facts |> map_filter (fn (name, (_, t)) =>
  1943                               make_fact ctxt format type_enc true (name, t))
  1944             |> pull_and_reorder_definitions
  1945     val fact_names =
  1946       facts |> map (fn {name, stature, ...} : ifact => (name, stature))
  1947     val lifted = lam_facts |> map (extract_lambda_def o snd o snd)
  1948     val lam_facts =
  1949       lam_facts |> map_filter (make_fact ctxt format type_enc true o apsnd snd)
  1950     val all_ts = concl_t :: hyp_ts @ fact_ts
  1951     val subs = tfree_classes_of_terms all_ts
  1952     val supers = tvar_classes_of_terms all_ts
  1953     val tycons = type_constrs_of_terms thy all_ts
  1954     val (supers, arity_clauses) =
  1955       if level_of_type_enc type_enc = No_Types then ([], [])
  1956       else make_arity_clauses thy tycons supers
  1957     val class_rel_clauses = make_class_rel_clauses thy subs supers
  1958   in
  1959     (fact_names |> map single, union (op =) subs supers, conjs,
  1960      facts @ lam_facts, class_rel_clauses, arity_clauses, lifted)
  1961   end
  1962 
  1963 val type_guard = `(make_fixed_const NONE) type_guard_name
  1964 
  1965 fun type_guard_iterm type_enc T tm =
  1966   IApp (IConst (type_guard, T --> @{typ bool}, [T])
  1967         |> mangle_type_args_in_iterm type_enc, tm)
  1968 
  1969 fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum
  1970   | is_var_positively_naked_in_term name _ (ATerm ((s, _), tms)) accum =
  1971     accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm (name, [])))
  1972   | is_var_positively_naked_in_term _ _ _ _ = true
  1973 
  1974 fun is_var_ghost_type_arg_in_term thy name pos tm accum =
  1975   is_var_positively_naked_in_term name pos tm accum orelse
  1976   let
  1977     val var = ATerm (name, [])
  1978     fun is_nasty_in_term (ATerm (_, [])) = false
  1979       | is_nasty_in_term (ATerm ((s, _), tms)) =
  1980         let
  1981           val ary = length tms
  1982           val cover = type_arg_cover thy s ary
  1983         in
  1984           exists (fn (j, tm) => tm = var andalso member (op =) cover j)
  1985                  (0 upto ary - 1 ~~ tms) orelse
  1986           exists is_nasty_in_term tms
  1987         end
  1988       | is_nasty_in_term _ = true
  1989   in is_nasty_in_term tm end
  1990 
  1991 fun should_guard_var_in_formula thy level pos phi (SOME true) name =
  1992     (case granularity_of_type_level level of
  1993        All_Vars => true
  1994      | Positively_Naked_Vars =>
  1995        formula_fold pos (is_var_positively_naked_in_term name) phi false
  1996      | Ghost_Type_Arg_Vars =>
  1997        formula_fold pos (is_var_ghost_type_arg_in_term thy name) phi false)
  1998   | should_guard_var_in_formula _ _ _ _ _ _ = true
  1999 
  2000 fun always_guard_var_in_formula _ _ _ _ _ _ = true
  2001 
  2002 fun should_generate_tag_bound_decl _ _ _ (SOME true) _ = false
  2003   | should_generate_tag_bound_decl ctxt mono (Tags (_, level)) _ T =
  2004     granularity_of_type_level level <> All_Vars andalso
  2005     should_encode_type ctxt mono level T
  2006   | should_generate_tag_bound_decl _ _ _ _ _ = false
  2007 
  2008 fun mk_aterm type_enc name T_args args =
  2009   ATerm (name, map_filter (ho_term_for_type_arg type_enc) T_args @ args)
  2010 
  2011 fun do_bound_type ctxt mono type_enc =
  2012   case type_enc of
  2013     Native (_, _, level) =>
  2014     fused_type ctxt mono level 0 #> ho_type_from_typ type_enc false 0 #> SOME
  2015   | _ => K NONE
  2016 
  2017 fun tag_with_type ctxt mono type_enc pos T tm =
  2018   IConst (type_tag, T --> T, [T])
  2019   |> mangle_type_args_in_iterm type_enc
  2020   |> ho_term_from_iterm ctxt mono type_enc pos
  2021   |> (fn ATerm (s, tms) => ATerm (s, tms @ [tm])
  2022        | _ => raise Fail "unexpected lambda-abstraction")
  2023 and ho_term_from_iterm ctxt mono type_enc pos =
  2024   let
  2025     fun term site u =
  2026       let
  2027         val (head, args) = strip_iterm_comb u
  2028         val pos =
  2029           case site of
  2030             Top_Level pos => pos
  2031           | Eq_Arg pos => pos
  2032           | _ => NONE
  2033         val t =
  2034           case head of
  2035             IConst (name as (s, _), _, T_args) =>
  2036             let
  2037               val arg_site = if is_tptp_equal s then Eq_Arg pos else Elsewhere
  2038             in map (term arg_site) args |> mk_aterm type_enc name T_args end
  2039           | IVar (name, _) =>
  2040             map (term Elsewhere) args |> mk_aterm type_enc name []
  2041           | IAbs ((name, T), tm) =>
  2042             if is_type_enc_higher_order type_enc then
  2043               AAbs (((name, ho_type_from_typ type_enc true 0 T),
  2044                      term Elsewhere tm), map (term Elsewhere) args)
  2045             else
  2046               raise Fail "unexpected lambda-abstraction"
  2047           | IApp _ => raise Fail "impossible \"IApp\""
  2048         val T = ityp_of u
  2049       in
  2050         if should_tag_with_type ctxt mono type_enc site u T then
  2051           tag_with_type ctxt mono type_enc pos T t
  2052         else
  2053           t
  2054       end
  2055   in term (Top_Level pos) end
  2056 and formula_from_iformula ctxt mono type_enc should_guard_var =
  2057   let
  2058     val thy = Proof_Context.theory_of ctxt
  2059     val level = level_of_type_enc type_enc
  2060     val do_term = ho_term_from_iterm ctxt mono type_enc
  2061     fun do_out_of_bound_type pos phi universal (name, T) =
  2062       if should_guard_type ctxt mono type_enc
  2063              (fn () => should_guard_var thy level pos phi universal name) T then
  2064         IVar (name, T)
  2065         |> type_guard_iterm type_enc T
  2066         |> do_term pos |> AAtom |> SOME
  2067       else if should_generate_tag_bound_decl ctxt mono type_enc universal T then
  2068         let
  2069           val var = ATerm (name, [])
  2070           val tagged_var = tag_with_type ctxt mono type_enc pos T var
  2071         in SOME (AAtom (ATerm (`I tptp_equal, [tagged_var, var]))) end
  2072       else
  2073         NONE
  2074     fun do_formula pos (AQuant (q, xs, phi)) =
  2075         let
  2076           val phi = phi |> do_formula pos
  2077           val universal = Option.map (q = AExists ? not) pos
  2078           val do_bound_type = do_bound_type ctxt mono type_enc
  2079         in
  2080           AQuant (q, xs |> map (apsnd (fn NONE => NONE
  2081                                         | SOME T => do_bound_type T)),
  2082                   (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))
  2083                       (map_filter
  2084                            (fn (_, NONE) => NONE
  2085                              | (s, SOME T) =>
  2086                                do_out_of_bound_type pos phi universal (s, T))
  2087                            xs)
  2088                       phi)
  2089         end
  2090       | do_formula pos (AConn conn) = aconn_map pos do_formula conn
  2091       | do_formula pos (AAtom tm) = AAtom (do_term pos tm)
  2092   in do_formula end
  2093 
  2094 (* Each fact is given a unique fact number to avoid name clashes (e.g., because
  2095    of monomorphization). The TPTP explicitly forbids name clashes, and some of
  2096    the remote provers might care. *)
  2097 fun formula_line_for_fact ctxt prefix encode freshen pos mono type_enc rank
  2098                           (j, {name, stature, role, iformula, atomic_types}) =
  2099   (prefix ^ (if freshen then string_of_int j ^ "_" else "") ^ encode name, role,
  2100    iformula
  2101    |> formula_from_iformula ctxt mono type_enc should_guard_var_in_formula
  2102                             (if pos then SOME true else NONE)
  2103    |> close_formula_universally
  2104    |> bound_tvars type_enc true atomic_types,
  2105    NONE,
  2106    let val rank = rank j in
  2107      case snd stature of
  2108        Intro => isabelle_info introN rank
  2109      | Inductive => isabelle_info inductiveN rank
  2110      | Elim => isabelle_info elimN rank
  2111      | Simp => isabelle_info simpN rank
  2112      | Def => isabelle_info defN rank
  2113      | _ => isabelle_info "" rank
  2114    end)
  2115   |> Formula
  2116 
  2117 fun formula_line_for_class_rel_clause type_enc
  2118         ({name, subclass, superclass, ...} : class_rel_clause) =
  2119   let val ty_arg = ATerm (tvar_a_name, []) in
  2120     Formula (class_rel_clause_prefix ^ ascii_of name, Axiom,
  2121              AConn (AImplies,
  2122                     [type_class_formula type_enc subclass ty_arg,
  2123                      type_class_formula type_enc superclass ty_arg])
  2124              |> mk_aquant AForall
  2125                           [(tvar_a_name, atype_of_type_vars type_enc)],
  2126              NONE, isabelle_info inductiveN helper_rank)
  2127   end
  2128 
  2129 fun formula_from_arity_atom type_enc (class, t, args) =
  2130   ATerm (t, map (fn arg => ATerm (arg, [])) args)
  2131   |> type_class_formula type_enc class
  2132 
  2133 fun formula_line_for_arity_clause type_enc
  2134         ({name, prem_atoms, concl_atom} : arity_clause) =
  2135   Formula (arity_clause_prefix ^ name, Axiom,
  2136            mk_ahorn (map (formula_from_arity_atom type_enc) prem_atoms)
  2137                     (formula_from_arity_atom type_enc concl_atom)
  2138            |> mk_aquant AForall
  2139                   (map (rpair (atype_of_type_vars type_enc)) (#3 concl_atom)),
  2140            NONE, isabelle_info inductiveN helper_rank)
  2141 
  2142 fun formula_line_for_conjecture ctxt mono type_enc
  2143         ({name, role, iformula, atomic_types, ...} : ifact) =
  2144   Formula (conjecture_prefix ^ name, role,
  2145            iformula
  2146            |> formula_from_iformula ctxt mono type_enc
  2147                   should_guard_var_in_formula (SOME false)
  2148            |> close_formula_universally
  2149            |> bound_tvars type_enc true atomic_types, NONE, [])
  2150 
  2151 fun type_enc_needs_free_types (Native (_, Polymorphic, _)) = true
  2152   | type_enc_needs_free_types (Native _) = false
  2153   | type_enc_needs_free_types _ = true
  2154 
  2155 fun formula_line_for_free_type j phi =
  2156   Formula (tfree_clause_prefix ^ string_of_int j, Hypothesis, phi, NONE, [])
  2157 fun formula_lines_for_free_types type_enc (facts : ifact list) =
  2158   if type_enc_needs_free_types type_enc then
  2159     let
  2160       val phis =
  2161         fold (union (op =)) (map #atomic_types facts) []
  2162         |> formulas_for_types type_enc add_sorts_on_tfree
  2163     in map2 formula_line_for_free_type (0 upto length phis - 1) phis end
  2164   else
  2165     []
  2166 
  2167 (** Symbol declarations **)
  2168 
  2169 fun decl_line_for_class order s =
  2170   let val name as (s, _) = `make_type_class s in
  2171     Decl (sym_decl_prefix ^ s, name,
  2172           if order = First_Order then
  2173             ATyAbs ([tvar_a_name],
  2174                     if avoid_first_order_phantom_type_vars then
  2175                       AFun (a_itself_atype, bool_atype)
  2176                     else
  2177                       bool_atype)
  2178           else
  2179             AFun (atype_of_types, bool_atype))
  2180   end
  2181 
  2182 fun decl_lines_for_classes type_enc classes =
  2183   case type_enc of
  2184     Native (order, Polymorphic, _) => map (decl_line_for_class order) classes
  2185   | _ => []
  2186 
  2187 fun sym_decl_table_for_facts thy type_enc sym_tab (conjs, facts, extra_tms) =
  2188   let
  2189     fun add_iterm_syms tm =
  2190       let val (head, args) = strip_iterm_comb tm in
  2191         (case head of
  2192            IConst ((s, s'), T, T_args) =>
  2193            let
  2194              val (pred_sym, in_conj) =
  2195                case Symtab.lookup sym_tab s of
  2196                  SOME ({pred_sym, in_conj, ...} : sym_info) =>
  2197                  (pred_sym, in_conj)
  2198                | NONE => (false, false)
  2199              val decl_sym =
  2200                (case type_enc of
  2201                   Guards _ => not pred_sym
  2202                 | _ => true) andalso
  2203                is_tptp_user_symbol s
  2204            in
  2205              if decl_sym then
  2206                Symtab.map_default (s, [])
  2207                    (insert_type thy #3 (s', T_args, T, pred_sym, length args,
  2208                                         in_conj))
  2209              else
  2210                I
  2211            end
  2212          | IAbs (_, tm) => add_iterm_syms tm
  2213          | _ => I)
  2214         #> fold add_iterm_syms args
  2215       end
  2216     val add_fact_syms = K add_iterm_syms |> formula_fold NONE |> ifact_lift
  2217     fun add_formula_var_types (AQuant (_, xs, phi)) =
  2218         fold (fn (_, SOME T) => insert_type thy I T | _ => I) xs
  2219         #> add_formula_var_types phi
  2220       | add_formula_var_types (AConn (_, phis)) =
  2221         fold add_formula_var_types phis
  2222       | add_formula_var_types _ = I
  2223     fun var_types () =
  2224       if polymorphism_of_type_enc type_enc = Polymorphic then [tvar_a]
  2225       else fold (ifact_lift add_formula_var_types) (conjs @ facts) []
  2226     fun add_undefined_const T =
  2227       let
  2228         val (s, s') =
  2229           `(make_fixed_const NONE) @{const_name undefined}
  2230           |> (case type_arg_policy [] type_enc @{const_name undefined} of
  2231                 Mangled_Type_Args => mangled_const_name type_enc [T]
  2232               | _ => I)
  2233       in
  2234         Symtab.map_default (s, [])
  2235                            (insert_type thy #3 (s', [T], T, false, 0, false))
  2236       end
  2237     fun add_TYPE_const () =
  2238       let val (s, s') = TYPE_name in
  2239         Symtab.map_default (s, [])
  2240             (insert_type thy #3
  2241                          (s', [tvar_a], @{typ "'a itself"}, false, 0, false))
  2242       end
  2243   in
  2244     Symtab.empty
  2245     |> is_type_enc_fairly_sound type_enc
  2246        ? (fold (fold add_fact_syms) [conjs, facts]
  2247           #> fold add_iterm_syms extra_tms
  2248           #> (case type_enc of
  2249                 Native (First_Order, Polymorphic, _) =>
  2250                 if avoid_first_order_phantom_type_vars then add_TYPE_const ()
  2251                 else I
  2252               | Native _ => I
  2253               | _ => fold add_undefined_const (var_types ())))
  2254   end
  2255 
  2256 (* We add "bool" in case the helper "True_or_False" is included later. *)
  2257 fun default_mono level =
  2258   {maybe_finite_Ts = [@{typ bool}],
  2259    surely_finite_Ts = [@{typ bool}],
  2260    maybe_infinite_Ts = known_infinite_types,
  2261    surely_infinite_Ts =
  2262      case level of
  2263        Noninf_Nonmono_Types (Strict, _) => []
  2264      | _ => known_infinite_types,
  2265    maybe_nonmono_Ts = [@{typ bool}]}
  2266 
  2267 (* This inference is described in section 2.3 of Claessen et al.'s "Sorting it
  2268    out with monotonicity" paper presented at CADE 2011. *)
  2269 fun add_iterm_mononotonicity_info _ _ (SOME false) _ mono = mono
  2270   | add_iterm_mononotonicity_info ctxt level _
  2271         (IApp (IApp (IConst ((s, _), Type (_, [T, _]), _), tm1), tm2))
  2272         (mono as {maybe_finite_Ts, surely_finite_Ts, maybe_infinite_Ts,
  2273                   surely_infinite_Ts, maybe_nonmono_Ts}) =
  2274     let val thy = Proof_Context.theory_of ctxt in
  2275       if is_tptp_equal s andalso exists is_maybe_universal_var [tm1, tm2] then
  2276         case level of
  2277           Noninf_Nonmono_Types (strictness, _) =>
  2278           if exists (type_instance thy T) surely_infinite_Ts orelse
  2279              member (type_equiv thy) maybe_finite_Ts T then
  2280             mono
  2281           else if is_type_kind_of_surely_infinite ctxt strictness
  2282                                                   surely_infinite_Ts T then
  2283             {maybe_finite_Ts = maybe_finite_Ts,
  2284              surely_finite_Ts = surely_finite_Ts,
  2285              maybe_infinite_Ts = maybe_infinite_Ts,
  2286              surely_infinite_Ts = surely_infinite_Ts |> insert_type thy I T,
  2287              maybe_nonmono_Ts = maybe_nonmono_Ts}
  2288           else
  2289             {maybe_finite_Ts = maybe_finite_Ts |> insert (type_equiv thy) T,
  2290              surely_finite_Ts = surely_finite_Ts,
  2291              maybe_infinite_Ts = maybe_infinite_Ts,
  2292              surely_infinite_Ts = surely_infinite_Ts,
  2293              maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type thy I T}
  2294         | Fin_Nonmono_Types _ =>
  2295           if exists (type_instance thy T) surely_finite_Ts orelse
  2296              member (type_equiv thy) maybe_infinite_Ts T then
  2297             mono
  2298           else if is_type_surely_finite ctxt T then
  2299             {maybe_finite_Ts = maybe_finite_Ts,
  2300              surely_finite_Ts = surely_finite_Ts |> insert_type thy I T,
  2301              maybe_infinite_Ts = maybe_infinite_Ts,
  2302              surely_infinite_Ts = surely_infinite_Ts,
  2303              maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type thy I T}
  2304           else
  2305             {maybe_finite_Ts = maybe_finite_Ts,
  2306              surely_finite_Ts = surely_finite_Ts,
  2307              maybe_infinite_Ts = maybe_infinite_Ts |> insert (type_equiv thy) T,
  2308              surely_infinite_Ts = surely_infinite_Ts,
  2309              maybe_nonmono_Ts = maybe_nonmono_Ts}
  2310         | _ => mono
  2311       else
  2312         mono
  2313     end
  2314   | add_iterm_mononotonicity_info _ _ _ _ mono = mono
  2315 fun add_fact_mononotonicity_info ctxt level ({role, iformula, ...} : ifact) =
  2316   formula_fold (SOME (role <> Conjecture))
  2317                (add_iterm_mononotonicity_info ctxt level) iformula
  2318 fun mononotonicity_info_for_facts ctxt type_enc facts =
  2319   let val level = level_of_type_enc type_enc in
  2320     default_mono level
  2321     |> is_type_level_monotonicity_based level
  2322        ? fold (add_fact_mononotonicity_info ctxt level) facts
  2323   end
  2324 
  2325 fun add_iformula_monotonic_types ctxt mono type_enc =
  2326   let
  2327     val thy = Proof_Context.theory_of ctxt
  2328     val level = level_of_type_enc type_enc
  2329     val should_encode = should_encode_type ctxt mono level
  2330     fun add_type T = not (should_encode T) ? insert_type thy I T
  2331     fun add_args (IApp (tm1, tm2)) = add_args tm1 #> add_term tm2
  2332       | add_args _ = I
  2333     and add_term tm = add_type (ityp_of tm) #> add_args tm
  2334   in formula_fold NONE (K add_term) end
  2335 fun add_fact_monotonic_types ctxt mono type_enc =
  2336   add_iformula_monotonic_types ctxt mono type_enc |> ifact_lift
  2337 fun monotonic_types_for_facts ctxt mono type_enc facts =
  2338   let val level = level_of_type_enc type_enc in
  2339     [] |> (polymorphism_of_type_enc type_enc = Polymorphic andalso
  2340            is_type_level_monotonicity_based level andalso
  2341            granularity_of_type_level level <> Ghost_Type_Arg_Vars)
  2342           ? fold (add_fact_monotonic_types ctxt mono type_enc) facts
  2343   end
  2344 
  2345 fun formula_line_for_guards_mono_type ctxt mono type_enc T =
  2346   Formula (guards_sym_formula_prefix ^
  2347            ascii_of (mangled_type type_enc T),
  2348            Axiom,
  2349            IConst (`make_bound_var "X", T, [])
  2350            |> type_guard_iterm type_enc T
  2351            |> AAtom
  2352            |> formula_from_iformula ctxt mono type_enc
  2353                                     always_guard_var_in_formula (SOME true)
  2354            |> close_formula_universally
  2355            |> bound_tvars type_enc true (atomic_types_of T),
  2356            NONE, isabelle_info inductiveN helper_rank)
  2357 
  2358 fun formula_line_for_tags_mono_type ctxt mono type_enc T =
  2359   let val x_var = ATerm (`make_bound_var "X", []) in
  2360     Formula (tags_sym_formula_prefix ^
  2361              ascii_of (mangled_type type_enc T),
  2362              Axiom,
  2363              eq_formula type_enc (atomic_types_of T) [] false
  2364                   (tag_with_type ctxt mono type_enc NONE T x_var) x_var,
  2365              NONE, isabelle_info defN helper_rank)
  2366   end
  2367 
  2368 fun problem_lines_for_mono_types ctxt mono type_enc Ts =
  2369   case type_enc of
  2370     Native _ => []
  2371   | Guards _ => map (formula_line_for_guards_mono_type ctxt mono type_enc) Ts
  2372   | Tags _ => map (formula_line_for_tags_mono_type ctxt mono type_enc) Ts
  2373 
  2374 fun decl_line_for_sym ctxt mono type_enc s
  2375                       (s', T_args, T, pred_sym, ary, _) =
  2376   let
  2377     val thy = Proof_Context.theory_of ctxt
  2378     val (T, T_args) =
  2379       if null T_args then
  2380         (T, [])
  2381       else case unprefix_and_unascii const_prefix s of
  2382         SOME s' =>
  2383         let
  2384           val s' = s' |> invert_const
  2385           val T = s' |> robust_const_type thy
  2386         in (T, robust_const_typargs thy (s', T)) end
  2387       | NONE => raise Fail "unexpected type arguments"
  2388   in
  2389     Decl (sym_decl_prefix ^ s, (s, s'),
  2390           T |> fused_type ctxt mono (level_of_type_enc type_enc) ary
  2391             |> ho_type_from_typ type_enc pred_sym ary
  2392             |> not (null T_args)
  2393                ? curry ATyAbs (map (tvar_name o fst o dest_TVar) T_args))
  2394   end
  2395 
  2396 fun honor_conj_sym_role in_conj =
  2397   if in_conj then (Hypothesis, I) else (Axiom, I)
  2398 
  2399 fun formula_line_for_guards_sym_decl ctxt mono type_enc n s j
  2400                                      (s', T_args, T, _, ary, in_conj) =
  2401   let
  2402     val thy = Proof_Context.theory_of ctxt
  2403     val (role, maybe_negate) = honor_conj_sym_role in_conj
  2404     val (arg_Ts, res_T) = chop_fun ary T
  2405     val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)
  2406     val bounds =
  2407       bound_names ~~ arg_Ts |> map (fn (name, T) => IConst (name, T, []))
  2408     val bound_Ts =
  2409       if exists (curry (op =) dummyT) T_args then
  2410         let val cover = type_arg_cover thy s ary in
  2411           map2 (fn j => if member (op =) cover j then SOME else K NONE)
  2412                (0 upto ary - 1) arg_Ts
  2413         end
  2414       else
  2415         replicate ary NONE
  2416   in
  2417     Formula (guards_sym_formula_prefix ^ s ^
  2418              (if n > 1 then "_" ^ string_of_int j else ""), role,
  2419              IConst ((s, s'), T, T_args)
  2420              |> fold (curry (IApp o swap)) bounds
  2421              |> type_guard_iterm type_enc res_T
  2422              |> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)
  2423              |> formula_from_iformula ctxt mono type_enc
  2424                                       always_guard_var_in_formula (SOME true)
  2425              |> close_formula_universally
  2426              |> bound_tvars type_enc (n > 1) (atomic_types_of T)
  2427              |> maybe_negate,
  2428              NONE, isabelle_info inductiveN helper_rank)
  2429   end
  2430 
  2431 fun formula_lines_for_tags_sym_decl ctxt mono type_enc n s
  2432         (j, (s', T_args, T, pred_sym, ary, in_conj)) =
  2433   let
  2434     val thy = Proof_Context.theory_of ctxt
  2435     val level = level_of_type_enc type_enc
  2436     val grain = granularity_of_type_level level
  2437     val ident_base =
  2438       tags_sym_formula_prefix ^ s ^
  2439       (if n > 1 then "_" ^ string_of_int j else "")
  2440     val (role, maybe_negate) = honor_conj_sym_role in_conj
  2441     val (arg_Ts, res_T) = chop_fun ary T
  2442     val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)
  2443     val bounds = bound_names |> map (fn name => ATerm (name, []))
  2444     val cst = mk_aterm type_enc (s, s') T_args
  2445     val eq = maybe_negate oo eq_formula type_enc (atomic_types_of T) [] pred_sym
  2446     val should_encode = should_encode_type ctxt mono level
  2447     val tag_with = tag_with_type ctxt mono type_enc NONE
  2448     val add_formula_for_res =
  2449       if should_encode res_T then
  2450         let
  2451           val tagged_bounds =
  2452             if grain = Ghost_Type_Arg_Vars then
  2453               let val cover = type_arg_cover thy s ary in
  2454                 map2 (fn (j, arg_T) => member (op =) cover j ? tag_with arg_T)
  2455                      (0 upto ary - 1 ~~ arg_Ts) bounds
  2456               end
  2457             else
  2458               bounds
  2459         in
  2460           cons (Formula (ident_base ^ "_res", role,
  2461                          eq (tag_with res_T (cst bounds)) (cst tagged_bounds),
  2462                          NONE, isabelle_info defN helper_rank))
  2463         end
  2464       else
  2465         I
  2466   in [] |> not pred_sym ? add_formula_for_res end
  2467 
  2468 fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd
  2469 
  2470 fun rationalize_decls thy (decls as decl :: (decls' as _ :: _)) =
  2471     let
  2472       val T = result_type_of_decl decl
  2473               |> map_type_tvar (fn (z, _) => TVar (z, HOLogic.typeS))
  2474     in
  2475       if forall (type_generalization thy T o result_type_of_decl) decls' then
  2476         [decl]
  2477       else
  2478         decls
  2479     end
  2480   | rationalize_decls _ decls = decls
  2481 
  2482 fun problem_lines_for_sym_decls ctxt mono type_enc (s, decls) =
  2483   case type_enc of
  2484     Native _ => [decl_line_for_sym ctxt mono type_enc s (hd decls)]
  2485   | Guards (_, level) =>
  2486     let
  2487       val thy = Proof_Context.theory_of ctxt
  2488       val decls = decls |> rationalize_decls thy
  2489       val n = length decls
  2490       val decls =
  2491         decls |> filter (should_encode_type ctxt mono level
  2492                          o result_type_of_decl)
  2493     in
  2494       (0 upto length decls - 1, decls)
  2495       |-> map2 (formula_line_for_guards_sym_decl ctxt mono type_enc n s)
  2496     end
  2497   | Tags (_, level) =>
  2498     if granularity_of_type_level level = All_Vars then
  2499       []
  2500     else
  2501       let val n = length decls in
  2502         (0 upto n - 1 ~~ decls)
  2503         |> maps (formula_lines_for_tags_sym_decl ctxt mono type_enc n s)
  2504       end
  2505 
  2506 fun problem_lines_for_sym_decl_table ctxt mono type_enc mono_Ts sym_decl_tab =
  2507   let
  2508     val syms = sym_decl_tab |> Symtab.dest |> sort_wrt fst
  2509     val mono_lines = problem_lines_for_mono_types ctxt mono type_enc mono_Ts
  2510     val decl_lines =
  2511       fold_rev (append o problem_lines_for_sym_decls ctxt mono type_enc) syms []
  2512   in mono_lines @ decl_lines end
  2513 
  2514 fun pair_append (xs1, xs2) (ys1, ys2) = (xs1 @ ys1, xs2 @ ys2)
  2515 
  2516 fun do_uncurried_alias_lines_for_sym ctxt monom_constrs mono type_enc sym_tab0
  2517                                      sym_tab base_s0 types in_conj =
  2518   let
  2519     fun do_alias ary =
  2520       let
  2521         val thy = Proof_Context.theory_of ctxt
  2522         val (role, maybe_negate) = honor_conj_sym_role in_conj
  2523         val base_name = base_s0 |> `(make_fixed_const (SOME type_enc))
  2524         val T = case types of [T] => T | _ => robust_const_type thy base_s0
  2525         val T_args = robust_const_typargs thy (base_s0, T)
  2526         val (base_name as (base_s, _), T_args) =
  2527           mangle_type_args_in_const type_enc base_name T_args
  2528         val base_ary = min_ary_of sym_tab0 base_s
  2529         fun do_const name = IConst (name, T, T_args)
  2530         val filter_ty_args =
  2531           filter_type_args_in_iterm thy monom_constrs type_enc
  2532         val ho_term_of = ho_term_from_iterm ctxt mono type_enc (SOME true)
  2533         val name1 as (s1, _) =
  2534           base_name |> ary - 1 > base_ary ? aliased_uncurried (ary - 1)
  2535         val name2 as (s2, _) = base_name |> aliased_uncurried ary
  2536         val (arg_Ts, _) = chop_fun ary T
  2537         val bound_names =
  2538           1 upto ary |> map (`I o make_bound_var o string_of_int)
  2539         val bounds = bound_names ~~ arg_Ts
  2540         val (first_bounds, last_bound) =
  2541           bounds |> map (fn (name, T) => IConst (name, T, [])) |> split_last
  2542         val tm1 =
  2543           mk_app_op type_enc (list_app (do_const name1) first_bounds) last_bound
  2544           |> filter_ty_args
  2545         val tm2 =
  2546           list_app (do_const name2) (first_bounds @ [last_bound])
  2547           |> filter_ty_args
  2548         val do_bound_type = do_bound_type ctxt mono type_enc
  2549         val eq =
  2550           eq_formula type_enc (atomic_types_of T)
  2551                      (map (apsnd do_bound_type) bounds) false
  2552                      (ho_term_of tm1) (ho_term_of tm2)
  2553       in
  2554         ([tm1, tm2],
  2555          [Formula (uncurried_alias_eq_prefix ^ s2, role, eq |> maybe_negate,
  2556                    NONE, isabelle_info defN helper_rank)])
  2557         |> (if ary - 1 = base_ary orelse Symtab.defined sym_tab s1 then I
  2558             else pair_append (do_alias (ary - 1)))
  2559       end
  2560   in do_alias end
  2561 fun uncurried_alias_lines_for_sym ctxt monom_constrs mono type_enc sym_tab0
  2562         sym_tab (s, {min_ary, types, in_conj, ...} : sym_info) =
  2563   case unprefix_and_unascii const_prefix s of
  2564     SOME mangled_s =>
  2565     if String.isSubstring uncurried_alias_sep mangled_s then
  2566       let
  2567         val base_s0 = mangled_s |> unmangled_const_name |> hd |> invert_const
  2568       in
  2569         do_uncurried_alias_lines_for_sym ctxt monom_constrs mono type_enc
  2570             sym_tab0 sym_tab base_s0 types in_conj min_ary
  2571       end
  2572     else
  2573       ([], [])
  2574   | NONE => ([], [])
  2575 fun uncurried_alias_lines_for_sym_table ctxt monom_constrs mono type_enc
  2576                                         uncurried_aliases sym_tab0 sym_tab =
  2577   ([], [])
  2578   |> uncurried_aliases
  2579      ? Symtab.fold_rev
  2580            (pair_append
  2581             o uncurried_alias_lines_for_sym ctxt monom_constrs mono type_enc
  2582                                             sym_tab0 sym_tab) sym_tab
  2583 
  2584 val implicit_declsN = "Should-be-implicit typings"
  2585 val explicit_declsN = "Explicit typings"
  2586 val uncurried_alias_eqsN = "Uncurried aliases"
  2587 val factsN = "Relevant facts"
  2588 val class_relsN = "Class relationships"
  2589 val aritiesN = "Arities"
  2590 val helpersN = "Helper facts"
  2591 val conjsN = "Conjectures"
  2592 val free_typesN = "Type variables"
  2593 
  2594 (* TFF allows implicit declarations of types, function symbols, and predicate
  2595    symbols (with "$i" as the type of individuals), but some provers (e.g.,
  2596    SNARK) require explicit declarations. The situation is similar for THF. *)
  2597 
  2598 fun default_type type_enc pred_sym s =
  2599   let
  2600     val ind =
  2601       case type_enc of
  2602         Native _ =>
  2603         if String.isPrefix type_const_prefix s orelse
  2604            String.isPrefix tfree_prefix s then
  2605           atype_of_types
  2606         else
  2607           individual_atype
  2608       | _ => individual_atype
  2609     fun typ 0 = if pred_sym then bool_atype else ind
  2610       | typ ary = AFun (ind, typ (ary - 1))
  2611   in typ end
  2612 
  2613 fun nary_type_constr_type n =
  2614   funpow n (curry AFun atype_of_types) atype_of_types
  2615 
  2616 fun undeclared_syms_in_problem type_enc problem =
  2617   let
  2618     fun do_sym (name as (s, _)) ty =
  2619       if is_tptp_user_symbol s then
  2620         Symtab.default (s, (name, ty))
  2621       else
  2622         I
  2623     fun do_type (AType (name, tys)) =
  2624         do_sym name (fn () => nary_type_constr_type (length tys))
  2625         #> fold do_type tys
  2626       | do_type (AFun (ty1, ty2)) = do_type ty1 #> do_type ty2
  2627       | do_type (ATyAbs (_, ty)) = do_type ty
  2628     fun do_term pred_sym (ATerm (name as (s, _), tms)) =
  2629         do_sym name (fn _ => default_type type_enc pred_sym s (length tms))
  2630         #> fold (do_term false) tms
  2631       | do_term _ (AAbs (((_, ty), tm), args)) =
  2632         do_type ty #> do_term false tm #> fold (do_term false) args
  2633     fun do_formula (AQuant (_, xs, phi)) =
  2634         fold do_type (map_filter snd xs) #> do_formula phi
  2635       | do_formula (AConn (_, phis)) = fold do_formula phis
  2636       | do_formula (AAtom tm) = do_term true tm
  2637     fun do_problem_line (Decl (_, _, ty)) = do_type ty
  2638       | do_problem_line (Formula (_, _, phi, _, _)) = do_formula phi
  2639   in
  2640     Symtab.empty
  2641     |> fold (fn (s, _) => Symtab.default (s, (("", ""), K tvar_a_atype)))
  2642             (declared_syms_in_problem problem)
  2643     |> fold (fold do_problem_line o snd) problem
  2644   end
  2645 
  2646 fun declare_undeclared_syms_in_atp_problem type_enc problem =
  2647   let
  2648     val decls =
  2649       Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *)
  2650                     | (s, (sym, ty)) =>
  2651                       cons (Decl (type_decl_prefix ^ s, sym, ty ())))
  2652                   (undeclared_syms_in_problem type_enc problem) []
  2653   in (implicit_declsN, decls) :: problem end
  2654 
  2655 fun exists_subdtype P =
  2656   let
  2657     fun ex U = P U orelse
  2658       (case U of Datatype.DtType (_, Us) => exists ex Us | _ => false)
  2659   in ex end
  2660 
  2661 fun is_poly_constr (_, Us) =
  2662   exists (exists_subdtype (fn Datatype.DtTFree _ => true | _ => false)) Us
  2663 
  2664 fun all_constrs_of_polymorphic_datatypes thy =
  2665   Symtab.fold (snd
  2666                #> #descr
  2667                #> maps (snd #> #3)
  2668                #> (fn cs => exists is_poly_constr cs ? append cs))
  2669               (Datatype.get_all thy) []
  2670   |> List.partition is_poly_constr
  2671   |> pairself (map fst)
  2672 
  2673 val app_op_and_predicator_threshold = 50
  2674 
  2675 fun prepare_atp_problem ctxt format prem_role type_enc mode lam_trans
  2676                         uncurried_aliases readable_names preproc hyp_ts concl_t
  2677                         facts =
  2678   let
  2679     val thy = Proof_Context.theory_of ctxt
  2680     val type_enc = type_enc |> adjust_type_enc format
  2681     (* Forcing explicit applications is expensive for polymorphic encodings,
  2682        because it takes only one existential variable ranging over "'a => 'b" to
  2683        ruin everything. Hence we do it only if there are few facts (which is
  2684        normally the case for "metis" and the minimizer). *)
  2685     val app_op_level =
  2686       if mode = Sledgehammer_Aggressive then
  2687         Full_App_Op_And_Predicator
  2688       else if length facts + length hyp_ts
  2689               > app_op_and_predicator_threshold then
  2690         if polymorphism_of_type_enc type_enc = Polymorphic then Min_App_Op
  2691         else Sufficient_App_Op
  2692       else
  2693         Sufficient_App_Op_And_Predicator
  2694     val exporter = (mode = Exporter)
  2695     val aggressive = (mode = Sledgehammer_Aggressive)
  2696     val lam_trans =
  2697       if lam_trans = keep_lamsN andalso
  2698          not (is_type_enc_higher_order type_enc) then
  2699         error ("Lambda translation scheme incompatible with first-order \
  2700                \encoding.")
  2701       else
  2702         lam_trans
  2703     val (fact_names, classes, conjs, facts, class_rel_clauses, arity_clauses,
  2704          lifted) =
  2705       translate_formulas ctxt prem_role format type_enc lam_trans preproc hyp_ts
  2706                          concl_t facts
  2707     val (_, sym_tab0) =
  2708       sym_table_for_facts ctxt type_enc app_op_level conjs facts
  2709     val mono = conjs @ facts |> mononotonicity_info_for_facts ctxt type_enc
  2710     val (_, monom_constrs) =
  2711       all_constrs_of_polymorphic_datatypes thy
  2712       |>> map (make_fixed_const (SOME type_enc))
  2713     fun firstorderize in_helper =
  2714       firstorderize_fact thy monom_constrs type_enc sym_tab0
  2715           (uncurried_aliases andalso not in_helper) aggressive
  2716     val (conjs, facts) = (conjs, facts) |> pairself (map (firstorderize false))
  2717     val (ho_stuff, sym_tab) =
  2718       sym_table_for_facts ctxt type_enc Min_App_Op conjs facts
  2719     val (uncurried_alias_eq_tms, uncurried_alias_eq_lines) =
  2720       uncurried_alias_lines_for_sym_table ctxt monom_constrs mono type_enc
  2721                                           uncurried_aliases sym_tab0 sym_tab
  2722     val (_, sym_tab) =
  2723       (ho_stuff, sym_tab)
  2724       |> fold (add_iterm_syms_to_sym_table ctxt Min_App_Op false false)
  2725               uncurried_alias_eq_tms
  2726     val helpers =
  2727       sym_tab |> helper_facts_for_sym_table ctxt format type_enc aggressive
  2728               |> map (firstorderize true)
  2729     val mono_Ts =
  2730       helpers @ conjs @ facts |> monotonic_types_for_facts ctxt mono type_enc
  2731     val class_decl_lines = decl_lines_for_classes type_enc classes
  2732     val sym_decl_lines =
  2733       (conjs, helpers @ facts, uncurried_alias_eq_tms)
  2734       |> sym_decl_table_for_facts thy type_enc sym_tab
  2735       |> problem_lines_for_sym_decl_table ctxt mono type_enc mono_Ts
  2736     val num_facts = length facts
  2737     val fact_lines =
  2738       map (formula_line_for_fact ctxt fact_prefix ascii_of (not exporter)
  2739                (not exporter) mono type_enc (rank_of_fact_num num_facts))
  2740           (0 upto num_facts - 1 ~~ facts)
  2741     val helper_lines =
  2742       0 upto length helpers - 1 ~~ helpers
  2743       |> map (formula_line_for_fact ctxt helper_prefix I false true mono
  2744                                     type_enc (K default_rank))
  2745     (* Reordering these might confuse the proof reconstruction code. *)
  2746     val problem =
  2747       [(explicit_declsN, class_decl_lines @ sym_decl_lines),
  2748        (uncurried_alias_eqsN, uncurried_alias_eq_lines),
  2749        (factsN, fact_lines),
  2750        (class_relsN,
  2751         map (formula_line_for_class_rel_clause type_enc) class_rel_clauses),
  2752        (aritiesN, map (formula_line_for_arity_clause type_enc) arity_clauses),
  2753        (helpersN, helper_lines),
  2754        (free_typesN, formula_lines_for_free_types type_enc (facts @ conjs)),
  2755        (conjsN, map (formula_line_for_conjecture ctxt mono type_enc) conjs)]
  2756     val problem =
  2757       problem
  2758       |> (case format of
  2759             CNF => ensure_cnf_problem
  2760           | CNF_UEQ => filter_cnf_ueq_problem
  2761           | FOF => I
  2762           | TFF (_, TPTP_Implicit) => I
  2763           | THF (_, TPTP_Implicit, _, _) => I
  2764           | _ => declare_undeclared_syms_in_atp_problem type_enc)
  2765     val (problem, pool) = problem |> nice_atp_problem readable_names format
  2766     fun add_sym_ary (s, {min_ary, ...} : sym_info) =
  2767       min_ary > 0 ? Symtab.insert (op =) (s, min_ary)
  2768   in
  2769     (problem,
  2770      case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
  2771      fact_names |> Vector.fromList,
  2772      lifted,
  2773      Symtab.empty |> Symtab.fold add_sym_ary sym_tab)
  2774   end
  2775 
  2776 (* FUDGE *)
  2777 val conj_weight = 0.0
  2778 val hyp_weight = 0.1
  2779 val fact_min_weight = 0.2
  2780 val fact_max_weight = 1.0
  2781 val type_info_default_weight = 0.8
  2782 
  2783 (* Weights are from 0.0 (most important) to 1.0 (least important). *)
  2784 fun atp_problem_selection_weights problem =
  2785   let
  2786     fun add_term_weights weight (ATerm (s, tms)) =
  2787         is_tptp_user_symbol s ? Symtab.default (s, weight)
  2788         #> fold (add_term_weights weight) tms
  2789       | add_term_weights weight (AAbs ((_, tm), args)) =
  2790         add_term_weights weight tm #> fold (add_term_weights weight) args
  2791     fun add_line_weights weight (Formula (_, _, phi, _, _)) =
  2792         formula_fold NONE (K (add_term_weights weight)) phi
  2793       | add_line_weights _ _ = I
  2794     fun add_conjectures_weights [] = I
  2795       | add_conjectures_weights conjs =
  2796         let val (hyps, conj) = split_last conjs in
  2797           add_line_weights conj_weight conj
  2798           #> fold (add_line_weights hyp_weight) hyps
  2799         end
  2800     fun add_facts_weights facts =
  2801       let
  2802         val num_facts = length facts
  2803         fun weight_of j =
  2804           fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j
  2805                             / Real.fromInt num_facts
  2806       in
  2807         map weight_of (0 upto num_facts - 1) ~~ facts
  2808         |> fold (uncurry add_line_weights)
  2809       end
  2810     val get = these o AList.lookup (op =) problem
  2811   in
  2812     Symtab.empty
  2813     |> add_conjectures_weights (get free_typesN @ get conjsN)
  2814     |> add_facts_weights (get factsN)
  2815     |> fold (fold (add_line_weights type_info_default_weight) o get)
  2816             [explicit_declsN, class_relsN, aritiesN]
  2817     |> Symtab.dest
  2818     |> sort (prod_ord Real.compare string_ord o pairself swap)
  2819   end
  2820 
  2821 (* Ugly hack: may make innocent victims (collateral damage) *)
  2822 fun may_be_app s args = String.isPrefix app_op_name s andalso length args = 2
  2823 fun may_be_predicator s =
  2824   member (op =) [predicator_name, prefixed_predicator_name] s
  2825 
  2826 fun strip_predicator (tm as ATerm (s, [tm'])) =
  2827     if may_be_predicator s then tm' else tm
  2828   | strip_predicator tm = tm
  2829 
  2830 fun make_head_roll (ATerm (s, tms)) =
  2831     if may_be_app s tms then make_head_roll (hd tms) ||> append (tl tms)
  2832     else (s, tms)
  2833   | make_head_roll _ = ("", [])
  2834 
  2835 fun strip_up_to_predicator (AQuant (_, _, phi)) = strip_up_to_predicator phi
  2836   | strip_up_to_predicator (AConn (_, phis)) = maps strip_up_to_predicator phis
  2837   | strip_up_to_predicator (AAtom tm) = [strip_predicator tm]
  2838 
  2839 fun strip_ahorn_etc (AQuant (_, _, phi)) = strip_ahorn_etc phi
  2840   | strip_ahorn_etc (AConn (AImplies, [phi1, phi2])) =
  2841     strip_ahorn_etc phi2 |>> append (strip_up_to_predicator phi1)
  2842   | strip_ahorn_etc phi = ([], hd (strip_up_to_predicator phi))
  2843 
  2844 fun strip_iff_etc (AQuant (_, _, phi)) = strip_iff_etc phi
  2845   | strip_iff_etc (AConn (AIff, [phi1, phi2])) =
  2846     pairself strip_up_to_predicator (phi1, phi2)
  2847   | strip_iff_etc _ = ([], [])
  2848 
  2849 val max_term_order_weight = 2147483647
  2850 
  2851 fun atp_problem_term_order_info problem =
  2852   let
  2853     fun add_edge s s' =
  2854       Graph.default_node (s, ())
  2855       #> Graph.default_node (s', ())
  2856       #> Graph.add_edge_acyclic (s, s')
  2857     fun add_term_deps head (ATerm (s, args)) =
  2858         if is_tptp_user_symbol head then
  2859           (if is_tptp_user_symbol s then perhaps (try (add_edge s head)) else I)
  2860           #> fold (add_term_deps head) args
  2861         else
  2862           I
  2863       | add_term_deps head (AAbs ((_, tm), args)) =
  2864         add_term_deps head tm #> fold (add_term_deps head) args
  2865     fun add_intro_deps pred (Formula (_, role, phi, _, info)) =
  2866         if pred (role, info) then
  2867           let val (hyps, concl) = strip_ahorn_etc phi in
  2868             case make_head_roll concl of
  2869               (head, args as _ :: _) => fold (add_term_deps head) (hyps @ args)
  2870             | _ => I
  2871           end
  2872         else
  2873           I
  2874       | add_intro_deps _ _ = I
  2875     fun add_atom_eq_deps (SOME true) (ATerm (s, [lhs as _, rhs])) =
  2876         if is_tptp_equal s then
  2877           case make_head_roll lhs of
  2878             (head, args as _ :: _) => fold (add_term_deps head) (rhs :: args)
  2879           | _ => I
  2880         else
  2881           I
  2882       | add_atom_eq_deps _ _ = I
  2883     fun add_eq_deps pred (Formula (_, role, phi, _, info)) =
  2884         if pred (role, info) then
  2885           case strip_iff_etc phi of
  2886             ([lhs], rhs) =>
  2887             (case make_head_roll lhs of
  2888                (head, args as _ :: _) => fold (add_term_deps head) (rhs @ args)
  2889              | _ => I)
  2890           | _ => formula_fold (SOME (role <> Conjecture)) add_atom_eq_deps phi
  2891         else
  2892           I
  2893       | add_eq_deps _ _ = I
  2894     fun has_status status (_, info) = extract_isabelle_status info = SOME status
  2895     fun is_conj (role, _) = (role = Conjecture orelse role = Hypothesis)
  2896     val graph =
  2897       Graph.empty
  2898       |> fold (fold (add_eq_deps (has_status defN)) o snd) problem
  2899       |> fold (fold (add_eq_deps (has_status simpN orf is_conj)) o snd) problem
  2900       |> fold (fold (add_intro_deps (has_status inductiveN)) o snd) problem
  2901       |> fold (fold (add_intro_deps (has_status introN)) o snd) problem
  2902     fun next_weight w = if w + w <= max_term_order_weight then w + w else w + 1
  2903     fun add_weights _ [] = I
  2904       | add_weights weight syms =
  2905         fold (AList.update (op =) o rpair weight) syms
  2906         #> add_weights (next_weight weight)
  2907                (fold (union (op =) o Graph.immediate_succs graph) syms [])
  2908   in
  2909     (* Sorting is not just for aesthetics: It specifies the precedence order
  2910        for the term ordering (KBO or LPO), from smaller to larger values. *)
  2911     [] |> add_weights 1 (Graph.minimals graph) |> sort (int_ord o pairself snd)
  2912   end
  2913 
  2914 end;