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