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