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