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