src/HOL/Tools/ATP/atp_translate.ML
author nik
Tue, 05 Jul 2011 17:09:59 +0100
changeset 44536 2cd0b478d1b6
parent 44535 3b0b448b4d69
child 44537 56d352659500
permissions -rw-r--r--
added generation of lambdas in THF
     1 (*  Title:      HOL/Tools/ATP/atp_translate.ML
     2     Author:     Fabian Immler, TU Muenchen
     3     Author:     Makarius
     4     Author:     Jasmin Blanchette, TU Muenchen
     5 
     6 Translation of HOL to FOL for Sledgehammer.
     7 *)
     8 
     9 signature ATP_TRANSLATE =
    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 format = ATP_Problem.format
    15   type formula_kind = ATP_Problem.formula_kind
    16   type 'a problem = 'a ATP_Problem.problem
    17 
    18   datatype locality =
    19     General | Helper | Extensionality | Intro | Elim | Simp | Local | Assum |
    20     Chained
    21 
    22   datatype order = First_Order | Higher_Order
    23   datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic
    24   datatype type_level =
    25     All_Types | Noninf_Nonmono_Types | Fin_Nonmono_Types | Const_Arg_Types |
    26     No_Types
    27   datatype type_heaviness = Heavyweight | Lightweight
    28 
    29   datatype type_enc =
    30     Simple_Types of order * type_level |
    31     Preds of polymorphism * type_level * type_heaviness |
    32     Tags of polymorphism * type_level * type_heaviness
    33 
    34   val bound_var_prefix : string
    35   val schematic_var_prefix : string
    36   val fixed_var_prefix : string
    37   val tvar_prefix : string
    38   val tfree_prefix : string
    39   val const_prefix : string
    40   val type_const_prefix : string
    41   val class_prefix : string
    42   val skolem_const_prefix : string
    43   val old_skolem_const_prefix : string
    44   val new_skolem_const_prefix : string
    45   val type_decl_prefix : string
    46   val sym_decl_prefix : string
    47   val preds_sym_formula_prefix : string
    48   val lightweight_tags_sym_formula_prefix : string
    49   val fact_prefix : string
    50   val conjecture_prefix : string
    51   val helper_prefix : string
    52   val class_rel_clause_prefix : string
    53   val arity_clause_prefix : string
    54   val tfree_clause_prefix : string
    55   val typed_helper_suffix : string
    56   val untyped_helper_suffix : string
    57   val type_tag_idempotence_helper_name : string
    58   val predicator_name : string
    59   val app_op_name : string
    60   val type_tag_name : string
    61   val type_pred_name : string
    62   val simple_type_prefix : string
    63   val prefixed_predicator_name : string
    64   val prefixed_app_op_name : string
    65   val prefixed_type_tag_name : string
    66   val ascii_of : string -> string
    67   val unascii_of : string -> string
    68   val strip_prefix_and_unascii : string -> string -> string option
    69   val proxy_table : (string * (string * (thm * (string * string)))) list
    70   val proxify_const : string -> (string * string) option
    71   val invert_const : string -> string
    72   val unproxify_const : string -> string
    73   val new_skolem_var_name_from_const : string -> string
    74   val num_type_args : theory -> string -> int
    75   val atp_irrelevant_consts : string list
    76   val atp_schematic_consts_of : term -> typ list Symtab.table
    77   val is_locality_global : locality -> bool
    78   val type_enc_from_string : string -> type_enc
    79   val polymorphism_of_type_enc : type_enc -> polymorphism
    80   val level_of_type_enc : type_enc -> type_level
    81   val is_type_enc_virtually_sound : type_enc -> bool
    82   val is_type_enc_fairly_sound : type_enc -> bool
    83   val choose_format : format list -> type_enc -> format * type_enc
    84   val mk_aconns :
    85     connective -> ('a, 'b, 'c) formula list -> ('a, 'b, 'c) formula
    86   val unmangled_const : string -> string * (string, 'b) ho_term list
    87   val unmangled_const_name : string -> string
    88   val helper_table : ((string * bool) * thm list) list
    89   val factsN : string
    90   val prepare_atp_problem :
    91     Proof.context -> format -> formula_kind -> formula_kind -> type_enc -> bool
    92     -> bool -> bool -> bool -> term list -> term
    93     -> ((string * locality) * term) list
    94     -> string problem * string Symtab.table * int * int
    95        * (string * locality) list vector * int list * int Symtab.table
    96   val atp_problem_weights : string problem -> (string * real) list
    97 end;
    98 
    99 structure ATP_Translate : ATP_TRANSLATE =
   100 struct
   101 
   102 open ATP_Util
   103 open ATP_Problem
   104 
   105 type name = string * string
   106 
   107 (* experimental *)
   108 val generate_useful_info = false
   109 
   110 fun useful_isabelle_info s =
   111   if generate_useful_info then
   112     SOME (ATerm ("[]", [ATerm ("isabelle_" ^ s, [])]))
   113   else
   114     NONE
   115 
   116 val intro_info = useful_isabelle_info "intro"
   117 val elim_info = useful_isabelle_info "elim"
   118 val simp_info = useful_isabelle_info "simp"
   119 
   120 val bound_var_prefix = "B_"
   121 val schematic_var_prefix = "V_"
   122 val fixed_var_prefix = "v_"
   123 
   124 val tvar_prefix = "T_"
   125 val tfree_prefix = "t_"
   126 
   127 val const_prefix = "c_"
   128 val type_const_prefix = "tc_"
   129 val class_prefix = "cl_"
   130 
   131 val skolem_const_prefix = "Sledgehammer" ^ Long_Name.separator ^ "Sko"
   132 val old_skolem_const_prefix = skolem_const_prefix ^ "o"
   133 val new_skolem_const_prefix = skolem_const_prefix ^ "n"
   134 
   135 val type_decl_prefix = "ty_"
   136 val sym_decl_prefix = "sy_"
   137 val preds_sym_formula_prefix = "psy_"
   138 val lightweight_tags_sym_formula_prefix = "tsy_"
   139 val fact_prefix = "fact_"
   140 val conjecture_prefix = "conj_"
   141 val helper_prefix = "help_"
   142 val class_rel_clause_prefix = "clar_"
   143 val arity_clause_prefix = "arity_"
   144 val tfree_clause_prefix = "tfree_"
   145 
   146 val typed_helper_suffix = "_T"
   147 val untyped_helper_suffix = "_U"
   148 val type_tag_idempotence_helper_name = helper_prefix ^ "ti_idem"
   149 
   150 val predicator_name = "hBOOL"
   151 val app_op_name = "hAPP"
   152 val type_tag_name = "ti"
   153 val type_pred_name = "is"
   154 val simple_type_prefix = "ty_"
   155 
   156 val prefixed_predicator_name = const_prefix ^ predicator_name
   157 val prefixed_app_op_name = const_prefix ^ app_op_name
   158 val prefixed_type_tag_name = const_prefix ^ type_tag_name
   159 
   160 (* Freshness almost guaranteed! *)
   161 val sledgehammer_weak_prefix = "Sledgehammer:"
   162 
   163 (*Escaping of special characters.
   164   Alphanumeric characters are left unchanged.
   165   The character _ goes to __
   166   Characters in the range ASCII space to / go to _A to _P, respectively.
   167   Other characters go to _nnn where nnn is the decimal ASCII code.*)
   168 val upper_a_minus_space = Char.ord #"A" - Char.ord #" "
   169 
   170 fun stringN_of_int 0 _ = ""
   171   | stringN_of_int k n =
   172     stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10)
   173 
   174 fun ascii_of_char c =
   175   if Char.isAlphaNum c then
   176     String.str c
   177   else if c = #"_" then
   178     "__"
   179   else if #" " <= c andalso c <= #"/" then
   180     "_" ^ String.str (Char.chr (Char.ord c + upper_a_minus_space))
   181   else
   182     (* fixed width, in case more digits follow *)
   183     "_" ^ stringN_of_int 3 (Char.ord c)
   184 
   185 val ascii_of = String.translate ascii_of_char
   186 
   187 (** Remove ASCII armoring from names in proof files **)
   188 
   189 (* We don't raise error exceptions because this code can run inside a worker
   190    thread. Also, the errors are impossible. *)
   191 val unascii_of =
   192   let
   193     fun un rcs [] = String.implode(rev rcs)
   194       | un rcs [#"_"] = un (#"_" :: rcs) [] (* ERROR *)
   195         (* Three types of _ escapes: __, _A to _P, _nnn *)
   196       | un rcs (#"_" :: #"_" :: cs) = un (#"_" :: rcs) cs
   197       | un rcs (#"_" :: c :: cs) =
   198         if #"A" <= c andalso c<= #"P" then
   199           (* translation of #" " to #"/" *)
   200           un (Char.chr (Char.ord c - upper_a_minus_space) :: rcs) cs
   201         else
   202           let val digits = List.take (c :: cs, 3) handle General.Subscript => [] in
   203             case Int.fromString (String.implode digits) of
   204               SOME n => un (Char.chr n :: rcs) (List.drop (cs, 2))
   205             | NONE => un (c :: #"_" :: rcs) cs (* ERROR *)
   206           end
   207       | un rcs (c :: cs) = un (c :: rcs) cs
   208   in un [] o String.explode end
   209 
   210 (* If string s has the prefix s1, return the result of deleting it,
   211    un-ASCII'd. *)
   212 fun strip_prefix_and_unascii s1 s =
   213   if String.isPrefix s1 s then
   214     SOME (unascii_of (String.extract (s, size s1, NONE)))
   215   else
   216     NONE
   217 
   218 val proxy_table =
   219   [("c_False", (@{const_name False}, (@{thm fFalse_def},
   220        ("fFalse", @{const_name ATP.fFalse})))),
   221    ("c_True", (@{const_name True}, (@{thm fTrue_def},
   222        ("fTrue", @{const_name ATP.fTrue})))),
   223    ("c_Not", (@{const_name Not}, (@{thm fNot_def},
   224        ("fNot", @{const_name ATP.fNot})))),
   225    ("c_conj", (@{const_name conj}, (@{thm fconj_def},
   226        ("fconj", @{const_name ATP.fconj})))),
   227    ("c_disj", (@{const_name disj}, (@{thm fdisj_def},
   228        ("fdisj", @{const_name ATP.fdisj})))),
   229    ("c_implies", (@{const_name implies}, (@{thm fimplies_def},
   230        ("fimplies", @{const_name ATP.fimplies})))),
   231    ("equal", (@{const_name HOL.eq}, (@{thm fequal_def},
   232        ("fequal", @{const_name ATP.fequal}))))]
   233 
   234 val proxify_const = AList.lookup (op =) proxy_table #> Option.map (snd o snd)
   235 
   236 (* Readable names for the more common symbolic functions. Do not mess with the
   237    table unless you know what you are doing. *)
   238 val const_trans_table =
   239   [(@{type_name Product_Type.prod}, "prod"),
   240    (@{type_name Sum_Type.sum}, "sum"),
   241    (@{const_name False}, "False"),
   242    (@{const_name True}, "True"),
   243    (@{const_name Not}, "Not"),
   244    (@{const_name conj}, "conj"),
   245    (@{const_name disj}, "disj"),
   246    (@{const_name implies}, "implies"),
   247    (@{const_name HOL.eq}, "equal"),
   248    (@{const_name If}, "If"),
   249    (@{const_name Set.member}, "member"),
   250    (@{const_name Meson.COMBI}, "COMBI"),
   251    (@{const_name Meson.COMBK}, "COMBK"),
   252    (@{const_name Meson.COMBB}, "COMBB"),
   253    (@{const_name Meson.COMBC}, "COMBC"),
   254    (@{const_name Meson.COMBS}, "COMBS")]
   255   |> Symtab.make
   256   |> fold (Symtab.update o swap o snd o snd o snd) proxy_table
   257 
   258 (* Invert the table of translations between Isabelle and ATPs. *)
   259 val const_trans_table_inv =
   260   const_trans_table |> Symtab.dest |> map swap |> Symtab.make
   261 val const_trans_table_unprox =
   262   Symtab.empty
   263   |> fold (fn (_, (isa, (_, (_, atp)))) => Symtab.update (atp, isa)) proxy_table
   264 
   265 val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
   266 val unproxify_const = perhaps (Symtab.lookup const_trans_table_unprox)
   267 
   268 fun lookup_const c =
   269   case Symtab.lookup const_trans_table c of
   270     SOME c' => c'
   271   | NONE => ascii_of c
   272 
   273 fun ascii_of_indexname (v, 0) = ascii_of v
   274   | ascii_of_indexname (v, i) = ascii_of v ^ "_" ^ string_of_int i
   275 
   276 fun make_bound_var x = bound_var_prefix ^ ascii_of x
   277 fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v
   278 fun make_fixed_var x = fixed_var_prefix ^ ascii_of x
   279 
   280 fun make_schematic_type_var (x, i) =
   281       tvar_prefix ^ (ascii_of_indexname (unprefix "'" x, i))
   282 fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (unprefix "'" x))
   283 
   284 (* "HOL.eq" is mapped to the ATP's equality. *)
   285 fun make_fixed_const @{const_name HOL.eq} = tptp_old_equal
   286   | make_fixed_const c = const_prefix ^ lookup_const c
   287 
   288 fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
   289 
   290 fun make_type_class clas = class_prefix ^ ascii_of clas
   291 
   292 fun new_skolem_var_name_from_const s =
   293   let val ss = s |> space_explode Long_Name.separator in
   294     nth ss (length ss - 2)
   295   end
   296 
   297 (* The number of type arguments of a constant, zero if it's monomorphic. For
   298    (instances of) Skolem pseudoconstants, this information is encoded in the
   299    constant name. *)
   300 fun num_type_args thy s =
   301   if String.isPrefix skolem_const_prefix s then
   302     s |> space_explode Long_Name.separator |> List.last |> Int.fromString |> the
   303   else
   304     (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
   305 
   306 (* These are either simplified away by "Meson.presimplify" (most of the time) or
   307    handled specially via "fFalse", "fTrue", ..., "fequal". *)
   308 val atp_irrelevant_consts =
   309   [@{const_name False}, @{const_name True}, @{const_name Not},
   310    @{const_name conj}, @{const_name disj}, @{const_name implies},
   311    @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
   312 
   313 val atp_monomorph_bad_consts =
   314   atp_irrelevant_consts @
   315   (* These are ignored anyway by the relevance filter (unless they appear in
   316      higher-order places) but not by the monomorphizer. *)
   317   [@{const_name all}, @{const_name "==>"}, @{const_name "=="},
   318    @{const_name Trueprop}, @{const_name All}, @{const_name Ex},
   319    @{const_name Ex1}, @{const_name Ball}, @{const_name Bex}]
   320 
   321 fun add_schematic_const (x as (_, T)) =
   322   Monomorph.typ_has_tvars T ? Symtab.insert_list (op =) x
   323 val add_schematic_consts_of =
   324   Term.fold_aterms (fn Const (x as (s, _)) =>
   325                        not (member (op =) atp_monomorph_bad_consts s)
   326                        ? add_schematic_const x
   327                       | _ => I)
   328 fun atp_schematic_consts_of t = add_schematic_consts_of t Symtab.empty
   329 
   330 (** Definitions and functions for FOL clauses and formulas for TPTP **)
   331 
   332 (* The first component is the type class; the second is a "TVar" or "TFree". *)
   333 datatype type_literal =
   334   TyLitVar of name * name |
   335   TyLitFree of name * name
   336 
   337 
   338 (** Isabelle arities **)
   339 
   340 datatype arity_literal =
   341   TConsLit of name * name * name list |
   342   TVarLit of name * name
   343 
   344 fun gen_TVars 0 = []
   345   | gen_TVars n = ("T_" ^ string_of_int n) :: gen_TVars (n-1)
   346 
   347 val type_class = the_single @{sort type}
   348 
   349 fun add_packed_sort tvar =
   350   fold (fn s => s <> type_class ? cons (`make_type_class s, `I tvar))
   351 
   352 type arity_clause =
   353   {name : string,
   354    prem_lits : arity_literal list,
   355    concl_lits : arity_literal}
   356 
   357 (* Arity of type constructor "tcon :: (arg1, ..., argN) res" *)
   358 fun make_axiom_arity_clause (tcons, name, (cls, args)) =
   359   let
   360     val tvars = gen_TVars (length args)
   361     val tvars_srts = ListPair.zip (tvars, args)
   362   in
   363     {name = name,
   364      prem_lits = [] |> fold (uncurry add_packed_sort) tvars_srts |> map TVarLit,
   365      concl_lits = TConsLit (`make_type_class cls,
   366                             `make_fixed_type_const tcons,
   367                             tvars ~~ tvars)}
   368   end
   369 
   370 fun arity_clause _ _ (_, []) = []
   371   | arity_clause seen n (tcons, ("HOL.type", _) :: ars) =  (* ignore *)
   372     arity_clause seen n (tcons, ars)
   373   | arity_clause seen n (tcons, (ar as (class, _)) :: ars) =
   374     if member (op =) seen class then
   375       (* multiple arities for the same (tycon, class) pair *)
   376       make_axiom_arity_clause (tcons,
   377           lookup_const tcons ^ "___" ^ ascii_of class ^ "_" ^ string_of_int n,
   378           ar) ::
   379       arity_clause seen (n + 1) (tcons, ars)
   380     else
   381       make_axiom_arity_clause (tcons, lookup_const tcons ^ "___" ^
   382                                ascii_of class, ar) ::
   383       arity_clause (class :: seen) n (tcons, ars)
   384 
   385 fun multi_arity_clause [] = []
   386   | multi_arity_clause ((tcons, ars) :: tc_arlists) =
   387       arity_clause [] 1 (tcons, ars) @ multi_arity_clause tc_arlists
   388 
   389 (* Generate all pairs (tycon, class, sorts) such that tycon belongs to class in
   390    theory thy provided its arguments have the corresponding sorts. *)
   391 fun type_class_pairs thy tycons classes =
   392   let
   393     val alg = Sign.classes_of thy
   394     fun domain_sorts tycon = Sorts.mg_domain alg tycon o single
   395     fun add_class tycon class =
   396       cons (class, domain_sorts tycon class)
   397       handle Sorts.CLASS_ERROR _ => I
   398     fun try_classes tycon = (tycon, fold (add_class tycon) classes [])
   399   in map try_classes tycons end
   400 
   401 (*Proving one (tycon, class) membership may require proving others, so iterate.*)
   402 fun iter_type_class_pairs _ _ [] = ([], [])
   403   | iter_type_class_pairs thy tycons classes =
   404       let
   405         fun maybe_insert_class s =
   406           (s <> type_class andalso not (member (op =) classes s))
   407           ? insert (op =) s
   408         val cpairs = type_class_pairs thy tycons classes
   409         val newclasses =
   410           [] |> fold (fold (fold (fold maybe_insert_class) o snd) o snd) cpairs
   411         val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
   412       in (classes' @ classes, union (op =) cpairs' cpairs) end
   413 
   414 fun make_arity_clauses thy tycons =
   415   iter_type_class_pairs thy tycons ##> multi_arity_clause
   416 
   417 
   418 (** Isabelle class relations **)
   419 
   420 type class_rel_clause =
   421   {name : string,
   422    subclass : name,
   423    superclass : name}
   424 
   425 (* Generate all pairs (sub, super) such that sub is a proper subclass of super
   426    in theory "thy". *)
   427 fun class_pairs _ [] _ = []
   428   | class_pairs thy subs supers =
   429       let
   430         val class_less = Sorts.class_less (Sign.classes_of thy)
   431         fun add_super sub super = class_less (sub, super) ? cons (sub, super)
   432         fun add_supers sub = fold (add_super sub) supers
   433       in fold add_supers subs [] end
   434 
   435 fun make_class_rel_clause (sub, super) =
   436   {name = sub ^ "_" ^ super, subclass = `make_type_class sub,
   437    superclass = `make_type_class super}
   438 
   439 fun make_class_rel_clauses thy subs supers =
   440   map make_class_rel_clause (class_pairs thy subs supers)
   441 
   442 datatype combterm =
   443   CombConst of name * typ * typ list |
   444   CombVar of name * typ |
   445   CombApp of combterm * combterm |
   446   CombAbs of (name * typ) * combterm
   447 
   448 fun combtyp_of (CombConst (_, T, _)) = T
   449   | combtyp_of (CombVar (_, T)) = T
   450   | combtyp_of (CombApp (t1, _)) = snd (dest_funT (combtyp_of t1))
   451   | combtyp_of (CombAbs ((_, T), tm)) = T --> combtyp_of tm
   452 
   453 (*gets the head of a combinator application, along with the list of arguments*)
   454 fun strip_combterm_comb u =
   455   let
   456     fun stripc (CombApp (t, u), ts) = stripc (t, u :: ts)
   457       | stripc x = x
   458   in stripc (u, []) end
   459 
   460 fun atyps_of T = fold_atyps (insert (op =)) T []
   461 
   462 fun new_skolem_const_name s num_T_args =
   463   [new_skolem_const_prefix, s, string_of_int num_T_args]
   464   |> space_implode Long_Name.separator
   465 
   466 (* Converts a term (with combinators) into a combterm. Also accumulates sort
   467    infomation. *)
   468 fun combterm_from_term thy bs (P $ Q) =
   469     let
   470       val (P', P_atomics_Ts) = combterm_from_term thy bs P
   471       val (Q', Q_atomics_Ts) = combterm_from_term thy bs Q
   472     in (CombApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end
   473   | combterm_from_term thy _ (Const (c, T)) =
   474     let
   475       val tvar_list =
   476         (if String.isPrefix old_skolem_const_prefix c then
   477            [] |> Term.add_tvarsT T |> map TVar
   478          else
   479            (c, T) |> Sign.const_typargs thy)
   480       val c' = CombConst (`make_fixed_const c, T, tvar_list)
   481     in (c', atyps_of T) end
   482   | combterm_from_term _ _ (Free (v, T)) =
   483     (CombConst (`make_fixed_var v, T, []), atyps_of T)
   484   | combterm_from_term _ _ (Var (v as (s, _), T)) =
   485     (if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then
   486        let
   487          val Ts = T |> strip_type |> swap |> op ::
   488          val s' = new_skolem_const_name s (length Ts)
   489        in CombConst (`make_fixed_const s', T, Ts) end
   490      else
   491        CombVar ((make_schematic_var v, s), T), atyps_of T)
   492   | combterm_from_term _ bs (Bound j) =
   493     nth bs j
   494     |> (fn (s, T) => (CombConst (`make_bound_var s, T, []), atyps_of T))
   495   | combterm_from_term thy bs (Abs (s, T, t)) =
   496     let val (tm, atomic_Ts) = combterm_from_term thy ((s, T) :: bs) t in
   497       (CombAbs ((`make_bound_var s, T), tm),
   498        union (op =) atomic_Ts (atyps_of T))
   499     end
   500 
   501 datatype locality =
   502   General | Helper | Extensionality | Intro | Elim | Simp | Local | Assum |
   503   Chained
   504 
   505 (* (quasi-)underapproximation of the truth *)
   506 fun is_locality_global Local = false
   507   | is_locality_global Assum = false
   508   | is_locality_global Chained = false
   509   | is_locality_global _ = true
   510 
   511 datatype order = First_Order | Higher_Order
   512 datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic
   513 datatype type_level =
   514   All_Types | Noninf_Nonmono_Types | Fin_Nonmono_Types | Const_Arg_Types |
   515   No_Types
   516 datatype type_heaviness = Heavyweight | Lightweight
   517 
   518 datatype type_enc =
   519   Simple_Types of order * type_level |
   520   Preds of polymorphism * type_level * type_heaviness |
   521   Tags of polymorphism * type_level * type_heaviness
   522 
   523 fun try_unsuffixes ss s =
   524   fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE
   525 
   526 fun type_enc_from_string s =
   527   (case try (unprefix "poly_") s of
   528      SOME s => (SOME Polymorphic, s)
   529    | NONE =>
   530      case try (unprefix "mono_") s of
   531        SOME s => (SOME Monomorphic, s)
   532      | NONE =>
   533        case try (unprefix "mangled_") s of
   534          SOME s => (SOME Mangled_Monomorphic, s)
   535        | NONE => (NONE, s))
   536   ||> (fn s =>
   537           (* "_query" and "_bang" are for the ASCII-challenged Metis and
   538              Mirabelle. *)
   539           case try_unsuffixes ["?", "_query"] s of
   540             SOME s => (Noninf_Nonmono_Types, s)
   541           | NONE =>
   542             case try_unsuffixes ["!", "_bang"] s of
   543               SOME s => (Fin_Nonmono_Types, s)
   544             | NONE => (All_Types, s))
   545   ||> apsnd (fn s =>
   546                 case try (unsuffix "_heavy") s of
   547                   SOME s => (Heavyweight, s)
   548                 | NONE => (Lightweight, s))
   549   |> (fn (poly, (level, (heaviness, core))) =>
   550          case (core, (poly, level, heaviness)) of
   551            ("simple", (NONE, _, Lightweight)) =>
   552            Simple_Types (First_Order, level)
   553          | ("simple_higher", (NONE, _, Lightweight)) =>
   554            if level = Noninf_Nonmono_Types then raise Same.SAME
   555            else Simple_Types (Higher_Order, level)
   556          | ("preds", (SOME poly, _, _)) => Preds (poly, level, heaviness)
   557          | ("tags", (SOME Polymorphic, _, _)) =>
   558            Tags (Polymorphic, level, heaviness)
   559          | ("tags", (SOME poly, _, _)) => Tags (poly, level, heaviness)
   560          | ("args", (SOME poly, All_Types (* naja *), Lightweight)) =>
   561            Preds (poly, Const_Arg_Types, Lightweight)
   562          | ("erased", (NONE, All_Types (* naja *), Lightweight)) =>
   563            Preds (Polymorphic, No_Types, Lightweight)
   564          | _ => raise Same.SAME)
   565   handle Same.SAME => error ("Unknown type system: " ^ quote s ^ ".")
   566 
   567 fun is_type_enc_higher_order (Simple_Types (Higher_Order, _)) = true
   568   | is_type_enc_higher_order _ = false
   569 
   570 fun polymorphism_of_type_enc (Simple_Types _) = Mangled_Monomorphic
   571   | polymorphism_of_type_enc (Preds (poly, _, _)) = poly
   572   | polymorphism_of_type_enc (Tags (poly, _, _)) = poly
   573 
   574 fun level_of_type_enc (Simple_Types (_, level)) = level
   575   | level_of_type_enc (Preds (_, level, _)) = level
   576   | level_of_type_enc (Tags (_, level, _)) = level
   577 
   578 fun heaviness_of_type_enc (Simple_Types _) = Heavyweight
   579   | heaviness_of_type_enc (Preds (_, _, heaviness)) = heaviness
   580   | heaviness_of_type_enc (Tags (_, _, heaviness)) = heaviness
   581 
   582 fun is_type_level_virtually_sound level =
   583   level = All_Types orelse level = Noninf_Nonmono_Types
   584 val is_type_enc_virtually_sound =
   585   is_type_level_virtually_sound o level_of_type_enc
   586 
   587 fun is_type_level_fairly_sound level =
   588   is_type_level_virtually_sound level orelse level = Fin_Nonmono_Types
   589 val is_type_enc_fairly_sound = is_type_level_fairly_sound o level_of_type_enc
   590 
   591 fun choose_format formats (Simple_Types (order, level)) =
   592     if member (op =) formats THF then
   593       (THF, Simple_Types (order, level))
   594     else if member (op =) formats TFF then
   595       (TFF, Simple_Types (First_Order, level))
   596     else
   597       choose_format formats (Preds (Mangled_Monomorphic, level, Heavyweight))
   598   | choose_format formats type_enc =
   599     (case hd formats of
   600        CNF_UEQ =>
   601        (CNF_UEQ, case type_enc of
   602                    Preds stuff =>
   603                    (if is_type_enc_fairly_sound type_enc then Tags else Preds)
   604                        stuff
   605                  | _ => type_enc)
   606      | format => (format, type_enc))
   607 
   608 type translated_formula =
   609   {name : string,
   610    locality : locality,
   611    kind : formula_kind,
   612    combformula : (name, typ, combterm) formula,
   613    atomic_types : typ list}
   614 
   615 fun update_combformula f ({name, locality, kind, combformula, atomic_types}
   616                           : translated_formula) =
   617   {name = name, locality = locality, kind = kind, combformula = f combformula,
   618    atomic_types = atomic_types} : translated_formula
   619 
   620 fun fact_lift f ({combformula, ...} : translated_formula) = f combformula
   621 
   622 val type_instance = Sign.typ_instance o Proof_Context.theory_of
   623 
   624 fun insert_type ctxt get_T x xs =
   625   let val T = get_T x in
   626     if exists (curry (type_instance ctxt) T o get_T) xs then xs
   627     else x :: filter_out (curry (type_instance ctxt o swap) T o get_T) xs
   628   end
   629 
   630 (* The Booleans indicate whether all type arguments should be kept. *)
   631 datatype type_arg_policy =
   632   Explicit_Type_Args of bool |
   633   Mangled_Type_Args of bool |
   634   No_Type_Args
   635 
   636 fun should_drop_arg_type_args (Simple_Types _) =
   637     false (* since TFF doesn't support overloading *)
   638   | should_drop_arg_type_args type_enc =
   639     level_of_type_enc type_enc = All_Types andalso
   640     heaviness_of_type_enc type_enc = Heavyweight
   641 
   642 fun type_arg_policy type_enc s =
   643   if s = type_tag_name then
   644     (if polymorphism_of_type_enc type_enc = Mangled_Monomorphic then
   645        Mangled_Type_Args
   646      else
   647        Explicit_Type_Args) false
   648   else case type_enc of
   649     Tags (_, All_Types, Heavyweight) => No_Type_Args
   650   | _ =>
   651     if level_of_type_enc type_enc = No_Types orelse
   652        s = @{const_name HOL.eq} orelse
   653        (s = app_op_name andalso
   654         level_of_type_enc type_enc = Const_Arg_Types) then
   655       No_Type_Args
   656     else
   657       should_drop_arg_type_args type_enc
   658       |> (if polymorphism_of_type_enc type_enc = Mangled_Monomorphic then
   659             Mangled_Type_Args
   660           else
   661             Explicit_Type_Args)
   662 
   663 (* Make literals for sorted type variables. *)
   664 fun generic_add_sorts_on_type (_, []) = I
   665   | generic_add_sorts_on_type ((x, i), s :: ss) =
   666     generic_add_sorts_on_type ((x, i), ss)
   667     #> (if s = the_single @{sort HOL.type} then
   668           I
   669         else if i = ~1 then
   670           insert (op =) (TyLitFree (`make_type_class s, `make_fixed_type_var x))
   671         else
   672           insert (op =) (TyLitVar (`make_type_class s,
   673                                    (make_schematic_type_var (x, i), x))))
   674 fun add_sorts_on_tfree (TFree (s, S)) = generic_add_sorts_on_type ((s, ~1), S)
   675   | add_sorts_on_tfree _ = I
   676 fun add_sorts_on_tvar (TVar z) = generic_add_sorts_on_type z
   677   | add_sorts_on_tvar _ = I
   678 
   679 fun type_literals_for_types type_enc add_sorts_on_typ Ts =
   680   [] |> level_of_type_enc type_enc <> No_Types ? fold add_sorts_on_typ Ts
   681 
   682 fun mk_aconns c phis =
   683   let val (phis', phi') = split_last phis in
   684     fold_rev (mk_aconn c) phis' phi'
   685   end
   686 fun mk_ahorn [] phi = phi
   687   | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])
   688 fun mk_aquant _ [] phi = phi
   689   | mk_aquant q xs (phi as AQuant (q', xs', phi')) =
   690     if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)
   691   | mk_aquant q xs phi = AQuant (q, xs, phi)
   692 
   693 fun close_universally atom_vars phi =
   694   let
   695     fun formula_vars bounds (AQuant (_, xs, phi)) =
   696         formula_vars (map fst xs @ bounds) phi
   697       | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis
   698       | formula_vars bounds (AAtom tm) =
   699         union (op =) (atom_vars tm []
   700                       |> filter_out (member (op =) bounds o fst))
   701   in mk_aquant AForall (formula_vars [] phi []) phi end
   702 
   703 fun combterm_vars (CombApp (tm1, tm2)) = fold combterm_vars [tm1, tm2]
   704   | combterm_vars (CombConst _) = I
   705   | combterm_vars (CombVar (name, T)) = insert (op =) (name, SOME T)
   706   | combterm_vars (CombAbs (_, tm)) = combterm_vars tm
   707 fun close_combformula_universally phi = close_universally combterm_vars phi
   708 
   709 fun term_vars bounds (ATerm (name as (s, _), tms)) =
   710     (is_tptp_variable s andalso not (member (op =) bounds name))
   711     ? insert (op =) (name, NONE) #> fold (term_vars bounds) tms
   712   | term_vars bounds (AAbs ((name, _), tm)) = term_vars (name :: bounds) tm
   713 fun close_formula_universally phi = close_universally (term_vars []) phi
   714 
   715 val homo_infinite_type_name = @{type_name ind} (* any infinite type *)
   716 val homo_infinite_type = Type (homo_infinite_type_name, [])
   717 
   718 fun ho_term_from_typ format type_enc =
   719   let
   720     fun term (Type (s, Ts)) =
   721       ATerm (case (is_type_enc_higher_order type_enc, s) of
   722                (true, @{type_name bool}) => `I tptp_bool_type
   723              | (true, @{type_name fun}) => `I tptp_fun_type
   724              | _ => if s = homo_infinite_type_name andalso
   725                        (format = TFF orelse format = THF) then
   726                       `I tptp_individual_type
   727                     else
   728                       `make_fixed_type_const s,
   729              map term Ts)
   730     | term (TFree (s, _)) = ATerm (`make_fixed_type_var s, [])
   731     | term (TVar ((x as (s, _)), _)) =
   732       ATerm ((make_schematic_type_var x, s), [])
   733   in term end
   734 
   735 fun ho_term_for_type_arg format type_enc T =
   736   if T = dummyT then NONE else SOME (ho_term_from_typ format type_enc T)
   737 
   738 (* This shouldn't clash with anything else. *)
   739 val mangled_type_sep = "\000"
   740 
   741 fun generic_mangled_type_name f (ATerm (name, [])) = f name
   742   | generic_mangled_type_name f (ATerm (name, tys)) =
   743     f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys)
   744     ^ ")"
   745   | generic_mangled_type_name f _ = raise Fail "unexpected type abstraction"
   746 
   747 val bool_atype = AType (`I tptp_bool_type)
   748 
   749 fun make_simple_type s =
   750   if s = tptp_bool_type orelse s = tptp_fun_type orelse
   751      s = tptp_individual_type then
   752     s
   753   else
   754     simple_type_prefix ^ ascii_of s
   755 
   756 fun ho_type_from_ho_term type_enc pred_sym ary =
   757   let
   758     fun to_atype ty =
   759       AType ((make_simple_type (generic_mangled_type_name fst ty),
   760               generic_mangled_type_name snd ty))
   761     fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1))
   762     fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty
   763       | to_fo ary (ATerm (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys
   764       | to_fo ary _ = raise Fail "unexpected type abstraction"
   765     fun to_ho (ty as ATerm ((s, _), tys)) =
   766         if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty
   767       | to_ho _ = raise Fail "unexpected type abstraction"
   768   in if is_type_enc_higher_order type_enc then to_ho else to_fo ary end
   769 
   770 fun ho_type_from_typ format type_enc pred_sym ary =
   771   ho_type_from_ho_term type_enc pred_sym ary
   772   o ho_term_from_typ format type_enc
   773 
   774 fun mangled_const_name format type_enc T_args (s, s') =
   775   let
   776     val ty_args = T_args |> map_filter (ho_term_for_type_arg format type_enc)
   777     fun type_suffix f g =
   778       fold_rev (curry (op ^) o g o prefix mangled_type_sep
   779                 o generic_mangled_type_name f) ty_args ""
   780   in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end
   781 
   782 val parse_mangled_ident =
   783   Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
   784 
   785 fun parse_mangled_type x =
   786   (parse_mangled_ident
   787    -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
   788                     [] >> ATerm) x
   789 and parse_mangled_types x =
   790   (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x
   791 
   792 fun unmangled_type s =
   793   s |> suffix ")" |> raw_explode
   794     |> Scan.finite Symbol.stopper
   795            (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^
   796                                                 quote s)) parse_mangled_type))
   797     |> fst
   798 
   799 val unmangled_const_name = space_explode mangled_type_sep #> hd
   800 fun unmangled_const s =
   801   let val ss = space_explode mangled_type_sep s in
   802     (hd ss, map unmangled_type (tl ss))
   803   end
   804 
   805 fun introduce_proxies type_enc =
   806   let
   807     fun intro top_level (CombApp (tm1, tm2)) =
   808         CombApp (intro top_level tm1, intro false tm2)
   809       | intro top_level (CombConst (name as (s, _), T, T_args)) =
   810         (case proxify_const s of
   811            SOME proxy_base =>
   812            if top_level orelse is_type_enc_higher_order type_enc then
   813              case (top_level, s) of
   814                (_, "c_False") => (`I tptp_false, [])
   815              | (_, "c_True") => (`I tptp_true, [])
   816              | (false, "c_Not") => (`I tptp_not, [])
   817              | (false, "c_conj") => (`I tptp_and, [])
   818              | (false, "c_disj") => (`I tptp_or, [])
   819              | (false, "c_implies") => (`I tptp_implies, [])
   820              | (false, s) =>
   821                if is_tptp_equal s then (`I tptp_equal, [])
   822                else (proxy_base |>> prefix const_prefix, T_args)
   823              | _ => (name, [])
   824            else
   825              (proxy_base |>> prefix const_prefix, T_args)
   826           | NONE => (name, T_args))
   827         |> (fn (name, T_args) => CombConst (name, T, T_args))
   828       | intro _ (CombAbs (bound, tm)) = CombAbs (bound, intro false tm)
   829       | intro _ tm = tm
   830   in intro true end
   831 
   832 fun combformula_from_prop thy type_enc eq_as_iff =
   833   let
   834     fun do_term bs t atomic_types =
   835       combterm_from_term thy bs (Envir.eta_contract t)
   836       |>> (introduce_proxies type_enc #> AAtom)
   837       ||> union (op =) atomic_types
   838     fun do_quant bs q s T t' =
   839       let val s = singleton (Name.variant_list (map fst bs)) s in
   840         do_formula ((s, T) :: bs) t'
   841         #>> mk_aquant q [(`make_bound_var s, SOME T)]
   842       end
   843     and do_conn bs c t1 t2 =
   844       do_formula bs t1 ##>> do_formula bs t2 #>> uncurry (mk_aconn c)
   845     and do_formula bs t =
   846       case t of
   847         @{const Trueprop} $ t1 => do_formula bs t1
   848       | @{const Not} $ t1 => do_formula bs t1 #>> mk_anot
   849       | Const (@{const_name All}, _) $ Abs (s, T, t') =>
   850         do_quant bs AForall s T t'
   851       | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
   852         do_quant bs AExists s T t'
   853       | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2
   854       | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2
   855       | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2
   856       | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
   857         if eq_as_iff then do_conn bs AIff t1 t2 else do_term bs t
   858       | _ => do_term bs t
   859   in do_formula [] end
   860 
   861 fun presimplify_term _ [] t = t
   862   | presimplify_term ctxt presimp_consts t =
   863     t |> exists_Const (member (op =) presimp_consts o fst) t
   864          ? (Skip_Proof.make_thm (Proof_Context.theory_of ctxt)
   865             #> Meson.presimplify ctxt
   866             #> prop_of)
   867 
   868 fun concealed_bound_name j = sledgehammer_weak_prefix ^ string_of_int j
   869 fun conceal_bounds Ts t =
   870   subst_bounds (map (Free o apfst concealed_bound_name)
   871                     (0 upto length Ts - 1 ~~ Ts), t)
   872 fun reveal_bounds Ts =
   873   subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
   874                     (0 upto length Ts - 1 ~~ Ts))
   875 
   876 fun is_fun_equality (@{const_name HOL.eq},
   877                      Type (_, [Type (@{type_name fun}, _), _])) = true
   878   | is_fun_equality _ = false
   879 
   880 fun extensionalize_term ctxt t =
   881   if exists_Const is_fun_equality t then
   882     let val thy = Proof_Context.theory_of ctxt in
   883       t |> cterm_of thy |> Meson.extensionalize_conv ctxt
   884         |> prop_of |> Logic.dest_equals |> snd
   885     end
   886   else
   887     t
   888 
   889 fun introduce_combinators_in_term ctxt kind t =
   890   let val thy = Proof_Context.theory_of ctxt in
   891     if Meson.is_fol_term thy t then
   892       t
   893     else
   894       let
   895         fun aux Ts t =
   896           case t of
   897             @{const Not} $ t1 => @{const Not} $ aux Ts t1
   898           | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
   899             t0 $ Abs (s, T, aux (T :: Ts) t')
   900           | (t0 as Const (@{const_name All}, _)) $ t1 =>
   901             aux Ts (t0 $ eta_expand Ts t1 1)
   902           | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
   903             t0 $ Abs (s, T, aux (T :: Ts) t')
   904           | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
   905             aux Ts (t0 $ eta_expand Ts t1 1)
   906           | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
   907           | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
   908           | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
   909           | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
   910               $ t1 $ t2 =>
   911             t0 $ aux Ts t1 $ aux Ts t2
   912           | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
   913                    t
   914                  else
   915                    t |> conceal_bounds Ts
   916                      |> Envir.eta_contract
   917                      |> cterm_of thy
   918                      |> Meson_Clausify.introduce_combinators_in_cterm
   919                      |> prop_of |> Logic.dest_equals |> snd
   920                      |> reveal_bounds Ts
   921         val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
   922       in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
   923       handle THM _ =>
   924              (* A type variable of sort "{}" will make abstraction fail. *)
   925              if kind = Conjecture then HOLogic.false_const
   926              else HOLogic.true_const
   927   end
   928 
   929 (* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
   930    same in Sledgehammer to prevent the discovery of unreplayable proofs. *)
   931 fun freeze_term t =
   932   let
   933     fun aux (t $ u) = aux t $ aux u
   934       | aux (Abs (s, T, t)) = Abs (s, T, aux t)
   935       | aux (Var ((s, i), T)) =
   936         Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
   937       | aux t = t
   938   in t |> exists_subterm is_Var t ? aux end
   939 
   940 fun preprocess_prop ctxt type_enc presimp_consts kind t =
   941   let
   942     val thy = Proof_Context.theory_of ctxt
   943     val t = t |> Envir.beta_eta_contract
   944               |> transform_elim_prop
   945               |> Object_Logic.atomize_term thy
   946     val need_trueprop = (fastype_of t = @{typ bool})
   947   in
   948     t |> need_trueprop ? HOLogic.mk_Trueprop
   949       |> Raw_Simplifier.rewrite_term thy (Meson.unfold_set_const_simps ctxt) []
   950       |> extensionalize_term ctxt
   951       |> presimplify_term ctxt presimp_consts
   952       |> perhaps (try (HOLogic.dest_Trueprop))
   953       |> not (is_type_enc_higher_order type_enc)
   954          ? introduce_combinators_in_term ctxt kind
   955   end
   956 
   957 (* making fact and conjecture formulas *)
   958 fun make_formula thy type_enc eq_as_iff name loc kind t =
   959   let
   960     val (combformula, atomic_types) =
   961       combformula_from_prop thy type_enc eq_as_iff t []
   962   in
   963     {name = name, locality = loc, kind = kind, combformula = combformula,
   964      atomic_types = atomic_types}
   965   end
   966 
   967 fun make_fact ctxt format type_enc eq_as_iff preproc presimp_consts
   968               ((name, loc), t) =
   969   let val thy = Proof_Context.theory_of ctxt in
   970     case t |> preproc ? preprocess_prop ctxt type_enc presimp_consts Axiom
   971            |> make_formula thy type_enc (eq_as_iff andalso format <> CNF) name
   972                            loc Axiom of
   973       formula as {combformula = AAtom (CombConst ((s, _), _, _)), ...} =>
   974       if s = tptp_true then NONE else SOME formula
   975     | formula => SOME formula
   976   end
   977 
   978 fun make_conjecture ctxt format prem_kind type_enc preproc presimp_consts ts =
   979   let
   980     val thy = Proof_Context.theory_of ctxt
   981     val last = length ts - 1
   982   in
   983     map2 (fn j => fn t =>
   984              let
   985                val (kind, maybe_negate) =
   986                  if j = last then
   987                    (Conjecture, I)
   988                  else
   989                    (prem_kind,
   990                     if prem_kind = Conjecture then update_combformula mk_anot
   991                     else I)
   992               in
   993                 t |> preproc ?
   994                      (preprocess_prop ctxt type_enc presimp_consts kind #> freeze_term)
   995                   |> make_formula thy type_enc (format <> CNF) (string_of_int j)
   996                                   Local kind
   997                   |> maybe_negate
   998               end)
   999          (0 upto last) ts
  1000   end
  1001 
  1002 (** Finite and infinite type inference **)
  1003 
  1004 fun deep_freeze_atyp (TVar (_, S)) = TFree ("v", S)
  1005   | deep_freeze_atyp T = T
  1006 val deep_freeze_type = map_atyps deep_freeze_atyp
  1007 
  1008 (* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are
  1009    dangerous because their "exhaust" properties can easily lead to unsound ATP
  1010    proofs. On the other hand, all HOL infinite types can be given the same
  1011    models in first-order logic (via Löwenheim-Skolem). *)
  1012 
  1013 fun should_encode_type ctxt (nonmono_Ts as _ :: _) _ T =
  1014     exists (curry (type_instance ctxt) (deep_freeze_type T)) nonmono_Ts
  1015   | should_encode_type _ _ All_Types _ = true
  1016   | should_encode_type ctxt _ Fin_Nonmono_Types T =
  1017     is_type_surely_finite ctxt false T
  1018   | should_encode_type _ _ _ _ = false
  1019 
  1020 fun should_predicate_on_type ctxt nonmono_Ts (Preds (_, level, heaviness))
  1021                              should_predicate_on_var T =
  1022     (heaviness = Heavyweight orelse should_predicate_on_var ()) andalso
  1023     should_encode_type ctxt nonmono_Ts level T
  1024   | should_predicate_on_type _ _ _ _ _ = false
  1025 
  1026 fun is_var_or_bound_var (CombConst ((s, _), _, _)) =
  1027     String.isPrefix bound_var_prefix s
  1028   | is_var_or_bound_var (CombVar _) = true
  1029   | is_var_or_bound_var _ = false
  1030 
  1031 datatype tag_site =
  1032   Top_Level of bool option |
  1033   Eq_Arg of bool option |
  1034   Elsewhere
  1035 
  1036 fun should_tag_with_type _ _ _ (Top_Level _) _ _ = false
  1037   | should_tag_with_type ctxt nonmono_Ts (Tags (poly, level, heaviness)) site
  1038                          u T =
  1039     (case heaviness of
  1040        Heavyweight => should_encode_type ctxt nonmono_Ts level T
  1041      | Lightweight =>
  1042        case (site, is_var_or_bound_var u) of
  1043          (Eq_Arg pos, true) =>
  1044          (* The first disjunct prevents a subtle soundness issue explained in
  1045             Blanchette's Ph.D. thesis. See also
  1046             "formula_lines_for_lightweight_tags_sym_decl". *)
  1047          (pos <> SOME false andalso poly = Polymorphic andalso
  1048           level <> All_Types andalso heaviness = Lightweight andalso
  1049           exists (fn T' => type_instance ctxt (T', T)) nonmono_Ts) orelse
  1050          should_encode_type ctxt nonmono_Ts level T
  1051        | _ => false)
  1052   | should_tag_with_type _ _ _ _ _ _ = false
  1053 
  1054 fun homogenized_type ctxt nonmono_Ts level =
  1055   let
  1056     val should_encode = should_encode_type ctxt nonmono_Ts level
  1057     fun homo 0 T = if should_encode T then T else homo_infinite_type
  1058       | homo ary (Type (@{type_name fun}, [T1, T2])) =
  1059         homo 0 T1 --> homo (ary - 1) T2
  1060       | homo _ _ = raise Fail "expected function type"
  1061   in homo end
  1062 
  1063 (** "hBOOL" and "hAPP" **)
  1064 
  1065 type sym_info =
  1066   {pred_sym : bool, min_ary : int, max_ary : int, types : typ list}
  1067 
  1068 fun add_combterm_syms_to_table ctxt explicit_apply =
  1069   let
  1070     fun consider_var_arity const_T var_T max_ary =
  1071       let
  1072         fun iter ary T =
  1073           if ary = max_ary orelse type_instance ctxt (var_T, T) orelse
  1074              type_instance ctxt (T, var_T) then
  1075             ary
  1076           else
  1077             iter (ary + 1) (range_type T)
  1078       in iter 0 const_T end
  1079     fun add_var_or_bound_var T (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
  1080       if explicit_apply = NONE andalso
  1081          (can dest_funT T orelse T = @{typ bool}) then
  1082         let
  1083           val bool_vars' = bool_vars orelse body_type T = @{typ bool}
  1084           fun repair_min_arity {pred_sym, min_ary, max_ary, types} =
  1085             {pred_sym = pred_sym andalso not bool_vars',
  1086              min_ary = fold (fn T' => consider_var_arity T' T) types min_ary,
  1087              max_ary = max_ary, types = types}
  1088           val fun_var_Ts' =
  1089             fun_var_Ts |> can dest_funT T ? insert_type ctxt I T
  1090         in
  1091           if bool_vars' = bool_vars andalso
  1092              pointer_eq (fun_var_Ts', fun_var_Ts) then
  1093             accum
  1094           else
  1095             ((bool_vars', fun_var_Ts'), Symtab.map (K repair_min_arity) sym_tab)
  1096         end
  1097       else
  1098         accum
  1099     fun add top_level tm (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
  1100       let val (head, args) = strip_combterm_comb tm in
  1101         (case head of
  1102            CombConst ((s, _), T, _) =>
  1103            if String.isPrefix bound_var_prefix s then
  1104              add_var_or_bound_var T accum
  1105            else
  1106              let val ary = length args in
  1107                ((bool_vars, fun_var_Ts),
  1108                 case Symtab.lookup sym_tab s of
  1109                   SOME {pred_sym, min_ary, max_ary, types} =>
  1110                   let
  1111                     val pred_sym =
  1112                       pred_sym andalso top_level andalso not bool_vars
  1113                     val types' = types |> insert_type ctxt I T
  1114                     val min_ary =
  1115                       if is_some explicit_apply orelse
  1116                          pointer_eq (types', types) then
  1117                         min_ary
  1118                       else
  1119                         fold (consider_var_arity T) fun_var_Ts min_ary
  1120                   in
  1121                     Symtab.update (s, {pred_sym = pred_sym,
  1122                                        min_ary = Int.min (ary, min_ary),
  1123                                        max_ary = Int.max (ary, max_ary),
  1124                                        types = types'})
  1125                                   sym_tab
  1126                   end
  1127                 | NONE =>
  1128                   let
  1129                     val pred_sym = top_level andalso not bool_vars
  1130                     val min_ary =
  1131                       case explicit_apply of
  1132                         SOME true => 0
  1133                       | SOME false => ary
  1134                       | NONE => fold (consider_var_arity T) fun_var_Ts ary
  1135                   in
  1136                     Symtab.update_new (s, {pred_sym = pred_sym,
  1137                                            min_ary = min_ary, max_ary = ary,
  1138                                            types = [T]})
  1139                                       sym_tab
  1140                   end)
  1141              end
  1142          | CombVar (_, T) => add_var_or_bound_var T accum
  1143          | _ => accum)
  1144         |> fold (add false) args
  1145       end
  1146   in add true end
  1147 fun add_fact_syms_to_table ctxt explicit_apply =
  1148   fact_lift (formula_fold NONE
  1149                           (K (add_combterm_syms_to_table ctxt explicit_apply)))
  1150 
  1151 val default_sym_tab_entries : (string * sym_info) list =
  1152   (prefixed_predicator_name,
  1153    {pred_sym = true, min_ary = 1, max_ary = 1, types = []}) ::
  1154   ([tptp_false, tptp_true]
  1155    |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, types = []})) @
  1156   ([tptp_equal, tptp_old_equal]
  1157    |> map (rpair {pred_sym = true, min_ary = 2, max_ary = 2, types = []}))
  1158 
  1159 fun sym_table_for_facts ctxt explicit_apply facts =
  1160   ((false, []), Symtab.empty)
  1161   |> fold (add_fact_syms_to_table ctxt explicit_apply) facts |> snd
  1162   |> fold Symtab.update default_sym_tab_entries
  1163 
  1164 fun min_arity_of sym_tab s =
  1165   case Symtab.lookup sym_tab s of
  1166     SOME ({min_ary, ...} : sym_info) => min_ary
  1167   | NONE =>
  1168     case strip_prefix_and_unascii const_prefix s of
  1169       SOME s =>
  1170       let val s = s |> unmangled_const_name |> invert_const in
  1171         if s = predicator_name then 1
  1172         else if s = app_op_name then 2
  1173         else if s = type_pred_name then 1
  1174         else 0
  1175       end
  1176     | NONE => 0
  1177 
  1178 (* True if the constant ever appears outside of the top-level position in
  1179    literals, or if it appears with different arities (e.g., because of different
  1180    type instantiations). If false, the constant always receives all of its
  1181    arguments and is used as a predicate. *)
  1182 fun is_pred_sym sym_tab s =
  1183   case Symtab.lookup sym_tab s of
  1184     SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
  1185     pred_sym andalso min_ary = max_ary
  1186   | NONE => false
  1187 
  1188 val predicator_combconst =
  1189   CombConst (`make_fixed_const predicator_name, @{typ "bool => bool"}, [])
  1190 fun predicator tm = CombApp (predicator_combconst, tm)
  1191 
  1192 fun introduce_predicators_in_combterm sym_tab tm =
  1193   case strip_combterm_comb tm of
  1194     (CombConst ((s, _), _, _), _) =>
  1195     if is_pred_sym sym_tab s then tm else predicator tm
  1196   | _ => predicator tm
  1197 
  1198 fun list_app head args = fold (curry (CombApp o swap)) args head
  1199 
  1200 val app_op = `make_fixed_const app_op_name
  1201 
  1202 fun explicit_app arg head =
  1203   let
  1204     val head_T = combtyp_of head
  1205     val (arg_T, res_T) = dest_funT head_T
  1206     val explicit_app =
  1207       CombConst (app_op, head_T --> head_T, [arg_T, res_T])
  1208   in list_app explicit_app [head, arg] end
  1209 fun list_explicit_app head args = fold explicit_app args head
  1210 
  1211 fun introduce_explicit_apps_in_combterm sym_tab =
  1212   let
  1213     fun aux tm =
  1214       case strip_combterm_comb tm of
  1215         (head as CombConst ((s, _), _, _), args) =>
  1216         args |> map aux
  1217              |> chop (min_arity_of sym_tab s)
  1218              |>> list_app head
  1219              |-> list_explicit_app
  1220       | (head, args) => list_explicit_app head (map aux args)
  1221   in aux end
  1222 
  1223 fun chop_fun 0 T = ([], T)
  1224   | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
  1225     chop_fun (n - 1) ran_T |>> cons dom_T
  1226   | chop_fun _ _ = raise Fail "unexpected non-function"
  1227 
  1228 fun filter_type_args _ _ _ [] = []
  1229   | filter_type_args thy s arity T_args =
  1230     let
  1231       (* will throw "TYPE" for pseudo-constants *)
  1232       val U = if s = app_op_name then
  1233                 @{typ "('a => 'b) => 'a => 'b"} |> Logic.varifyT_global
  1234               else
  1235                 s |> Sign.the_const_type thy
  1236     in
  1237       case Term.add_tvarsT (U |> chop_fun arity |> snd) [] of
  1238         [] => []
  1239       | res_U_vars =>
  1240         let val U_args = (s, U) |> Sign.const_typargs thy in
  1241           U_args ~~ T_args
  1242           |> map (fn (U, T) =>
  1243                      if member (op =) res_U_vars (dest_TVar U) then T
  1244                      else dummyT)
  1245         end
  1246     end
  1247     handle TYPE _ => T_args
  1248 
  1249 fun enforce_type_arg_policy_in_combterm ctxt format type_enc =
  1250   let
  1251     val thy = Proof_Context.theory_of ctxt
  1252     fun aux arity (CombApp (tm1, tm2)) =
  1253         CombApp (aux (arity + 1) tm1, aux 0 tm2)
  1254       | aux arity (CombConst (name as (s, _), T, T_args)) =
  1255         (case strip_prefix_and_unascii const_prefix s of
  1256            NONE => (name, T_args)
  1257          | SOME s'' =>
  1258            let
  1259              val s'' = invert_const s''
  1260              fun filtered_T_args false = T_args
  1261                | filtered_T_args true = filter_type_args thy s'' arity T_args
  1262            in
  1263              case type_arg_policy type_enc s'' of
  1264                Explicit_Type_Args drop_args =>
  1265                (name, filtered_T_args drop_args)
  1266              | Mangled_Type_Args drop_args =>
  1267                (mangled_const_name format type_enc (filtered_T_args drop_args)
  1268                                    name, [])
  1269              | No_Type_Args => (name, [])
  1270            end)
  1271         |> (fn (name, T_args) => CombConst (name, T, T_args))
  1272       | aux _ (CombAbs (bound, tm)) = CombAbs (bound, aux 0 tm)
  1273       | aux _ tm = tm
  1274   in aux 0 end
  1275 
  1276 fun repair_combterm ctxt format type_enc sym_tab =
  1277   not (is_type_enc_higher_order type_enc)
  1278   ? (introduce_explicit_apps_in_combterm sym_tab
  1279      #> introduce_predicators_in_combterm sym_tab)
  1280   #> enforce_type_arg_policy_in_combterm ctxt format type_enc
  1281 fun repair_fact ctxt format type_enc sym_tab =
  1282   update_combformula (formula_map
  1283       (repair_combterm ctxt format type_enc sym_tab))
  1284 
  1285 (** Helper facts **)
  1286 
  1287 (* The Boolean indicates that a fairly sound type encoding is needed. *)
  1288 val helper_table =
  1289   [(("COMBI", false), @{thms Meson.COMBI_def}),
  1290    (("COMBK", false), @{thms Meson.COMBK_def}),
  1291    (("COMBB", false), @{thms Meson.COMBB_def}),
  1292    (("COMBC", false), @{thms Meson.COMBC_def}),
  1293    (("COMBS", false), @{thms Meson.COMBS_def}),
  1294    (("fequal", true),
  1295     (* This is a lie: Higher-order equality doesn't need a sound type encoding.
  1296        However, this is done so for backward compatibility: Including the
  1297        equality helpers by default in Metis breaks a few existing proofs. *)
  1298     @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1]
  1299            fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),
  1300    (("fFalse", false), [@{lemma "~ fFalse" by (unfold fFalse_def) fast}]),
  1301    (("fFalse", true), @{thms True_or_False}),
  1302    (("fTrue", false), [@{lemma "fTrue" by (unfold fTrue_def) fast}]),
  1303    (("fTrue", true), @{thms True_or_False}),
  1304    (("fNot", false),
  1305     @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1]
  1306            fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),
  1307    (("fconj", false),
  1308     @{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q"
  1309         by (unfold fconj_def) fast+}),
  1310    (("fdisj", false),
  1311     @{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q"
  1312         by (unfold fdisj_def) fast+}),
  1313    (("fimplies", false),
  1314     @{lemma "P | fimplies P Q" "~ Q | fimplies P Q" "~ fimplies P Q | ~ P | Q"
  1315         by (unfold fimplies_def) fast+}),
  1316    (("If", true), @{thms if_True if_False True_or_False})]
  1317   |> map (apsnd (map zero_var_indexes))
  1318 
  1319 val type_tag = `make_fixed_const type_tag_name
  1320 
  1321 fun type_tag_idempotence_fact () =
  1322   let
  1323     fun var s = ATerm (`I s, [])
  1324     fun tag tm = ATerm (type_tag, [var "T", tm])
  1325     val tagged_a = tag (var "A")
  1326   in
  1327     Formula (type_tag_idempotence_helper_name, Axiom,
  1328              AAtom (ATerm (`I tptp_equal, [tag tagged_a, tagged_a]))
  1329              |> close_formula_universally, simp_info, NONE)
  1330   end
  1331 
  1332 fun should_specialize_helper type_enc t =
  1333   polymorphism_of_type_enc type_enc = Mangled_Monomorphic andalso
  1334   level_of_type_enc type_enc <> No_Types andalso
  1335   not (null (Term.hidden_polymorphism t))
  1336 
  1337 fun helper_facts_for_sym ctxt format type_enc (s, {types, ...} : sym_info) =
  1338   case strip_prefix_and_unascii const_prefix s of
  1339     SOME mangled_s =>
  1340     let
  1341       val thy = Proof_Context.theory_of ctxt
  1342       val unmangled_s = mangled_s |> unmangled_const_name
  1343       fun dub needs_fairly_sound j k =
  1344         (unmangled_s ^ "_" ^ string_of_int j ^ "_" ^ string_of_int k ^
  1345          (if mangled_s = unmangled_s then "" else "_" ^ ascii_of mangled_s) ^
  1346          (if needs_fairly_sound then typed_helper_suffix
  1347           else untyped_helper_suffix),
  1348          Helper)
  1349       fun dub_and_inst needs_fairly_sound (th, j) =
  1350         let val t = prop_of th in
  1351           if should_specialize_helper type_enc t then
  1352             map (fn T => specialize_type thy (invert_const unmangled_s, T) t)
  1353                 types
  1354           else
  1355             [t]
  1356         end
  1357         |> map (fn (k, t) => (dub needs_fairly_sound j k, t)) o tag_list 1
  1358       val make_facts =
  1359         map_filter (make_fact ctxt format type_enc false false [])
  1360       val fairly_sound = is_type_enc_fairly_sound type_enc
  1361     in
  1362       helper_table
  1363       |> maps (fn ((helper_s, needs_fairly_sound), ths) =>
  1364                   if helper_s <> unmangled_s orelse
  1365                      (needs_fairly_sound andalso not fairly_sound) then
  1366                     []
  1367                   else
  1368                     ths ~~ (1 upto length ths)
  1369                     |> maps (dub_and_inst needs_fairly_sound)
  1370                     |> make_facts)
  1371     end
  1372   | NONE => []
  1373 fun helper_facts_for_sym_table ctxt format type_enc sym_tab =
  1374   Symtab.fold_rev (append o helper_facts_for_sym ctxt format type_enc) sym_tab
  1375                   []
  1376 
  1377 (***************************************************************)
  1378 (* Type Classes Present in the Axiom or Conjecture Clauses     *)
  1379 (***************************************************************)
  1380 
  1381 fun set_insert (x, s) = Symtab.update (x, ()) s
  1382 
  1383 fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
  1384 
  1385 (* Remove this trivial type class (FIXME: similar code elsewhere) *)
  1386 fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset
  1387 
  1388 fun classes_of_terms get_Ts =
  1389   map (map snd o get_Ts)
  1390   #> List.foldl add_classes Symtab.empty
  1391   #> delete_type #> Symtab.keys
  1392 
  1393 val tfree_classes_of_terms = classes_of_terms OldTerm.term_tfrees
  1394 val tvar_classes_of_terms = classes_of_terms OldTerm.term_tvars
  1395 
  1396 fun fold_type_constrs f (Type (s, Ts)) x =
  1397     fold (fold_type_constrs f) Ts (f (s, x))
  1398   | fold_type_constrs _ _ x = x
  1399 
  1400 (*Type constructors used to instantiate overloaded constants are the only ones needed.*)
  1401 fun add_type_constrs_in_term thy =
  1402   let
  1403     fun add (Const (@{const_name Meson.skolem}, _) $ _) = I
  1404       | add (t $ u) = add t #> add u
  1405       | add (Const (x as (s, _))) =
  1406         if String.isPrefix skolem_const_prefix s then I
  1407         else x |> Sign.const_typargs thy |> fold (fold_type_constrs set_insert)
  1408       | add (Abs (_, _, u)) = add u
  1409       | add _ = I
  1410   in add end
  1411 
  1412 fun type_constrs_of_terms thy ts =
  1413   Symtab.keys (fold (add_type_constrs_in_term thy) ts Symtab.empty)
  1414 
  1415 fun translate_formulas ctxt format prem_kind type_enc preproc hyp_ts concl_t
  1416                        facts =
  1417   let
  1418     val thy = Proof_Context.theory_of ctxt
  1419     val fact_ts = facts |> map snd
  1420     val presimp_consts = Meson.presimplified_consts ctxt
  1421     val make_fact = make_fact ctxt format type_enc true preproc presimp_consts
  1422     val (facts, fact_names) =
  1423       facts |> map (fn (name, t) => (name, t) |> make_fact |> rpair name)
  1424             |> map_filter (try (apfst the))
  1425             |> ListPair.unzip
  1426     (* Remove existing facts from the conjecture, as this can dramatically
  1427        boost an ATP's performance (for some reason). *)
  1428     val hyp_ts =
  1429       hyp_ts
  1430       |> map (fn t => if member (op aconv) fact_ts t then @{prop True} else t)
  1431     val goal_t = Logic.list_implies (hyp_ts, concl_t)
  1432     val all_ts = goal_t :: fact_ts
  1433     val subs = tfree_classes_of_terms all_ts
  1434     val supers = tvar_classes_of_terms all_ts
  1435     val tycons = type_constrs_of_terms thy all_ts
  1436     val conjs =
  1437       hyp_ts @ [concl_t]
  1438       |> make_conjecture ctxt format prem_kind type_enc preproc presimp_consts
  1439     val (supers', arity_clauses) =
  1440       if level_of_type_enc type_enc = No_Types then ([], [])
  1441       else make_arity_clauses thy tycons supers
  1442     val class_rel_clauses = make_class_rel_clauses thy subs supers'
  1443   in
  1444     (fact_names |> map single, (conjs, facts, class_rel_clauses, arity_clauses))
  1445   end
  1446 
  1447 fun fo_literal_from_type_literal (TyLitVar (class, name)) =
  1448     (true, ATerm (class, [ATerm (name, [])]))
  1449   | fo_literal_from_type_literal (TyLitFree (class, name)) =
  1450     (true, ATerm (class, [ATerm (name, [])]))
  1451 
  1452 fun formula_from_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
  1453 
  1454 val type_pred = `make_fixed_const type_pred_name
  1455 
  1456 fun type_pred_combterm ctxt format type_enc T tm =
  1457   CombApp (CombConst (type_pred, T --> @{typ bool}, [T])
  1458            |> enforce_type_arg_policy_in_combterm ctxt format type_enc, tm)
  1459 
  1460 fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum
  1461   | is_var_positively_naked_in_term name _ (ATerm ((s, _), tms)) accum =
  1462     accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm (name, [])))
  1463   | is_var_positively_naked_in_term name _ _ _ = true
  1464 fun should_predicate_on_var_in_formula pos phi (SOME true) name =
  1465     formula_fold pos (is_var_positively_naked_in_term name) phi false
  1466   | should_predicate_on_var_in_formula _ _ _ _ = true
  1467 
  1468 fun mk_aterm format type_enc name T_args args =
  1469   ATerm (name, map_filter (ho_term_for_type_arg format type_enc) T_args @ args)
  1470 
  1471 fun tag_with_type ctxt format nonmono_Ts type_enc pos T tm =
  1472   CombConst (type_tag, T --> T, [T])
  1473   |> enforce_type_arg_policy_in_combterm ctxt format type_enc
  1474   |> ho_term_from_combterm ctxt format nonmono_Ts type_enc (Top_Level pos)
  1475   |> (fn ATerm (s, tms) => ATerm (s, tms @ [tm]))
  1476 and ho_term_from_combterm ctxt format nonmono_Ts type_enc =
  1477   let
  1478     fun aux site u =
  1479       let
  1480         val (head, args) = strip_combterm_comb u
  1481         val pos =
  1482           case site of
  1483             Top_Level pos => pos
  1484           | Eq_Arg pos => pos
  1485           | Elsewhere => NONE
  1486         val t =
  1487           case head of
  1488             CombConst (name as (s, _), _, T_args) =>
  1489             let
  1490               val arg_site = if is_tptp_equal s then Eq_Arg pos else Elsewhere
  1491             in
  1492               mk_aterm format type_enc name T_args (map (aux arg_site) args)
  1493             end
  1494           | CombVar (name, _) => mk_aterm format type_enc name [] []
  1495           | CombAbs ((name, T), tm) =>
  1496             AAbs ((name, ho_type_from_typ format type_enc true 0 T), aux Elsewhere tm)
  1497           | CombApp _ => raise Fail "impossible \"CombApp\""
  1498         val T = combtyp_of u
  1499       in
  1500         t |> (if should_tag_with_type ctxt nonmono_Ts type_enc site u T then
  1501                 tag_with_type ctxt format nonmono_Ts type_enc pos T
  1502               else
  1503                 I)
  1504       end
  1505   in aux end
  1506 and formula_from_combformula ctxt format nonmono_Ts type_enc
  1507                              should_predicate_on_var =
  1508   let
  1509     fun do_term pos =
  1510       ho_term_from_combterm ctxt format nonmono_Ts type_enc (Top_Level pos)
  1511     val do_bound_type =
  1512       case type_enc of
  1513         Simple_Types (_, level) =>
  1514         homogenized_type ctxt nonmono_Ts level 0
  1515         #> ho_type_from_typ format type_enc false 0 #> SOME
  1516       | _ => K NONE
  1517     fun do_out_of_bound_type pos phi universal (name, T) =
  1518       if should_predicate_on_type ctxt nonmono_Ts type_enc
  1519              (fn () => should_predicate_on_var pos phi universal name) T then
  1520         CombVar (name, T)
  1521         |> type_pred_combterm ctxt format type_enc T
  1522         |> do_term pos |> AAtom |> SOME
  1523       else
  1524         NONE
  1525     fun do_formula pos (AQuant (q, xs, phi)) =
  1526         let
  1527           val phi = phi |> do_formula pos
  1528           val universal = Option.map (q = AExists ? not) pos
  1529         in
  1530           AQuant (q, xs |> map (apsnd (fn NONE => NONE
  1531                                         | SOME T => do_bound_type T)),
  1532                   (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))
  1533                       (map_filter
  1534                            (fn (_, NONE) => NONE
  1535                              | (s, SOME T) =>
  1536                                do_out_of_bound_type pos phi universal (s, T))
  1537                            xs)
  1538                       phi)
  1539         end
  1540       | do_formula pos (AConn conn) = aconn_map pos do_formula conn
  1541       | do_formula pos (AAtom tm) = AAtom (do_term pos tm)
  1542   in do_formula end
  1543 
  1544 fun bound_tvars type_enc Ts =
  1545   mk_ahorn (map (formula_from_fo_literal o fo_literal_from_type_literal)
  1546                 (type_literals_for_types type_enc add_sorts_on_tvar Ts))
  1547 
  1548 (* Each fact is given a unique fact number to avoid name clashes (e.g., because
  1549    of monomorphization). The TPTP explicitly forbids name clashes, and some of
  1550    the remote provers might care. *)
  1551 fun formula_line_for_fact ctxt format prefix encode freshen pos nonmono_Ts
  1552         type_enc (j, {name, locality, kind, combformula, atomic_types}) =
  1553   (prefix ^ (if freshen then string_of_int j ^ "_" else "") ^ encode name,
  1554    kind,
  1555    combformula
  1556    |> close_combformula_universally
  1557    |> formula_from_combformula ctxt format nonmono_Ts type_enc
  1558                                should_predicate_on_var_in_formula
  1559                                (if pos then SOME true else NONE)
  1560    |> bound_tvars type_enc atomic_types
  1561    |> close_formula_universally,
  1562    NONE,
  1563    case locality of
  1564      Intro => intro_info
  1565    | Elim => elim_info
  1566    | Simp => simp_info
  1567    | _ => NONE)
  1568   |> Formula
  1569 
  1570 fun formula_line_for_class_rel_clause ({name, subclass, superclass, ...}
  1571                                        : class_rel_clause) =
  1572   let val ty_arg = ATerm (`I "T", []) in
  1573     Formula (class_rel_clause_prefix ^ ascii_of name, Axiom,
  1574              AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
  1575                                AAtom (ATerm (superclass, [ty_arg]))])
  1576              |> close_formula_universally, intro_info, NONE)
  1577   end
  1578 
  1579 fun fo_literal_from_arity_literal (TConsLit (c, t, args)) =
  1580     (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
  1581   | fo_literal_from_arity_literal (TVarLit (c, sort)) =
  1582     (false, ATerm (c, [ATerm (sort, [])]))
  1583 
  1584 fun formula_line_for_arity_clause ({name, prem_lits, concl_lits, ...}
  1585                                    : arity_clause) =
  1586   Formula (arity_clause_prefix ^ name, Axiom,
  1587            mk_ahorn (map (formula_from_fo_literal o apfst not
  1588                           o fo_literal_from_arity_literal) prem_lits)
  1589                     (formula_from_fo_literal
  1590                          (fo_literal_from_arity_literal concl_lits))
  1591            |> close_formula_universally, intro_info, NONE)
  1592 
  1593 fun formula_line_for_conjecture ctxt format nonmono_Ts type_enc
  1594         ({name, kind, combformula, atomic_types, ...} : translated_formula) =
  1595   Formula (conjecture_prefix ^ name, kind,
  1596            formula_from_combformula ctxt format nonmono_Ts type_enc
  1597                should_predicate_on_var_in_formula (SOME false)
  1598                (close_combformula_universally combformula)
  1599            |> bound_tvars type_enc atomic_types
  1600            |> close_formula_universally, NONE, NONE)
  1601 
  1602 fun free_type_literals type_enc ({atomic_types, ...} : translated_formula) =
  1603   atomic_types |> type_literals_for_types type_enc add_sorts_on_tfree
  1604                |> map fo_literal_from_type_literal
  1605 
  1606 fun formula_line_for_free_type j lit =
  1607   Formula (tfree_clause_prefix ^ string_of_int j, Hypothesis,
  1608            formula_from_fo_literal lit, NONE, NONE)
  1609 fun formula_lines_for_free_types type_enc facts =
  1610   let
  1611     val litss = map (free_type_literals type_enc) facts
  1612     val lits = fold (union (op =)) litss []
  1613   in map2 formula_line_for_free_type (0 upto length lits - 1) lits end
  1614 
  1615 (** Symbol declarations **)
  1616 
  1617 fun should_declare_sym type_enc pred_sym s =
  1618   is_tptp_user_symbol s andalso not (String.isPrefix bound_var_prefix s) andalso
  1619   (case type_enc of
  1620      Simple_Types _ => true
  1621    | Tags (_, _, Lightweight) => true
  1622    | _ => not pred_sym)
  1623 
  1624 fun sym_decl_table_for_facts ctxt type_enc repaired_sym_tab (conjs, facts) =
  1625   let
  1626     fun add_combterm in_conj tm =
  1627       let val (head, args) = strip_combterm_comb tm in
  1628         (case head of
  1629            CombConst ((s, s'), T, T_args) =>
  1630            let val pred_sym = is_pred_sym repaired_sym_tab s in
  1631              if should_declare_sym type_enc pred_sym s then
  1632                Symtab.map_default (s, [])
  1633                    (insert_type ctxt #3 (s', T_args, T, pred_sym, length args,
  1634                                          in_conj))
  1635              else
  1636                I
  1637            end
  1638          | _ => I)
  1639         #> fold (add_combterm in_conj) args
  1640       end
  1641     fun add_fact in_conj =
  1642       fact_lift (formula_fold NONE (K (add_combterm in_conj)))
  1643   in
  1644     Symtab.empty
  1645     |> is_type_enc_fairly_sound type_enc
  1646        ? (fold (add_fact true) conjs #> fold (add_fact false) facts)
  1647   end
  1648 
  1649 (* This inference is described in section 2.3 of Claessen et al.'s "Sorting it
  1650    out with monotonicity" paper presented at CADE 2011. *)
  1651 fun add_combterm_nonmonotonic_types _ _ _ _ (SOME false) _ = I
  1652   | add_combterm_nonmonotonic_types ctxt level sound locality _
  1653         (CombApp (CombApp (CombConst ((s, _), Type (_, [T, _]), _), tm1),
  1654                            tm2)) =
  1655     (is_tptp_equal s andalso exists is_var_or_bound_var [tm1, tm2] andalso
  1656      (case level of
  1657         Noninf_Nonmono_Types =>
  1658         not (is_locality_global locality) orelse
  1659         not (is_type_surely_infinite ctxt sound T)
  1660       | Fin_Nonmono_Types => is_type_surely_finite ctxt false T
  1661       | _ => true)) ? insert_type ctxt I (deep_freeze_type T)
  1662   | add_combterm_nonmonotonic_types _ _ _ _ _ _ = I
  1663 fun add_fact_nonmonotonic_types ctxt level sound
  1664         ({kind, locality, combformula, ...} : translated_formula) =
  1665   formula_fold (SOME (kind <> Conjecture))
  1666                (add_combterm_nonmonotonic_types ctxt level sound locality)
  1667                combformula
  1668 fun nonmonotonic_types_for_facts ctxt type_enc sound facts =
  1669   let val level = level_of_type_enc type_enc in
  1670     if level = Noninf_Nonmono_Types orelse level = Fin_Nonmono_Types then
  1671       [] |> fold (add_fact_nonmonotonic_types ctxt level sound) facts
  1672          (* We must add "bool" in case the helper "True_or_False" is added
  1673             later. In addition, several places in the code rely on the list of
  1674             nonmonotonic types not being empty. *)
  1675          |> insert_type ctxt I @{typ bool}
  1676     else
  1677       []
  1678   end
  1679 
  1680 fun decl_line_for_sym ctxt format nonmono_Ts type_enc s
  1681                       (s', T_args, T, pred_sym, ary, _) =
  1682   let
  1683     val (T_arg_Ts, level) =
  1684       case type_enc of
  1685         Simple_Types (_, level) => ([], level)
  1686       | _ => (replicate (length T_args) homo_infinite_type, No_Types)
  1687   in
  1688     Decl (sym_decl_prefix ^ s, (s, s'),
  1689           (T_arg_Ts ---> (T |> homogenized_type ctxt nonmono_Ts level ary))
  1690           |> ho_type_from_typ format type_enc pred_sym (length T_arg_Ts + ary))
  1691   end
  1692 
  1693 fun formula_line_for_preds_sym_decl ctxt format conj_sym_kind nonmono_Ts
  1694         poly_nonmono_Ts type_enc n s j (s', T_args, T, _, ary, in_conj) =
  1695   let
  1696     val (kind, maybe_negate) =
  1697       if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
  1698       else (Axiom, I)
  1699     val (arg_Ts, res_T) = chop_fun ary T
  1700     val num_args = length arg_Ts
  1701     val bound_names =
  1702       1 upto num_args |> map (`I o make_bound_var o string_of_int)
  1703     val bounds =
  1704       bound_names ~~ arg_Ts |> map (fn (name, T) => CombConst (name, T, []))
  1705     val sym_needs_arg_types = n > 1 orelse exists (curry (op =) dummyT) T_args
  1706     fun should_keep_arg_type T =
  1707       sym_needs_arg_types orelse
  1708       not (should_predicate_on_type ctxt nonmono_Ts type_enc (K false) T)
  1709     val bound_Ts =
  1710       arg_Ts |> map (fn T => if should_keep_arg_type T then SOME T else NONE)
  1711   in
  1712     Formula (preds_sym_formula_prefix ^ s ^
  1713              (if n > 1 then "_" ^ string_of_int j else ""), kind,
  1714              CombConst ((s, s'), T, T_args)
  1715              |> fold (curry (CombApp o swap)) bounds
  1716              |> type_pred_combterm ctxt format type_enc res_T
  1717              |> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)
  1718              |> formula_from_combformula ctxt format poly_nonmono_Ts type_enc
  1719                                          (K (K (K (K true)))) (SOME true)
  1720              |> n > 1 ? bound_tvars type_enc (atyps_of T)
  1721              |> close_formula_universally
  1722              |> maybe_negate,
  1723              intro_info, NONE)
  1724   end
  1725 
  1726 fun formula_lines_for_lightweight_tags_sym_decl ctxt format conj_sym_kind
  1727         poly_nonmono_Ts type_enc n s
  1728         (j, (s', T_args, T, pred_sym, ary, in_conj)) =
  1729   let
  1730     val ident_base =
  1731       lightweight_tags_sym_formula_prefix ^ s ^
  1732       (if n > 1 then "_" ^ string_of_int j else "")
  1733     val (kind, maybe_negate) =
  1734       if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
  1735       else (Axiom, I)
  1736     val (arg_Ts, res_T) = chop_fun ary T
  1737     val bound_names =
  1738       1 upto length arg_Ts |> map (`I o make_bound_var o string_of_int)
  1739     val bounds = bound_names |> map (fn name => ATerm (name, []))
  1740     val cst = mk_aterm format type_enc (s, s') T_args
  1741     val atomic_Ts = atyps_of T
  1742     fun eq tms =
  1743       (if pred_sym then AConn (AIff, map AAtom tms)
  1744        else AAtom (ATerm (`I tptp_equal, tms)))
  1745       |> bound_tvars type_enc atomic_Ts
  1746       |> close_formula_universally
  1747       |> maybe_negate
  1748     (* See also "should_tag_with_type". *)
  1749     fun should_encode T =
  1750       should_encode_type ctxt poly_nonmono_Ts All_Types T orelse
  1751       (case type_enc of
  1752          Tags (Polymorphic, level, Lightweight) =>
  1753          level <> All_Types andalso Monomorph.typ_has_tvars T
  1754        | _ => false)
  1755     val tag_with = tag_with_type ctxt format poly_nonmono_Ts type_enc NONE
  1756     val add_formula_for_res =
  1757       if should_encode res_T then
  1758         cons (Formula (ident_base ^ "_res", kind,
  1759                        eq [tag_with res_T (cst bounds), cst bounds],
  1760                        simp_info, NONE))
  1761       else
  1762         I
  1763     fun add_formula_for_arg k =
  1764       let val arg_T = nth arg_Ts k in
  1765         if should_encode arg_T then
  1766           case chop k bounds of
  1767             (bounds1, bound :: bounds2) =>
  1768             cons (Formula (ident_base ^ "_arg" ^ string_of_int (k + 1), kind,
  1769                            eq [cst (bounds1 @ tag_with arg_T bound :: bounds2),
  1770                                cst bounds],
  1771                            simp_info, NONE))
  1772           | _ => raise Fail "expected nonempty tail"
  1773         else
  1774           I
  1775       end
  1776   in
  1777     [] |> not pred_sym ? add_formula_for_res
  1778        |> fold add_formula_for_arg (ary - 1 downto 0)
  1779   end
  1780 
  1781 fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd
  1782 
  1783 fun problem_lines_for_sym_decls ctxt format conj_sym_kind nonmono_Ts
  1784                                 poly_nonmono_Ts type_enc (s, decls) =
  1785   case type_enc of
  1786     Simple_Types _ =>
  1787     decls |> map (decl_line_for_sym ctxt format nonmono_Ts type_enc s)
  1788   | Preds _ =>
  1789     let
  1790       val decls =
  1791         case decls of
  1792           decl :: (decls' as _ :: _) =>
  1793           let val T = result_type_of_decl decl in
  1794             if forall (curry (type_instance ctxt o swap) T
  1795                        o result_type_of_decl) decls' then
  1796               [decl]
  1797             else
  1798               decls
  1799           end
  1800         | _ => decls
  1801       val n = length decls
  1802       val decls =
  1803         decls |> filter (should_predicate_on_type ctxt poly_nonmono_Ts type_enc
  1804                                                   (K true)
  1805                          o result_type_of_decl)
  1806     in
  1807       (0 upto length decls - 1, decls)
  1808       |-> map2 (formula_line_for_preds_sym_decl ctxt format conj_sym_kind
  1809                    nonmono_Ts poly_nonmono_Ts type_enc n s)
  1810     end
  1811   | Tags (_, _, heaviness) =>
  1812     (case heaviness of
  1813        Heavyweight => []
  1814      | Lightweight =>
  1815        let val n = length decls in
  1816          (0 upto n - 1 ~~ decls)
  1817          |> maps (formula_lines_for_lightweight_tags_sym_decl ctxt format
  1818                       conj_sym_kind poly_nonmono_Ts type_enc n s)
  1819        end)
  1820 
  1821 fun problem_lines_for_sym_decl_table ctxt format conj_sym_kind nonmono_Ts
  1822                                      poly_nonmono_Ts type_enc sym_decl_tab =
  1823   sym_decl_tab
  1824   |> Symtab.dest
  1825   |> sort_wrt fst
  1826   |> rpair []
  1827   |-> fold_rev (append o problem_lines_for_sym_decls ctxt format conj_sym_kind
  1828                              nonmono_Ts poly_nonmono_Ts type_enc)
  1829 
  1830 fun needs_type_tag_idempotence (Tags (poly, level, heaviness)) =
  1831     poly <> Mangled_Monomorphic andalso
  1832     ((level = All_Types andalso heaviness = Lightweight) orelse
  1833      level = Noninf_Nonmono_Types orelse level = Fin_Nonmono_Types)
  1834   | needs_type_tag_idempotence _ = false
  1835 
  1836 fun offset_of_heading_in_problem _ [] j = j
  1837   | offset_of_heading_in_problem needle ((heading, lines) :: problem) j =
  1838     if heading = needle then j
  1839     else offset_of_heading_in_problem needle problem (j + length lines)
  1840 
  1841 val implicit_declsN = "Should-be-implicit typings"
  1842 val explicit_declsN = "Explicit typings"
  1843 val factsN = "Relevant facts"
  1844 val class_relsN = "Class relationships"
  1845 val aritiesN = "Arities"
  1846 val helpersN = "Helper facts"
  1847 val conjsN = "Conjectures"
  1848 val free_typesN = "Type variables"
  1849 
  1850 val explicit_apply = NONE (* for experimental purposes *)
  1851 
  1852 fun prepare_atp_problem ctxt format conj_sym_kind prem_kind type_enc sound
  1853         exporter readable_names preproc hyp_ts concl_t facts =
  1854   let
  1855     val (format, type_enc) = choose_format [format] type_enc
  1856     val (fact_names, (conjs, facts, class_rel_clauses, arity_clauses)) =
  1857       translate_formulas ctxt format prem_kind type_enc preproc hyp_ts concl_t
  1858                          facts
  1859     val sym_tab = conjs @ facts |> sym_table_for_facts ctxt explicit_apply
  1860     val nonmono_Ts =
  1861       conjs @ facts |> nonmonotonic_types_for_facts ctxt type_enc sound
  1862     val repair = repair_fact ctxt format type_enc sym_tab
  1863     val (conjs, facts) = (conjs, facts) |> pairself (map repair)
  1864     val repaired_sym_tab =
  1865       conjs @ facts |> sym_table_for_facts ctxt (SOME false)
  1866     val helpers =
  1867       repaired_sym_tab |> helper_facts_for_sym_table ctxt format type_enc
  1868                        |> map repair
  1869     val poly_nonmono_Ts =
  1870       if null nonmono_Ts orelse nonmono_Ts = [@{typ bool}] orelse
  1871          polymorphism_of_type_enc type_enc <> Polymorphic then
  1872         nonmono_Ts
  1873       else
  1874         [TVar (("'a", 0), HOLogic.typeS)]
  1875     val sym_decl_lines =
  1876       (conjs, helpers @ facts)
  1877       |> sym_decl_table_for_facts ctxt type_enc repaired_sym_tab
  1878       |> problem_lines_for_sym_decl_table ctxt format conj_sym_kind nonmono_Ts
  1879                                                poly_nonmono_Ts type_enc
  1880     val helper_lines =
  1881       0 upto length helpers - 1 ~~ helpers
  1882       |> map (formula_line_for_fact ctxt format helper_prefix I false true
  1883                                     poly_nonmono_Ts type_enc)
  1884       |> (if needs_type_tag_idempotence type_enc then
  1885             cons (type_tag_idempotence_fact ())
  1886           else
  1887             I)
  1888     (* Reordering these might confuse the proof reconstruction code or the SPASS
  1889        FLOTTER hack. *)
  1890     val problem =
  1891       [(explicit_declsN, sym_decl_lines),
  1892        (factsN,
  1893         map (formula_line_for_fact ctxt format fact_prefix ascii_of
  1894                                    (not exporter) (not exporter) nonmono_Ts
  1895                                    type_enc)
  1896             (0 upto length facts - 1 ~~ facts)),
  1897        (class_relsN, map formula_line_for_class_rel_clause class_rel_clauses),
  1898        (aritiesN, map formula_line_for_arity_clause arity_clauses),
  1899        (helpersN, helper_lines),
  1900        (conjsN,
  1901         map (formula_line_for_conjecture ctxt format nonmono_Ts type_enc)
  1902             conjs),
  1903        (free_typesN, formula_lines_for_free_types type_enc (facts @ conjs))]
  1904     val problem =
  1905       problem
  1906       |> (case format of
  1907             CNF => ensure_cnf_problem
  1908           | CNF_UEQ => filter_cnf_ueq_problem
  1909           | _ => I)
  1910       |> (if is_format_typed format then
  1911             declare_undeclared_syms_in_atp_problem type_decl_prefix
  1912                                                    implicit_declsN
  1913           else
  1914             I)
  1915     val (problem, pool) = problem |> nice_atp_problem readable_names
  1916     val helpers_offset = offset_of_heading_in_problem helpersN problem 0
  1917     val typed_helpers =
  1918       map_filter (fn (j, {name, ...}) =>
  1919                      if String.isSuffix typed_helper_suffix name then SOME j
  1920                      else NONE)
  1921                  ((helpers_offset + 1 upto helpers_offset + length helpers)
  1922                   ~~ helpers)
  1923     fun add_sym_arity (s, {min_ary, ...} : sym_info) =
  1924       if min_ary > 0 then
  1925         case strip_prefix_and_unascii const_prefix s of
  1926           SOME s => Symtab.insert (op =) (s, min_ary)
  1927         | NONE => I
  1928       else
  1929         I
  1930   in
  1931     (problem,
  1932      case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
  1933      offset_of_heading_in_problem conjsN problem 0,
  1934      offset_of_heading_in_problem factsN problem 0,
  1935      fact_names |> Vector.fromList,
  1936      typed_helpers,
  1937      Symtab.empty |> Symtab.fold add_sym_arity sym_tab)
  1938   end
  1939 
  1940 (* FUDGE *)
  1941 val conj_weight = 0.0
  1942 val hyp_weight = 0.1
  1943 val fact_min_weight = 0.2
  1944 val fact_max_weight = 1.0
  1945 val type_info_default_weight = 0.8
  1946 
  1947 fun add_term_weights weight (ATerm (s, tms)) =
  1948     is_tptp_user_symbol s ? Symtab.default (s, weight)
  1949     #> fold (add_term_weights weight) tms
  1950   | add_term_weights weight (AAbs (_, tm)) = add_term_weights weight tm
  1951 fun add_problem_line_weights weight (Formula (_, _, phi, _, _)) =
  1952     formula_fold NONE (K (add_term_weights weight)) phi
  1953   | add_problem_line_weights _ _ = I
  1954 
  1955 fun add_conjectures_weights [] = I
  1956   | add_conjectures_weights conjs =
  1957     let val (hyps, conj) = split_last conjs in
  1958       add_problem_line_weights conj_weight conj
  1959       #> fold (add_problem_line_weights hyp_weight) hyps
  1960     end
  1961 
  1962 fun add_facts_weights facts =
  1963   let
  1964     val num_facts = length facts
  1965     fun weight_of j =
  1966       fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j
  1967                         / Real.fromInt num_facts
  1968   in
  1969     map weight_of (0 upto num_facts - 1) ~~ facts
  1970     |> fold (uncurry add_problem_line_weights)
  1971   end
  1972 
  1973 (* Weights are from 0.0 (most important) to 1.0 (least important). *)
  1974 fun atp_problem_weights problem =
  1975   let val get = these o AList.lookup (op =) problem in
  1976     Symtab.empty
  1977     |> add_conjectures_weights (get free_typesN @ get conjsN)
  1978     |> add_facts_weights (get factsN)
  1979     |> fold (fold (add_problem_line_weights type_info_default_weight) o get)
  1980             [explicit_declsN, class_relsN, aritiesN]
  1981     |> Symtab.dest
  1982     |> sort (prod_ord Real.compare string_ord o pairself swap)
  1983   end
  1984 
  1985 end;