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