src/HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML
author blanchet
Tue, 17 May 2011 15:11:36 +0200
changeset 43702 c9b0968484fb
parent 43701 1068d8fc1331
child 43703 06afb3070af6
permissions -rw-r--r--
more work on "shallow" encoding + adjustments to other encodings
blanchet@40358
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML
blanchet@38506
     2
    Author:     Fabian Immler, TU Muenchen
blanchet@38506
     3
    Author:     Makarius
blanchet@38506
     4
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@38506
     5
blanchet@39734
     6
Translation of HOL to FOL for Sledgehammer.
blanchet@38506
     7
*)
blanchet@38506
     8
blanchet@40249
     9
signature SLEDGEHAMMER_ATP_TRANSLATE =
blanchet@38506
    10
sig
blanchet@43088
    11
  type 'a fo_term = 'a ATP_Problem.fo_term
blanchet@43580
    12
  type formula_kind = ATP_Problem.formula_kind
blanchet@38506
    13
  type 'a problem = 'a ATP_Problem.problem
blanchet@43511
    14
  type locality = Sledgehammer_Filter.locality
blanchet@43484
    15
blanchet@43484
    16
  datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic
blanchet@43484
    17
  datatype type_level =
blanchet@43484
    18
    All_Types | Nonmonotonic_Types | Finite_Types | Const_Arg_Types | No_Types
blanchet@43699
    19
  datatype type_depth = Deep | Shallow
blanchet@43484
    20
blanchet@43484
    21
  datatype type_system =
blanchet@43587
    22
    Simple_Types of type_level |
blanchet@43699
    23
    Preds of polymorphism * type_level * type_depth |
blanchet@43699
    24
    Tags of polymorphism * type_level * type_depth
blanchet@43484
    25
blanchet@40358
    26
  type translated_formula
blanchet@38506
    27
blanchet@43517
    28
  val readable_names : bool Config.T
blanchet@40445
    29
  val fact_prefix : string
blanchet@38506
    30
  val conjecture_prefix : string
blanchet@43439
    31
  val predicator_base : string
blanchet@43415
    32
  val explicit_app_base : string
blanchet@43420
    33
  val type_pred_base : string
blanchet@43433
    34
  val tff_type_prefix : string
blanchet@43484
    35
  val type_sys_from_string : string -> type_system
blanchet@43484
    36
  val polymorphism_of_type_sys : type_system -> polymorphism
blanchet@43484
    37
  val level_of_type_sys : type_system -> type_level
blanchet@43484
    38
  val is_type_sys_virtually_sound : type_system -> bool
blanchet@43484
    39
  val is_type_sys_fairly_sound : type_system -> bool
blanchet@43413
    40
  val unmangled_const : string -> string * string fo_term list
blanchet@41336
    41
  val translate_atp_fact :
blanchet@43511
    42
    Proof.context -> bool -> (string * locality) * thm
blanchet@43511
    43
    -> translated_formula option * ((string * locality) * thm)
blanchet@40240
    44
  val prepare_atp_problem :
blanchet@43580
    45
    Proof.context -> formula_kind -> formula_kind -> type_system -> bool
blanchet@43580
    46
    -> term list -> term
blanchet@41339
    47
    -> (translated_formula option * ((string * 'a) * thm)) list
blanchet@43412
    48
    -> string problem * string Symtab.table * int * int
blanchet@43620
    49
       * (string * 'a) list vector * int Symtab.table
blanchet@41561
    50
  val atp_problem_weights : string problem -> (string * real) list
blanchet@38506
    51
end;
blanchet@38506
    52
blanchet@41388
    53
structure Sledgehammer_ATP_Translate : SLEDGEHAMMER_ATP_TRANSLATE =
blanchet@38506
    54
struct
blanchet@38506
    55
blanchet@38506
    56
open ATP_Problem
blanchet@39734
    57
open Metis_Translate
blanchet@38506
    58
open Sledgehammer_Util
blanchet@43511
    59
open Sledgehammer_Filter
blanchet@43511
    60
blanchet@43511
    61
(* experimental *)
blanchet@43511
    62
val generate_useful_info = false
blanchet@38506
    63
blanchet@43439
    64
(* Readable names are often much shorter, especially if types are mangled in
blanchet@43460
    65
   names. Also, the logic for generating legal SNARK sort names is only
blanchet@43460
    66
   implemented for readable names. Finally, readable names are, well, more
blanchet@43460
    67
   readable. For these reason, they are enabled by default. *)
blanchet@43517
    68
val readable_names =
blanchet@43517
    69
  Attrib.setup_config_bool @{binding sledgehammer_atp_readable_names} (K true)
blanchet@43439
    70
blanchet@43414
    71
val type_decl_prefix = "type_"
blanchet@43414
    72
val sym_decl_prefix = "sym_"
blanchet@40445
    73
val fact_prefix = "fact_"
blanchet@38506
    74
val conjecture_prefix = "conj_"
blanchet@38506
    75
val helper_prefix = "help_"
blanchet@43414
    76
val class_rel_clause_prefix = "crel_";
blanchet@38506
    77
val arity_clause_prefix = "arity_"
blanchet@40156
    78
val tfree_prefix = "tfree_"
blanchet@38506
    79
blanchet@43439
    80
val predicator_base = "hBOOL"
blanchet@43415
    81
val explicit_app_base = "hAPP"
blanchet@43413
    82
val type_pred_base = "is"
blanchet@43433
    83
val tff_type_prefix = "ty_"
blanchet@43402
    84
blanchet@43433
    85
fun make_tff_type s = tff_type_prefix ^ ascii_of s
blanchet@43402
    86
blanchet@38506
    87
(* Freshness almost guaranteed! *)
blanchet@38506
    88
val sledgehammer_weak_prefix = "Sledgehammer:"
blanchet@38506
    89
blanchet@43484
    90
datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic
blanchet@43484
    91
datatype type_level =
blanchet@43484
    92
  All_Types | Nonmonotonic_Types | Finite_Types | Const_Arg_Types | No_Types
blanchet@43699
    93
datatype type_depth = Deep | Shallow
blanchet@43484
    94
blanchet@43484
    95
datatype type_system =
blanchet@43587
    96
  Simple_Types of type_level |
blanchet@43699
    97
  Preds of polymorphism * type_level * type_depth |
blanchet@43699
    98
  Tags of polymorphism * type_level * type_depth
blanchet@43484
    99
blanchet@43559
   100
fun try_unsuffixes ss s =
blanchet@43559
   101
  fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE
blanchet@43559
   102
blanchet@43484
   103
fun type_sys_from_string s =
blanchet@43587
   104
  (case try (unprefix "poly_") s of
blanchet@43587
   105
     SOME s => (SOME Polymorphic, s)
blanchet@43484
   106
   | NONE =>
blanchet@43484
   107
     case try (unprefix "mono_") s of
blanchet@43587
   108
       SOME s => (SOME Monomorphic, s)
blanchet@43587
   109
     | NONE =>
blanchet@43587
   110
       case try (unprefix "mangled_") s of
blanchet@43587
   111
         SOME s => (SOME Mangled_Monomorphic, s)
blanchet@43587
   112
       | NONE => (NONE, s))
blanchet@43484
   113
  ||> (fn s =>
blanchet@43559
   114
          (* "_query" and "_bang" are for the ASCII-challenged Mirabelle. *)
blanchet@43559
   115
          case try_unsuffixes ["?", "_query"] s of
blanchet@43484
   116
            SOME s => (Nonmonotonic_Types, s)
blanchet@43484
   117
          | NONE =>
blanchet@43559
   118
            case try_unsuffixes ["!", "_bang"] s of
blanchet@43484
   119
              SOME s => (Finite_Types, s)
blanchet@43484
   120
            | NONE => (All_Types, s))
blanchet@43699
   121
  ||> apsnd (fn s =>
blanchet@43699
   122
                case try (unsuffix "_shallow") s of
blanchet@43699
   123
                  SOME s => (Shallow, s)
blanchet@43699
   124
                | NONE => (Deep, s))
blanchet@43699
   125
  |> (fn (poly, (level, (depth, core))) =>
blanchet@43699
   126
         case (core, (poly, level, depth)) of
blanchet@43699
   127
           ("simple_types", (NONE, level, Deep (* naja *))) =>
blanchet@43699
   128
           Simple_Types level
blanchet@43699
   129
         | ("preds", (SOME Polymorphic, level, depth)) =>
blanchet@43699
   130
           if level = All_Types then Preds (Polymorphic, level, depth)
blanchet@43618
   131
           else raise Same.SAME
blanchet@43699
   132
         | ("preds", (SOME poly, level, depth)) => Preds (poly, level, depth)
blanchet@43699
   133
         | ("tags", (SOME Polymorphic, level, depth)) =>
blanchet@43618
   134
           if level = All_Types orelse level = Finite_Types then
blanchet@43699
   135
             Tags (Polymorphic, level, depth)
blanchet@43618
   136
           else
blanchet@43618
   137
             raise Same.SAME
blanchet@43699
   138
         | ("tags", (SOME poly, level, depth)) => Tags (poly, level, depth)
blanchet@43699
   139
         | ("args", (SOME poly, All_Types (* naja *), Deep (* naja *))) =>
blanchet@43699
   140
           Preds (poly, Const_Arg_Types, Shallow)
blanchet@43699
   141
         | ("erased", (NONE, All_Types (* naja *), Deep (* naja *))) =>
blanchet@43699
   142
           Preds (Polymorphic, No_Types, Shallow)
blanchet@43618
   143
         | _ => raise Same.SAME)
blanchet@43618
   144
  handle Same.SAME => error ("Unknown type system: " ^ quote s ^ ".")
blanchet@43484
   145
blanchet@43587
   146
fun polymorphism_of_type_sys (Simple_Types _) = Mangled_Monomorphic
blanchet@43699
   147
  | polymorphism_of_type_sys (Preds (poly, _, _)) = poly
blanchet@43699
   148
  | polymorphism_of_type_sys (Tags (poly, _, _)) = poly
blanchet@43484
   149
blanchet@43587
   150
fun level_of_type_sys (Simple_Types level) = level
blanchet@43699
   151
  | level_of_type_sys (Preds (_, level, _)) = level
blanchet@43699
   152
  | level_of_type_sys (Tags (_, level, _)) = level
blanchet@43699
   153
blanchet@43702
   154
fun depth_of_type_sys (Simple_Types _) = Deep
blanchet@43702
   155
  | depth_of_type_sys (Preds (_, _, depth)) = depth
blanchet@43702
   156
  | depth_of_type_sys (Tags (_, _, depth)) = depth
blanchet@43702
   157
blanchet@43557
   158
fun is_type_level_virtually_sound level =
blanchet@43557
   159
  level = All_Types orelse level = Nonmonotonic_Types
blanchet@43484
   160
val is_type_sys_virtually_sound =
blanchet@43484
   161
  is_type_level_virtually_sound o level_of_type_sys
blanchet@43484
   162
blanchet@43484
   163
fun is_type_level_fairly_sound level =
blanchet@43484
   164
  is_type_level_virtually_sound level orelse level = Finite_Types
blanchet@43484
   165
val is_type_sys_fairly_sound = is_type_level_fairly_sound o level_of_type_sys
blanchet@43484
   166
blanchet@43444
   167
fun formula_map f (AQuant (q, xs, phi)) = AQuant (q, xs, formula_map f phi)
blanchet@43444
   168
  | formula_map f (AConn (c, phis)) = AConn (c, map (formula_map f) phis)
blanchet@43444
   169
  | formula_map f (AAtom tm) = AAtom (f tm)
blanchet@43444
   170
blanchet@43550
   171
fun formula_fold pos f =
blanchet@43547
   172
  let
blanchet@43547
   173
    fun aux pos (AQuant (_, _, phi)) = aux pos phi
blanchet@43550
   174
      | aux pos (AConn (ANot, [phi])) = aux (Option.map not pos) phi
blanchet@43547
   175
      | aux pos (AConn (AImplies, [phi1, phi2])) =
blanchet@43550
   176
        aux (Option.map not pos) phi1 #> aux pos phi2
blanchet@43550
   177
      | aux pos (AConn (AAnd, phis)) = fold (aux pos) phis
blanchet@43550
   178
      | aux pos (AConn (AOr, phis)) = fold (aux pos) phis
blanchet@43550
   179
      | aux _ (AConn (_, phis)) = fold (aux NONE) phis
blanchet@43547
   180
      | aux pos (AAtom tm) = f pos tm
blanchet@43550
   181
  in aux (SOME pos) end
blanchet@43444
   182
blanchet@40358
   183
type translated_formula =
blanchet@38991
   184
  {name: string,
blanchet@43511
   185
   locality: locality,
blanchet@43396
   186
   kind: formula_kind,
blanchet@43433
   187
   combformula: (name, typ, combterm) formula,
blanchet@43433
   188
   atomic_types: typ list}
blanchet@38506
   189
blanchet@43511
   190
fun update_combformula f ({name, locality, kind, combformula, atomic_types}
blanchet@43511
   191
                          : translated_formula) =
blanchet@43511
   192
  {name = name, locality = locality, kind = kind, combformula = f combformula,
blanchet@43433
   193
   atomic_types = atomic_types} : translated_formula
blanchet@43413
   194
blanchet@43429
   195
fun fact_lift f ({combformula, ...} : translated_formula) = f combformula
blanchet@43429
   196
blanchet@43443
   197
val boring_consts = [explicit_app_base, @{const_name Metis.fequal}]
blanchet@43443
   198
blanchet@43443
   199
fun should_omit_type_args type_sys s =
blanchet@43460
   200
  s <> type_pred_base andalso s <> type_tag_name andalso
blanchet@43460
   201
  (s = @{const_name HOL.eq} orelse level_of_type_sys type_sys = No_Types orelse
blanchet@43460
   202
   (case type_sys of
blanchet@43702
   203
      Tags (_, _, Shallow) => false
blanchet@43702
   204
    | Tags (_, All_Types, Deep) => true
blanchet@43460
   205
    | _ => polymorphism_of_type_sys type_sys <> Mangled_Monomorphic andalso
blanchet@43460
   206
           member (op =) boring_consts s))
blanchet@43547
   207
blanchet@43618
   208
(* The Booleans indicate whether all type arguments should be kept. *)
blanchet@43618
   209
datatype type_arg_policy =
blanchet@43618
   210
  Explicit_Type_Args of bool |
blanchet@43618
   211
  Mangled_Type_Args of bool |
blanchet@43618
   212
  No_Type_Args
blanchet@41384
   213
blanchet@43702
   214
fun should_drop_arg_type_args type_sys =
blanchet@43702
   215
  level_of_type_sys type_sys = All_Types andalso
blanchet@43702
   216
  depth_of_type_sys type_sys = Deep
blanchet@43702
   217
blanchet@43460
   218
fun general_type_arg_policy type_sys =
blanchet@43460
   219
  if level_of_type_sys type_sys = No_Types then
blanchet@43460
   220
    No_Type_Args
blanchet@43460
   221
  else if polymorphism_of_type_sys type_sys = Mangled_Monomorphic then
blanchet@43702
   222
    Mangled_Type_Args (should_drop_arg_type_args type_sys)
blanchet@43460
   223
  else
blanchet@43702
   224
    Explicit_Type_Args (should_drop_arg_type_args type_sys)
blanchet@43434
   225
blanchet@43395
   226
fun type_arg_policy type_sys s =
blanchet@43443
   227
  if should_omit_type_args type_sys s then No_Type_Args
blanchet@43434
   228
  else general_type_arg_policy type_sys
blanchet@43088
   229
blanchet@43224
   230
fun atp_type_literals_for_types type_sys kind Ts =
blanchet@43460
   231
  if level_of_type_sys type_sys = No_Types then
blanchet@43224
   232
    []
blanchet@43224
   233
  else
blanchet@43224
   234
    Ts |> type_literals_for_types
blanchet@43224
   235
       |> filter (fn TyLitVar _ => kind <> Conjecture
blanchet@43224
   236
                   | TyLitFree _ => kind = Conjecture)
blanchet@41385
   237
blanchet@43580
   238
fun mk_anot phi = AConn (ANot, [phi])
blanchet@38506
   239
fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2])
blanchet@43405
   240
fun mk_aconns c phis =
blanchet@43405
   241
  let val (phis', phi') = split_last phis in
blanchet@43405
   242
    fold_rev (mk_aconn c) phis' phi'
blanchet@43405
   243
  end
blanchet@38506
   244
fun mk_ahorn [] phi = phi
blanchet@43405
   245
  | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])
