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