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