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