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