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