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