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