blanchet@43393
   246
fun mk_aquant _ [] phi = phi
blanchet@43393
   247
  | mk_aquant q xs (phi as AQuant (q', xs', phi')) =
blanchet@43393
   248
    if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)
blanchet@43393
   249
  | mk_aquant q xs phi = AQuant (q, xs, phi)
blanchet@38506
   250
blanchet@43393
   251
fun close_universally atom_vars phi =
blanchet@41393
   252
  let
blanchet@41393
   253
    fun formula_vars bounds (AQuant (_, xs, phi)) =
blanchet@43397
   254
        formula_vars (map fst xs @ bounds) phi
blanchet@41393
   255
      | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis
blanchet@43393
   256
      | formula_vars bounds (AAtom tm) =
blanchet@43397
   257
        union (op =) (atom_vars tm []
blanchet@43397
   258
                      |> filter_out (member (op =) bounds o fst))
blanchet@43393
   259
  in mk_aquant AForall (formula_vars [] phi []) phi end
blanchet@43393
   260
blanchet@43402
   261
fun combterm_vars (CombApp (tm1, tm2)) = fold combterm_vars [tm1, tm2]
blanchet@43393
   262
  | combterm_vars (CombConst _) = I
blanchet@43445
   263
  | combterm_vars (CombVar (name, T)) = insert (op =) (name, SOME T)
blanchet@43545
   264
fun close_combformula_universally phi = close_universally combterm_vars phi
blanchet@43393
   265
blanchet@43393
   266
fun term_vars (ATerm (name as (s, _), tms)) =
blanchet@43402
   267
  is_atp_variable s ? insert (op =) (name, NONE)
blanchet@43397
   268
  #> fold term_vars tms
blanchet@43545
   269
fun close_formula_universally phi = close_universally term_vars phi
blanchet@41393
   270
blanchet@43433
   271
fun fo_term_from_typ (Type (s, Ts)) =
blanchet@43433
   272
    ATerm (`make_fixed_type_const s, map fo_term_from_typ Ts)
blanchet@43433
   273
  | fo_term_from_typ (TFree (s, _)) =
blanchet@43433
   274
    ATerm (`make_fixed_type_var s, [])
blanchet@43433
   275
  | fo_term_from_typ (TVar ((x as (s, _)), _)) =
blanchet@43433
   276
    ATerm ((make_schematic_type_var x, s), [])
blanchet@43433
   277
blanchet@43433
   278
(* This shouldn't clash with anything else. *)
blanchet@43413
   279
val mangled_type_sep = "\000"
blanchet@43413
   280
blanchet@43433
   281
fun generic_mangled_type_name f (ATerm (name, [])) = f name
blanchet@43433
   282
  | generic_mangled_type_name f (ATerm (name, tys)) =
blanchet@43626
   283
    f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys)
blanchet@43626
   284
    ^ ")"
blanchet@43433
   285
val mangled_type_name =
blanchet@43433
   286
  fo_term_from_typ
blanchet@43433
   287
  #> (fn ty => (make_tff_type (generic_mangled_type_name fst ty),
blanchet@43433
   288
                generic_mangled_type_name snd ty))
blanchet@43413
   289
blanchet@43445
   290
fun generic_mangled_type_suffix f g Ts =
blanchet@43413
   291
  fold_rev (curry (op ^) o g o prefix mangled_type_sep
blanchet@43445
   292
            o generic_mangled_type_name f) Ts ""
blanchet@43433
   293
fun mangled_const_name T_args (s, s') =
blanchet@43433
   294
  let val ty_args = map fo_term_from_typ T_args in
blanchet@43433
   295
    (s ^ generic_mangled_type_suffix fst ascii_of ty_args,
blanchet@43433
   296
     s' ^ generic_mangled_type_suffix snd I ty_args)
blanchet@43433
   297
  end
blanchet@43413
   298
blanchet@43413
   299
val parse_mangled_ident =
blanchet@43413
   300
  Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
blanchet@43413
   301
blanchet@43413
   302
fun parse_mangled_type x =
blanchet@43413
   303
  (parse_mangled_ident
blanchet@43413
   304
   -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
blanchet@43413
   305
                    [] >> ATerm) x
blanchet@43413
   306
and parse_mangled_types x =
blanchet@43413
   307
  (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x
blanchet@43413
   308
blanchet@43413
   309
fun unmangled_type s =
blanchet@43413
   310
  s |> suffix ")" |> raw_explode
blanchet@43413
   311
    |> Scan.finite Symbol.stopper
blanchet@43413
   312
           (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^
blanchet@43413
   313
                                                quote s)) parse_mangled_type))
blanchet@43413
   314
    |> fst
blanchet@43413
   315
blanchet@43432
   316
val unmangled_const_name = space_explode mangled_type_sep #> hd
blanchet@43413
   317
