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