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