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