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