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