src/HOL/Tools/ATP/atp_translate.ML
author nik
Tue, 05 Jul 2011 17:09:59 +0100
changeset 44535 3b0b448b4d69
parent 44495 996b2022ff78
child 44536 2cd0b478d1b6
permissions -rw-r--r--
add support for lambdas in TPTP THF generator + killed an unsound type encoding (because the monotonicity calculus assumes first-order)
     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 
   447 fun combtyp_of (CombConst (_, T, _)) = T
   448   | combtyp_of (CombVar (_, T)) = T
   449   | combtyp_of (CombApp (t1, _)) = snd (dest_funT (combtyp_of t1))
   450 
   451 (*gets the head of a combinator application, along with the list of arguments*)
   452 fun strip_combterm_comb u =
   453   let
   454     fun stripc (CombApp (t, u), ts) = stripc (t, u :: ts)
   455       | stripc x = x
   456   in stripc (u, []) end
   457 
   458 fun atyps_of T = fold_atyps (insert (op =)) T []
   459 
   460 fun new_skolem_const_name s num_T_args =
   461   [new_skolem_const_prefix, s, string_of_int num_T_args]
   462   |> space_implode Long_Name.separator
   463 
   464 (* Converts a term (with combinators) into a combterm. Also accumulates sort
   465    infomation. *)
   466 fun combterm_from_term thy bs (P $ Q) =
   467     let
   468       val (P', P_atomics_Ts) = combterm_from_term thy bs P
   469       val (Q', Q_atomics_Ts) = combterm_from_term thy bs Q
   470     in (CombApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end
   471   | combterm_from_term thy _ (Const (c, T)) =
   472     let
   473       val tvar_list =
   474         (if String.isPrefix old_skolem_const_prefix c then
   475            [] |> Term.add_tvarsT T |> map TVar
   476          else
   477            (c, T) |> Sign.const_typargs thy)
   478       val c' = CombConst (`make_fixed_const c, T, tvar_list)
   479     in (c', atyps_of T) end
   480   | combterm_from_term _ _ (Free (v, T)) =
   481     (CombConst (`make_fixed_var v, T, []), atyps_of T)
   482   | combterm_from_term _ _ (Var (v as (s, _), T)) =
   483     (if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then
   484        let
   485          val Ts = T |> strip_type |> swap |> op ::
   486          val s' = new_skolem_const_name s (length Ts)
   487        in CombConst (`make_fixed_const s', T, Ts) end
   488      else
   489        CombVar ((make_schematic_var v, s), T), atyps_of T)
   490   | combterm_from_term _ bs (Bound j) =
   491     nth bs j
   492     |> (fn (s, T) => (CombConst (`make_bound_var s, T, []), atyps_of T))
   493   | combterm_from_term _ _ (Abs _) = raise Fail "HOL clause: Abs"
   494 
   495 datatype locality =
   496   General | Helper | Extensionality | Intro | Elim | Simp | Local | Assum |
   497   Chained
   498 
   499 (* (quasi-)underapproximation of the truth *)
   500 fun is_locality_global Local = false
   501   | is_locality_global Assum = false
   502   | is_locality_global Chained = false
   503   | is_locality_global _ = true
   504 
   505 datatype order = First_Order | Higher_Order
   506 datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic
   507 datatype type_level =
   508   All_Types | Noninf_Nonmono_Types | Fin_Nonmono_Types | Const_Arg_Types |
   509   No_Types
   510 datatype type_heaviness = Heavyweight | Lightweight
   511 
   512 datatype type_enc =
   513   Simple_Types of order * type_level |
   514   Preds of polymorphism * type_level * type_heaviness |
   515   Tags of polymorphism * type_level * type_heaviness
   516 
   517 fun try_unsuffixes ss s =
   518   fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE
   519 
   520 fun type_enc_from_string s =
   521   (case try (unprefix "poly_") s of
   522      SOME s => (SOME Polymorphic, s)
   523    | NONE =>
   524      case try (unprefix "mono_") s of
   525        SOME s => (SOME Monomorphic, s)
   526      | NONE =>
   527        case try (unprefix "mangled_") s of
   528          SOME s => (SOME Mangled_Monomorphic, s)
   529        | NONE => (NONE, s))
   530   ||> (fn s =>
   531           (* "_query" and "_bang" are for the ASCII-challenged Metis and
   532              Mirabelle. *)
   533           case try_unsuffixes ["?", "_query"] s of
   534             SOME s => (Noninf_Nonmono_Types, s)
   535           | NONE =>
   536             case try_unsuffixes ["!", "_bang"] s of
   537               SOME s => (Fin_Nonmono_Types, s)
   538             | NONE => (All_Types, s))
   539   ||> apsnd (fn s =>
   540                 case try (unsuffix "_heavy") s of
   541                   SOME s => (Heavyweight, s)
   542                 | NONE => (Lightweight, s))
   543   |> (fn (poly, (level, (heaviness, core))) =>
   544          case (core, (poly, level, heaviness)) of
   545            ("simple", (NONE, _, Lightweight)) =>
   546            Simple_Types (First_Order, level)
   547          | ("simple_higher", (NONE, _, Lightweight)) =>
   548            if level = Noninf_Nonmono_Types then raise Same.SAME
   549            else Simple_Types (Higher_Order, level)
   550          | ("preds", (SOME poly, _, _)) => Preds (poly, level, heaviness)
   551          | ("tags", (SOME Polymorphic, _, _)) =>
   552            Tags (Polymorphic, level, heaviness)
   553          | ("tags", (SOME poly, _, _)) => Tags (poly, level, heaviness)
   554          | ("args", (SOME poly, All_Types (* naja *), Lightweight)) =>
   555            Preds (poly, Const_Arg_Types, Lightweight)
   556          | ("erased", (NONE, All_Types (* naja *), Lightweight)) =>
   557            Preds (Polymorphic, No_Types, Lightweight)
   558          | _ => raise Same.SAME)
   559   handle Same.SAME => error ("Unknown type system: " ^ quote s ^ ".")
   560 
   561 fun is_type_enc_higher_order (Simple_Types (Higher_Order, _)) = true
   562   | is_type_enc_higher_order _ = false
   563 
   564 fun polymorphism_of_type_enc (Simple_Types _) = Mangled_Monomorphic
   565   | polymorphism_of_type_enc (Preds (poly, _, _)) = poly
   566   | polymorphism_of_type_enc (Tags (poly, _, _)) = poly
   567 
   568 fun level_of_type_enc (Simple_Types (_, level)) = level
   569   | level_of_type_enc (Preds (_, level, _)) = level
   570   | level_of_type_enc (Tags (_, level, _)) = level
   571 
   572 fun heaviness_of_type_enc (Simple_Types _) = Heavyweight
   573   | heaviness_of_type_enc (Preds (_, _, heaviness)) = heaviness
   574   | heaviness_of_type_enc (Tags (_, _, heaviness)) = heaviness
   575 
   576 fun is_type_level_virtually_sound level =
   577   level = All_Types orelse level = Noninf_Nonmono_Types
   578 val is_type_enc_virtually_sound =
   579   is_type_level_virtually_sound o level_of_type_enc
   580 
   581 fun is_type_level_fairly_sound level =
   582   is_type_level_virtually_sound level orelse level = Fin_Nonmono_Types
   583 val is_type_enc_fairly_sound = is_type_level_fairly_sound o level_of_type_enc
   584 
   585 fun choose_format formats (Simple_Types (order, level)) =
   586     if member (op =) formats THF then
   587       (THF, Simple_Types (order, level))
   588     else if member (op =) formats TFF then
   589       (TFF, Simple_Types (First_Order, level))
   590     else
   591       choose_format formats (Preds (Mangled_Monomorphic, level, Heavyweight))
   592   | choose_format formats type_enc =
   593     (case hd formats of
   594        CNF_UEQ =>
   595        (CNF_UEQ, case type_enc of
   596                    Preds stuff =>
   597                    (if is_type_enc_fairly_sound type_enc then Tags else Preds)
   598                        stuff
   599                  | _ => type_enc)
   600      | format => (format, type_enc))
   601 
   602 type translated_formula =
   603   {name : string,
   604    locality : locality,
   605    kind : formula_kind,
   606    combformula : (name, typ, combterm) formula,
   607    atomic_types : typ list}
   608 
   609 fun update_combformula f ({name, locality, kind, combformula, atomic_types}
   610                           : translated_formula) =
   611   {name = name, locality = locality, kind = kind, combformula = f combformula,
   612    atomic_types = atomic_types} : translated_formula
   613 
   614 fun fact_lift f ({combformula, ...} : translated_formula) = f combformula
   615 
   616 val type_instance = Sign.typ_instance o Proof_Context.theory_of
   617 
   618 fun insert_type ctxt get_T x xs =
   619   let val T = get_T x in
   620     if exists (curry (type_instance ctxt) T o get_T) xs then xs
   621     else x :: filter_out (curry (type_instance ctxt o swap) T o get_T) xs
   622   end
   623 
   624 (* The Booleans indicate whether all type arguments should be kept. *)
   625 datatype type_arg_policy =
   626   Explicit_Type_Args of bool |
   627   Mangled_Type_Args of bool |
   628   No_Type_Args
   629 
   630 fun should_drop_arg_type_args (Simple_Types _) =
   631     false (* since TFF doesn't support overloading *)
   632   | should_drop_arg_type_args type_enc =
   633     level_of_type_enc type_enc = All_Types andalso
   634     heaviness_of_type_enc type_enc = Heavyweight
   635 
   636 fun type_arg_policy type_enc s =
   637   if s = type_tag_name then
   638     (if polymorphism_of_type_enc type_enc = Mangled_Monomorphic then
   639        Mangled_Type_Args
   640      else
   641        Explicit_Type_Args) false
   642   else case type_enc of
   643     Tags (_, All_Types, Heavyweight) => No_Type_Args
   644   | _ =>
   645     if level_of_type_enc type_enc = No_Types orelse
   646        s = @{const_name HOL.eq} orelse
   647        (s = app_op_name andalso
   648         level_of_type_enc type_enc = Const_Arg_Types) then
   649       No_Type_Args
   650     else
   651       should_drop_arg_type_args type_enc
   652       |> (if polymorphism_of_type_enc type_enc = Mangled_Monomorphic then
   653             Mangled_Type_Args
   654           else
   655             Explicit_Type_Args)
   656 
   657 (* Make literals for sorted type variables. *)
   658 fun generic_add_sorts_on_type (_, []) = I
   659   | generic_add_sorts_on_type ((x, i), s :: ss) =
   660     generic_add_sorts_on_type ((x, i), ss)
   661     #> (if s = the_single @{sort HOL.type} then
   662           I
   663         else if i = ~1 then
   664           insert (op =) (TyLitFree (`make_type_class s, `make_fixed_type_var x))
   665         else
   666           insert (op =) (TyLitVar (`make_type_class s,
   667                                    (make_schematic_type_var (x, i), x))))
   668 fun add_sorts_on_tfree (TFree (s, S)) = generic_add_sorts_on_type ((s, ~1), S)
   669   | add_sorts_on_tfree _ = I
   670 fun add_sorts_on_tvar (TVar z) = generic_add_sorts_on_type z
   671   | add_sorts_on_tvar _ = I
   672 
   673 fun type_literals_for_types type_enc add_sorts_on_typ Ts =
   674   [] |> level_of_type_enc type_enc <> No_Types ? fold add_sorts_on_typ Ts
   675 
   676 fun mk_aconns c phis =
   677   let val (phis', phi') = split_last phis in
   678     fold_rev (mk_aconn c) phis' phi'
   679   end
   680 fun mk_ahorn [] phi = phi
   681   | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])
   682 fun mk_aquant _ [] phi = phi
   683   | mk_aquant q xs (phi as AQuant (q', xs', phi')) =
   684     if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)
   685   | mk_aquant q xs phi = AQuant (q, xs, phi)
   686 
   687 fun close_universally atom_vars phi =
   688   let
   689     fun formula_vars bounds (AQuant (_, xs, phi)) =
   690         formula_vars (map fst xs @ bounds) phi
   691       | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis
   692       | formula_vars bounds (AAtom tm) =
   693         union (op =) (atom_vars tm []
   694                       |> filter_out (member (op =) bounds o fst))
   695   in mk_aquant AForall (formula_vars [] phi []) phi end
   696 
   697 fun combterm_vars (CombApp (tm1, tm2)) = fold combterm_vars [tm1, tm2]
   698   | combterm_vars (CombConst _) = I
   699   | combterm_vars (CombVar (name, T)) = insert (op =) (name, SOME T)
   700 fun close_combformula_universally phi = close_universally combterm_vars phi
   701 
   702 fun term_vars bounds (ATerm (name as (s, _), tms)) =
   703     (is_tptp_variable s andalso not (member (op =) bounds name))
   704     ? insert (op =) (name, NONE) #> fold (term_vars bounds) tms
   705   | term_vars bounds (AAbs ((name, _), tm)) = term_vars (name :: bounds) tm
   706 fun close_formula_universally phi = close_universally (term_vars []) phi
   707 
   708 val homo_infinite_type_name = @{type_name ind} (* any infinite type *)
   709 val homo_infinite_type = Type (homo_infinite_type_name, [])
   710 
   711 fun ho_term_from_typ format type_enc =
   712   let
   713     fun term (Type (s, Ts)) =
   714       ATerm (case (is_type_enc_higher_order type_enc, s) of
   715                (true, @{type_name bool}) => `I tptp_bool_type
   716              | (true, @{type_name fun}) => `I tptp_fun_type
   717              | _ => if s = homo_infinite_type_name andalso
   718                        (format = TFF orelse format = THF) then
   719                       `I tptp_individual_type
   720                     else
   721                       `make_fixed_type_const s,
   722              map term Ts)
   723     | term (TFree (s, _)) = ATerm (`make_fixed_type_var s, [])
   724     | term (TVar ((x as (s, _)), _)) =
   725       ATerm ((make_schematic_type_var x, s), [])
   726   in term end
   727 
   728 fun ho_term_for_type_arg format type_enc T =
   729   if T = dummyT then NONE else SOME (ho_term_from_typ format type_enc T)
   730 
   731 (* This shouldn't clash with anything else. *)
   732 val mangled_type_sep = "\000"
   733 
   734 fun generic_mangled_type_name f (ATerm (name, [])) = f name
   735   | generic_mangled_type_name f (ATerm (name, tys)) =
   736     f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys)
   737     ^ ")"
   738   | generic_mangled_type_name f _ = raise Fail "unexpected type abstraction"
   739 
   740 val bool_atype = AType (`I tptp_bool_type)
   741 
   742 fun make_simple_type s =
   743   if s = tptp_bool_type orelse s = tptp_fun_type orelse
   744      s = tptp_individual_type then
   745     s
   746   else
   747     simple_type_prefix ^ ascii_of s
   748 
   749 fun ho_type_from_ho_term type_enc pred_sym ary =
   750   let
   751     fun to_atype ty =
   752       AType ((make_simple_type (generic_mangled_type_name fst ty),
   753               generic_mangled_type_name snd ty))
   754     fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1))
   755     fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty
   756       | to_fo ary (ATerm (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys
   757       | to_fo ary _ = raise Fail "unexpected type abstraction"
   758     fun to_ho (ty as ATerm ((s, _), tys)) =
   759         if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty
   760       | to_ho _ = raise Fail "unexpected type abstraction"
   761   in if is_type_enc_higher_order type_enc then to_ho else to_fo ary end
   762 
   763 fun mangled_type format type_enc pred_sym ary =
   764   ho_type_from_ho_term type_enc pred_sym ary
   765   o ho_term_from_typ format type_enc
   766 
   767 fun mangled_const_name format type_enc T_args (s, s') =
   768   let
   769     val ty_args = T_args |> map_filter (ho_term_for_type_arg format type_enc)
   770     fun type_suffix f g =
   771       fold_rev (curry (op ^) o g o prefix mangled_type_sep
   772                 o generic_mangled_type_name f) ty_args ""
   773   in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end
   774 
   775 val parse_mangled_ident =
   776   Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
   777 
   778 fun parse_mangled_type x =
   779   (parse_mangled_ident
   780    -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
   781                     [] >> ATerm) x
   782 and parse_mangled_types x =
   783   (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x
   784 
   785 fun unmangled_type s =
   786   s |> suffix ")" |> raw_explode
   787     |> Scan.finite Symbol.stopper
   788            (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^
   789                                                 quote s)) parse_mangled_type))
   790     |> fst
   791 
   792 val unmangled_const_name = space_explode mangled_type_sep #> hd
   793 fun unmangled_const s =
   794   let val ss = space_explode mangled_type_sep s in
   795     (hd ss, map unmangled_type (tl ss))
   796   end
   797 
   798 fun introduce_proxies type_enc =
   799   let
   800     fun intro top_level (CombApp (tm1, tm2)) =
   801         CombApp (intro top_level tm1, intro false tm2)
   802       | intro top_level (CombConst (name as (s, _), T, T_args)) =
   803         (case proxify_const s of
   804            SOME proxy_base =>
   805            if top_level orelse is_type_enc_higher_order type_enc then
   806              case (top_level, s) of
   807                (_, "c_False") => (`I tptp_false, [])
   808              | (_, "c_True") => (`I tptp_true, [])
   809              | (false, "c_Not") => (`I tptp_not, [])
   810              | (false, "c_conj") => (`I tptp_and, [])
   811              | (false, "c_disj") => (`I tptp_or, [])
   812              | (false, "c_implies") => (`I tptp_implies, [])
   813              | (false, s) =>
   814                if is_tptp_equal s then (`I tptp_equal, [])
   815                else (proxy_base |>> prefix const_prefix, T_args)
   816              | _ => (name, [])
   817            else
   818              (proxy_base |>> prefix const_prefix, T_args)
   819           | NONE => (name, T_args))
   820         |> (fn (name, T_args) => CombConst (name, T, T_args))
   821       | intro _ tm = tm
   822   in intro true end
   823 
   824 fun combformula_from_prop thy type_enc eq_as_iff =
   825   let
   826     fun do_term bs t atomic_types =
   827       combterm_from_term thy bs (Envir.eta_contract t)
   828       |>> (introduce_proxies type_enc #> AAtom)
   829       ||> union (op =) atomic_types
   830     fun do_quant bs q s T t' =
   831       let val s = singleton (Name.variant_list (map fst bs)) s in
   832         do_formula ((s, T) :: bs) t'
   833         #>> mk_aquant q [(`make_bound_var s, SOME T)]
   834       end
   835     and do_conn bs c t1 t2 =
   836       do_formula bs t1 ##>> do_formula bs t2 #>> uncurry (mk_aconn c)
   837     and do_formula bs t =
   838       case t of
   839         @{const Trueprop} $ t1 => do_formula bs t1
   840       | @{const Not} $ t1 => do_formula bs t1 #>> mk_anot
   841       | Const (@{const_name All}, _) $ Abs (s, T, t') =>
   842         do_quant bs AForall s T t'
   843       | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
   844         do_quant bs AExists s T t'
   845       | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2
   846       | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2
   847       | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2
   848       | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
   849         if eq_as_iff then do_conn bs AIff t1 t2 else do_term bs t
   850       | _ => do_term bs t
   851   in do_formula [] end
   852 
   853 fun presimplify_term _ [] t = t
   854   | presimplify_term ctxt presimp_consts t =
   855     t |> exists_Const (member (op =) presimp_consts o fst) t
   856          ? (Skip_Proof.make_thm (Proof_Context.theory_of ctxt)
   857             #> Meson.presimplify ctxt
   858             #> prop_of)
   859 
   860 fun concealed_bound_name j = sledgehammer_weak_prefix ^ string_of_int j
   861 fun conceal_bounds Ts t =
   862   subst_bounds (map (Free o apfst concealed_bound_name)
   863                     (0 upto length Ts - 1 ~~ Ts), t)
   864 fun reveal_bounds Ts =
   865   subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
   866                     (0 upto length Ts - 1 ~~ Ts))
   867 
   868 fun is_fun_equality (@{const_name HOL.eq},
   869                      Type (_, [Type (@{type_name fun}, _), _])) = true
   870   | is_fun_equality _ = false
   871 
   872 fun extensionalize_term ctxt t =
   873   if exists_Const is_fun_equality t then
   874     let val thy = Proof_Context.theory_of ctxt in
   875       t |> cterm_of thy |> Meson.extensionalize_conv ctxt
   876         |> prop_of |> Logic.dest_equals |> snd
   877     end
   878   else
   879     t
   880 
   881 fun introduce_combinators_in_term ctxt kind t =
   882   let val thy = Proof_Context.theory_of ctxt in
   883     if Meson.is_fol_term thy t then
   884       t
   885     else
   886       let
   887         fun aux Ts t =
   888           case t of
   889             @{const Not} $ t1 => @{const Not} $ aux Ts t1
   890           | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
   891             t0 $ Abs (s, T, aux (T :: Ts) t')
   892           | (t0 as Const (@{const_name All}, _)) $ t1 =>
   893             aux Ts (t0 $ eta_expand Ts t1 1)
   894           | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
   895             t0 $ Abs (s, T, aux (T :: Ts) t')
   896           | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
   897             aux Ts (t0 $ eta_expand Ts t1 1)
   898           | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
   899           | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
   900           | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
   901           | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
   902               $ t1 $ t2 =>
   903             t0 $ aux Ts t1 $ aux Ts t2
   904           | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
   905                    t
   906                  else
   907                    t |> conceal_bounds Ts
   908                      |> Envir.eta_contract
   909                      |> cterm_of thy
   910                      |> Meson_Clausify.introduce_combinators_in_cterm
   911                      |> prop_of |> Logic.dest_equals |> snd
   912                      |> reveal_bounds Ts
   913         val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
   914       in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
   915       handle THM _ =>
   916              (* A type variable of sort "{}" will make abstraction fail. *)
   917              if kind = Conjecture then HOLogic.false_const
   918              else HOLogic.true_const
   919   end
   920 
   921 (* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
   922    same in Sledgehammer to prevent the discovery of unreplayable proofs. *)
   923 fun freeze_term t =
   924   let
   925     fun aux (t $ u) = aux t $ aux u
   926       | aux (Abs (s, T, t)) = Abs (s, T, aux t)
   927       | aux (Var ((s, i), T)) =
   928         Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
   929       | aux t = t
   930   in t |> exists_subterm is_Var t ? aux end
   931 
   932 fun preprocess_prop ctxt presimp_consts kind t =
   933   let
   934     val thy = Proof_Context.theory_of ctxt
   935     val t = t |> Envir.beta_eta_contract
   936               |> transform_elim_prop
   937               |> Object_Logic.atomize_term thy
   938     val need_trueprop = (fastype_of t = @{typ bool})
   939   in
   940     t |> need_trueprop ? HOLogic.mk_Trueprop
   941       |> Raw_Simplifier.rewrite_term thy (Meson.unfold_set_const_simps ctxt) []
   942       |> extensionalize_term ctxt
   943       |> presimplify_term ctxt presimp_consts
   944       |> perhaps (try (HOLogic.dest_Trueprop))
   945       |> introduce_combinators_in_term ctxt kind
   946   end
   947 
   948 (* making fact and conjecture formulas *)
   949 fun make_formula thy type_enc eq_as_iff name loc kind t =
   950   let
   951     val (combformula, atomic_types) =
   952       combformula_from_prop thy type_enc eq_as_iff t []
   953   in
   954     {name = name, locality = loc, kind = kind, combformula = combformula,
   955      atomic_types = atomic_types}
   956   end
   957 
   958 fun make_fact ctxt format type_enc eq_as_iff preproc presimp_consts
   959               ((name, loc), t) =
   960   let val thy = Proof_Context.theory_of ctxt in
   961     case t |> preproc ? preprocess_prop ctxt presimp_consts Axiom
   962            |> make_formula thy type_enc (eq_as_iff andalso format <> CNF) name
   963                            loc Axiom of
   964       formula as {combformula = AAtom (CombConst ((s, _), _, _)), ...} =>
   965       if s = tptp_true then NONE else SOME formula
   966     | formula => SOME formula
   967   end
   968 
   969 fun make_conjecture ctxt format prem_kind type_enc preproc presimp_consts ts =
   970   let
   971     val thy = Proof_Context.theory_of ctxt
   972     val last = length ts - 1
   973   in
   974     map2 (fn j => fn t =>
   975              let
   976                val (kind, maybe_negate) =
   977                  if j = last then
   978                    (Conjecture, I)
   979                  else
   980                    (prem_kind,
   981                     if prem_kind = Conjecture then update_combformula mk_anot
   982                     else I)
   983               in
   984                 t |> preproc ?
   985                      (preprocess_prop ctxt presimp_consts kind #> freeze_term)
   986                   |> make_formula thy type_enc (format <> CNF) (string_of_int j)
   987                                   Local kind
   988                   |> maybe_negate
   989               end)
   990          (0 upto last) ts
   991   end
   992 
   993 (** Finite and infinite type inference **)
   994 
   995 fun deep_freeze_atyp (TVar (_, S)) = TFree ("v", S)
   996   | deep_freeze_atyp T = T
   997 val deep_freeze_type = map_atyps deep_freeze_atyp
   998 
   999 (* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are
  1000    dangerous because their "exhaust" properties can easily lead to unsound ATP
  1001    proofs. On the other hand, all HOL infinite types can be given the same
  1002    models in first-order logic (via Löwenheim-Skolem). *)
  1003 
  1004 fun should_encode_type ctxt (nonmono_Ts as _ :: _) _ T =
  1005     exists (curry (type_instance ctxt) (deep_freeze_type T)) nonmono_Ts
  1006   | should_encode_type _ _ All_Types _ = true
  1007   | should_encode_type ctxt _ Fin_Nonmono_Types T =
  1008     is_type_surely_finite ctxt false T
  1009   | should_encode_type _ _ _ _ = false
  1010 
  1011 fun should_predicate_on_type ctxt nonmono_Ts (Preds (_, level, heaviness))
  1012                              should_predicate_on_var T =
  1013     (heaviness = Heavyweight orelse should_predicate_on_var ()) andalso
  1014     should_encode_type ctxt nonmono_Ts level T
  1015   | should_predicate_on_type _ _ _ _ _ = false
  1016 
  1017 fun is_var_or_bound_var (CombConst ((s, _), _, _)) =
  1018     String.isPrefix bound_var_prefix s
  1019   | is_var_or_bound_var (CombVar _) = true
  1020   | is_var_or_bound_var _ = false
  1021 
  1022 datatype tag_site =
  1023   Top_Level of bool option |
  1024   Eq_Arg of bool option |
  1025   Elsewhere
  1026 
  1027 fun should_tag_with_type _ _ _ (Top_Level _) _ _ = false
  1028   | should_tag_with_type ctxt nonmono_Ts (Tags (poly, level, heaviness)) site
  1029                          u T =
  1030     (case heaviness of
  1031        Heavyweight => should_encode_type ctxt nonmono_Ts level T
  1032      | Lightweight =>
  1033        case (site, is_var_or_bound_var u) of
  1034          (Eq_Arg pos, true) =>
  1035          (* The first disjunct prevents a subtle soundness issue explained in
  1036             Blanchette's Ph.D. thesis. See also
  1037             "formula_lines_for_lightweight_tags_sym_decl". *)
  1038          (pos <> SOME false andalso poly = Polymorphic andalso
  1039           level <> All_Types andalso heaviness = Lightweight andalso
  1040           exists (fn T' => type_instance ctxt (T', T)) nonmono_Ts) orelse
  1041          should_encode_type ctxt nonmono_Ts level T
  1042        | _ => false)
  1043   | should_tag_with_type _ _ _ _ _ _ = false
  1044 
  1045 fun homogenized_type ctxt nonmono_Ts level =
  1046   let
  1047     val should_encode = should_encode_type ctxt nonmono_Ts level
  1048     fun homo 0 T = if should_encode T then T else homo_infinite_type
  1049       | homo ary (Type (@{type_name fun}, [T1, T2])) =
  1050         homo 0 T1 --> homo (ary - 1) T2
  1051       | homo _ _ = raise Fail "expected function type"
  1052   in homo end
  1053 
  1054 (** "hBOOL" and "hAPP" **)
  1055 
  1056 type sym_info =
  1057   {pred_sym : bool, min_ary : int, max_ary : int, types : typ list}
  1058 
  1059 fun add_combterm_syms_to_table ctxt explicit_apply =
  1060   let
  1061     fun consider_var_arity const_T var_T max_ary =
  1062       let
  1063         fun iter ary T =
  1064           if ary = max_ary orelse type_instance ctxt (var_T, T) orelse
  1065              type_instance ctxt (T, var_T) then
  1066             ary
  1067           else
  1068             iter (ary + 1) (range_type T)
  1069       in iter 0 const_T end
  1070     fun add_var_or_bound_var T (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
  1071       if explicit_apply = NONE andalso
  1072          (can dest_funT T orelse T = @{typ bool}) then
  1073         let
  1074           val bool_vars' = bool_vars orelse body_type T = @{typ bool}
  1075           fun repair_min_arity {pred_sym, min_ary, max_ary, types} =
  1076             {pred_sym = pred_sym andalso not bool_vars',
  1077              min_ary = fold (fn T' => consider_var_arity T' T) types min_ary,
  1078              max_ary = max_ary, types = types}
  1079           val fun_var_Ts' =
  1080             fun_var_Ts |> can dest_funT T ? insert_type ctxt I T
  1081         in
  1082           if bool_vars' = bool_vars andalso
  1083              pointer_eq (fun_var_Ts', fun_var_Ts) then
  1084             accum
  1085           else
  1086             ((bool_vars', fun_var_Ts'), Symtab.map (K repair_min_arity) sym_tab)
  1087         end
  1088       else
  1089         accum
  1090     fun add top_level tm (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
  1091       let val (head, args) = strip_combterm_comb tm in
  1092         (case head of
  1093            CombConst ((s, _), T, _) =>
  1094            if String.isPrefix bound_var_prefix s then
  1095              add_var_or_bound_var T accum
  1096            else
  1097              let val ary = length args in
  1098                ((bool_vars, fun_var_Ts),
  1099                 case Symtab.lookup sym_tab s of
  1100                   SOME {pred_sym, min_ary, max_ary, types} =>
  1101                   let
  1102                     val pred_sym =
  1103                       pred_sym andalso top_level andalso not bool_vars
  1104                     val types' = types |> insert_type ctxt I T
  1105                     val min_ary =
  1106                       if is_some explicit_apply orelse
  1107                          pointer_eq (types', types) then
  1108                         min_ary
  1109                       else
  1110                         fold (consider_var_arity T) fun_var_Ts min_ary
  1111                   in
  1112                     Symtab.update (s, {pred_sym = pred_sym,
  1113                                        min_ary = Int.min (ary, min_ary),
  1114                                        max_ary = Int.max (ary, max_ary),
  1115                                        types = types'})
  1116                                   sym_tab
  1117                   end
  1118                 | NONE =>
  1119                   let
  1120                     val pred_sym = top_level andalso not bool_vars
  1121                     val min_ary =
  1122                       case explicit_apply of
  1123                         SOME true => 0
  1124                       | SOME false => ary
  1125                       | NONE => fold (consider_var_arity T) fun_var_Ts ary
  1126                   in
  1127                     Symtab.update_new (s, {pred_sym = pred_sym,
  1128                                            min_ary = min_ary, max_ary = ary,
  1129                                            types = [T]})
  1130                                       sym_tab
  1131                   end)
  1132              end
  1133          | CombVar (_, T) => add_var_or_bound_var T accum
  1134          | _ => accum)
  1135         |> fold (add false) args
  1136       end
  1137   in add true end
  1138 fun add_fact_syms_to_table ctxt explicit_apply =
  1139   fact_lift (formula_fold NONE
  1140                           (K (add_combterm_syms_to_table ctxt explicit_apply)))
  1141 
  1142 val default_sym_tab_entries : (string * sym_info) list =
  1143   (prefixed_predicator_name,
  1144    {pred_sym = true, min_ary = 1, max_ary = 1, types = []}) ::
  1145   ([tptp_false, tptp_true]
  1146    |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, types = []})) @
  1147   ([tptp_equal, tptp_old_equal]
  1148    |> map (rpair {pred_sym = true, min_ary = 2, max_ary = 2, types = []}))
  1149 
  1150 fun sym_table_for_facts ctxt explicit_apply facts =
  1151   ((false, []), Symtab.empty)
  1152   |> fold (add_fact_syms_to_table ctxt explicit_apply) facts |> snd
  1153   |> fold Symtab.update default_sym_tab_entries
  1154 
  1155 fun min_arity_of sym_tab s =
  1156   case Symtab.lookup sym_tab s of
  1157     SOME ({min_ary, ...} : sym_info) => min_ary
  1158   | NONE =>
  1159     case strip_prefix_and_unascii const_prefix s of
  1160       SOME s =>
  1161       let val s = s |> unmangled_const_name |> invert_const in
  1162         if s = predicator_name then 1
  1163         else if s = app_op_name then 2
  1164         else if s = type_pred_name then 1
  1165         else 0
  1166       end
  1167     | NONE => 0
  1168 
  1169 (* True if the constant ever appears outside of the top-level position in
  1170    literals, or if it appears with different arities (e.g., because of different
  1171    type instantiations). If false, the constant always receives all of its
  1172    arguments and is used as a predicate. *)
  1173 fun is_pred_sym sym_tab s =
  1174   case Symtab.lookup sym_tab s of
  1175     SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
  1176     pred_sym andalso min_ary = max_ary
  1177   | NONE => false
  1178 
  1179 val predicator_combconst =
  1180   CombConst (`make_fixed_const predicator_name, @{typ "bool => bool"}, [])
  1181 fun predicator tm = CombApp (predicator_combconst, tm)
  1182 
  1183 fun introduce_predicators_in_combterm sym_tab tm =
  1184   case strip_combterm_comb tm of
  1185     (CombConst ((s, _), _, _), _) =>
  1186     if is_pred_sym sym_tab s then tm else predicator tm
  1187   | _ => predicator tm
  1188 
  1189 fun list_app head args = fold (curry (CombApp o swap)) args head
  1190 
  1191 val app_op = `make_fixed_const app_op_name
  1192 
  1193 fun explicit_app arg head =
  1194   let
  1195     val head_T = combtyp_of head
  1196     val (arg_T, res_T) = dest_funT head_T
  1197     val explicit_app =
  1198       CombConst (app_op, head_T --> head_T, [arg_T, res_T])
  1199   in list_app explicit_app [head, arg] end
  1200 fun list_explicit_app head args = fold explicit_app args head
  1201 
  1202 fun introduce_explicit_apps_in_combterm sym_tab =
  1203   let
  1204     fun aux tm =
  1205       case strip_combterm_comb tm of
  1206         (head as CombConst ((s, _), _, _), args) =>
  1207         args |> map aux
  1208              |> chop (min_arity_of sym_tab s)
  1209              |>> list_app head
  1210              |-> list_explicit_app
  1211       | (head, args) => list_explicit_app head (map aux args)
  1212   in aux end
  1213 
  1214 fun chop_fun 0 T = ([], T)
  1215   | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
  1216     chop_fun (n - 1) ran_T |>> cons dom_T
  1217   | chop_fun _ _ = raise Fail "unexpected non-function"
  1218 
  1219 fun filter_type_args _ _ _ [] = []
  1220   | filter_type_args thy s arity T_args =
  1221     let
  1222       (* will throw "TYPE" for pseudo-constants *)
  1223       val U = if s = app_op_name then
  1224                 @{typ "('a => 'b) => 'a => 'b"} |> Logic.varifyT_global
  1225               else
  1226                 s |> Sign.the_const_type thy
  1227     in
  1228       case Term.add_tvarsT (U |> chop_fun arity |> snd) [] of
  1229         [] => []
  1230       | res_U_vars =>
  1231         let val U_args = (s, U) |> Sign.const_typargs thy in
  1232           U_args ~~ T_args
  1233           |> map (fn (U, T) =>
  1234                      if member (op =) res_U_vars (dest_TVar U) then T
  1235                      else dummyT)
  1236         end
  1237     end
  1238     handle TYPE _ => T_args
  1239 
  1240 fun enforce_type_arg_policy_in_combterm ctxt format type_enc =
  1241   let
  1242     val thy = Proof_Context.theory_of ctxt
  1243     fun aux arity (CombApp (tm1, tm2)) =
  1244         CombApp (aux (arity + 1) tm1, aux 0 tm2)
  1245       | aux arity (CombConst (name as (s, _), T, T_args)) =
  1246         (case strip_prefix_and_unascii const_prefix s of
  1247            NONE => (name, T_args)
  1248          | SOME s'' =>
  1249            let
  1250              val s'' = invert_const s''
  1251              fun filtered_T_args false = T_args
  1252                | filtered_T_args true = filter_type_args thy s'' arity T_args
  1253            in
  1254              case type_arg_policy type_enc s'' of
  1255                Explicit_Type_Args drop_args =>
  1256                (name, filtered_T_args drop_args)
  1257              | Mangled_Type_Args drop_args =>
  1258                (mangled_const_name format type_enc (filtered_T_args drop_args)
  1259                                    name, [])
  1260              | No_Type_Args => (name, [])
  1261            end)
  1262         |> (fn (name, T_args) => CombConst (name, T, T_args))
  1263       | aux _ tm = tm
  1264   in aux 0 end
  1265 
  1266 fun repair_combterm ctxt format type_enc sym_tab =
  1267   not (is_type_enc_higher_order type_enc)
  1268   ? (introduce_explicit_apps_in_combterm sym_tab
  1269      #> introduce_predicators_in_combterm sym_tab)
  1270   #> enforce_type_arg_policy_in_combterm ctxt format type_enc
  1271 fun repair_fact ctxt format type_enc sym_tab =
  1272   update_combformula (formula_map
  1273       (repair_combterm ctxt format type_enc sym_tab))
  1274 
  1275 (** Helper facts **)
  1276 
  1277 (* The Boolean indicates that a fairly sound type encoding is needed. *)
  1278 val helper_table =
  1279   [(("COMBI", false), @{thms Meson.COMBI_def}),
  1280    (("COMBK", false), @{thms Meson.COMBK_def}),
  1281    (("COMBB", false), @{thms Meson.COMBB_def}),
  1282    (("COMBC", false), @{thms Meson.COMBC_def}),
  1283    (("COMBS", false), @{thms Meson.COMBS_def}),
  1284    (("fequal", true),
  1285     (* This is a lie: Higher-order equality doesn't need a sound type encoding.
  1286        However, this is done so for backward compatibility: Including the
  1287        equality helpers by default in Metis breaks a few existing proofs. *)
  1288     @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1]
  1289            fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),
  1290    (("fFalse", false), [@{lemma "~ fFalse" by (unfold fFalse_def) fast}]),
  1291    (("fFalse", true), @{thms True_or_False}),
  1292    (("fTrue", false), [@{lemma "fTrue" by (unfold fTrue_def) fast}]),
  1293    (("fTrue", true), @{thms True_or_False}),
  1294    (("fNot", false),
  1295     @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1]
  1296            fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),
  1297    (("fconj", false),
  1298     @{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q"
  1299         by (unfold fconj_def) fast+}),
  1300    (("fdisj", false),
  1301     @{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q"
  1302         by (unfold fdisj_def) fast+}),
  1303    (("fimplies", false),
  1304     @{lemma "P | fimplies P Q" "~ Q | fimplies P Q" "~ fimplies P Q | ~ P | Q"
  1305         by (unfold fimplies_def) fast+}),
  1306    (("If", true), @{thms if_True if_False True_or_False})]
  1307   |> map (apsnd (map zero_var_indexes))
  1308 
  1309 val type_tag = `make_fixed_const type_tag_name
  1310 
  1311 fun type_tag_idempotence_fact () =
  1312   let
  1313     fun var s = ATerm (`I s, [])
  1314     fun tag tm = ATerm (type_tag, [var "T", tm])
  1315     val tagged_a = tag (var "A")
  1316   in
  1317     Formula (type_tag_idempotence_helper_name, Axiom,
  1318              AAtom (ATerm (`I tptp_equal, [tag tagged_a, tagged_a]))
  1319              |> close_formula_universally, simp_info, NONE)
  1320   end
  1321 
  1322 fun should_specialize_helper type_enc t =
  1323   polymorphism_of_type_enc type_enc = Mangled_Monomorphic andalso
  1324   level_of_type_enc type_enc <> No_Types andalso
  1325   not (null (Term.hidden_polymorphism t))
  1326 
  1327 fun helper_facts_for_sym ctxt format type_enc (s, {types, ...} : sym_info) =
  1328   case strip_prefix_and_unascii const_prefix s of
  1329     SOME mangled_s =>
  1330     let
  1331       val thy = Proof_Context.theory_of ctxt
  1332       val unmangled_s = mangled_s |> unmangled_const_name
  1333       fun dub needs_fairly_sound j k =
  1334         (unmangled_s ^ "_" ^ string_of_int j ^ "_" ^ string_of_int k ^
  1335          (if mangled_s = unmangled_s then "" else "_" ^ ascii_of mangled_s) ^
  1336          (if needs_fairly_sound then typed_helper_suffix
  1337           else untyped_helper_suffix),
  1338          Helper)
  1339       fun dub_and_inst needs_fairly_sound (th, j) =
  1340         let val t = prop_of th in
  1341           if should_specialize_helper type_enc t then
  1342             map (fn T => specialize_type thy (invert_const unmangled_s, T) t)
  1343                 types
  1344           else
  1345             [t]
  1346         end
  1347         |> map (fn (k, t) => (dub needs_fairly_sound j k, t)) o tag_list 1
  1348       val make_facts =
  1349         map_filter (make_fact ctxt format type_enc false false [])
  1350       val fairly_sound = is_type_enc_fairly_sound type_enc
  1351     in
  1352       helper_table
  1353       |> maps (fn ((helper_s, needs_fairly_sound), ths) =>
  1354                   if helper_s <> unmangled_s orelse
  1355                      (needs_fairly_sound andalso not fairly_sound) then
  1356                     []
  1357                   else
  1358                     ths ~~ (1 upto length ths)
  1359                     |> maps (dub_and_inst needs_fairly_sound)
  1360                     |> make_facts)
  1361     end
  1362   | NONE => []
  1363 fun helper_facts_for_sym_table ctxt format type_enc sym_tab =
  1364   Symtab.fold_rev (append o helper_facts_for_sym ctxt format type_enc) sym_tab
  1365                   []
  1366 
  1367 (***************************************************************)
  1368 (* Type Classes Present in the Axiom or Conjecture Clauses     *)
  1369 (***************************************************************)
  1370 
  1371 fun set_insert (x, s) = Symtab.update (x, ()) s
  1372 
  1373 fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
  1374 
  1375 (* Remove this trivial type class (FIXME: similar code elsewhere) *)
  1376 fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset
  1377 
  1378 fun classes_of_terms get_Ts =
  1379   map (map snd o get_Ts)
  1380   #> List.foldl add_classes Symtab.empty
  1381   #> delete_type #> Symtab.keys
  1382 
  1383 val tfree_classes_of_terms = classes_of_terms OldTerm.term_tfrees
  1384 val tvar_classes_of_terms = classes_of_terms OldTerm.term_tvars
  1385 
  1386 fun fold_type_constrs f (Type (s, Ts)) x =
  1387     fold (fold_type_constrs f) Ts (f (s, x))
  1388   | fold_type_constrs _ _ x = x
  1389 
  1390 (*Type constructors used to instantiate overloaded constants are the only ones needed.*)
  1391 fun add_type_constrs_in_term thy =
  1392   let
  1393     fun add (Const (@{const_name Meson.skolem}, _) $ _) = I
  1394       | add (t $ u) = add t #> add u
  1395       | add (Const (x as (s, _))) =
  1396         if String.isPrefix skolem_const_prefix s then I
  1397         else x |> Sign.const_typargs thy |> fold (fold_type_constrs set_insert)
  1398       | add (Abs (_, _, u)) = add u
  1399       | add _ = I
  1400   in add end
  1401 
  1402 fun type_constrs_of_terms thy ts =
  1403   Symtab.keys (fold (add_type_constrs_in_term thy) ts Symtab.empty)
  1404 
  1405 fun translate_formulas ctxt format prem_kind type_enc preproc hyp_ts concl_t
  1406                        facts =
  1407   let
  1408     val thy = Proof_Context.theory_of ctxt
  1409     val fact_ts = facts |> map snd
  1410     val presimp_consts = Meson.presimplified_consts ctxt
  1411     val make_fact = make_fact ctxt format type_enc true preproc presimp_consts
  1412     val (facts, fact_names) =
  1413       facts |> map (fn (name, t) => (name, t) |> make_fact |> rpair name)
  1414             |> map_filter (try (apfst the))
  1415             |> ListPair.unzip
  1416     (* Remove existing facts from the conjecture, as this can dramatically
  1417        boost an ATP's performance (for some reason). *)
  1418     val hyp_ts =
  1419       hyp_ts
  1420       |> map (fn t => if member (op aconv) fact_ts t then @{prop True} else t)
  1421     val goal_t = Logic.list_implies (hyp_ts, concl_t)
  1422     val all_ts = goal_t :: fact_ts
  1423     val subs = tfree_classes_of_terms all_ts
  1424     val supers = tvar_classes_of_terms all_ts
  1425     val tycons = type_constrs_of_terms thy all_ts
  1426     val conjs =
  1427       hyp_ts @ [concl_t]
  1428       |> make_conjecture ctxt format prem_kind type_enc preproc presimp_consts
  1429     val (supers', arity_clauses) =
  1430       if level_of_type_enc type_enc = No_Types then ([], [])
  1431       else make_arity_clauses thy tycons supers
  1432     val class_rel_clauses = make_class_rel_clauses thy subs supers'
  1433   in
  1434     (fact_names |> map single, (conjs, facts, class_rel_clauses, arity_clauses))
  1435   end
  1436 
  1437 fun fo_literal_from_type_literal (TyLitVar (class, name)) =
  1438     (true, ATerm (class, [ATerm (name, [])]))
  1439   | fo_literal_from_type_literal (TyLitFree (class, name)) =
  1440     (true, ATerm (class, [ATerm (name, [])]))
  1441 
  1442 fun formula_from_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
  1443 
  1444 val type_pred = `make_fixed_const type_pred_name
  1445 
  1446 fun type_pred_combterm ctxt format type_enc T tm =
  1447   CombApp (CombConst (type_pred, T --> @{typ bool}, [T])
  1448            |> enforce_type_arg_policy_in_combterm ctxt format type_enc, tm)
  1449 
  1450 fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum
  1451   | is_var_positively_naked_in_term name _ (ATerm ((s, _), tms)) accum =
  1452     accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm (name, [])))
  1453   | is_var_positively_naked_in_term name _ _ _ = true
  1454 fun should_predicate_on_var_in_formula pos phi (SOME true) name =
  1455     formula_fold pos (is_var_positively_naked_in_term name) phi false
  1456   | should_predicate_on_var_in_formula _ _ _ _ = true
  1457 
  1458 fun mk_const_aterm format type_enc x T_args args =
  1459   ATerm (x, map_filter (ho_term_for_type_arg format type_enc) T_args @ args)
  1460 
  1461 fun tag_with_type ctxt format nonmono_Ts type_enc pos T tm =
  1462   CombConst (type_tag, T --> T, [T])
  1463   |> enforce_type_arg_policy_in_combterm ctxt format type_enc
  1464   |> term_from_combterm ctxt format nonmono_Ts type_enc (Top_Level pos)
  1465   |> (fn ATerm (s, tms) => ATerm (s, tms @ [tm]))
  1466 and term_from_combterm ctxt format nonmono_Ts type_enc =
  1467   let
  1468     fun aux site u =
  1469       let
  1470         val (head, args) = strip_combterm_comb u
  1471         val (x as (s, _), T_args) =
  1472           case head of
  1473             CombConst (name, _, T_args) => (name, T_args)
  1474           | CombVar (name, _) => (name, [])
  1475           | CombApp _ => raise Fail "impossible \"CombApp\""
  1476         val (pos, arg_site) =
  1477           case site of
  1478             Top_Level pos =>
  1479             (pos, if is_tptp_equal s then Eq_Arg pos else Elsewhere)
  1480           | Eq_Arg pos => (pos, Elsewhere)
  1481           | Elsewhere => (NONE, Elsewhere)
  1482         val t = mk_const_aterm format type_enc x T_args
  1483                     (map (aux arg_site) args)
  1484         val T = combtyp_of u
  1485       in
  1486         t |> (if should_tag_with_type ctxt nonmono_Ts type_enc site u T then
  1487                 tag_with_type ctxt format nonmono_Ts type_enc pos T
  1488               else
  1489                 I)
  1490       end
  1491   in aux end
  1492 and formula_from_combformula ctxt format nonmono_Ts type_enc
  1493                              should_predicate_on_var =
  1494   let
  1495     fun do_term pos =
  1496       term_from_combterm ctxt format nonmono_Ts type_enc (Top_Level pos)
  1497     val do_bound_type =
  1498       case type_enc of
  1499         Simple_Types (_, level) =>
  1500         homogenized_type ctxt nonmono_Ts level 0
  1501         #> mangled_type format type_enc false 0 #> SOME
  1502       | _ => K NONE
  1503     fun do_out_of_bound_type pos phi universal (name, T) =
  1504       if should_predicate_on_type ctxt nonmono_Ts type_enc
  1505              (fn () => should_predicate_on_var pos phi universal name) T then
  1506         CombVar (name, T)
  1507         |> type_pred_combterm ctxt format type_enc T
  1508         |> do_term pos |> AAtom |> SOME
  1509       else
  1510         NONE
  1511     fun do_formula pos (AQuant (q, xs, phi)) =
  1512         let
  1513           val phi = phi |> do_formula pos
  1514           val universal = Option.map (q = AExists ? not) pos
  1515         in
  1516           AQuant (q, xs |> map (apsnd (fn NONE => NONE
  1517                                         | SOME T => do_bound_type T)),
  1518                   (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))
  1519                       (map_filter
  1520                            (fn (_, NONE) => NONE
  1521                              | (s, SOME T) =>
  1522                                do_out_of_bound_type pos phi universal (s, T))
  1523                            xs)
  1524                       phi)
  1525         end
  1526       | do_formula pos (AConn conn) = aconn_map pos do_formula conn
  1527       | do_formula pos (AAtom tm) = AAtom (do_term pos tm)
  1528   in do_formula end
  1529 
  1530 fun bound_tvars type_enc Ts =
  1531   mk_ahorn (map (formula_from_fo_literal o fo_literal_from_type_literal)
  1532                 (type_literals_for_types type_enc add_sorts_on_tvar Ts))
  1533 
  1534 (* Each fact is given a unique fact number to avoid name clashes (e.g., because
  1535    of monomorphization). The TPTP explicitly forbids name clashes, and some of
  1536    the remote provers might care. *)
  1537 fun formula_line_for_fact ctxt format prefix encode freshen pos nonmono_Ts
  1538         type_enc (j, {name, locality, kind, combformula, atomic_types}) =
  1539   (prefix ^ (if freshen then string_of_int j ^ "_" else "") ^ encode name,
  1540    kind,
  1541    combformula
  1542    |> close_combformula_universally
  1543    |> formula_from_combformula ctxt format nonmono_Ts type_enc
  1544                                should_predicate_on_var_in_formula
  1545                                (if pos then SOME true else NONE)
  1546    |> bound_tvars type_enc atomic_types
  1547    |> close_formula_universally,
  1548    NONE,
  1549    case locality of
  1550      Intro => intro_info
  1551    | Elim => elim_info
  1552    | Simp => simp_info
  1553    | _ => NONE)
  1554   |> Formula
  1555 
  1556 fun formula_line_for_class_rel_clause ({name, subclass, superclass, ...}
  1557                                        : class_rel_clause) =
  1558   let val ty_arg = ATerm (`I "T", []) in
  1559     Formula (class_rel_clause_prefix ^ ascii_of name, Axiom,
  1560              AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
  1561                                AAtom (ATerm (superclass, [ty_arg]))])
  1562              |> close_formula_universally, intro_info, NONE)
  1563   end
  1564 
  1565 fun fo_literal_from_arity_literal (TConsLit (c, t, args)) =
  1566     (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
  1567   | fo_literal_from_arity_literal (TVarLit (c, sort)) =
  1568     (false, ATerm (c, [ATerm (sort, [])]))
  1569 
  1570 fun formula_line_for_arity_clause ({name, prem_lits, concl_lits, ...}
  1571                                    : arity_clause) =
  1572   Formula (arity_clause_prefix ^ name, Axiom,
  1573            mk_ahorn (map (formula_from_fo_literal o apfst not
  1574                           o fo_literal_from_arity_literal) prem_lits)
  1575                     (formula_from_fo_literal
  1576                          (fo_literal_from_arity_literal concl_lits))
  1577            |> close_formula_universally, intro_info, NONE)
  1578 
  1579 fun formula_line_for_conjecture ctxt format nonmono_Ts type_enc
  1580         ({name, kind, combformula, atomic_types, ...} : translated_formula) =
  1581   Formula (conjecture_prefix ^ name, kind,
  1582            formula_from_combformula ctxt format nonmono_Ts type_enc
  1583                should_predicate_on_var_in_formula (SOME false)
  1584                (close_combformula_universally combformula)
  1585            |> bound_tvars type_enc atomic_types
  1586            |> close_formula_universally, NONE, NONE)
  1587 
  1588 fun free_type_literals type_enc ({atomic_types, ...} : translated_formula) =
  1589   atomic_types |> type_literals_for_types type_enc add_sorts_on_tfree
  1590                |> map fo_literal_from_type_literal
  1591 
  1592 fun formula_line_for_free_type j lit =
  1593   Formula (tfree_clause_prefix ^ string_of_int j, Hypothesis,
  1594            formula_from_fo_literal lit, NONE, NONE)
  1595 fun formula_lines_for_free_types type_enc facts =
  1596   let
  1597     val litss = map (free_type_literals type_enc) facts
  1598     val lits = fold (union (op =)) litss []
  1599   in map2 formula_line_for_free_type (0 upto length lits - 1) lits end
  1600 
  1601 (** Symbol declarations **)
  1602 
  1603 fun should_declare_sym type_enc pred_sym s =
  1604   is_tptp_user_symbol s andalso not (String.isPrefix bound_var_prefix s) andalso
  1605   (case type_enc of
  1606      Simple_Types _ => true
  1607    | Tags (_, _, Lightweight) => true
  1608    | _ => not pred_sym)
  1609 
  1610 fun sym_decl_table_for_facts ctxt type_enc repaired_sym_tab (conjs, facts) =
  1611   let
  1612     fun add_combterm in_conj tm =
  1613       let val (head, args) = strip_combterm_comb tm in
  1614         (case head of
  1615            CombConst ((s, s'), T, T_args) =>
  1616            let val pred_sym = is_pred_sym repaired_sym_tab s in
  1617              if should_declare_sym type_enc pred_sym s then
  1618                Symtab.map_default (s, [])
  1619                    (insert_type ctxt #3 (s', T_args, T, pred_sym, length args,
  1620                                          in_conj))
  1621              else
  1622                I
  1623            end
  1624          | _ => I)
  1625         #> fold (add_combterm in_conj) args
  1626       end
  1627     fun add_fact in_conj =
  1628       fact_lift (formula_fold NONE (K (add_combterm in_conj)))
  1629   in
  1630     Symtab.empty
  1631     |> is_type_enc_fairly_sound type_enc
  1632        ? (fold (add_fact true) conjs #> fold (add_fact false) facts)
  1633   end
  1634 
  1635 (* This inference is described in section 2.3 of Claessen et al.'s "Sorting it
  1636    out with monotonicity" paper presented at CADE 2011. *)
  1637 fun add_combterm_nonmonotonic_types _ _ _ _ (SOME false) _ = I
  1638   | add_combterm_nonmonotonic_types ctxt level sound locality _
  1639         (CombApp (CombApp (CombConst ((s, _), Type (_, [T, _]), _), tm1),
  1640                            tm2)) =
  1641     (is_tptp_equal s andalso exists is_var_or_bound_var [tm1, tm2] andalso
  1642      (case level of
  1643         Noninf_Nonmono_Types =>
  1644         not (is_locality_global locality) orelse
  1645         not (is_type_surely_infinite ctxt sound T)
  1646       | Fin_Nonmono_Types => is_type_surely_finite ctxt false T
  1647       | _ => true)) ? insert_type ctxt I (deep_freeze_type T)
  1648   | add_combterm_nonmonotonic_types _ _ _ _ _ _ = I
  1649 fun add_fact_nonmonotonic_types ctxt level sound
  1650         ({kind, locality, combformula, ...} : translated_formula) =
  1651   formula_fold (SOME (kind <> Conjecture))
  1652                (add_combterm_nonmonotonic_types ctxt level sound locality)
  1653                combformula
  1654 fun nonmonotonic_types_for_facts ctxt type_enc sound facts =
  1655   let val level = level_of_type_enc type_enc in
  1656     if level = Noninf_Nonmono_Types orelse level = Fin_Nonmono_Types then
  1657       [] |> fold (add_fact_nonmonotonic_types ctxt level sound) facts
  1658          (* We must add "bool" in case the helper "True_or_False" is added
  1659             later. In addition, several places in the code rely on the list of
  1660             nonmonotonic types not being empty. *)
  1661          |> insert_type ctxt I @{typ bool}
  1662     else
  1663       []
  1664   end
  1665 
  1666 fun decl_line_for_sym ctxt format nonmono_Ts type_enc s
  1667                       (s', T_args, T, pred_sym, ary, _) =
  1668   let
  1669     val (T_arg_Ts, level) =
  1670       case type_enc of
  1671         Simple_Types (_, level) => ([], level)
  1672       | _ => (replicate (length T_args) homo_infinite_type, No_Types)
  1673   in
  1674     Decl (sym_decl_prefix ^ s, (s, s'),
  1675           (T_arg_Ts ---> (T |> homogenized_type ctxt nonmono_Ts level ary))
  1676           |> mangled_type format type_enc pred_sym (length T_arg_Ts + ary))
  1677   end
  1678 
  1679 fun formula_line_for_preds_sym_decl ctxt format conj_sym_kind nonmono_Ts
  1680         poly_nonmono_Ts type_enc n s j (s', T_args, T, _, ary, in_conj) =
  1681   let
  1682     val (kind, maybe_negate) =
  1683       if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
  1684       else (Axiom, I)
  1685     val (arg_Ts, res_T) = chop_fun ary T
  1686     val num_args = length arg_Ts
  1687     val bound_names =
  1688       1 upto num_args |> map (`I o make_bound_var o string_of_int)
  1689     val bounds =
  1690       bound_names ~~ arg_Ts |> map (fn (name, T) => CombConst (name, T, []))
  1691     val sym_needs_arg_types = n > 1 orelse exists (curry (op =) dummyT) T_args
  1692     fun should_keep_arg_type T =
  1693       sym_needs_arg_types orelse
  1694       not (should_predicate_on_type ctxt nonmono_Ts type_enc (K false) T)
  1695     val bound_Ts =
  1696       arg_Ts |> map (fn T => if should_keep_arg_type T then SOME T else NONE)
  1697   in
  1698     Formula (preds_sym_formula_prefix ^ s ^
  1699              (if n > 1 then "_" ^ string_of_int j else ""), kind,
  1700              CombConst ((s, s'), T, T_args)
  1701              |> fold (curry (CombApp o swap)) bounds
  1702              |> type_pred_combterm ctxt format type_enc res_T
  1703              |> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)
  1704              |> formula_from_combformula ctxt format poly_nonmono_Ts type_enc
  1705                                          (K (K (K (K true)))) (SOME true)
  1706              |> n > 1 ? bound_tvars type_enc (atyps_of T)
  1707              |> close_formula_universally
  1708              |> maybe_negate,
  1709              intro_info, NONE)
  1710   end
  1711 
  1712 fun formula_lines_for_lightweight_tags_sym_decl ctxt format conj_sym_kind
  1713         poly_nonmono_Ts type_enc n s
  1714         (j, (s', T_args, T, pred_sym, ary, in_conj)) =
  1715   let
  1716     val ident_base =
  1717       lightweight_tags_sym_formula_prefix ^ s ^
  1718       (if n > 1 then "_" ^ string_of_int j else "")
  1719     val (kind, maybe_negate) =
  1720       if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
  1721       else (Axiom, I)
  1722     val (arg_Ts, res_T) = chop_fun ary T
  1723     val bound_names =
  1724       1 upto length arg_Ts |> map (`I o make_bound_var o string_of_int)
  1725     val bounds = bound_names |> map (fn name => ATerm (name, []))
  1726     val cst = mk_const_aterm format type_enc (s, s') T_args
  1727     val atomic_Ts = atyps_of T
  1728     fun eq tms =
  1729       (if pred_sym then AConn (AIff, map AAtom tms)
  1730        else AAtom (ATerm (`I tptp_equal, tms)))
  1731       |> bound_tvars type_enc atomic_Ts
  1732       |> close_formula_universally
  1733       |> maybe_negate
  1734     (* See also "should_tag_with_type". *)
  1735     fun should_encode T =
  1736       should_encode_type ctxt poly_nonmono_Ts All_Types T orelse
  1737       (case type_enc of
  1738          Tags (Polymorphic, level, Lightweight) =>
  1739          level <> All_Types andalso Monomorph.typ_has_tvars T
  1740        | _ => false)
  1741     val tag_with = tag_with_type ctxt format poly_nonmono_Ts type_enc NONE
  1742     val add_formula_for_res =
  1743       if should_encode res_T then
  1744         cons (Formula (ident_base ^ "_res", kind,
  1745                        eq [tag_with res_T (cst bounds), cst bounds],
  1746                        simp_info, NONE))
  1747       else
  1748         I
  1749     fun add_formula_for_arg k =
  1750       let val arg_T = nth arg_Ts k in
  1751         if should_encode arg_T then
  1752           case chop k bounds of
  1753             (bounds1, bound :: bounds2) =>
  1754             cons (Formula (ident_base ^ "_arg" ^ string_of_int (k + 1), kind,
  1755                            eq [cst (bounds1 @ tag_with arg_T bound :: bounds2),
  1756                                cst bounds],
  1757                            simp_info, NONE))
  1758           | _ => raise Fail "expected nonempty tail"
  1759         else
  1760           I
  1761       end
  1762   in
  1763     [] |> not pred_sym ? add_formula_for_res
  1764        |> fold add_formula_for_arg (ary - 1 downto 0)
  1765   end
  1766 
  1767 fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd
  1768 
  1769 fun problem_lines_for_sym_decls ctxt format conj_sym_kind nonmono_Ts
  1770                                 poly_nonmono_Ts type_enc (s, decls) =
  1771   case type_enc of
  1772     Simple_Types _ =>
  1773     decls |> map (decl_line_for_sym ctxt format nonmono_Ts type_enc s)
  1774   | Preds _ =>
  1775     let
  1776       val decls =
  1777         case decls of
  1778           decl :: (decls' as _ :: _) =>
  1779           let val T = result_type_of_decl decl in
  1780             if forall (curry (type_instance ctxt o swap) T
  1781                        o result_type_of_decl) decls' then
  1782               [decl]
  1783             else
  1784               decls
  1785           end
  1786         | _ => decls
  1787       val n = length decls
  1788       val decls =
  1789         decls |> filter (should_predicate_on_type ctxt poly_nonmono_Ts type_enc
  1790                                                   (K true)
  1791                          o result_type_of_decl)
  1792     in
  1793       (0 upto length decls - 1, decls)
  1794       |-> map2 (formula_line_for_preds_sym_decl ctxt format conj_sym_kind
  1795                    nonmono_Ts poly_nonmono_Ts type_enc n s)
  1796     end
  1797   | Tags (_, _, heaviness) =>
  1798     (case heaviness of
  1799        Heavyweight => []
  1800      | Lightweight =>
  1801        let val n = length decls in
  1802          (0 upto n - 1 ~~ decls)
  1803          |> maps (formula_lines_for_lightweight_tags_sym_decl ctxt format
  1804                       conj_sym_kind poly_nonmono_Ts type_enc n s)
  1805        end)
  1806 
  1807 fun problem_lines_for_sym_decl_table ctxt format conj_sym_kind nonmono_Ts
  1808                                      poly_nonmono_Ts type_enc sym_decl_tab =
  1809   sym_decl_tab
  1810   |> Symtab.dest
  1811   |> sort_wrt fst
  1812   |> rpair []
  1813   |-> fold_rev (append o problem_lines_for_sym_decls ctxt format conj_sym_kind
  1814                              nonmono_Ts poly_nonmono_Ts type_enc)
  1815 
  1816 fun needs_type_tag_idempotence (Tags (poly, level, heaviness)) =
  1817     poly <> Mangled_Monomorphic andalso
  1818     ((level = All_Types andalso heaviness = Lightweight) orelse
  1819      level = Noninf_Nonmono_Types orelse level = Fin_Nonmono_Types)
  1820   | needs_type_tag_idempotence _ = false
  1821 
  1822 fun offset_of_heading_in_problem _ [] j = j
  1823   | offset_of_heading_in_problem needle ((heading, lines) :: problem) j =
  1824     if heading = needle then j
  1825     else offset_of_heading_in_problem needle problem (j + length lines)
  1826 
  1827 val implicit_declsN = "Should-be-implicit typings"
  1828 val explicit_declsN = "Explicit typings"
  1829 val factsN = "Relevant facts"
  1830 val class_relsN = "Class relationships"
  1831 val aritiesN = "Arities"
  1832 val helpersN = "Helper facts"
  1833 val conjsN = "Conjectures"
  1834 val free_typesN = "Type variables"
  1835 
  1836 val explicit_apply = NONE (* for experimental purposes *)
  1837 
  1838 fun prepare_atp_problem ctxt format conj_sym_kind prem_kind type_enc sound
  1839         exporter readable_names preproc hyp_ts concl_t facts =
  1840   let
  1841     val (format, type_enc) = choose_format [format] type_enc
  1842     val (fact_names, (conjs, facts, class_rel_clauses, arity_clauses)) =
  1843       translate_formulas ctxt format prem_kind type_enc preproc hyp_ts concl_t
  1844                          facts
  1845     val sym_tab = conjs @ facts |> sym_table_for_facts ctxt explicit_apply
  1846     val nonmono_Ts =
  1847       conjs @ facts |> nonmonotonic_types_for_facts ctxt type_enc sound
  1848     val repair = repair_fact ctxt format type_enc sym_tab
  1849     val (conjs, facts) = (conjs, facts) |> pairself (map repair)
  1850     val repaired_sym_tab =
  1851       conjs @ facts |> sym_table_for_facts ctxt (SOME false)
  1852     val helpers =
  1853       repaired_sym_tab |> helper_facts_for_sym_table ctxt format type_enc
  1854                        |> map repair
  1855     val poly_nonmono_Ts =
  1856       if null nonmono_Ts orelse nonmono_Ts = [@{typ bool}] orelse
  1857          polymorphism_of_type_enc type_enc <> Polymorphic then
  1858         nonmono_Ts
  1859       else
  1860         [TVar (("'a", 0), HOLogic.typeS)]
  1861     val sym_decl_lines =
  1862       (conjs, helpers @ facts)
  1863       |> sym_decl_table_for_facts ctxt type_enc repaired_sym_tab
  1864       |> problem_lines_for_sym_decl_table ctxt format conj_sym_kind nonmono_Ts
  1865                                                poly_nonmono_Ts type_enc
  1866     val helper_lines =
  1867       0 upto length helpers - 1 ~~ helpers
  1868       |> map (formula_line_for_fact ctxt format helper_prefix I false true
  1869                                     poly_nonmono_Ts type_enc)
  1870       |> (if needs_type_tag_idempotence type_enc then
  1871             cons (type_tag_idempotence_fact ())
  1872           else
  1873             I)
  1874     (* Reordering these might confuse the proof reconstruction code or the SPASS
  1875        FLOTTER hack. *)
  1876     val problem =
  1877       [(explicit_declsN, sym_decl_lines),
  1878        (factsN,
  1879         map (formula_line_for_fact ctxt format fact_prefix ascii_of
  1880                                    (not exporter) (not exporter) nonmono_Ts
  1881                                    type_enc)
  1882             (0 upto length facts - 1 ~~ facts)),
  1883        (class_relsN, map formula_line_for_class_rel_clause class_rel_clauses),
  1884        (aritiesN, map formula_line_for_arity_clause arity_clauses),
  1885        (helpersN, helper_lines),
  1886        (conjsN,
  1887         map (formula_line_for_conjecture ctxt format nonmono_Ts type_enc)
  1888             conjs),
  1889        (free_typesN, formula_lines_for_free_types type_enc (facts @ conjs))]
  1890     val problem =
  1891       problem
  1892       |> (case format of
  1893             CNF => ensure_cnf_problem
  1894           | CNF_UEQ => filter_cnf_ueq_problem
  1895           | _ => I)
  1896       |> (if is_format_typed format then
  1897             declare_undeclared_syms_in_atp_problem type_decl_prefix
  1898                                                    implicit_declsN
  1899           else
  1900             I)
  1901     val (problem, pool) = problem |> nice_atp_problem readable_names
  1902     val helpers_offset = offset_of_heading_in_problem helpersN problem 0
  1903     val typed_helpers =
  1904       map_filter (fn (j, {name, ...}) =>
  1905                      if String.isSuffix typed_helper_suffix name then SOME j
  1906                      else NONE)
  1907                  ((helpers_offset + 1 upto helpers_offset + length helpers)
  1908                   ~~ helpers)
  1909     fun add_sym_arity (s, {min_ary, ...} : sym_info) =
  1910       if min_ary > 0 then
  1911         case strip_prefix_and_unascii const_prefix s of
  1912           SOME s => Symtab.insert (op =) (s, min_ary)
  1913         | NONE => I
  1914       else
  1915         I
  1916   in
  1917     (problem,
  1918      case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
  1919      offset_of_heading_in_problem conjsN problem 0,
  1920      offset_of_heading_in_problem factsN problem 0,
  1921      fact_names |> Vector.fromList,
  1922      typed_helpers,
  1923      Symtab.empty |> Symtab.fold add_sym_arity sym_tab)
  1924   end
  1925 
  1926 (* FUDGE *)
  1927 val conj_weight = 0.0
  1928 val hyp_weight = 0.1
  1929 val fact_min_weight = 0.2
  1930 val fact_max_weight = 1.0
  1931 val type_info_default_weight = 0.8
  1932 
  1933 fun add_term_weights weight (ATerm (s, tms)) =
  1934     is_tptp_user_symbol s ? Symtab.default (s, weight)
  1935     #> fold (add_term_weights weight) tms
  1936   | add_term_weights weight (AAbs (_, tm)) = add_term_weights weight tm
  1937 fun add_problem_line_weights weight (Formula (_, _, phi, _, _)) =
  1938     formula_fold NONE (K (add_term_weights weight)) phi
  1939   | add_problem_line_weights _ _ = I
  1940 
  1941 fun add_conjectures_weights [] = I
  1942   | add_conjectures_weights conjs =
  1943     let val (hyps, conj) = split_last conjs in
  1944       add_problem_line_weights conj_weight conj
  1945       #> fold (add_problem_line_weights hyp_weight) hyps
  1946     end
  1947 
  1948 fun add_facts_weights facts =
  1949   let
  1950     val num_facts = length facts
  1951     fun weight_of j =
  1952       fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j
  1953                         / Real.fromInt num_facts
  1954   in
  1955     map weight_of (0 upto num_facts - 1) ~~ facts
  1956     |> fold (uncurry add_problem_line_weights)
  1957   end
  1958 
  1959 (* Weights are from 0.0 (most important) to 1.0 (least important). *)
  1960 fun atp_problem_weights problem =
  1961   let val get = these o AList.lookup (op =) problem in
  1962     Symtab.empty
  1963     |> add_conjectures_weights (get free_typesN @ get conjsN)
  1964     |> add_facts_weights (get factsN)
  1965     |> fold (fold (add_problem_line_weights type_info_default_weight) o get)
  1966             [explicit_declsN, class_relsN, aritiesN]
  1967     |> Symtab.dest
  1968     |> sort (prod_ord Real.compare string_ord o pairself swap)
  1969   end
  1970 
  1971 end;