fun unmangled_const s =
blanchet@43413
   318
  let val ss = space_explode mangled_type_sep s in
blanchet@43413
   319
    (hd ss, map unmangled_type (tl ss))
blanchet@43413
   320
  end
blanchet@43413
   321
blanchet@43545
   322
fun introduce_proxies tm =
blanchet@43439
   323
  let
blanchet@43439
   324
    fun aux top_level (CombApp (tm1, tm2)) =
blanchet@43439
   325
        CombApp (aux top_level tm1, aux false tm2)
blanchet@43445
   326
      | aux top_level (CombConst (name as (s, s'), T, T_args)) =
blanchet@43441
   327
        (case proxify_const s of
blanchet@43439
   328
           SOME proxy_base =>
blanchet@43439
   329
           if top_level then
blanchet@43439
   330
             (case s of
blanchet@43439
   331
                "c_False" => (tptp_false, s')
blanchet@43439
   332
              | "c_True" => (tptp_true, s')
blanchet@43439
   333
              | _ => name, [])
blanchet@43440
   334
           else
blanchet@43445
   335
             (proxy_base |>> prefix const_prefix, T_args)
blanchet@43445
   336
          | NONE => (name, T_args))
blanchet@43445
   337
        |> (fn (name, T_args) => CombConst (name, T, T_args))
blanchet@43439
   338
      | aux _ tm = tm
blanchet@43545
   339
  in aux true tm end
blanchet@43439
   340
blanchet@43433
   341
fun combformula_from_prop thy eq_as_iff =
blanchet@38506
   342
  let
blanchet@43439
   343
    fun do_term bs t atomic_types =
blanchet@41388
   344
      combterm_from_term thy bs (Envir.eta_contract t)
blanchet@43439
   345
      |>> (introduce_proxies #> AAtom)
blanchet@43439
   346
      ||> union (op =) atomic_types
blanchet@38506
   347
    fun do_quant bs q s T t' =
blanchet@38743
   348
      let val s = Name.variant (map fst bs) s in
blanchet@38743
   349
        do_formula ((s, T) :: bs) t'
blanchet@43433
   350
        #>> mk_aquant q [(`make_bound_var s, SOME T)]
blanchet@38743
   351
      end
blanchet@38506
   352
    and do_conn bs c t1 t2 =
blanchet@38506
   353
      do_formula bs t1 ##>> do_formula bs t2
blanchet@43402
   354
      #>> uncurry (mk_aconn c)
blanchet@38506
   355
    and do_formula bs t =
blanchet@38506
   356
      case t of
blanchet@43402
   357
        @{const Not} $ t1 => do_formula bs t1 #>> mk_anot
blanchet@38506
   358
      | Const (@{const_name All}, _) $ Abs (s, T, t') =>
blanchet@38506
   359
        do_quant bs AForall s T t'
blanchet@38506
   360
      | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
blanchet@38506
   361
        do_quant bs AExists s T t'
haftmann@39028
   362
      | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2
haftmann@39028
   363
      | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2
haftmann@39019
   364
      | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2
haftmann@39093
   365
      | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
blanchet@41388
   366
        if eq_as_iff then do_conn bs AIff t1 t2 else do_term bs t
blanchet@41388
   367
      | _ => do_term bs t
blanchet@38506
   368
  in do_formula [] end
blanchet@38506
   369
blanchet@43615
   370
fun presimplify_term ctxt =
blanchet@43615
   371
  Skip_Proof.make_thm (Proof_Context.theory_of ctxt)
blanchet@43615
   372
  #> Meson.presimplify ctxt
blanchet@43615
   373
  #> prop_of
blanchet@38506
   374
wenzelm@41739
   375
fun concealed_bound_name j = sledgehammer_weak_prefix ^ string_of_int j
blanchet@38506
   376
fun conceal_bounds Ts t =
blanchet@38506
   377
  subst_bounds (map (Free o apfst concealed_bound_name)
blanchet@38506
   378
                    (0 upto length Ts - 1 ~~ Ts), t)
blanchet@38506
   379
fun reveal_bounds Ts =
blanchet@38506
   380
  subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
blanchet@38506
   381
                    (0 upto length Ts - 1 ~~ Ts))
blanchet@38506
   382
blanchet@43612
   383
fun extensionalize_term ctxt t =
blanchet@43612
   384
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43612
   385
    t |> cterm_of thy |> Meson.extensionalize_conv ctxt
blanchet@43612
   386
      |> prop_of |> Logic.dest_equals |> snd
blanchet@43612
   387
  end
blanchet@38831
   388
blanchet@38506
   389
fun introduce_combinators_in_term ctxt kind t =
wenzelm@43232
   390
  let val thy = Proof_Context.theory_of ctxt in
blanchet@38716
   391
    if Meson.is_fol_term thy t then
blanchet@38716
   392
      t
blanchet@38716
   393
    else
blanchet@38716
   394
      let
blanchet@38716
   395
        fun aux Ts t =
blanchet@38716
   396
          case t of
blanchet@38716
   397
            @{const Not} $ t1 => @{const Not} $ aux Ts t1
blanchet@38716
   398
          | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
blanchet@38716
   399
            t0 $ Abs (s, T, aux (T :: Ts) t')
blanchet@38890
   400
          | (t0 as Const (@{const_name All}, _)) $ t1 =>
blanchet@38890
   401
            aux Ts (t0 $ eta_expand Ts t1 1)
blanchet@38716
   402
          | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
blanchet@38716
   403
            t0 $ Abs (s, T, aux (T :: Ts) t')
blanchet@38890
   404
          | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
blanchet@38890
   405
            aux Ts (t0 $ eta_expand Ts t1 1)
haftmann@39028
   406
          | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39028
   407
          | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39019
   408
          | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39093
   409
          | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
blanchet@38716
   410
              $ t1 $ t2 =>
blanchet@38716
   411
            t0 $ aux Ts t1 $ aux Ts t2
blanchet@38716
   412
          | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
blanchet@38716
   413
                   t
blanchet@38716
   414
                 else
blanchet@38716
   415
                   t |> conceal_bounds Ts
blanchet@38716
   416
                     |> Envir.eta_contract
blanchet@38716
   417
                     |> cterm_of thy
blanchet@40071
   418
                     |> Meson_Clausify.introduce_combinators_in_cterm
blanchet@38716
   419
                     |> prop_of |> Logic.dest_equals |> snd
blanchet@38716
   420
                     |> reveal_bounds Ts
blanchet@39616
   421
        val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
blanchet@38716
   422
      in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
blanchet@38716
   423
      handle THM _ =>
blanchet@38716
   424
             (* A type variable of sort "{}" will make abstraction fail. *)
blanchet@38836
   425
             if kind = Conjecture then HOLogic.false_const
blanchet@38836
   426
             else HOLogic.true_const
blanchet@38716
   427
  end
blanchet@38506
   428
blanchet@38506
   429
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
blanchet@43224
   430
   same in Sledgehammer to prevent the discovery of unreplayable proofs. *)
blanchet@38506
   431
fun freeze_term t =
blanchet@38506
   432
  let
blanchet@38506
   433
    fun aux (t $ u) = aux t $ aux u
blanchet@38506
   434
      | aux (Abs (s, T, t)) = Abs (s, T, aux t)
blanchet@38506
   435
      | aux (Var ((s, i), T)) =
blanchet@38506
   436
        Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
blanchet@38506
   437
      | aux t = t
blanchet@38506
   438
  in t |> exists_subterm is_Var t ? aux end
blanchet@38506
   439
blanchet@40445
   440
(* making fact and conjecture formulas *)
blanchet@43511
   441
fun make_formula ctxt eq_as_iff presimp name loc kind t =
blanchet@38506
   442
  let
wenzelm@43232
   443
    val thy = Proof_Context.theory_of ctxt
blanchet@38831
   444
    val t = t |> Envir.beta_eta_contract
blanchet@38890
   445
              |> transform_elim_term
blanchet@41459
   446
              |> Object_Logic.atomize_term thy
blanchet@43434
   447
    val need_trueprop = (fastype_of t = @{typ bool})
blanchet@38890
   448
    val t = t |> need_trueprop ? HOLogic.mk_Trueprop
blanchet@43607
   449
              |> Raw_Simplifier.rewrite_term thy
blanchet@43607
   450
                     (Meson.unfold_set_const_simps ctxt) []
blanchet@43612
   451
              |> extensionalize_term ctxt
blanchet@43615
   452
              |> presimp ? presimplify_term ctxt
blanchet@38506
   453
              |> perhaps (try (HOLogic.dest_Trueprop))
blanchet@38506
   454
              |> introduce_combinators_in_term ctxt kind
blanchet@38836
   455
              |> kind <> Axiom ? freeze_term
blanchet@43433
   456
    val (combformula, atomic_types) =
blanchet@43433
   457
      combformula_from_prop thy eq_as_iff t []
blanchet@38506
   458
  in
blanchet@43511
   459
    {name = name, locality = loc, kind = kind, combformula = combformula,
blanchet@43433
   460
     atomic_types = atomic_types}
blanchet@38506
   461
  end
blanchet@38506
   462
blanchet@43511
   463
fun make_fact ctxt keep_trivial eq_as_iff presimp ((name, loc), t) =
blanchet@43511
   464
  case (keep_trivial, make_formula ctxt eq_as_iff presimp name loc Axiom t) of
blanchet@42861
   465
    (false, {combformula = AAtom (CombConst (("c_True", _), _, _)), ...}) =>
blanchet@42861
   466
    NONE
blanchet@42861
   467
  | (_, formula) => SOME formula
blanchet@43432
   468
blanchet@43580
   469
fun make_conjecture ctxt prem_kind ts =
blanchet@38836
   470
  let val last = length ts - 1 in
blanchet@43580
   471
    map2 (fn j => fn t =>
blanchet@43580
   472
             let
blanchet@43580
   473
               val (kind, maybe_negate) =
blanchet@43580
   474
                 if j = last then
blanchet@43580
   475
                   (Conjecture, I)
blanchet@43580
   476
                 else
blanchet@43580
   477
                   (prem_kind,
blanchet@43580
   478
                    if prem_kind = Conjecture then update_combformula mk_anot
blanchet@43580
   479
                    else I)
blanchet@43580
   480
              in
blanchet@43580
   481
                make_formula ctxt true true (string_of_int j) Chained kind t
blanchet@43580
   482
                |> maybe_negate
blanchet@43580
   483
              end)
blanchet@38836
   484
         (0 upto last) ts
blanchet@38836
   485
  end
blanchet@38506
   486
blanchet@43552
   487
(** Finite and infinite type inference **)
blanchet@43552
   488
blanchet@43552
   489
(* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are
blanchet@43552
   490
   dangerous because their "exhaust" properties can easily lead to unsound ATP
blanchet@43552
   491
   proofs. On the other hand, all HOL infinite types can be given the same
blanchet@43552
   492
   models in first-order logic (via Löwenheim-Skolem). *)
blanchet@43552
   493
blanchet@43552
   494
fun should_encode_type _ _ All_Types _ = true
blanchet@43552
   495
  | should_encode_type ctxt _ Finite_Types T = is_type_surely_finite ctxt T
blanchet@43552
   496
  | should_encode_type _ nonmono_Ts Nonmonotonic_Types T =
blanchet@43552
   497
    exists (curry Type.raw_instance T) nonmono_Ts
blanchet@43552
   498
  | should_encode_type _ _ _ _ = false
blanchet@43552
   499
blanchet@43700
   500
fun should_predicate_on_type ctxt nonmono_Ts (Preds (_, level, depth)) T =
blanchet@43700
   501
    (case depth of
blanchet@43700
   502
       Deep => should_encode_type ctxt nonmono_Ts level T
blanchet@43700
   503
     | Shallow => error "Not implemented yet.")
blanchet@43552
   504
  | should_predicate_on_type _ _ _ _ = false
blanchet@43552
   505
blanchet@43700
   506
datatype tag_site = Top_Level | Eq_Arg | Elsewhere
blanchet@43700
   507
blanchet@43700
   508
fun should_tag_with_type _ _ _ Top_Level _ _ = false
blanchet@43700
   509
  | should_tag_with_type ctxt nonmono_Ts (Tags (_, level, depth)) site u T =
blanchet@43700
   510
    (case depth of
blanchet@43700
   511
       Deep => should_encode_type ctxt nonmono_Ts level T
blanchet@43700
   512
     | Shallow =>
blanchet@43700
   513
       case (site, u) of
blanchet@43702
   514
         (Eq_Arg, CombVar _) =>
blanchet@43702
   515
         let
blanchet@43702
   516
           (* This trick ensures that "If" helpers are not unduely tagged, while
blanchet@43702
   517
              "True_or_False" is correctly tagged. *)
blanchet@43702
   518
           val level' = if null nonmono_Ts then level else Nonmonotonic_Types
blanchet@43702
   519
         in should_encode_type ctxt nonmono_Ts level' T end
blanchet@43700
   520
       | _ => false)
blanchet@43700
   521
  | should_tag_with_type _ _ _ _ _ _ = false
blanchet@43552
   522
blanchet@43552
   523
val homo_infinite_T = @{typ ind} (* any infinite type *)
blanchet@43552
   524
blanchet@43552
   525
fun homogenized_type ctxt nonmono_Ts level T =
blanchet@43552
   526
  if should_encode_type ctxt nonmono_Ts level T then T else homo_infinite_T
blanchet@43552
   527
blanchet@43444
   528
(** "hBOOL" and "hAPP" **)
blanchet@41561
   529
blanchet@43445
   530
type sym_info =
blanchet@43434
   531
  {pred_sym : bool, min_ary : int, max_ary : int, typ : typ option}
blanchet@43434
   532
blanchet@43445
   533
fun add_combterm_syms_to_table explicit_apply =
blanchet@43429
   534
  let
blanchet@43429
   535
    fun aux top_level tm =
blanchet@43429
   536
      let val (head, args) = strip_combterm_comb tm in
blanchet@43429
   537
        (case head of
blanchet@43434
   538
           CombConst ((s, _), T, _) =>
blanchet@43429
   539
           if String.isPrefix bound_var_prefix s then
blanchet@43429
   540
             I
blanchet@43429
   541
           else
blanchet@43434
   542
             let val ary = length args in
blanchet@43429
   543
               Symtab.map_default
blanchet@43429
   544
                   (s, {pred_sym = true,
blanchet@43434
   545
                        min_ary = if explicit_apply then 0 else ary,
blanchet@43434
   546
                        max_ary = 0, typ = SOME T})
blanchet@43434
   547
                   (fn {pred_sym, min_ary, max_ary, typ} =>
blanchet@43429
   548
                       {pred_sym = pred_sym andalso top_level,
blanchet@43434
   549
                        min_ary = Int.min (ary, min_ary),
blanchet@43434
   550
                        max_ary = Int.max (ary, max_ary),
blanchet@43434
   551
                        typ = if typ = SOME T then typ else NONE})
blanchet@43429
   552
            end
blanchet@43429
   553
         | _ => I)
blanchet@43429
   554
        #> fold (aux false) args
blanchet@43429
   555
      end
blanchet@43429
   556
  in aux true end
blanchet@43545
   557
fun add_fact_syms_to_table explicit_apply =
blanchet@43550
   558
  fact_lift (formula_fold true (K (add_combterm_syms_to_table explicit_apply)))
blanchet@38506
   559
blanchet@43546
   560
val default_sym_table_entries : (string * sym_info) list =
blanchet@43434
   561
  [("equal", {pred_sym = true, min_ary = 2, max_ary = 2, typ = NONE}),
blanchet@43439
   562
   (make_fixed_const predicator_base,
blanchet@43434
   563
    {pred_sym = true, min_ary = 1, max_ary = 1, typ = NONE})] @
blanchet@43439
   564
  ([tptp_false, tptp_true]
blanchet@43434
   565
   |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, typ = NONE}))
blanchet@41388
   566
blanchet@43415
   567
fun sym_table_for_facts explicit_apply facts =
blanchet@43439
   568
  Symtab.empty |> fold Symtab.default default_sym_table_entries
blanchet@43445
   569
               |> fold (add_fact_syms_to_table explicit_apply) facts
blanchet@38506
   570
blanchet@43429
   571
fun min_arity_of sym_tab s =
blanchet@43429
   572
  case Symtab.lookup sym_tab s of
blanchet@43445
   573
    SOME ({min_ary, ...} : sym_info) => min_ary
blanchet@43429
   574
  | NONE =>
blanchet@43429
   575
    case strip_prefix_and_unascii const_prefix s of
blanchet@43418
   576
      SOME s =>
blanchet@43441
   577
      let val s = s |> unmangled_const_name |> invert_const in
blanchet@43439
   578
        if s = predicator_base then 1
blanchet@43418
   579
        else if s = explicit_app_base then 2
blanchet@43418
   580
        else if s = type_pred_base then 1
blanchet@43428
   581
        else 0
blanchet@43418
   582
      end
blanchet@38506
   583
    | NONE => 0
blanchet@38506
   584
blanchet@38506
   585
(* True if the constant ever appears outside of the top-level position in
blanchet@38506
   586
   literals, or if it appears with different arities (e.g., because of different
blanchet@38506
   587
   type instantiations). If false, the constant always receives all of its
blanchet@38506
   588
   arguments and is used as a predicate. *)
blanchet@43429
   589
fun is_pred_sym sym_tab s =
blanchet@43429
   590
  case Symtab.lookup sym_tab s of
blanchet@43445
   591
    SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
blanchet@43445
   592
    pred_sym andalso min_ary = max_ary
blanchet@43429
   593
  | NONE => false
blanchet@38506
   594
blanchet@43439
   595
val predicator_combconst =
blanchet@43439
   596
  CombConst (`make_fixed_const predicator_base, @{typ "bool => bool"}, [])
blanchet@43439
   597
fun predicator tm = CombApp (predicator_combconst, tm)
blanchet@38506
   598
blanchet@43439
   599
fun introduce_predicators_in_combterm sym_tab tm =
blanchet@43413
   600
  case strip_combterm_comb tm of
blanchet@43413
   601
    (CombConst ((s, _), _, _), _) =>
blanchet@43439
   602
    if is_pred_sym sym_tab s then tm else predicator tm
blanchet@43439
   603
  | _ => predicator tm
blanchet@38506
   604
blanchet@43415
   605
fun list_app head args = fold (curry (CombApp o swap)) args head
blanchet@38506
   606
blanchet@43415
   607
fun explicit_app arg head =
blanchet@43415
   608
  let
blanchet@43433
   609
    val head_T = combtyp_of head
blanchet@43563
   610
    val (arg_T, res_T) = dest_funT head_T
blanchet@43415
   611
    val explicit_app =
blanchet@43433
   612
      CombConst (`make_fixed_const explicit_app_base, head_T --> head_T,
blanchet@43563
   613
                 [arg_T, res_T])
blanchet@43415
   614
  in list_app explicit_app [head, arg] end
blanchet@43415
   615
fun list_explicit_app head args = fold explicit_app args head
blanchet@43415
   616
blanchet@43436
   617
fun introduce_explicit_apps_in_combterm sym_tab =
blanchet@43415
   618
  let
blanchet@43415
   619
    fun aux tm =
blanchet@43415
   620
      case strip_combterm_comb tm of
blanchet@43415
   621
        (head as CombConst ((s, _), _, _), args) =>
blanchet@43415
   622
        args |> map aux
blanchet@43428
   623
             |> chop (min_arity_of sym_tab s)
blanchet@43415
   624
             |>> list_app head
blanchet@43415
   625
             |-> list_explicit_app
blanchet@43415
   626
      | (head, args) => list_explicit_app head (map aux args)
blanchet@43415
   627
  in aux end
blanchet@43415
   628
blanchet@43618
   629
fun chop_fun 0 T = ([], T)
blanchet@43618
   630
  | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
blanchet@43618
   631
    chop_fun (n - 1) ran_T |>> cons dom_T
blanchet@43618
   632
  | chop_fun _ _ = raise Fail "unexpected non-function"
blanchet@43618
   633
blanchet@43651
   634
fun filter_type_args _ _ _ [] = []
blanchet@43651
   635
  | filter_type_args thy s arity T_args =
blanchet@43652
   636
    let val U = s |> Sign.the_const_type thy (* may throw "TYPE" *) in
blanchet@43652
   637
      case Term.add_tvarsT (U |> chop_fun arity |> snd) [] of
blanchet@43652
   638
        [] => []
blanchet@43652
   639
      | res_U_vars =>
blanchet@43652
   640
        let val U_args = (s, U) |> Sign.const_typargs thy in
blanchet@43652
   641
          U_args ~~ T_args
blanchet@43652
   642
          |> map_filter (fn (U, T) =>
blanchet@43652
   643
                            if member (op =) res_U_vars (dest_TVar U) then
blanchet@43652
   644
                              SOME T
blanchet@43652
   645
                            else
blanchet@43652
   646
                              NONE)
blanchet@43652
   647
        end
blanchet@43651
   648
    end
blanchet@43651
   649
    handle TYPE _ => T_args
blanchet@43618
   650
blanchet@43618
   651
fun enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys =
blanchet@43618
   652
  let
blanchet@43618
   653
    val thy = Proof_Context.theory_of ctxt
blanchet@43618
   654
    fun aux arity (CombApp (tm1, tm2)) =
blanchet@43618
   655
        CombApp (aux (arity + 1) tm1, aux 0 tm2)
blanchet@43618
   656
      | aux arity (CombConst (name as (s, _), T, T_args)) =
blanchet@43571
   657
        let
blanchet@43571
   658
          val level = level_of_type_sys type_sys
blanchet@43571
   659
          val (T, T_args) =
blanchet@43571
   660
            (* Aggressively merge most "hAPPs" if the type system is unsound
blanchet@43571
   661
               anyway, by distinguishing overloads only on the homogenized
blanchet@43571
   662
               result type. *)
blanchet@43571
   663
            if s = const_prefix ^ explicit_app_base andalso
blanchet@43591
   664
               length T_args = 2 andalso
blanchet@43571
   665
               not (is_type_sys_virtually_sound type_sys) then
blanchet@43571
   666
              T_args |> map (homogenized_type ctxt nonmono_Ts level)
blanchet@43571
   667
                     |> (fn Ts => let val T = hd Ts --> nth Ts 1 in
blanchet@43702
   668
                                    (T --> T, tl Ts)
blanchet@43571
   669
                                  end)
blanchet@43571
   670
            else
blanchet@43571
   671
              (T, T_args)
blanchet@43571
   672
        in
blanchet@43571
   673
          (case strip_prefix_and_unascii const_prefix s of
blanchet@43571
   674
             NONE => (name, T_args)
blanchet@43571
   675
           | SOME s'' =>
blanchet@43618
   676
             let
blanchet@43618
   677
               val s'' = invert_const s''
blanchet@43702
   678
               fun filtered_T_args false = T_args
blanchet@43702
   679
                 | filtered_T_args true = filter_type_args thy s'' arity T_args
blanchet@43618
   680
             in
blanchet@43571
   681
               case type_arg_policy type_sys s'' of
blanchet@43702
   682
                 Explicit_Type_Args drop_args =>
blanchet@43702
   683
                 (name, filtered_T_args drop_args)
blanchet@43702
   684
               | Mangled_Type_Args drop_args =>
blanchet@43702
   685
                 (mangled_const_name (filtered_T_args drop_args) name, [])
blanchet@43618
   686
               | No_Type_Args => (name, [])
blanchet@43571
   687
             end)
blanchet@43571
   688
          |> (fn (name, T_args) => CombConst (name, T, T_args))
blanchet@43571
   689
        end
blanchet@43618
   690
      | aux _ tm = tm
blanchet@43618
   691
  in aux 0 end
blanchet@43444
   692
blanchet@43571
   693
fun repair_combterm ctxt nonmono_Ts type_sys sym_tab =
blanchet@43436
   694
  introduce_explicit_apps_in_combterm sym_tab
blanchet@43439
   695
  #> introduce_predicators_in_combterm sym_tab
blanchet@43618
   696
  #> enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys
blanchet@43571
   697
fun repair_fact ctxt nonmono_Ts type_sys sym_tab =
blanchet@43571
   698
  update_combformula (formula_map
blanchet@43571
   699
      (repair_combterm ctxt nonmono_Ts type_sys sym_tab))
blanchet@43444
   700
blanchet@43444
   701
(** Helper facts **)
blanchet@43444
   702
blanchet@43444
   703
fun ti_ti_helper_fact () =
blanchet@43444
   704
  let
blanchet@43444
   705
    fun var s = ATerm (`I s, [])
blanchet@43460
   706
    fun tag tm = ATerm (`make_fixed_const type_tag_name, [var "X", tm])
blanchet@43444
   707
  in
blanchet@43483
   708
    Formula (helper_prefix ^ "ti_ti", Axiom,
blanchet@43444
   709
             AAtom (ATerm (`I "equal", [tag (tag (var "Y")), tag (var "Y")]))
blanchet@43444
   710
             |> close_formula_universally, NONE, NONE)
blanchet@43444
   711
  end
blanchet@43444
   712
blanchet@43445
   713
fun helper_facts_for_sym ctxt type_sys (s, {typ, ...} : sym_info) =
blanchet@43444
   714
  case strip_prefix_and_unascii const_prefix s of
blanchet@43444
   715
    SOME mangled_s =>
blanchet@43444
   716
    let
blanchet@43444
   717
      val thy = Proof_Context.theory_of ctxt
blanchet@43444
   718
      val unmangled_s = mangled_s |> unmangled_const_name
blanchet@43450
   719
      fun dub_and_inst c needs_some_types (th, j) =
blanchet@43450
   720
        ((c ^ "_" ^ string_of_int j ^ (if needs_some_types then "T" else ""),
blanchet@43511
   721
          Chained),
blanchet@43444
   722
         let val t = th |> prop_of in
blanchet@43618
   723
           t |> ((case general_type_arg_policy type_sys of
blanchet@43618
   724
                    Mangled_Type_Args _ => true
blanchet@43618
   725
                  | _ => false) andalso
blanchet@43444
   726
                 not (null (Term.hidden_polymorphism t)))
blanchet@43444
   727
                ? (case typ of
blanchet@43444
   728
                     SOME T => specialize_type thy (invert_const unmangled_s, T)
blanchet@43444
   729
                   | NONE => I)
blanchet@43444
   730
         end)
blanchet@43444
   731
      fun make_facts eq_as_iff =
blanchet@43444
   732
        map_filter (make_fact ctxt false eq_as_iff false)
blanchet@43460
   733
      val has_some_types = is_type_sys_fairly_sound type_sys
blanchet@43444
   734
    in
blanchet@43444
   735
      metis_helpers
blanchet@43450
   736
      |> maps (fn (metis_s, (needs_some_types, ths)) =>
blanchet@43444
   737
                  if metis_s <> unmangled_s orelse
blanchet@43460
   738
                     (needs_some_types andalso not has_some_types) then
blanchet@43444
   739
                    []
blanchet@43444
   740
                  else
blanchet@43444
   741
                    ths ~~ (1 upto length ths)
blanchet@43450
   742
                    |> map (dub_and_inst mangled_s needs_some_types)
blanchet@43450
   743
                    |> make_facts (not needs_some_types))
blanchet@43444
   744
    end
blanchet@43444
   745
  | NONE => []
blanchet@43444
   746
fun helper_facts_for_sym_table ctxt type_sys sym_tab =
blanchet@43444
   747
  Symtab.fold_rev (append o helper_facts_for_sym ctxt type_sys) sym_tab []
blanchet@43444
   748
blanchet@43444
   749
fun translate_atp_fact ctxt keep_trivial =
blanchet@43444
   750
  `(make_fact ctxt keep_trivial true true o apsnd prop_of)
blanchet@43444
   751
blanchet@43580
   752
fun translate_formulas ctxt prem_kind type_sys hyp_ts concl_t rich_facts =
blanchet@43444
   753
  let
blanchet@43444
   754
    val thy = Proof_Context.theory_of ctxt
blanchet@43444
   755
    val fact_ts = map (prop_of o snd o snd) rich_facts
blanchet@43444
   756
    val (facts, fact_names) =
blanchet@43444
   757
      rich_facts
blanchet@43444
   758
      |> map_filter (fn (NONE, _) => NONE
blanchet@43444
   759
                      | (SOME fact, (name, _)) => SOME (fact, name))
blanchet@43444
   760
      |> ListPair.unzip
blanchet@43444
   761
    (* Remove existing facts from the conjecture, as this can dramatically
blanchet@43444
   762
       boost an ATP's performance (for some reason). *)
blanchet@43444
   763
    val hyp_ts = hyp_ts |> filter_out (member (op aconv) fact_ts)
blanchet@43444
   764
    val goal_t = Logic.list_implies (hyp_ts, concl_t)
blanchet@43444
   765
    val all_ts = goal_t :: fact_ts
blanchet@43444
   766
    val subs = tfree_classes_of_terms all_ts
blanchet@43444
   767
    val supers = tvar_classes_of_terms all_ts
blanchet@43444
   768
    val tycons = type_consts_of_terms thy all_ts
blanchet@43580
   769
    val conjs = make_conjecture ctxt prem_kind (hyp_ts @ [concl_t])
blanchet@43444
   770
    val (supers', arity_clauses) =
blanchet@43460
   771
      if level_of_type_sys type_sys = No_Types then ([], [])
blanchet@43444
   772
      else make_arity_clauses thy tycons supers
blanchet@43444
   773
    val class_rel_clauses = make_class_rel_clauses thy subs supers'
blanchet@43444
   774
  in
blanchet@43444
   775
    (fact_names |> map single, (conjs, facts, class_rel_clauses, arity_clauses))
blanchet@43444
   776
  end
blanchet@43444
   777
blanchet@43444
   778
fun fo_literal_from_type_literal (TyLitVar (class, name)) =
blanchet@43444
   779
    (true, ATerm (class, [ATerm (name, [])]))
blanchet@43444
   780
  | fo_literal_from_type_literal (TyLitFree (class, name)) =
blanchet@43444
   781
    (true, ATerm (class, [ATerm (name, [])]))
blanchet@43444
   782
blanchet@43444
   783
fun formula_from_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
blanchet@43444
   784
blanchet@43571
   785
fun type_pred_combatom ctxt nonmono_Ts type_sys T tm =
blanchet@43444
   786
  CombApp (CombConst (`make_fixed_const type_pred_base, T --> @{typ bool}, [T]),
blanchet@43444
   787
           tm)
blanchet@43618
   788
  |> enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys
blanchet@43444
   789
  |> AAtom
blanchet@43444
   790
blanchet@43700
   791
fun tag_with_type ctxt nonmono_Ts type_sys T tm =
blanchet@43700
   792
  CombConst (`make_fixed_const type_tag_name, T --> T, [T])
blanchet@43700
   793
  |> enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys
blanchet@43700
   794
  |> term_from_combterm ctxt nonmono_Ts type_sys Top_Level
blanchet@43700
   795
  |> (fn ATerm (s, tms) => ATerm (s, tms @ [tm]))
blanchet@43700
   796
and term_from_combterm ctxt nonmono_Ts type_sys site u =
blanchet@43444
   797
  let
blanchet@43700
   798
    val (head, args) = strip_combterm_comb u
blanchet@43700
   799
    val (x as (s, _), T_args) =
blanchet@43700
   800
      case head of
blanchet@43700
   801
        CombConst (name, _, T_args) => (name, T_args)
blanchet@43700
   802
      | CombVar (name, _) => (name, [])
blanchet@43700
   803
      | CombApp _ => raise Fail "impossible \"CombApp\""
blanchet@43700
   804
    val arg_site = if site = Top_Level andalso s = "equal" then Eq_Arg
blanchet@43700
   805
                   else Elsewhere
blanchet@43700
   806
    val t = ATerm (x, map fo_term_from_typ T_args @
blanchet@43700
   807
                      map (term_from_combterm ctxt nonmono_Ts type_sys arg_site)
blanchet@43700
   808
                          args)
blanchet@43700
   809
    val T = combtyp_of u
blanchet@43700
   810
  in
blanchet@43700
   811
    t |> (if should_tag_with_type ctxt nonmono_Ts type_sys site u T then
blanchet@43700
   812
            tag_with_type ctxt nonmono_Ts type_sys T
blanchet@43700
   813
          else
blanchet@43700
   814
            I)
blanchet@43700
   815
  end
blanchet@43700
   816
and formula_from_combformula ctxt nonmono_Ts type_sys =
blanchet@43700
   817
  let
blanchet@43444
   818
    val do_bound_type =
blanchet@43552
   819
      case type_sys of
blanchet@43587
   820
        Simple_Types level =>
blanchet@43552
   821
        SOME o mangled_type_name o homogenized_type ctxt nonmono_Ts level
blanchet@43552
   822
      | _ => K NONE
blanchet@43444
   823
    fun do_out_of_bound_type (s, T) =
blanchet@43550
   824
      if should_predicate_on_type ctxt nonmono_Ts type_sys T then
blanchet@43571
   825
        type_pred_combatom ctxt nonmono_Ts type_sys T (CombVar (s, T))
blanchet@43444
   826
        |> do_formula |> SOME
blanchet@43444
   827
      else
blanchet@43444
   828
        NONE
blanchet@43444
   829
    and do_formula (AQuant (q, xs, phi)) =
blanchet@43444
   830
        AQuant (q, xs |> map (apsnd (fn NONE => NONE
blanchet@43445
   831
                                      | SOME T => do_bound_type T)),
blanchet@43444
   832
                (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))
blanchet@43444
   833
                    (map_filter
blanchet@43444
   834
                         (fn (_, NONE) => NONE
blanchet@43445
   835
                           | (s, SOME T) => do_out_of_bound_type (s, T)) xs)
blanchet@43444
   836
                    (do_formula phi))
blanchet@43444
   837
      | do_formula (AConn (c, phis)) = AConn (c, map do_formula phis)
blanchet@43700
   838
      | do_formula (AAtom tm) =
blanchet@43700
   839
        AAtom (term_from_combterm ctxt nonmono_Ts type_sys Top_Level tm)
blanchet@43444
   840
  in do_formula end
blanchet@43444
   841
blanchet@43592
   842
fun bound_atomic_types type_sys Ts =
blanchet@43592
   843
  mk_ahorn (map (formula_from_fo_literal o fo_literal_from_type_literal)
blanchet@43592
   844
                (atp_type_literals_for_types type_sys Axiom Ts))
blanchet@43592
   845
blanchet@43550
   846
fun formula_for_fact ctxt nonmono_Ts type_sys
blanchet@43444
   847
                     ({combformula, atomic_types, ...} : translated_formula) =
blanchet@43592
   848
  combformula
blanchet@43592
   849
  |> close_combformula_universally
blanchet@43592
   850
  |> formula_from_combformula ctxt nonmono_Ts type_sys
blanchet@43592
   851
  |> bound_atomic_types type_sys atomic_types
blanchet@43444
   852
  |> close_formula_universally
blanchet@43444
   853
blanchet@43511
   854
fun useful_isabelle_info s = SOME (ATerm ("[]", [ATerm ("isabelle_" ^ s, [])]))
blanchet@43511
   855
blanchet@43444
   856
(* Each fact is given a unique fact number to avoid name clashes (e.g., because
blanchet@43444
   857
   of monomorphization). The TPTP explicitly forbids name clashes, and some of
blanchet@43444
   858
   the remote provers might care. *)
blanchet@43550
   859
fun formula_line_for_fact ctxt prefix nonmono_Ts type_sys
blanchet@43511
   860
                          (j, formula as {name, locality, kind, ...}) =
blanchet@43550
   861
  Formula (prefix ^ (if polymorphism_of_type_sys type_sys = Polymorphic then ""
blanchet@43550
   862
                     else string_of_int j ^ "_") ^
blanchet@43518
   863
           ascii_of name,
blanchet@43550
   864
           kind, formula_for_fact ctxt nonmono_Ts type_sys formula, NONE,
blanchet@43511
   865
           if generate_useful_info then
blanchet@43511
   866
             case locality of
blanchet@43511
   867
               Intro => useful_isabelle_info "intro"
blanchet@43511
   868
             | Elim => useful_isabelle_info "elim"
blanchet@43511
   869
             | Simp => useful_isabelle_info "simp"
blanchet@43511
   870
             | _ => NONE
blanchet@43511
   871
           else
blanchet@43511
   872
             NONE)
blanchet@43444
   873
blanchet@43444
   874
fun formula_line_for_class_rel_clause (ClassRelClause {name, subclass,
blanchet@43444
   875
                                                       superclass, ...}) =
blanchet@43444
   876
  let val ty_arg = ATerm (`I "T", []) in
blanchet@43448
   877
    Formula (class_rel_clause_prefix ^ ascii_of name, Axiom,
blanchet@43444
   878
             AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
blanchet@43444
   879
                               AAtom (ATerm (superclass, [ty_arg]))])
blanchet@43444
   880
             |> close_formula_universally, NONE, NONE)
blanchet@43444
   881
  end
blanchet@43444
   882
blanchet@43444
   883
fun fo_literal_from_arity_literal (TConsLit (c, t, args)) =
blanchet@43444
   884
    (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
blanchet@43444
   885
  | fo_literal_from_arity_literal (TVarLit (c, sort)) =
blanchet@43444
   886
    (false, ATerm (c, [ATerm (sort, [])]))
blanchet@43444
   887
blanchet@43444
   888
fun formula_line_for_arity_clause (ArityClause {name, conclLit, premLits,
blanchet@43444
   889
                                                ...}) =
blanchet@43448
   890
  Formula (arity_clause_prefix ^ ascii_of name, Axiom,
blanchet@43444
   891
           mk_ahorn (map (formula_from_fo_literal o apfst not
blanchet@43444
   892
                          o fo_literal_from_arity_literal) premLits)
blanchet@43444
   893
                    (formula_from_fo_literal
blanchet@43444
   894
                         (fo_literal_from_arity_literal conclLit))
blanchet@43444
   895
           |> close_formula_universally, NONE, NONE)
blanchet@43444
   896
blanchet@43550
   897
fun formula_line_for_conjecture ctxt nonmono_Ts type_sys
blanchet@43444
   898
        ({name, kind, combformula, ...} : translated_formula) =
blanchet@43448
   899
  Formula (conjecture_prefix ^ name, kind,
blanchet@43550
   900
           formula_from_combformula ctxt nonmono_Ts type_sys
blanchet@43444
   901
                                    (close_combformula_universally combformula)
blanchet@43444
   902
           |> close_formula_universally, NONE, NONE)
blanchet@43444
   903
blanchet@43444
   904
fun free_type_literals type_sys ({atomic_types, ...} : translated_formula) =
blanchet@43444
   905
  atomic_types |> atp_type_literals_for_types type_sys Conjecture
blanchet@43444
   906
               |> map fo_literal_from_type_literal
blanchet@43444
   907
blanchet@43444
   908
fun formula_line_for_free_type j lit =
blanchet@43448
   909
  Formula (tfree_prefix ^ string_of_int j, Hypothesis,
blanchet@43444
   910
           formula_from_fo_literal lit, NONE, NONE)
blanchet@43444
   911
fun formula_lines_for_free_types type_sys facts =
blanchet@43444
   912
  let
blanchet@43444
   913
    val litss = map (free_type_literals type_sys) facts
blanchet@43444
   914
    val lits = fold (union (op =)) litss []
blanchet@43444
   915
  in map2 formula_line_for_free_type (0 upto length lits - 1) lits end
blanchet@43444
   916
blanchet@43444
   917
(** Symbol declarations **)
blanchet@43415
   918
blanchet@43547
   919
fun insert_type get_T x xs =
blanchet@43547
   920
  let val T = get_T x in
blanchet@43547
   921
    if exists (curry Type.raw_instance T o get_T) xs then xs
blanchet@43547
   922
    else x :: filter_out ((fn T' => Type.raw_instance (T', T)) o get_T) xs
blanchet@43547
   923
  end
blanchet@43547
   924
blanchet@43445
   925
fun should_declare_sym type_sys pred_sym s =
blanchet@43413
   926
  not (String.isPrefix bound_var_prefix s) andalso s <> "equal" andalso
blanchet@43618
   927
  not (String.isPrefix tptp_special_prefix s) andalso
blanchet@43587
   928
  ((case type_sys of Simple_Types _ => true | _ => false) orelse not pred_sym)
blanchet@43413
   929
blanchet@43568
   930
fun sym_decl_table_for_facts type_sys repaired_sym_tab (conjs, facts) =
blanchet@43445
   931
  let
blanchet@43568
   932
    fun add_combterm in_conj tm =
blanchet@43445
   933
      let val (head, args) = strip_combterm_comb tm in
blanchet@43445
   934
        (case head of
blanchet@43445
   935
           CombConst ((s, s'), T, T_args) =>
blanchet@43445
   936
           let val pred_sym = is_pred_sym repaired_sym_tab s in
blanchet@43445
   937
             if should_declare_sym type_sys pred_sym s then
blanchet@43447
   938
               Symtab.map_default (s, [])
blanchet@43568
   939
                   (insert_type #3 (s', T_args, T, pred_sym, length args,
blanchet@43568
   940
                                    in_conj))
blanchet@43445
   941
             else
blanchet@43445
   942
               I
blanchet@43445
   943
           end
blanchet@43445
   944
         | _ => I)
blanchet@43568
   945
        #> fold (add_combterm in_conj) args
blanchet@43445
   946
      end
blanchet@43568
   947
    fun add_fact in_conj =
blanchet@43568
   948
      fact_lift (formula_fold true (K (add_combterm in_conj)))
blanchet@43568
   949
  in
blanchet@43568
   950
    Symtab.empty
blanchet@43568
   951
    |> is_type_sys_fairly_sound type_sys
blanchet@43568
   952
       ? (fold (add_fact true) conjs #> fold (add_fact false) facts)
blanchet@43568
   953
  end
blanchet@43445
   954
blanchet@43547
   955
fun is_var_or_bound_var (CombConst ((s, _), _, _)) =
blanchet@43547
   956
    String.isPrefix bound_var_prefix s
blanchet@43547
   957
  | is_var_or_bound_var (CombVar _) = true
blanchet@43547
   958
  | is_var_or_bound_var _ = false
blanchet@43547
   959
blanchet@43555
   960
(* This inference is described in section 2.3 of Claessen et al.'s "Sorting it
blanchet@43555
   961
   out with monotonicity" paper presented at CADE 2011. *)
blanchet@43700
   962
fun add_combterm_nonmonotonic_types _ _  (SOME false) _ = I
blanchet@43700
   963
  | add_combterm_nonmonotonic_types ctxt level _
blanchet@43550
   964
        (CombApp (CombApp (CombConst (("equal", _), Type (_, [T, _]), _), tm1),
blanchet@43550
   965
                  tm2)) =
blanchet@43550
   966
    (exists is_var_or_bound_var [tm1, tm2] andalso
blanchet@43700
   967
     (case level of
blanchet@43700
   968
        Nonmonotonic_Types => not (is_type_surely_infinite ctxt T)
blanchet@43700
   969
      | Finite_Types => is_type_surely_finite ctxt T
blanchet@43700
   970
      | _ => true)) ? insert_type I T
blanchet@43700
   971
  | add_combterm_nonmonotonic_types _ _ _ _ = I
blanchet@43700
   972
fun add_fact_nonmonotonic_types ctxt level ({kind, combformula, ...}
blanchet@43700
   973
                                            : translated_formula) =
blanchet@43700
   974
  formula_fold (kind <> Conjecture)
blanchet@43700
   975
               (add_combterm_nonmonotonic_types ctxt level) combformula
blanchet@43550
   976
fun add_nonmonotonic_types_for_facts ctxt type_sys facts =
blanchet@43700
   977
  let val level = level_of_type_sys type_sys in
blanchet@43700
   978
    (level = Nonmonotonic_Types orelse
blanchet@43700
   979
     (case type_sys of
blanchet@43700
   980
        Tags (poly, _, Shallow) => poly <> Polymorphic
blanchet@43700
   981
      | _ => false))
blanchet@43700
   982
    ? (fold (add_fact_nonmonotonic_types ctxt level) facts
blanchet@43700
   983
       (* in case helper "True_or_False" is included *)
blanchet@43700
   984
       #> insert_type I @{typ bool})
blanchet@43700
   985
  end
blanchet@43547
   986
blanchet@43618
   987
fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd
blanchet@43450
   988
blanchet@43619
   989
fun decl_line_for_sym ctxt nonmono_Ts level s (s', _, T, pred_sym, ary, _) =
blanchet@43619
   990
  let
blanchet@43619
   991
    val translate_type =
blanchet@43619
   992
      mangled_type_name o homogenized_type ctxt nonmono_Ts level
blanchet@43619
   993
    val (arg_tys, res_ty) =
blanchet@43619
   994
      T |> chop_fun ary |>> map translate_type ||> translate_type
blanchet@43619
   995
  in
blanchet@43619
   996
    Decl (sym_decl_prefix ^ s, (s, s'), arg_tys,
blanchet@43619
   997
          if pred_sym then `I tptp_tff_bool_type else res_ty)
blanchet@43450
   998
  end
blanchet@43450
   999
blanchet@43463
  1000
fun is_polymorphic_type T = fold_atyps (fn TVar _ => K true | _ => I) T false
blanchet@43463
  1001
blanchet@43700
  1002
fun formula_line_for_pred_sym_decl ctxt conj_sym_kind nonmono_Ts type_sys n s j
blanchet@43700
  1003
                                   (s', T_args, T, _, ary, in_conj) =
blanchet@43450
  1004
  let
blanchet@43580
  1005
    val (kind, maybe_negate) =
blanchet@43580
  1006
      if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
blanchet@43580
  1007
      else (Axiom, I)
blanchet@43618
  1008
    val (arg_Ts, res_T) = chop_fun ary T
blanchet@43450
  1009
    val bound_names =
blanchet@43450
  1010
      1 upto length arg_Ts |> map (`I o make_bound_var o string_of_int)
blanchet@43700
  1011
    val bounds =
blanchet@43450
  1012
      bound_names ~~ arg_Ts |> map (fn (name, T) => CombConst (name, T, []))
blanchet@43450
  1013
    val bound_Ts =
blanchet@43463
  1014
      arg_Ts |> map (fn T => if n > 1 orelse is_polymorphic_type T then SOME T
blanchet@43463
  1015
                             else NONE)
blanchet@43450
  1016
  in
blanchet@43483
  1017
    Formula (sym_decl_prefix ^ s ^
blanchet@43580
  1018
             (if n > 1 then "_" ^ string_of_int j else ""), kind,
blanchet@43450
  1019
             CombConst ((s, s'), T, T_args)
blanchet@43700
  1020
             |> fold (curry (CombApp o swap)) bounds
blanchet@43571
  1021
             |> type_pred_combatom ctxt nonmono_Ts type_sys res_T
blanchet@43450
  1022
             |> mk_aquant AForall (bound_names ~~ bound_Ts)
blanchet@43550
  1023
             |> formula_from_combformula ctxt nonmono_Ts type_sys
blanchet@43592
  1024
             |> n > 1 ? bound_atomic_types type_sys (atyps_of T)
blanchet@43580
  1025
             |> close_formula_universally
blanchet@43580
  1026
             |> maybe_negate,
blanchet@43450
  1027
             NONE, NONE)
blanchet@43450
  1028
  end
blanchet@43450
  1029
blanchet@43700
  1030
fun formula_lines_for_tag_sym_decl ctxt nonmono_Ts type_sys n s
blanchet@43700
  1031
                                   (j, (s', T_args, T, _, ary, _)) =
blanchet@43700
  1032
  let
blanchet@43700
  1033
    val ident_base =
blanchet@43700
  1034
      sym_decl_prefix ^ s ^ (if n > 1 then "_" ^ string_of_int j else "")
blanchet@43700
  1035
    val (arg_Ts, res_T) = chop_fun ary T
blanchet@43700
  1036
    val bound_names =
blanchet@43700
  1037
      1 upto length arg_Ts |> map (`I o make_bound_var o string_of_int)
blanchet@43700
  1038
    val bounds = bound_names |> map (fn name => ATerm (name, []))
blanchet@43700
  1039
    fun const args = ATerm ((s, s'), map fo_term_from_typ T_args @ args)
blanchet@43701
  1040
    val atomic_Ts = atyps_of T
blanchet@43701
  1041
    fun eq tm1 tm2 =
blanchet@43701
  1042
      ATerm (`I "equal", [tm1, tm2])
blanchet@43701
  1043
      |> AAtom
blanchet@43701
  1044
      |> bound_atomic_types type_sys atomic_Ts
blanchet@43701
  1045
      |> close_formula_universally
blanchet@43700
  1046
    val should_encode =
blanchet@43700
  1047
      should_encode_type ctxt nonmono_Ts
blanchet@43700
  1048
          (if polymorphism_of_type_sys type_sys = Polymorphic then All_Types
blanchet@43700
  1049
           else Nonmonotonic_Types)
blanchet@43700
  1050
    val tag_with = tag_with_type ctxt nonmono_Ts type_sys
blanchet@43700
  1051
    val add_formula_for_res =
blanchet@43700
  1052
      if should_encode res_T then
blanchet@43700
  1053
        cons (Formula (ident_base ^ "_res", Axiom,
blanchet@43701
  1054
                       eq (tag_with res_T (const bounds)) (const bounds),
blanchet@43700
  1055
                       NONE, NONE))
blanchet@43700
  1056
      else
blanchet@43700
  1057
        I
blanchet@43700
  1058
    fun add_formula_for_arg k =
blanchet@43700
  1059
      let val arg_T = nth arg_Ts k in
blanchet@43700
  1060
        if should_encode arg_T then
blanchet@43700
  1061
          case chop k bounds of
blanchet@43700
  1062
            (bounds1, bound :: bounds2) =>
blanchet@43700
  1063
            cons (Formula (ident_base ^ "_arg" ^ string_of_int (k + 1), Axiom,
blanchet@43701
  1064
                           eq (const (bounds1 @ tag_with arg_T bound ::
blanchet@43701
  1065
                                      bounds2))
blanchet@43701
  1066
                              (const bounds),
blanchet@43700
  1067
                           NONE, NONE))
blanchet@43700
  1068
          | _ => raise Fail "expected nonempty tail"
blanchet@43700
  1069
        else
blanchet@43700
  1070
          I
blanchet@43700
  1071
      end
blanchet@43700
  1072
  in
blanchet@43700
  1073
    [] |> add_formula_for_res
blanchet@43700
  1074
       |> fold add_formula_for_arg (ary - 1 downto 0)
blanchet@43700
  1075
  end
blanchet@43700
  1076
blanchet@43580
  1077
fun problem_lines_for_sym_decls ctxt conj_sym_kind nonmono_Ts type_sys
blanchet@43580
  1078
                                (s, decls) =
blanchet@43552
  1079
  case type_sys of
blanchet@43619
  1080
    Simple_Types level => map (decl_line_for_sym ctxt nonmono_Ts level s) decls
blanchet@43700
  1081
  | Preds _ =>
blanchet@43445
  1082
    let
blanchet@43450
  1083
      val decls =
blanchet@43450
  1084
        case decls of
blanchet@43450
  1085
          decl :: (decls' as _ :: _) =>
blanchet@43463
  1086
          let val T = result_type_of_decl decl in
blanchet@43463
  1087
            if forall ((fn T' => Type.raw_instance (T', T))
blanchet@43463
  1088
                       o result_type_of_decl) decls' then
blanchet@43463
  1089
              [decl]
blanchet@43463
  1090
            else
blanchet@43463
  1091
              decls
blanchet@43463
  1092
          end
blanchet@43450
  1093
        | _ => decls
blanchet@43450
  1094
      val n = length decls
blanchet@43450
  1095
      val decls =
blanchet@43550
  1096
        decls |> filter (should_predicate_on_type ctxt nonmono_Ts type_sys
blanchet@43450
  1097
                         o result_type_of_decl)
blanchet@43445
  1098
    in
blanchet@43580
  1099
      (0 upto length decls - 1, decls)
blanchet@43700
  1100
      |-> map2 (formula_line_for_pred_sym_decl ctxt conj_sym_kind nonmono_Ts
blanchet@43700
  1101
                                               type_sys n s)
blanchet@43445
  1102
    end
blanchet@43700
  1103
  | Tags (_, _, depth) =>
blanchet@43700
  1104
    (case depth of
blanchet@43700
  1105
       Deep => []
blanchet@43700
  1106
     | Shallow =>
blanchet@43700
  1107
       let val n = length decls in
blanchet@43700
  1108
         (0 upto n - 1 ~~ decls)
blanchet@43700
  1109
         |> maps (formula_lines_for_tag_sym_decl ctxt nonmono_Ts type_sys n s)
blanchet@43700
  1110
       end)
blanchet@43450
  1111
blanchet@43580
  1112
fun problem_lines_for_sym_decl_table ctxt conj_sym_kind nonmono_Ts type_sys
blanchet@43580
  1113
                                     sym_decl_tab =
blanchet@43580
  1114
  Symtab.fold_rev (append o problem_lines_for_sym_decls ctxt conj_sym_kind
blanchet@43580
  1115
                                                        nonmono_Ts type_sys)
blanchet@43445
  1116
                  sym_decl_tab []
blanchet@43410
  1117
blanchet@43414
  1118
fun add_tff_types_in_formula (AQuant (_, xs, phi)) =
blanchet@43414
  1119
    union (op =) (map_filter snd xs) #> add_tff_types_in_formula phi
blanchet@43414
  1120
  | add_tff_types_in_formula (AConn (_, phis)) =
blanchet@43414
  1121
    fold add_tff_types_in_formula phis
blanchet@43414
  1122
  | add_tff_types_in_formula (AAtom _) = I
blanchet@43414
  1123
blanchet@43433
  1124
fun add_tff_types_in_problem_line (Decl (_, _, arg_Ts, res_T)) =
blanchet@43433
  1125
    union (op =) (res_T :: arg_Ts)
blanchet@43448
  1126
  | add_tff_types_in_problem_line (Formula (_, _, phi, _, _)) =
blanchet@43414
  1127
    add_tff_types_in_formula phi
blanchet@43414
  1128
blanchet@43414
  1129
fun tff_types_in_problem problem =
blanchet@43414
  1130
  fold (fold add_tff_types_in_problem_line o snd) problem []
blanchet@43414
  1131
blanchet@43416
  1132
fun decl_line_for_tff_type (s, s') =
blanchet@43439
  1133
  Decl (type_decl_prefix ^ ascii_of s, (s, s'), [], `I tptp_tff_type_of_types)
blanchet@43414
  1134
blanchet@43702
  1135
fun should_add_ti_ti_helper (Tags (Polymorphic, level, Deep)) =
blanchet@43702
  1136
    level = Nonmonotonic_Types orelse level = Finite_Types
blanchet@43702
  1137
  | should_add_ti_ti_helper _ = false
blanchet@43702
  1138
blanchet@43414
  1139
val type_declsN = "Types"
blanchet@43415
  1140
val sym_declsN = "Symbol types"
blanchet@41405
  1141
val factsN = "Relevant facts"
blanchet@41405
  1142
val class_relsN = "Class relationships"
blanchet@43414
  1143
val aritiesN = "Arities"
blanchet@41405
  1144
val helpersN = "Helper facts"
blanchet@41405
  1145
val conjsN = "Conjectures"
blanchet@41561
  1146
val free_typesN = "Type variables"
blanchet@41405
  1147
blanchet@41405
  1148
fun offset_of_heading_in_problem _ [] j = j
blanchet@41405
  1149
  | offset_of_heading_in_problem needle ((heading, lines) :: problem) j =
blanchet@41405
  1150
    if heading = needle then j
blanchet@41405
  1151
    else offset_of_heading_in_problem needle problem (j + length lines)
blanchet@41405
  1152
blanchet@43580
  1153
fun prepare_atp_problem ctxt conj_sym_kind prem_kind type_sys explicit_apply
blanchet@43580
  1154
                        hyp_ts concl_t facts =
blanchet@38506
  1155
  let
blanchet@41561
  1156
    val (fact_names, (conjs, facts, class_rel_clauses, arity_clauses)) =
blanchet@43580
  1157
      translate_formulas ctxt prem_kind type_sys hyp_ts concl_t facts
blanchet@43434
  1158
    val sym_tab = conjs @ facts |> sym_table_for_facts explicit_apply
blanchet@43552
  1159
    val nonmono_Ts =
blanchet@43552
  1160
      [] |> fold (add_nonmonotonic_types_for_facts ctxt type_sys) [facts, conjs]
blanchet@43571
  1161
    val repair = repair_fact ctxt nonmono_Ts type_sys sym_tab
blanchet@43552
  1162
    val (conjs, facts) = (conjs, facts) |> pairself (map repair)
blanchet@43550
  1163
    val repaired_sym_tab = conjs @ facts |> sym_table_for_facts false
blanchet@43444
  1164
    val helpers =
blanchet@43552
  1165
      repaired_sym_tab |> helper_facts_for_sym_table ctxt type_sys |> map repair
blanchet@43550
  1166
    val sym_decl_lines =
blanchet@43596
  1167
      (conjs, helpers @ facts)
blanchet@43550
  1168
      |> sym_decl_table_for_facts type_sys repaired_sym_tab
blanchet@43580
  1169
      |> problem_lines_for_sym_decl_table ctxt conj_sym_kind nonmono_Ts type_sys
blanchet@43393
  1170
    (* Reordering these might confuse the proof reconstruction code or the SPASS
blanchet@43393
  1171
       Flotter hack. *)
blanchet@38506
  1172
    val problem =
blanchet@43432
  1173
      [(sym_declsN, sym_decl_lines),
blanchet@43550
  1174
       (factsN, map (formula_line_for_fact ctxt fact_prefix nonmono_Ts type_sys)
blanchet@43051
  1175
                    (0 upto length facts - 1 ~~ facts)),
blanchet@43416
  1176
       (class_relsN, map formula_line_for_class_rel_clause class_rel_clauses),
blanchet@43416
  1177
       (aritiesN, map formula_line_for_arity_clause arity_clauses),
blanchet@43550
  1178
       (helpersN, map (formula_line_for_fact ctxt helper_prefix nonmono_Ts
blanchet@43550
  1179
                                             type_sys)
blanchet@43434
  1180
                      (0 upto length helpers - 1 ~~ helpers)
blanchet@43702
  1181
                  |> should_add_ti_ti_helper type_sys
blanchet@43702
  1182
                     ? cons (ti_ti_helper_fact ())),
blanchet@43550
  1183
       (conjsN, map (formula_line_for_conjecture ctxt nonmono_Ts type_sys)
blanchet@43550
  1184
                    conjs),
blanchet@43416
  1185
       (free_typesN, formula_lines_for_free_types type_sys (facts @ conjs))]
blanchet@43414
  1186
    val problem =
blanchet@43432
  1187
      problem
blanchet@43552
  1188
      |> (case type_sys of
blanchet@43587
  1189
            Simple_Types _ =>
blanchet@43432
  1190
            cons (type_declsN,
blanchet@43432
  1191
                  map decl_line_for_tff_type (tff_types_in_problem problem))
blanchet@43552
  1192
          | _ => I)
blanchet@43517
  1193
    val (problem, pool) =
blanchet@43517
  1194
      problem |> nice_atp_problem (Config.get ctxt readable_names)
blanchet@43649
  1195
    fun add_sym_arity (s, {min_ary, ...} : sym_info) =
blanchet@43620
  1196
      if min_ary > 0 then
blanchet@43620
  1197
        case strip_prefix_and_unascii const_prefix s of
blanchet@43620
  1198
          SOME s => Symtab.insert (op =) (s, min_ary)
blanchet@43620
  1199
        | NONE => I
blanchet@43620
  1200
      else
blanchet@43620
  1201
        I
blanchet@38506
  1202
  in
blanchet@38506
  1203
    (problem,
blanchet@38506
  1204
     case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
blanchet@43456
  1205
     offset_of_heading_in_problem conjsN problem 0,
blanchet@43412
  1206
     offset_of_heading_in_problem factsN problem 0,
blanchet@43620
  1207
     fact_names |> Vector.fromList,
blanchet@43620
  1208
     Symtab.empty |> Symtab.fold add_sym_arity sym_tab)
blanchet@38506
  1209
  end
blanchet@38506
  1210
blanchet@41561
  1211
(* FUDGE *)
blanchet@41561
  1212
val conj_weight = 0.0
blanchet@42641
  1213
val hyp_weight = 0.1
blanchet@42641
  1214
val fact_min_weight = 0.2
blanchet@41561
  1215
val fact_max_weight = 1.0
blanchet@43479
  1216
val type_info_default_weight = 0.8
blanchet@41561
  1217
blanchet@41561
  1218
fun add_term_weights weight (ATerm (s, tms)) =
blanchet@43599
  1219
  (not (is_atp_variable s) andalso s <> "equal" andalso
blanchet@43618
  1220
   not (String.isPrefix tptp_special_prefix s)) ? Symtab.default (s, weight)
blanchet@41561
  1221
  #> fold (add_term_weights weight) tms
blanchet@43448
  1222
fun add_problem_line_weights weight (Formula (_, _, phi, _, _)) =
blanchet@43550
  1223
    formula_fold true (K (add_term_weights weight)) phi
blanchet@43399
  1224
  | add_problem_line_weights _ _ = I
blanchet@41561
  1225
blanchet@41561
  1226
fun add_conjectures_weights [] = I
blanchet@41561
  1227
  | add_conjectures_weights conjs =
blanchet@41561
  1228
    let val (hyps, conj) = split_last conjs in
blanchet@41561
  1229
      add_problem_line_weights conj_weight conj
blanchet@41561
  1230
      #> fold (add_problem_line_weights hyp_weight) hyps
blanchet@41561
  1231
    end
blanchet@41561
  1232
blanchet@41561
  1233
fun add_facts_weights facts =
blanchet@41561
  1234
  let
blanchet@41561
  1235
    val num_facts = length facts
blanchet@41561
  1236
    fun weight_of j =
blanchet@41561
  1237
      fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j
blanchet@41561
  1238
                        / Real.fromInt num_facts
blanchet@41561
  1239
  in
blanchet@41561
  1240
    map weight_of (0 upto num_facts - 1) ~~ facts
blanchet@41561
  1241
    |> fold (uncurry add_problem_line_weights)
blanchet@41561
  1242
  end
blanchet@41561
  1243
blanchet@41561
  1244
(* Weights are from 0.0 (most important) to 1.0 (least important). *)
blanchet@41561
  1245
fun atp_problem_weights problem =
blanchet@43479
  1246
  let val get = these o AList.lookup (op =) problem in
blanchet@43479
  1247
    Symtab.empty
blanchet@43479
  1248
    |> add_conjectures_weights (get free_typesN @ get conjsN)
blanchet@43479
  1249
    |> add_facts_weights (get factsN)
blanchet@43479
  1250
    |> fold (fold (add_problem_line_weights type_info_default_weight) o get)
blanchet@43479
  1251
            [sym_declsN, class_relsN, aritiesN]
blanchet@43479
  1252
    |> Symtab.dest
blanchet@43479
  1253
    |> sort (prod_ord Real.compare string_ord o pairself swap)
blanchet@43479
  1254
  end
blanchet@41561
  1255
blanchet@38506
  1256
end;