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