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