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