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