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