src/HOL/Tools/ATP/atp_translate.ML
author blanchet
Sun, 17 Jul 2011 14:11:35 +0200
changeset 44730 b7890554c424
parent 44729 be41d12de6fa
child 44731 57ef3cd4126e
permissions -rw-r--r--
renamed internal data structure
     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 (* intermediate terms *)
   451 datatype iterm =
   452   IConst of name * typ * typ list |
   453   IVar of name * typ |
   454   IApp of iterm * iterm |
   455   IAbs of (name * typ) * iterm
   456 
   457 fun ityp_of (IConst (_, T, _)) = T
   458   | ityp_of (IVar (_, T)) = T
   459   | ityp_of (IApp (t1, _)) = snd (dest_funT (ityp_of t1))
   460   | ityp_of (IAbs ((_, T), tm)) = T --> ityp_of tm
   461 
   462 (*gets the head of a combinator application, along with the list of arguments*)
   463 fun strip_iterm_comb u =
   464   let
   465     fun stripc (IApp (t, u), ts) = stripc (t, u :: ts)
   466       | stripc x = x
   467   in stripc (u, []) end
   468 
   469 fun atyps_of T = fold_atyps (insert (op =)) T []
   470 
   471 fun new_skolem_const_name s num_T_args =
   472   [new_skolem_const_prefix, s, string_of_int num_T_args]
   473   |> space_implode Long_Name.separator
   474 
   475 (* Converts an Isabelle/HOL term (with combinators) into an intermediate term.
   476    Also accumulates sort infomation. *)
   477 fun iterm_from_term thy bs (P $ Q) =
   478     let
   479       val (P', P_atomics_Ts) = iterm_from_term thy bs P
   480       val (Q', Q_atomics_Ts) = iterm_from_term thy bs Q
   481     in (IApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end
   482   | iterm_from_term thy _ (Const (c, T)) =
   483     let
   484       val tvar_list =
   485         (if String.isPrefix old_skolem_const_prefix c then
   486            [] |> Term.add_tvarsT T |> map TVar
   487          else
   488            (c, T) |> Sign.const_typargs thy)
   489       val c' = IConst (`make_fixed_const c, T, tvar_list)
   490     in (c', atyps_of T) end
   491   | iterm_from_term _ _ (Free (v, T)) =
   492     (IConst (`make_fixed_var v, T, []), atyps_of T)
   493   | iterm_from_term _ _ (Var (v as (s, _), T)) =
   494     (if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then
   495        let
   496          val Ts = T |> strip_type |> swap |> op ::
   497          val s' = new_skolem_const_name s (length Ts)
   498        in IConst (`make_fixed_const s', T, Ts) end
   499      else
   500        IVar ((make_schematic_var v, s), T), atyps_of T)
   501   | iterm_from_term _ bs (Bound j) =
   502     nth bs j |> (fn (s, T) => (IConst (`make_bound_var s, T, []), atyps_of T))
   503   | iterm_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) = iterm_from_term thy ((s, T) :: bs) t
   508     in
   509       (IAbs ((`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    iformula : (name, typ, iterm) formula,
   625    atomic_types : typ list}
   626 
   627 fun update_iformula f ({name, locality, kind, iformula, atomic_types}
   628                        : translated_formula) =
   629   {name = name, locality = locality, kind = kind, iformula = f iformula,
   630    atomic_types = atomic_types} : translated_formula
   631 
   632 fun fact_lift f ({iformula, ...} : translated_formula) = f iformula
   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 iterm_vars (IApp (tm1, tm2)) = fold iterm_vars [tm1, tm2]
   716   | iterm_vars (IConst _) = I
   717   | iterm_vars (IVar (name, T)) = insert (op =) (name, SOME T)
   718   | iterm_vars (IAbs (_, tm)) = iterm_vars tm
   719 fun close_iformula_universally phi = close_universally iterm_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 (IApp (tm1, tm2)) =
   820         IApp (intro top_level tm1, intro false tm2)
   821       | intro top_level (IConst (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) => IConst (name, T, T_args))
   842       | intro _ (IAbs (bound, tm)) = IAbs (bound, intro false tm)
   843       | intro _ tm = tm
   844   in intro true end
   845 
   846 fun iformula_from_prop thy type_enc eq_as_iff =
   847   let
   848     fun do_term bs t atomic_types =
   849       iterm_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 (iformula, atomic_types) =
   990       iformula_from_prop thy type_enc eq_as_iff t []
   991   in
   992     {name = name, locality = loc, kind = kind, iformula = iformula,
   993      atomic_types = atomic_types}
   994   end
   995 
   996 fun make_fact ctxt format type_enc trans_lambdas eq_as_iff preproc
   997               presimp_consts ((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 {iformula = AAtom (IConst ((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_iformula 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 (IConst ((s, _), _, _)) =
  1058     String.isPrefix bound_var_prefix s
  1059   | is_var_or_bound_var (IVar _) = 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_iterm_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_iterm_comb tm in
  1132         (case head of
  1133            IConst ((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          | IVar (_, T) => add_var_or_bound_var T accum
  1174          | IAbs ((_, T), tm) => accum |> add_var_or_bound_var T |> add false tm
  1175          | _ => accum)
  1176         |> fold (add false) args
  1177       end
  1178   in add true end
  1179 fun add_fact_syms_to_table ctxt explicit_apply =
  1180   fact_lift (formula_fold NONE
  1181                           (K (add_iterm_syms_to_table ctxt explicit_apply)))
  1182 
  1183 val default_sym_tab_entries : (string * sym_info) list =
  1184   (prefixed_predicator_name,
  1185    {pred_sym = true, min_ary = 1, max_ary = 1, types = []}) ::
  1186   ([tptp_false, tptp_true]
  1187    |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, types = []})) @
  1188   ([tptp_equal, tptp_old_equal]
  1189    |> map (rpair {pred_sym = true, min_ary = 2, max_ary = 2, types = []}))
  1190 
  1191 fun sym_table_for_facts ctxt explicit_apply facts =
  1192   ((false, []), Symtab.empty)
  1193   |> fold (add_fact_syms_to_table ctxt explicit_apply) facts |> snd
  1194   |> fold Symtab.update default_sym_tab_entries
  1195 
  1196 fun min_arity_of sym_tab s =
  1197   case Symtab.lookup sym_tab s of
  1198     SOME ({min_ary, ...} : sym_info) => min_ary
  1199   | NONE =>
  1200     case strip_prefix_and_unascii const_prefix s of
  1201       SOME s =>
  1202       let val s = s |> unmangled_const_name |> invert_const in
  1203         if s = predicator_name then 1
  1204         else if s = app_op_name then 2
  1205         else if s = type_pred_name then 1
  1206         else 0
  1207       end
  1208     | NONE => 0
  1209 
  1210 (* True if the constant ever appears outside of the top-level position in
  1211    literals, or if it appears with different arities (e.g., because of different
  1212    type instantiations). If false, the constant always receives all of its
  1213    arguments and is used as a predicate. *)
  1214 fun is_pred_sym sym_tab s =
  1215   case Symtab.lookup sym_tab s of
  1216     SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
  1217     pred_sym andalso min_ary = max_ary
  1218   | NONE => false
  1219 
  1220 val predicator_combconst =
  1221   IConst (`make_fixed_const predicator_name, @{typ "bool => bool"}, [])
  1222 fun predicator tm = IApp (predicator_combconst, tm)
  1223 
  1224 fun introduce_predicators_in_iterm sym_tab tm =
  1225   case strip_iterm_comb tm of
  1226     (IConst ((s, _), _, _), _) =>
  1227     if is_pred_sym sym_tab s then tm else predicator tm
  1228   | _ => predicator tm
  1229 
  1230 fun list_app head args = fold (curry (IApp o swap)) args head
  1231 
  1232 val app_op = `make_fixed_const app_op_name
  1233 
  1234 fun explicit_app arg head =
  1235   let
  1236     val head_T = ityp_of head
  1237     val (arg_T, res_T) = dest_funT head_T
  1238     val explicit_app = IConst (app_op, head_T --> head_T, [arg_T, res_T])
  1239   in list_app explicit_app [head, arg] end
  1240 fun list_explicit_app head args = fold explicit_app args head
  1241 
  1242 fun introduce_explicit_apps_in_iterm sym_tab =
  1243   let
  1244     fun aux tm =
  1245       case strip_iterm_comb tm of
  1246         (head as IConst ((s, _), _, _), args) =>
  1247         args |> map aux
  1248              |> chop (min_arity_of sym_tab s)
  1249              |>> list_app head
  1250              |-> list_explicit_app
  1251       | (head, args) => list_explicit_app head (map aux args)
  1252   in aux end
  1253 
  1254 fun chop_fun 0 T = ([], T)
  1255   | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
  1256     chop_fun (n - 1) ran_T |>> cons dom_T
  1257   | chop_fun _ _ = raise Fail "unexpected non-function"
  1258 
  1259 fun filter_type_args _ _ _ [] = []
  1260   | filter_type_args thy s arity T_args =
  1261     let
  1262       (* will throw "TYPE" for pseudo-constants *)
  1263       val U = if s = app_op_name then
  1264                 @{typ "('a => 'b) => 'a => 'b"} |> Logic.varifyT_global
  1265               else
  1266                 s |> Sign.the_const_type thy
  1267     in
  1268       case Term.add_tvarsT (U |> chop_fun arity |> snd) [] of
  1269         [] => []
  1270       | res_U_vars =>
  1271         let val U_args = (s, U) |> Sign.const_typargs thy in
  1272           U_args ~~ T_args
  1273           |> map (fn (U, T) =>
  1274                      if member (op =) res_U_vars (dest_TVar U) then T
  1275                      else dummyT)
  1276         end
  1277     end
  1278     handle TYPE _ => T_args
  1279 
  1280 fun enforce_type_arg_policy_in_iterm ctxt format type_enc =
  1281   let
  1282     val thy = Proof_Context.theory_of ctxt
  1283     fun aux arity (IApp (tm1, tm2)) = IApp (aux (arity + 1) tm1, aux 0 tm2)
  1284       | aux arity (IConst (name as (s, _), T, T_args)) =
  1285         (case strip_prefix_and_unascii const_prefix s of
  1286            NONE => (name, T_args)
  1287          | SOME s'' =>
  1288            let
  1289              val s'' = invert_const s''
  1290              fun filtered_T_args false = T_args
  1291                | filtered_T_args true = filter_type_args thy s'' arity T_args
  1292            in
  1293              case type_arg_policy type_enc s'' of
  1294                Explicit_Type_Args drop_args =>
  1295                (name, filtered_T_args drop_args)
  1296              | Mangled_Type_Args drop_args =>
  1297                (mangled_const_name format type_enc (filtered_T_args drop_args)
  1298                                    name, [])
  1299              | No_Type_Args => (name, [])
  1300            end)
  1301         |> (fn (name, T_args) => IConst (name, T, T_args))
  1302       | aux _ (IAbs (bound, tm)) = IAbs (bound, aux 0 tm)
  1303       | aux _ tm = tm
  1304   in aux 0 end
  1305 
  1306 fun repair_iterm ctxt format type_enc sym_tab =
  1307   not (is_type_enc_higher_order type_enc)
  1308   ? (introduce_explicit_apps_in_iterm sym_tab
  1309      #> introduce_predicators_in_iterm sym_tab)
  1310   #> enforce_type_arg_policy_in_iterm ctxt format type_enc
  1311 fun repair_fact ctxt format type_enc sym_tab =
  1312   update_iformula (formula_map (repair_iterm ctxt format type_enc sym_tab))
  1313 
  1314 (** Helper facts **)
  1315 
  1316 (* The Boolean indicates that a fairly sound type encoding is needed. *)
  1317 val helper_table =
  1318   [(("COMBI", false), @{thms Meson.COMBI_def}),
  1319    (("COMBK", false), @{thms Meson.COMBK_def}),
  1320    (("COMBB", false), @{thms Meson.COMBB_def}),
  1321    (("COMBC", false), @{thms Meson.COMBC_def}),
  1322    (("COMBS", false), @{thms Meson.COMBS_def}),
  1323    (("fFalse", false), [@{lemma "~ fFalse" by (unfold fFalse_def) fast}]),
  1324    (("fFalse", true), @{thms True_or_False}),
  1325    (("fTrue", false), [@{lemma "fTrue" by (unfold fTrue_def) fast}]),
  1326    (("fTrue", true), @{thms True_or_False}),
  1327    (("fNot", false),
  1328     @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1]
  1329            fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),
  1330    (("fconj", false),
  1331     @{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q"
  1332         by (unfold fconj_def) fast+}),
  1333    (("fdisj", false),
  1334     @{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q"
  1335         by (unfold fdisj_def) fast+}),
  1336    (("fimplies", false),
  1337     @{lemma "P | fimplies P Q" "~ Q | fimplies P Q" "~ fimplies P Q | ~ P | Q"
  1338         by (unfold fimplies_def) fast+}),
  1339    (("fequal", true),
  1340     (* This is a lie: Higher-order equality doesn't need a sound type encoding.
  1341        However, this is done so for backward compatibility: Including the
  1342        equality helpers by default in Metis breaks a few existing proofs. *)
  1343     @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1]
  1344            fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),
  1345    (("fAll", false), []), (*TODO: add helpers*)
  1346    (("fEx", false), []), (*TODO: add helpers*)
  1347    (("If", true), @{thms if_True if_False True_or_False})]
  1348   |> map (apsnd (map zero_var_indexes))
  1349 
  1350 val type_tag = `make_fixed_const type_tag_name
  1351 
  1352 fun type_tag_idempotence_fact () =
  1353   let
  1354     fun var s = ATerm (`I s, [])
  1355     fun tag tm = ATerm (type_tag, [var "T", tm])
  1356     val tagged_a = tag (var "A")
  1357   in
  1358     Formula (type_tag_idempotence_helper_name, Axiom,
  1359              AAtom (ATerm (`I tptp_equal, [tag tagged_a, tagged_a]))
  1360              |> close_formula_universally, isabelle_info simpN, NONE)
  1361   end
  1362 
  1363 fun should_specialize_helper type_enc t =
  1364   polymorphism_of_type_enc type_enc = Mangled_Monomorphic andalso
  1365   level_of_type_enc type_enc <> No_Types andalso
  1366   not (null (Term.hidden_polymorphism t))
  1367 
  1368 fun helper_facts_for_sym ctxt format type_enc (s, {types, ...} : sym_info) =
  1369   case strip_prefix_and_unascii const_prefix s of
  1370     SOME mangled_s =>
  1371     let
  1372       val thy = Proof_Context.theory_of ctxt
  1373       val unmangled_s = mangled_s |> unmangled_const_name
  1374       fun dub needs_fairly_sound j k =
  1375         (unmangled_s ^ "_" ^ string_of_int j ^ "_" ^ string_of_int k ^
  1376          (if mangled_s = unmangled_s then "" else "_" ^ ascii_of mangled_s) ^
  1377          (if needs_fairly_sound then typed_helper_suffix
  1378           else untyped_helper_suffix),
  1379          Helper)
  1380       fun dub_and_inst needs_fairly_sound (th, j) =
  1381         let val t = prop_of th in
  1382           if should_specialize_helper type_enc t then
  1383             map (fn T => specialize_type thy (invert_const unmangled_s, T) t)
  1384                 types
  1385           else
  1386             [t]
  1387         end
  1388         |> map (fn (k, t) => (dub needs_fairly_sound j k, t)) o tag_list 1
  1389       val make_facts =
  1390         map_filter (make_fact ctxt format type_enc I false false [])
  1391       val fairly_sound = is_type_enc_fairly_sound type_enc
  1392     in
  1393       helper_table
  1394       |> maps (fn ((helper_s, needs_fairly_sound), ths) =>
  1395                   if helper_s <> unmangled_s orelse
  1396                      (needs_fairly_sound andalso not fairly_sound) then
  1397                     []
  1398                   else
  1399                     ths ~~ (1 upto length ths)
  1400                     |> maps (dub_and_inst needs_fairly_sound)
  1401                     |> make_facts)
  1402     end
  1403   | NONE => []
  1404 fun helper_facts_for_sym_table ctxt format type_enc sym_tab =
  1405   Symtab.fold_rev (append o helper_facts_for_sym ctxt format type_enc) sym_tab
  1406                   []
  1407 
  1408 (***************************************************************)
  1409 (* Type Classes Present in the Axiom or Conjecture Clauses     *)
  1410 (***************************************************************)
  1411 
  1412 fun set_insert (x, s) = Symtab.update (x, ()) s
  1413 
  1414 fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
  1415 
  1416 (* Remove this trivial type class (FIXME: similar code elsewhere) *)
  1417 fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset
  1418 
  1419 fun classes_of_terms get_Ts =
  1420   map (map snd o get_Ts)
  1421   #> List.foldl add_classes Symtab.empty
  1422   #> delete_type #> Symtab.keys
  1423 
  1424 val tfree_classes_of_terms = classes_of_terms OldTerm.term_tfrees
  1425 val tvar_classes_of_terms = classes_of_terms OldTerm.term_tvars
  1426 
  1427 fun fold_type_constrs f (Type (s, Ts)) x =
  1428     fold (fold_type_constrs f) Ts (f (s, x))
  1429   | fold_type_constrs _ _ x = x
  1430 
  1431 (*Type constructors used to instantiate overloaded constants are the only ones needed.*)
  1432 fun add_type_constrs_in_term thy =
  1433   let
  1434     fun add (Const (@{const_name Meson.skolem}, _) $ _) = I
  1435       | add (t $ u) = add t #> add u
  1436       | add (Const (x as (s, _))) =
  1437         if String.isPrefix skolem_const_prefix s then I
  1438         else x |> Sign.const_typargs thy |> fold (fold_type_constrs set_insert)
  1439       | add (Abs (_, _, u)) = add u
  1440       | add _ = I
  1441   in add end
  1442 
  1443 fun type_constrs_of_terms thy ts =
  1444   Symtab.keys (fold (add_type_constrs_in_term thy) ts Symtab.empty)
  1445 
  1446 fun translate_formulas ctxt format prem_kind type_enc trans_lambdas preproc
  1447                        hyp_ts concl_t facts =
  1448   let
  1449     val thy = Proof_Context.theory_of ctxt
  1450     val fact_ts = facts |> map snd
  1451     val presimp_consts = Meson.presimplified_consts ctxt
  1452     val make_fact =
  1453       make_fact ctxt format type_enc trans_lambdas true preproc presimp_consts
  1454     val (facts, fact_names) =
  1455       facts |> map (fn (name, t) => (name, t) |> make_fact |> rpair name)
  1456             |> map_filter (try (apfst the))
  1457             |> ListPair.unzip
  1458     (* Remove existing facts from the conjecture, as this can dramatically
  1459        boost an ATP's performance (for some reason). *)
  1460     val hyp_ts =
  1461       hyp_ts
  1462       |> map (fn t => if member (op aconv) fact_ts t then @{prop True} else t)
  1463     val goal_t = Logic.list_implies (hyp_ts, concl_t)
  1464     val all_ts = goal_t :: fact_ts
  1465     val subs = tfree_classes_of_terms all_ts
  1466     val supers = tvar_classes_of_terms all_ts
  1467     val tycons = type_constrs_of_terms thy all_ts
  1468     val conjs =
  1469       hyp_ts @ [concl_t]
  1470       |> make_conjecture ctxt format prem_kind type_enc trans_lambdas preproc
  1471                          presimp_consts
  1472     val (supers', arity_clauses) =
  1473       if level_of_type_enc type_enc = No_Types then ([], [])
  1474       else make_arity_clauses thy tycons supers
  1475     val class_rel_clauses = make_class_rel_clauses thy subs supers'
  1476   in
  1477     (fact_names |> map single, (conjs, facts, class_rel_clauses, arity_clauses))
  1478   end
  1479 
  1480 fun fo_literal_from_type_literal (TyLitVar (class, name)) =
  1481     (true, ATerm (class, [ATerm (name, [])]))
  1482   | fo_literal_from_type_literal (TyLitFree (class, name)) =
  1483     (true, ATerm (class, [ATerm (name, [])]))
  1484 
  1485 fun formula_from_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
  1486 
  1487 val type_pred = `make_fixed_const type_pred_name
  1488 
  1489 fun type_pred_iterm ctxt format type_enc T tm =
  1490   IApp (IConst (type_pred, T --> @{typ bool}, [T])
  1491         |> enforce_type_arg_policy_in_iterm ctxt format type_enc, tm)
  1492 
  1493 fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum
  1494   | is_var_positively_naked_in_term name _ (ATerm ((s, _), tms)) accum =
  1495     accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm (name, [])))
  1496   | is_var_positively_naked_in_term _ _ _ _ = true
  1497 fun should_predicate_on_var_in_formula pos phi (SOME true) name =
  1498     formula_fold pos (is_var_positively_naked_in_term name) phi false
  1499   | should_predicate_on_var_in_formula _ _ _ _ = true
  1500 
  1501 fun mk_aterm format type_enc name T_args args =
  1502   ATerm (name, map_filter (ho_term_for_type_arg format type_enc) T_args @ args)
  1503 
  1504 fun tag_with_type ctxt format nonmono_Ts type_enc pos T tm =
  1505   IConst (type_tag, T --> T, [T])
  1506   |> enforce_type_arg_policy_in_iterm ctxt format type_enc
  1507   |> ho_term_from_iterm ctxt format nonmono_Ts type_enc (Top_Level pos)
  1508   |> (fn ATerm (s, tms) => ATerm (s, tms @ [tm])
  1509        | _ => raise Fail "unexpected lambda-abstraction")
  1510 and ho_term_from_iterm ctxt format nonmono_Ts type_enc =
  1511   let
  1512     fun aux site u =
  1513       let
  1514         val (head, args) = strip_iterm_comb u
  1515         val pos =
  1516           case site of
  1517             Top_Level pos => pos
  1518           | Eq_Arg pos => pos
  1519           | Elsewhere => NONE
  1520         val t =
  1521           case head of
  1522             IConst (name as (s, _), _, T_args) =>
  1523             let
  1524               val arg_site = if is_tptp_equal s then Eq_Arg pos else Elsewhere
  1525             in
  1526               mk_aterm format type_enc name T_args (map (aux arg_site) args)
  1527             end
  1528           | IVar (name, _) =>
  1529             mk_aterm format type_enc name [] (map (aux Elsewhere) args)
  1530           | IAbs ((name, T), tm) =>
  1531             AAbs ((name, ho_type_from_typ format type_enc true 0 T),
  1532                   aux Elsewhere tm)
  1533           | IApp _ => raise Fail "impossible \"IApp\""
  1534         val T = ityp_of u
  1535       in
  1536         t |> (if should_tag_with_type ctxt nonmono_Ts type_enc site u T then
  1537                 tag_with_type ctxt format nonmono_Ts type_enc pos T
  1538               else
  1539                 I)
  1540       end
  1541   in aux end
  1542 and formula_from_iformula ctxt format nonmono_Ts type_enc
  1543                           should_predicate_on_var =
  1544   let
  1545     fun do_term pos =
  1546       ho_term_from_iterm ctxt format nonmono_Ts type_enc (Top_Level pos)
  1547     val do_bound_type =
  1548       case type_enc of
  1549         Simple_Types (_, level) =>
  1550         homogenized_type ctxt nonmono_Ts level 0
  1551         #> ho_type_from_typ format type_enc false 0 #> SOME
  1552       | _ => K NONE
  1553     fun do_out_of_bound_type pos phi universal (name, T) =
  1554       if should_predicate_on_type ctxt nonmono_Ts type_enc
  1555              (fn () => should_predicate_on_var pos phi universal name) T then
  1556         IVar (name, T)
  1557         |> type_pred_iterm ctxt format type_enc T
  1558         |> do_term pos |> AAtom |> SOME
  1559       else
  1560         NONE
  1561     fun do_formula pos (AQuant (q, xs, phi)) =
  1562         let
  1563           val phi = phi |> do_formula pos
  1564           val universal = Option.map (q = AExists ? not) pos
  1565         in
  1566           AQuant (q, xs |> map (apsnd (fn NONE => NONE
  1567                                         | SOME T => do_bound_type T)),
  1568                   (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))
  1569                       (map_filter
  1570                            (fn (_, NONE) => NONE
  1571                              | (s, SOME T) =>
  1572                                do_out_of_bound_type pos phi universal (s, T))
  1573                            xs)
  1574                       phi)
  1575         end
  1576       | do_formula pos (AConn conn) = aconn_map pos do_formula conn
  1577       | do_formula pos (AAtom tm) = AAtom (do_term pos tm)
  1578   in do_formula end
  1579 
  1580 fun bound_tvars type_enc Ts =
  1581   mk_ahorn (map (formula_from_fo_literal o fo_literal_from_type_literal)
  1582                 (type_literals_for_types type_enc add_sorts_on_tvar Ts))
  1583 
  1584 (* Each fact is given a unique fact number to avoid name clashes (e.g., because
  1585    of monomorphization). The TPTP explicitly forbids name clashes, and some of
  1586    the remote provers might care. *)
  1587 fun formula_line_for_fact ctxt format prefix encode freshen pos nonmono_Ts
  1588         type_enc (j, {name, locality, kind, iformula, atomic_types}) =
  1589   (prefix ^ (if freshen then string_of_int j ^ "_" else "") ^ encode name,
  1590    kind,
  1591    iformula
  1592    |> close_iformula_universally
  1593    |> formula_from_iformula ctxt format nonmono_Ts type_enc
  1594                             should_predicate_on_var_in_formula
  1595                             (if pos then SOME true else NONE)
  1596    |> bound_tvars type_enc atomic_types
  1597    |> close_formula_universally,
  1598    NONE,
  1599    case locality of
  1600      Intro => isabelle_info introN
  1601    | Elim => isabelle_info elimN
  1602    | Simp => isabelle_info simpN
  1603    | _ => NONE)
  1604   |> Formula
  1605 
  1606 fun formula_line_for_class_rel_clause ({name, subclass, superclass, ...}
  1607                                        : class_rel_clause) =
  1608   let val ty_arg = ATerm (`I "T", []) in
  1609     Formula (class_rel_clause_prefix ^ ascii_of name, Axiom,
  1610              AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
  1611                                AAtom (ATerm (superclass, [ty_arg]))])
  1612              |> close_formula_universally, isabelle_info introN, NONE)
  1613   end
  1614 
  1615 fun fo_literal_from_arity_literal (TConsLit (c, t, args)) =
  1616     (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
  1617   | fo_literal_from_arity_literal (TVarLit (c, sort)) =
  1618     (false, ATerm (c, [ATerm (sort, [])]))
  1619 
  1620 fun formula_line_for_arity_clause ({name, prem_lits, concl_lits, ...}
  1621                                    : arity_clause) =
  1622   Formula (arity_clause_prefix ^ name, Axiom,
  1623            mk_ahorn (map (formula_from_fo_literal o apfst not
  1624                           o fo_literal_from_arity_literal) prem_lits)
  1625                     (formula_from_fo_literal
  1626                          (fo_literal_from_arity_literal concl_lits))
  1627            |> close_formula_universally, isabelle_info introN, NONE)
  1628 
  1629 fun formula_line_for_conjecture ctxt format nonmono_Ts type_enc
  1630         ({name, kind, iformula, atomic_types, ...} : translated_formula) =
  1631   Formula (conjecture_prefix ^ name, kind,
  1632            formula_from_iformula ctxt format nonmono_Ts type_enc
  1633                should_predicate_on_var_in_formula (SOME false)
  1634                (close_iformula_universally iformula)
  1635            |> bound_tvars type_enc atomic_types
  1636            |> close_formula_universally, NONE, NONE)
  1637 
  1638 fun free_type_literals type_enc ({atomic_types, ...} : translated_formula) =
  1639   atomic_types |> type_literals_for_types type_enc add_sorts_on_tfree
  1640                |> map fo_literal_from_type_literal
  1641 
  1642 fun formula_line_for_free_type j lit =
  1643   Formula (tfree_clause_prefix ^ string_of_int j, Hypothesis,
  1644            formula_from_fo_literal lit, NONE, NONE)
  1645 fun formula_lines_for_free_types type_enc facts =
  1646   let
  1647     val litss = map (free_type_literals type_enc) facts
  1648     val lits = fold (union (op =)) litss []
  1649   in map2 formula_line_for_free_type (0 upto length lits - 1) lits end
  1650 
  1651 (** Symbol declarations **)
  1652 
  1653 fun should_declare_sym type_enc pred_sym s =
  1654   is_tptp_user_symbol s andalso not (String.isPrefix bound_var_prefix s) andalso
  1655   (case type_enc of
  1656      Simple_Types _ => true
  1657    | Tags (_, _, Lightweight) => true
  1658    | _ => not pred_sym)
  1659 
  1660 fun sym_decl_table_for_facts ctxt type_enc repaired_sym_tab (conjs, facts) =
  1661   let
  1662     fun add_iterm in_conj tm =
  1663       let val (head, args) = strip_iterm_comb tm in
  1664         (case head of
  1665            IConst ((s, s'), T, T_args) =>
  1666            let val pred_sym = is_pred_sym repaired_sym_tab s in
  1667              if should_declare_sym type_enc pred_sym s then
  1668                Symtab.map_default (s, [])
  1669                    (insert_type ctxt #3 (s', T_args, T, pred_sym, length args,
  1670                                          in_conj))
  1671              else
  1672                I
  1673            end
  1674          | IAbs (_, tm) => add_iterm in_conj tm
  1675          | _ => I)
  1676         #> fold (add_iterm in_conj) args
  1677       end
  1678     fun add_fact in_conj = fact_lift (formula_fold NONE (K (add_iterm in_conj)))
  1679   in
  1680     Symtab.empty
  1681     |> is_type_enc_fairly_sound type_enc
  1682        ? (fold (add_fact true) conjs #> fold (add_fact false) facts)
  1683   end
  1684 
  1685 (* This inference is described in section 2.3 of Claessen et al.'s "Sorting it
  1686    out with monotonicity" paper presented at CADE 2011. *)
  1687 fun add_iterm_nonmonotonic_types _ _ _ _ (SOME false) _ = I
  1688   | add_iterm_nonmonotonic_types ctxt level sound locality _
  1689         (IApp (IApp (IConst ((s, _), Type (_, [T, _]), _), tm1), tm2)) =
  1690     (is_tptp_equal s andalso exists is_var_or_bound_var [tm1, tm2] andalso
  1691      (case level of
  1692         Noninf_Nonmono_Types =>
  1693         not (is_locality_global locality) orelse
  1694         not (is_type_surely_infinite ctxt sound T)
  1695       | Fin_Nonmono_Types => is_type_surely_finite ctxt false T
  1696       | _ => true)) ? insert_type ctxt I (deep_freeze_type T)
  1697   | add_iterm_nonmonotonic_types _ _ _ _ _ _ = I
  1698 fun add_fact_nonmonotonic_types ctxt level sound
  1699         ({kind, locality, iformula, ...} : translated_formula) =
  1700   formula_fold (SOME (kind <> Conjecture))
  1701                (add_iterm_nonmonotonic_types ctxt level sound locality)
  1702                iformula
  1703 fun nonmonotonic_types_for_facts ctxt type_enc sound facts =
  1704   let val level = level_of_type_enc type_enc in
  1705     if level = Noninf_Nonmono_Types orelse level = Fin_Nonmono_Types then
  1706       [] |> fold (add_fact_nonmonotonic_types ctxt level sound) facts
  1707          (* We must add "bool" in case the helper "True_or_False" is added
  1708             later. In addition, several places in the code rely on the list of
  1709             nonmonotonic types not being empty. *)
  1710          |> insert_type ctxt I @{typ bool}
  1711     else
  1712       []
  1713   end
  1714 
  1715 fun decl_line_for_sym ctxt format nonmono_Ts type_enc s
  1716                       (s', T_args, T, pred_sym, ary, _) =
  1717   let
  1718     val (T_arg_Ts, level) =
  1719       case type_enc of
  1720         Simple_Types (_, level) => ([], level)
  1721       | _ => (replicate (length T_args) homo_infinite_type, No_Types)
  1722   in
  1723     Decl (sym_decl_prefix ^ s, (s, s'),
  1724           (T_arg_Ts ---> (T |> homogenized_type ctxt nonmono_Ts level ary))
  1725           |> ho_type_from_typ format type_enc pred_sym (length T_arg_Ts + ary))
  1726   end
  1727 
  1728 fun formula_line_for_preds_sym_decl ctxt format conj_sym_kind nonmono_Ts
  1729         poly_nonmono_Ts type_enc n s j (s', T_args, T, _, ary, in_conj) =
  1730   let
  1731     val (kind, maybe_negate) =
  1732       if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
  1733       else (Axiom, I)
  1734     val (arg_Ts, res_T) = chop_fun ary T
  1735     val num_args = length arg_Ts
  1736     val bound_names =
  1737       1 upto num_args |> map (`I o make_bound_var o string_of_int)
  1738     val bounds =
  1739       bound_names ~~ arg_Ts |> map (fn (name, T) => IConst (name, T, []))
  1740     val sym_needs_arg_types = n > 1 orelse exists (curry (op =) dummyT) T_args
  1741     fun should_keep_arg_type T =
  1742       sym_needs_arg_types orelse
  1743       not (should_predicate_on_type ctxt nonmono_Ts type_enc (K false) T)
  1744     val bound_Ts =
  1745       arg_Ts |> map (fn T => if should_keep_arg_type T then SOME T else NONE)
  1746   in
  1747     Formula (preds_sym_formula_prefix ^ s ^
  1748              (if n > 1 then "_" ^ string_of_int j else ""), kind,
  1749              IConst ((s, s'), T, T_args)
  1750              |> fold (curry (IApp o swap)) bounds
  1751              |> type_pred_iterm ctxt format type_enc res_T
  1752              |> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)
  1753              |> formula_from_iformula ctxt format poly_nonmono_Ts type_enc
  1754                                       (K (K (K (K true)))) (SOME true)
  1755              |> n > 1 ? bound_tvars type_enc (atyps_of T)
  1756              |> close_formula_universally
  1757              |> maybe_negate,
  1758              isabelle_info introN, NONE)
  1759   end
  1760 
  1761 fun formula_lines_for_lightweight_tags_sym_decl ctxt format conj_sym_kind
  1762         poly_nonmono_Ts type_enc n s
  1763         (j, (s', T_args, T, pred_sym, ary, in_conj)) =
  1764   let
  1765     val ident_base =
  1766       lightweight_tags_sym_formula_prefix ^ s ^
  1767       (if n > 1 then "_" ^ string_of_int j else "")
  1768     val (kind, maybe_negate) =
  1769       if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
  1770       else (Axiom, I)
  1771     val (arg_Ts, res_T) = chop_fun ary T
  1772     val bound_names =
  1773       1 upto length arg_Ts |> map (`I o make_bound_var o string_of_int)
  1774     val bounds = bound_names |> map (fn name => ATerm (name, []))
  1775     val cst = mk_aterm format type_enc (s, s') T_args
  1776     val atomic_Ts = atyps_of T
  1777     fun eq tms =
  1778       (if pred_sym then AConn (AIff, map AAtom tms)
  1779        else AAtom (ATerm (`I tptp_equal, tms)))
  1780       |> bound_tvars type_enc atomic_Ts
  1781       |> close_formula_universally
  1782       |> maybe_negate
  1783     (* See also "should_tag_with_type". *)
  1784     fun should_encode T =
  1785       should_encode_type ctxt poly_nonmono_Ts All_Types T orelse
  1786       (case type_enc of
  1787          Tags (Polymorphic, level, Lightweight) =>
  1788          level <> All_Types andalso Monomorph.typ_has_tvars T
  1789        | _ => false)
  1790     val tag_with = tag_with_type ctxt format poly_nonmono_Ts type_enc NONE
  1791     val add_formula_for_res =
  1792       if should_encode res_T then
  1793         cons (Formula (ident_base ^ "_res", kind,
  1794                        eq [tag_with res_T (cst bounds), cst bounds],
  1795                        isabelle_info simpN, NONE))
  1796       else
  1797         I
  1798     fun add_formula_for_arg k =
  1799       let val arg_T = nth arg_Ts k in
  1800         if should_encode arg_T then
  1801           case chop k bounds of
  1802             (bounds1, bound :: bounds2) =>
  1803             cons (Formula (ident_base ^ "_arg" ^ string_of_int (k + 1), kind,
  1804                            eq [cst (bounds1 @ tag_with arg_T bound :: bounds2),
  1805                                cst bounds],
  1806                            isabelle_info simpN, NONE))
  1807           | _ => raise Fail "expected nonempty tail"
  1808         else
  1809           I
  1810       end
  1811   in
  1812     [] |> not pred_sym ? add_formula_for_res
  1813        |> fold add_formula_for_arg (ary - 1 downto 0)
  1814   end
  1815 
  1816 fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd
  1817 
  1818 fun problem_lines_for_sym_decls ctxt format conj_sym_kind nonmono_Ts
  1819                                 poly_nonmono_Ts type_enc (s, decls) =
  1820   case type_enc of
  1821     Simple_Types _ =>
  1822     decls |> map (decl_line_for_sym ctxt format nonmono_Ts type_enc s)
  1823   | Preds _ =>
  1824     let
  1825       val decls =
  1826         case decls of
  1827           decl :: (decls' as _ :: _) =>
  1828           let val T = result_type_of_decl decl in
  1829             if forall (curry (type_instance ctxt o swap) T
  1830                        o result_type_of_decl) decls' then
  1831               [decl]
  1832             else
  1833               decls
  1834           end
  1835         | _ => decls
  1836       val n = length decls
  1837       val decls =
  1838         decls |> filter (should_predicate_on_type ctxt poly_nonmono_Ts type_enc
  1839                                                   (K true)
  1840                          o result_type_of_decl)
  1841     in
  1842       (0 upto length decls - 1, decls)
  1843       |-> map2 (formula_line_for_preds_sym_decl ctxt format conj_sym_kind
  1844                    nonmono_Ts poly_nonmono_Ts type_enc n s)
  1845     end
  1846   | Tags (_, _, heaviness) =>
  1847     (case heaviness of
  1848        Heavyweight => []
  1849      | Lightweight =>
  1850        let val n = length decls in
  1851          (0 upto n - 1 ~~ decls)
  1852          |> maps (formula_lines_for_lightweight_tags_sym_decl ctxt format
  1853                       conj_sym_kind poly_nonmono_Ts type_enc n s)
  1854        end)
  1855 
  1856 fun problem_lines_for_sym_decl_table ctxt format conj_sym_kind nonmono_Ts
  1857                                      poly_nonmono_Ts type_enc sym_decl_tab =
  1858   sym_decl_tab
  1859   |> Symtab.dest
  1860   |> sort_wrt fst
  1861   |> rpair []
  1862   |-> fold_rev (append o problem_lines_for_sym_decls ctxt format conj_sym_kind
  1863                              nonmono_Ts poly_nonmono_Ts type_enc)
  1864 
  1865 fun needs_type_tag_idempotence (Tags (poly, level, heaviness)) =
  1866     poly <> Mangled_Monomorphic andalso
  1867     ((level = All_Types andalso heaviness = Lightweight) orelse
  1868      level = Noninf_Nonmono_Types orelse level = Fin_Nonmono_Types)
  1869   | needs_type_tag_idempotence _ = false
  1870 
  1871 fun offset_of_heading_in_problem _ [] j = j
  1872   | offset_of_heading_in_problem needle ((heading, lines) :: problem) j =
  1873     if heading = needle then j
  1874     else offset_of_heading_in_problem needle problem (j + length lines)
  1875 
  1876 val implicit_declsN = "Should-be-implicit typings"
  1877 val explicit_declsN = "Explicit typings"
  1878 val factsN = "Relevant facts"
  1879 val class_relsN = "Class relationships"
  1880 val aritiesN = "Arities"
  1881 val helpersN = "Helper facts"
  1882 val conjsN = "Conjectures"
  1883 val free_typesN = "Type variables"
  1884 
  1885 val explicit_apply = NONE (* for experiments *)
  1886 
  1887 fun prepare_atp_problem ctxt format conj_sym_kind prem_kind type_enc sound
  1888         exporter trans_lambdas readable_names preproc hyp_ts concl_t facts =
  1889   let
  1890     val (format, type_enc) = choose_format [format] type_enc
  1891     val (fact_names, (conjs, facts, class_rel_clauses, arity_clauses)) =
  1892       translate_formulas ctxt format prem_kind type_enc trans_lambdas preproc
  1893                          hyp_ts concl_t facts
  1894     val sym_tab = conjs @ facts |> sym_table_for_facts ctxt explicit_apply
  1895     val nonmono_Ts =
  1896       conjs @ facts |> nonmonotonic_types_for_facts ctxt type_enc sound
  1897     val repair = repair_fact ctxt format type_enc sym_tab
  1898     val (conjs, facts) = (conjs, facts) |> pairself (map repair)
  1899     val repaired_sym_tab =
  1900       conjs @ facts |> sym_table_for_facts ctxt (SOME false)
  1901     val helpers =
  1902       repaired_sym_tab |> helper_facts_for_sym_table ctxt format type_enc
  1903                        |> map repair
  1904     val poly_nonmono_Ts =
  1905       if null nonmono_Ts orelse nonmono_Ts = [@{typ bool}] orelse
  1906          polymorphism_of_type_enc type_enc <> Polymorphic then
  1907         nonmono_Ts
  1908       else
  1909         [TVar (("'a", 0), HOLogic.typeS)]
  1910     val sym_decl_lines =
  1911       (conjs, helpers @ facts)
  1912       |> sym_decl_table_for_facts ctxt type_enc repaired_sym_tab
  1913       |> problem_lines_for_sym_decl_table ctxt format conj_sym_kind nonmono_Ts
  1914                                                poly_nonmono_Ts type_enc
  1915     val helper_lines =
  1916       0 upto length helpers - 1 ~~ helpers
  1917       |> map (formula_line_for_fact ctxt format helper_prefix I false true
  1918                                     poly_nonmono_Ts type_enc)
  1919       |> (if needs_type_tag_idempotence type_enc then
  1920             cons (type_tag_idempotence_fact ())
  1921           else
  1922             I)
  1923     (* Reordering these might confuse the proof reconstruction code or the SPASS
  1924        FLOTTER hack. *)
  1925     val problem =
  1926       [(explicit_declsN, sym_decl_lines),
  1927        (factsN,
  1928         map (formula_line_for_fact ctxt format fact_prefix ascii_of
  1929                                    (not exporter) (not exporter) nonmono_Ts
  1930                                    type_enc)
  1931             (0 upto length facts - 1 ~~ facts)),
  1932        (class_relsN, map formula_line_for_class_rel_clause class_rel_clauses),
  1933        (aritiesN, map formula_line_for_arity_clause arity_clauses),
  1934        (helpersN, helper_lines),
  1935        (conjsN,
  1936         map (formula_line_for_conjecture ctxt format nonmono_Ts type_enc)
  1937             conjs),
  1938        (free_typesN, formula_lines_for_free_types type_enc (facts @ conjs))]
  1939     val problem =
  1940       problem
  1941       |> (case format of
  1942             CNF => ensure_cnf_problem
  1943           | CNF_UEQ => filter_cnf_ueq_problem
  1944           | _ => I)
  1945       |> (if is_format_typed format then
  1946             declare_undeclared_syms_in_atp_problem type_decl_prefix
  1947                                                    implicit_declsN
  1948           else
  1949             I)
  1950     val (problem, pool) = problem |> nice_atp_problem readable_names
  1951     val helpers_offset = offset_of_heading_in_problem helpersN problem 0
  1952     val typed_helpers =
  1953       map_filter (fn (j, {name, ...}) =>
  1954                      if String.isSuffix typed_helper_suffix name then SOME j
  1955                      else NONE)
  1956                  ((helpers_offset + 1 upto helpers_offset + length helpers)
  1957                   ~~ helpers)
  1958     fun add_sym_arity (s, {min_ary, ...} : sym_info) =
  1959       if min_ary > 0 then
  1960         case strip_prefix_and_unascii const_prefix s of
  1961           SOME s => Symtab.insert (op =) (s, min_ary)
  1962         | NONE => I
  1963       else
  1964         I
  1965   in
  1966     (problem,
  1967      case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
  1968      offset_of_heading_in_problem conjsN problem 0,
  1969      offset_of_heading_in_problem factsN problem 0,
  1970      fact_names |> Vector.fromList,
  1971      typed_helpers,
  1972      Symtab.empty |> Symtab.fold add_sym_arity sym_tab)
  1973   end
  1974 
  1975 (* FUDGE *)
  1976 val conj_weight = 0.0
  1977 val hyp_weight = 0.1
  1978 val fact_min_weight = 0.2
  1979 val fact_max_weight = 1.0
  1980 val type_info_default_weight = 0.8
  1981 
  1982 fun add_term_weights weight (ATerm (s, tms)) =
  1983     is_tptp_user_symbol s ? Symtab.default (s, weight)
  1984     #> fold (add_term_weights weight) tms
  1985   | add_term_weights weight (AAbs (_, tm)) = add_term_weights weight tm
  1986 fun add_problem_line_weights weight (Formula (_, _, phi, _, _)) =
  1987     formula_fold NONE (K (add_term_weights weight)) phi
  1988   | add_problem_line_weights _ _ = I
  1989 
  1990 fun add_conjectures_weights [] = I
  1991   | add_conjectures_weights conjs =
  1992     let val (hyps, conj) = split_last conjs in
  1993       add_problem_line_weights conj_weight conj
  1994       #> fold (add_problem_line_weights hyp_weight) hyps
  1995     end
  1996 
  1997 fun add_facts_weights facts =
  1998   let
  1999     val num_facts = length facts
  2000     fun weight_of j =
  2001       fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j
  2002                         / Real.fromInt num_facts
  2003   in
  2004     map weight_of (0 upto num_facts - 1) ~~ facts
  2005     |> fold (uncurry add_problem_line_weights)
  2006   end
  2007 
  2008 (* Weights are from 0.0 (most important) to 1.0 (least important). *)
  2009 fun atp_problem_weights problem =
  2010   let val get = these o AList.lookup (op =) problem in
  2011     Symtab.empty
  2012     |> add_conjectures_weights (get free_typesN @ get conjsN)
  2013     |> add_facts_weights (get factsN)
  2014     |> fold (fold (add_problem_line_weights type_info_default_weight) o get)
  2015             [explicit_declsN, class_relsN, aritiesN]
  2016     |> Symtab.dest
  2017     |> sort (prod_ord Real.compare string_ord o pairself swap)
  2018   end
  2019 
  2020 end;