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