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