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