src/HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML
author blanchet
Wed, 15 Dec 2010 11:26:28 +0100
changeset 41388 9c68004b8c9d
parent 41386 eb80538166b6
child 41393 a5ee3b8e5a90
permissions -rw-r--r--
added Sledgehammer support for higher-order propositional reasoning
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@38506
    11
  type 'a problem = 'a ATP_Problem.problem
blanchet@40358
    12
  type translated_formula
blanchet@38506
    13
blanchet@41382
    14
  datatype type_system =
blanchet@41382
    15
    Tags of bool |
blanchet@41382
    16
    Preds of bool |
blanchet@41382
    17
    Const_Args |
blanchet@41382
    18
    No_Types
blanchet@41382
    19
blanchet@41386
    20
  val precise_overloaded_args : bool Unsynchronized.ref
blanchet@40445
    21
  val fact_prefix : string
blanchet@38506
    22
  val conjecture_prefix : string
blanchet@41386
    23
  val types_dangerous_types : type_system -> bool
blanchet@41384
    24
  val num_atp_type_args : theory -> type_system -> string -> int
blanchet@41336
    25
  val translate_atp_fact :
blanchet@39249
    26
    Proof.context -> (string * 'a) * thm
blanchet@41339
    27
    -> translated_formula option * ((string * 'a) * thm)
blanchet@40240
    28
  val prepare_atp_problem :
blanchet@41382
    29
    Proof.context -> bool -> bool -> type_system -> bool -> term list -> term
blanchet@41339
    30
    -> (translated_formula option * ((string * 'a) * thm)) list
blanchet@39053
    31
    -> string problem * string Symtab.table * int * (string * 'a) list vector
blanchet@38506
    32
end;
blanchet@38506
    33
blanchet@41388
    34
structure Sledgehammer_ATP_Translate : SLEDGEHAMMER_ATP_TRANSLATE =
blanchet@38506
    35
struct
blanchet@38506
    36
blanchet@38506
    37
open ATP_Problem
blanchet@39734
    38
open Metis_Translate
blanchet@38506
    39
open Sledgehammer_Util
blanchet@38506
    40
blanchet@41386
    41
(* FIXME: Remove references once appropriate defaults have been determined
blanchet@41386
    42
   empirically. *)
blanchet@41386
    43
val precise_overloaded_args = Unsynchronized.ref false
blanchet@41386
    44
blanchet@40445
    45
val fact_prefix = "fact_"
blanchet@38506
    46
val conjecture_prefix = "conj_"
blanchet@38506
    47
val helper_prefix = "help_"
blanchet@38506
    48
val class_rel_clause_prefix = "clrel_";
blanchet@38506
    49
val arity_clause_prefix = "arity_"
blanchet@40156
    50
val tfree_prefix = "tfree_"
blanchet@38506
    51
blanchet@38506
    52
(* Freshness almost guaranteed! *)
blanchet@38506
    53
val sledgehammer_weak_prefix = "Sledgehammer:"
blanchet@38506
    54
blanchet@40358
    55
type translated_formula =
blanchet@38991
    56
  {name: string,
blanchet@38991
    57
   kind: kind,
blanchet@38991
    58
   combformula: (name, combterm) formula,
blanchet@38991
    59
   ctypes_sorts: typ list}
blanchet@38506
    60
blanchet@41382
    61
datatype type_system =
blanchet@41382
    62
  Tags of bool |
blanchet@41382
    63
  Preds of bool |
blanchet@41382
    64
  Const_Args |
blanchet@41382
    65
  No_Types
blanchet@41382
    66
blanchet@41386
    67
fun types_dangerous_types (Tags _) = true
blanchet@41386
    68
  | types_dangerous_types (Preds _) = true
blanchet@41386
    69
  | types_dangerous_types _ = false
blanchet@41386
    70
blanchet@41384
    71
(* This is an approximation. If it returns "true" for a constant that isn't
blanchet@41384
    72
   overloaded (i.e., that has one uniform definition), needless clutter is
blanchet@41384
    73
   generated; if it returns "false" for an overloaded constant, the ATP gets a
blanchet@41384
    74
   license to do unsound reasoning if the type system is "overloaded_args". *)
blanchet@41384
    75
fun is_overloaded thy s =
blanchet@41386
    76
  not (!precise_overloaded_args) orelse
blanchet@41384
    77
  length (Defs.specifications_of (Theory.defs_of thy) s) > 1
blanchet@41384
    78
blanchet@41384
    79
fun needs_type_args thy type_sys s =
blanchet@41384
    80
  case type_sys of
blanchet@41386
    81
    Tags full_types => not full_types andalso is_overloaded thy s
blanchet@41388
    82
  | Preds _ => is_overloaded thy s (* FIXME: could be more precise *)
blanchet@41386
    83
  | Const_Args => is_overloaded thy s
blanchet@41384
    84
  | No_Types => false
blanchet@41384
    85
blanchet@41384
    86
fun num_atp_type_args thy type_sys s =
blanchet@41384
    87
  if needs_type_args thy type_sys s then num_type_args thy s else 0
blanchet@41384
    88
blanchet@41385
    89
fun atp_type_literals_for_types type_sys Ts =
blanchet@41385
    90
  if type_sys = No_Types then [] else type_literals_for_types Ts
blanchet@41385
    91
blanchet@38506
    92
fun mk_anot phi = AConn (ANot, [phi])
blanchet@38506
    93
fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2])
blanchet@38506
    94
fun mk_ahorn [] phi = phi
blanchet@38506
    95
  | mk_ahorn (phi :: phis) psi =
blanchet@38506
    96
    AConn (AImplies, [fold (mk_aconn AAnd) phis phi, psi])
blanchet@38506
    97
blanchet@41388
    98
fun combformula_for_prop thy eq_as_iff =
blanchet@38506
    99
  let
blanchet@41388
   100
    fun do_term bs t ts =
blanchet@41388
   101
      combterm_from_term thy bs (Envir.eta_contract t)
blanchet@41388
   102
      |>> AAtom ||> union (op =) ts
blanchet@38506
   103
    fun do_quant bs q s T t' =
blanchet@38743
   104
      let val s = Name.variant (map fst bs) s in
blanchet@38743
   105
        do_formula ((s, T) :: bs) t'
blanchet@38743
   106
        #>> (fn phi => AQuant (q, [`make_bound_var s], phi))
blanchet@38743
   107
      end
blanchet@38506
   108
    and do_conn bs c t1 t2 =
blanchet@38506
   109
      do_formula bs t1 ##>> do_formula bs t2
blanchet@38506
   110
      #>> (fn (phi1, phi2) => AConn (c, [phi1, phi2]))
blanchet@38506
   111
    and do_formula bs t =
blanchet@38506
   112
      case t of
blanchet@38506
   113
        @{const Not} $ t1 =>
blanchet@38506
   114
        do_formula bs t1 #>> (fn phi => AConn (ANot, [phi]))
blanchet@38506
   115
      | Const (@{const_name All}, _) $ Abs (s, T, t') =>
blanchet@38506
   116
        do_quant bs AForall s T t'
blanchet@38506
   117
      | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
blanchet@38506
   118
        do_quant bs AExists s T t'
haftmann@39028
   119
      | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2
haftmann@39028
   120
      | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2
haftmann@39019
   121
      | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2
haftmann@39093
   122
      | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
blanchet@41388
   123
        if eq_as_iff then do_conn bs AIff t1 t2 else do_term bs t
blanchet@41388
   124
      | _ => do_term bs t
blanchet@38506
   125
  in do_formula [] end
blanchet@38506
   126
blanchet@38841
   127
val presimplify_term = prop_of o Meson.presimplify oo Skip_Proof.make_thm
blanchet@38506
   128
blanchet@38506
   129
fun concealed_bound_name j = sledgehammer_weak_prefix ^ Int.toString j
blanchet@38506
   130
fun conceal_bounds Ts t =
blanchet@38506
   131
  subst_bounds (map (Free o apfst concealed_bound_name)
blanchet@38506
   132
                    (0 upto length Ts - 1 ~~ Ts), t)
blanchet@38506
   133
fun reveal_bounds Ts =
blanchet@38506
   134
  subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
blanchet@38506
   135
                    (0 upto length Ts - 1 ~~ Ts))
blanchet@38506
   136
blanchet@38831
   137
(* Removes the lambdas from an equation of the form "t = (%x. u)".
blanchet@40071
   138
   (Cf. "extensionalize_theorem" in "Meson_Clausify".) *)
blanchet@38831
   139
fun extensionalize_term t =
blanchet@38831
   140
  let
blanchet@38831
   141
    fun aux j (@{const Trueprop} $ t') = @{const Trueprop} $ aux j t'
blanchet@38831
   142
      | aux j (t as Const (s, Type (_, [Type (_, [_, T']),
blanchet@38831
   143
                                        Type (_, [_, res_T])]))
blanchet@38831
   144
                    $ t2 $ Abs (var_s, var_T, t')) =
haftmann@39093
   145
        if s = @{const_name HOL.eq} orelse s = @{const_name "=="} then
blanchet@38831
   146
          let val var_t = Var ((var_s, j), var_T) in
blanchet@38831
   147
            Const (s, T' --> T' --> res_T)
blanchet@38831
   148
              $ betapply (t2, var_t) $ subst_bound (var_t, t')
blanchet@38831
   149
            |> aux (j + 1)
blanchet@38831
   150
          end
blanchet@38831
   151
        else
blanchet@38831
   152
          t
blanchet@38831
   153
      | aux _ t = t
blanchet@38831
   154
  in aux (maxidx_of_term t + 1) t end
blanchet@38831
   155
blanchet@38506
   156
fun introduce_combinators_in_term ctxt kind t =
blanchet@38716
   157
  let val thy = ProofContext.theory_of ctxt in
blanchet@38716
   158
    if Meson.is_fol_term thy t then
blanchet@38716
   159
      t
blanchet@38716
   160
    else
blanchet@38716
   161
      let
blanchet@38716
   162
        fun aux Ts t =
blanchet@38716
   163
          case t of
blanchet@38716
   164
            @{const Not} $ t1 => @{const Not} $ aux Ts t1
blanchet@38716
   165
          | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
blanchet@38716
   166
            t0 $ Abs (s, T, aux (T :: Ts) t')
blanchet@38890
   167
          | (t0 as Const (@{const_name All}, _)) $ t1 =>
blanchet@38890
   168
            aux Ts (t0 $ eta_expand Ts t1 1)
blanchet@38716
   169
          | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
blanchet@38716
   170
            t0 $ Abs (s, T, aux (T :: Ts) t')
blanchet@38890
   171
          | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
blanchet@38890
   172
            aux Ts (t0 $ eta_expand Ts t1 1)
haftmann@39028
   173
          | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39028
   174
          | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39019
   175
          | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39093
   176
          | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
blanchet@38716
   177
              $ t1 $ t2 =>
blanchet@38716
   178
            t0 $ aux Ts t1 $ aux Ts t2
blanchet@38716
   179
          | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
blanchet@38716
   180
                   t
blanchet@38716
   181
                 else
blanchet@38716
   182
                   t |> conceal_bounds Ts
blanchet@38716
   183
                     |> Envir.eta_contract
blanchet@38716
   184
                     |> cterm_of thy
blanchet@40071
   185
                     |> Meson_Clausify.introduce_combinators_in_cterm
blanchet@38716
   186
                     |> prop_of |> Logic.dest_equals |> snd
blanchet@38716
   187
                     |> reveal_bounds Ts
blanchet@39616
   188
        val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
blanchet@38716
   189
      in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
blanchet@38716
   190
      handle THM _ =>
blanchet@38716
   191
             (* A type variable of sort "{}" will make abstraction fail. *)
blanchet@38836
   192
             if kind = Conjecture then HOLogic.false_const
blanchet@38836
   193
             else HOLogic.true_const
blanchet@38716
   194
  end
blanchet@38506
   195
blanchet@38506
   196
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
blanchet@38506
   197
   same in Sledgehammer to prevent the discovery of unreplable proofs. *)
blanchet@38506
   198
fun freeze_term t =
blanchet@38506
   199
  let
blanchet@38506
   200
    fun aux (t $ u) = aux t $ aux u
blanchet@38506
   201
      | aux (Abs (s, T, t)) = Abs (s, T, aux t)
blanchet@38506
   202
      | aux (Var ((s, i), T)) =
blanchet@38506
   203
        Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
blanchet@38506
   204
      | aux t = t
blanchet@38506
   205
  in t |> exists_subterm is_Var t ? aux end
blanchet@38506
   206
blanchet@38827
   207
(* "Object_Logic.atomize_term" isn't as powerful as it could be; for example,
blanchet@38827
   208
    it leaves metaequalities over "prop"s alone. *)
blanchet@38828
   209
val atomize_term =
blanchet@38828
   210
  let
blanchet@38828
   211
    fun aux (@{const Trueprop} $ t1) = t1
blanchet@38828
   212
      | aux (Const (@{const_name all}, _) $ Abs (s, T, t')) =
blanchet@38828
   213
        HOLogic.all_const T $ Abs (s, T, aux t')
blanchet@38828
   214
      | aux (@{const "==>"} $ t1 $ t2) = HOLogic.mk_imp (pairself aux (t1, t2))
blanchet@38828
   215
      | aux (Const (@{const_name "=="}, Type (_, [@{typ prop}, _])) $ t1 $ t2) =
blanchet@38828
   216
        HOLogic.eq_const HOLogic.boolT $ aux t1 $ aux t2
blanchet@38828
   217
      | aux (Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2) =
blanchet@38828
   218
        HOLogic.eq_const T $ t1 $ t2
blanchet@38828
   219
      | aux _ = raise Fail "aux"
blanchet@38828
   220
  in perhaps (try aux) end
blanchet@38827
   221
blanchet@40445
   222
(* making fact and conjecture formulas *)
blanchet@41388
   223
fun make_formula ctxt eq_as_iff presimp name kind t =
blanchet@38506
   224
  let
blanchet@38506
   225
    val thy = ProofContext.theory_of ctxt
blanchet@38831
   226
    val t = t |> Envir.beta_eta_contract
blanchet@38890
   227
              |> transform_elim_term
blanchet@38827
   228
              |> atomize_term
blanchet@38890
   229
    val need_trueprop = (fastype_of t = HOLogic.boolT)
blanchet@38890
   230
    val t = t |> need_trueprop ? HOLogic.mk_Trueprop
blanchet@38506
   231
              |> extensionalize_term
blanchet@38506
   232
              |> presimp ? presimplify_term thy
blanchet@38506
   233
              |> perhaps (try (HOLogic.dest_Trueprop))
blanchet@38506
   234
              |> introduce_combinators_in_term ctxt kind
blanchet@38836
   235
              |> kind <> Axiom ? freeze_term
blanchet@41388
   236
    val (combformula, ctypes_sorts) = combformula_for_prop thy eq_as_iff t []
blanchet@38506
   237
  in
blanchet@38991
   238
    {name = name, combformula = combformula, kind = kind,
blanchet@38991
   239
     ctypes_sorts = ctypes_sorts}
blanchet@38506
   240
  end
blanchet@38506
   241
blanchet@41388
   242
fun make_fact ctxt eq_as_iff presimp ((name, _), th) =
blanchet@41388
   243
  case make_formula ctxt eq_as_iff presimp name Axiom (prop_of th) of
blanchet@38991
   244
    {combformula = AAtom (CombConst (("c_True", _), _, _)), ...} => NONE
blanchet@41339
   245
  | formula => SOME formula
blanchet@38836
   246
fun make_conjecture ctxt ts =
blanchet@38836
   247
  let val last = length ts - 1 in
blanchet@41388
   248
    map2 (fn j => make_formula ctxt true true (Int.toString j)
blanchet@38836
   249
                               (if j = last then Conjecture else Hypothesis))
blanchet@38836
   250
         (0 upto last) ts
blanchet@38836
   251
  end
blanchet@38506
   252
blanchet@38506
   253
(** Helper facts **)
blanchet@38506
   254
blanchet@41388
   255
fun count_term (ATerm ((s, _), tms)) =
blanchet@41388
   256
  (if is_atp_variable s then I
blanchet@41388
   257
   else Symtab.map_entry s (Integer.add 1))
blanchet@41388
   258
  #> fold count_term tms
blanchet@41388
   259
fun count_formula (AQuant (_, _, phi)) = count_formula phi
blanchet@41388
   260
  | count_formula (AConn (_, phis)) = fold count_formula phis
blanchet@41388
   261
  | count_formula (AAtom tm) = count_term tm
blanchet@38506
   262
blanchet@38506
   263
val init_counters =
blanchet@41388
   264
  metis_helpers |> map fst |> sort_distinct string_ord |> map (rpair 0)
blanchet@41388
   265
  |> Symtab.make
blanchet@38506
   266
blanchet@41388
   267
fun get_helper_facts ctxt type_sys formulas =
blanchet@38506
   268
  let
blanchet@41388
   269
    val no_dangerous_types = types_dangerous_types type_sys
blanchet@41388
   270
    val ct = init_counters |> fold count_formula formulas
blanchet@41388
   271
    fun is_used s = the (Symtab.lookup ct s) > 0
blanchet@41388
   272
    fun dub c needs_full_types (th, j) =
blanchet@41388
   273
      ((c ^ "_" ^ string_of_int j ^ (if needs_full_types then "ft" else ""),
blanchet@41388
   274
        false), th)
blanchet@41388
   275
    fun make_facts eq_as_iff = map_filter (make_fact ctxt eq_as_iff false)
blanchet@38506
   276
  in
blanchet@41388
   277
    metis_helpers
blanchet@41388
   278
    |> filter (is_used o fst)
blanchet@41388
   279
    |> maps (fn (c, (needs_full_types, ths)) =>
blanchet@41388
   280
                if needs_full_types andalso not no_dangerous_types then
blanchet@41388
   281
                  []
blanchet@41388
   282
                else
blanchet@41388
   283
                  ths ~~ (1 upto length ths)
blanchet@41388
   284
                  |> map (dub c needs_full_types)
blanchet@41388
   285
                  |> make_facts (not needs_full_types))
blanchet@38506
   286
  end
blanchet@38506
   287
blanchet@41388
   288
fun translate_atp_fact ctxt = `(make_fact ctxt true true)
blanchet@39248
   289
blanchet@41382
   290
fun translate_formulas ctxt type_sys hyp_ts concl_t rich_facts =
blanchet@38506
   291
  let
blanchet@38506
   292
    val thy = ProofContext.theory_of ctxt
blanchet@41339
   293
    val fact_ts = map (prop_of o snd o snd) rich_facts
blanchet@41339
   294
    val (facts, fact_names) =
blanchet@41339
   295
      rich_facts
blanchet@41339
   296
      |> map_filter (fn (NONE, _) => NONE
blanchet@41339
   297
                      | (SOME fact, (name, _)) => SOME (fact, name))
blanchet@41339
   298
      |> ListPair.unzip
blanchet@40445
   299
    (* Remove existing facts from the conjecture, as this can dramatically
blanchet@39249
   300
       boost an ATP's performance (for some reason). *)
blanchet@40445
   301
    val hyp_ts = hyp_ts |> filter_out (member (op aconv) fact_ts)
blanchet@38506
   302
    val goal_t = Logic.list_implies (hyp_ts, concl_t)
blanchet@38506
   303
    val subs = tfree_classes_of_terms [goal_t]
blanchet@40445
   304
    val supers = tvar_classes_of_terms fact_ts
blanchet@40445
   305
    val tycons = type_consts_of_terms thy (goal_t :: fact_ts)
blanchet@40445
   306
    (* TFrees in the conjecture; TVars in the facts *)
blanchet@38836
   307
    val conjectures = make_conjecture ctxt (hyp_ts @ [concl_t])
blanchet@41385
   308
    val (supers', arity_clauses) =
blanchet@41385
   309
      if type_sys = No_Types then ([], [])
blanchet@41385
   310
      else make_arity_clauses thy tycons supers
blanchet@38506
   311
    val class_rel_clauses = make_class_rel_clauses thy subs supers'
blanchet@38506
   312
  in
blanchet@40445
   313
    (fact_names |> map single |> Vector.fromList,
blanchet@41388
   314
     (conjectures, facts, class_rel_clauses, arity_clauses))
blanchet@38506
   315
  end
blanchet@38506
   316
blanchet@41386
   317
fun tag_with_type ty t = ATerm (`I type_tag_name, [ty, t])
blanchet@38506
   318
blanchet@38506
   319
fun fo_term_for_combtyp (CombTVar name) = ATerm (name, [])
blanchet@38506
   320
  | fo_term_for_combtyp (CombTFree name) = ATerm (name, [])
blanchet@38506
   321
  | fo_term_for_combtyp (CombType (name, tys)) =
blanchet@38506
   322
    ATerm (name, map fo_term_for_combtyp tys)
blanchet@38506
   323
blanchet@38506
   324
fun fo_literal_for_type_literal (TyLitVar (class, name)) =
blanchet@38506
   325
    (true, ATerm (class, [ATerm (name, [])]))
blanchet@38506
   326
  | fo_literal_for_type_literal (TyLitFree (class, name)) =
blanchet@38506
   327
    (true, ATerm (class, [ATerm (name, [])]))
blanchet@38506
   328
blanchet@38506
   329
fun formula_for_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
blanchet@38506
   330
blanchet@41386
   331
(* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are
blanchet@41386
   332
   considered dangerous because their "exhaust" properties can easily lead to
blanchet@41386
   333
   unsound ATP proofs. The checks below are an (unsound) approximation of
blanchet@41386
   334
   finiteness. *)
blanchet@41386
   335
blanchet@41386
   336
fun is_dtyp_dangerous _ (Datatype_Aux.DtTFree _) = true
blanchet@41386
   337
  | is_dtyp_dangerous ctxt (Datatype_Aux.DtType (s, Us)) =
blanchet@41386
   338
    is_type_constr_dangerous ctxt s andalso forall (is_dtyp_dangerous ctxt) Us
blanchet@41386
   339
  | is_dtyp_dangerous _ (Datatype_Aux.DtRec _) = false
blanchet@41386
   340
and is_type_dangerous ctxt (Type (s, Ts)) =
blanchet@41386
   341
    is_type_constr_dangerous ctxt s andalso forall (is_type_dangerous ctxt) Ts
blanchet@41388
   342
  | is_type_dangerous _ _ = false
blanchet@41386
   343
and is_type_constr_dangerous ctxt s =
blanchet@41386
   344
  let val thy = ProofContext.theory_of ctxt in
blanchet@41386
   345
    case Datatype_Data.get_info thy s of
blanchet@41386
   346
      SOME {descr, ...} =>
blanchet@41386
   347
      forall (fn (_, (_, _, constrs)) =>
blanchet@41386
   348
                 forall (forall (is_dtyp_dangerous ctxt) o snd) constrs) descr
blanchet@41386
   349
    | NONE =>
blanchet@41386
   350
      case Typedef.get_info ctxt s of
blanchet@41386
   351
        ({rep_type, ...}, _) :: _ => is_type_dangerous ctxt rep_type
blanchet@41386
   352
      | [] => true
blanchet@41386
   353
  end
blanchet@41386
   354
blanchet@41386
   355
fun is_combtyp_dangerous ctxt (CombType ((s, _), tys)) =
blanchet@41386
   356
    (case strip_prefix_and_unascii type_const_prefix s of
blanchet@41386
   357
       SOME s' => forall (is_combtyp_dangerous ctxt) tys andalso
blanchet@41386
   358
                  is_type_constr_dangerous ctxt (invert_const s')
blanchet@41386
   359
     | NONE => false)
blanchet@41386
   360
  | is_combtyp_dangerous _ _ = false
blanchet@41386
   361
blanchet@41386
   362
fun should_tag_with_type ctxt (Tags full_types) ty =
blanchet@41386
   363
    full_types orelse is_combtyp_dangerous ctxt ty
blanchet@41386
   364
  | should_tag_with_type _ _ _ = false
blanchet@41386
   365
blanchet@41388
   366
val fname_table =
blanchet@41388
   367
  [("c_False", (0, ("c_fFalse", @{const_name Metis.fFalse}))),
blanchet@41388
   368
   ("c_True", (0, ("c_fTrue", @{const_name Metis.fTrue}))),
blanchet@41388
   369
   ("c_Not", (1, ("c_fNot", @{const_name Metis.fNot}))),
blanchet@41388
   370
   ("c_conj", (2, ("c_fconj", @{const_name Metis.fconj}))),
blanchet@41388
   371
   ("c_disj", (2, ("c_fdisj", @{const_name Metis.fdisj}))),
blanchet@41388
   372
   ("c_implies", (2, ("c_fimplies", @{const_name Metis.fimplies}))),
blanchet@41388
   373
   ("equal", (2, ("c_fequal", @{const_name Metis.fequal})))]
blanchet@41388
   374
blanchet@41386
   375
fun fo_term_for_combterm ctxt type_sys =
blanchet@38506
   376
  let
blanchet@41386
   377
    val thy = ProofContext.theory_of ctxt
blanchet@38506
   378
    fun aux top_level u =
blanchet@38506
   379
      let
blanchet@38506
   380
        val (head, args) = strip_combterm_comb u
blanchet@38506
   381
        val (x, ty_args) =
blanchet@38506
   382
          case head of
blanchet@38506
   383
            CombConst (name as (s, s'), _, ty_args) =>
blanchet@41388
   384
            (case AList.lookup (op =) fname_table s of
blanchet@41388
   385
               SOME (n, fname) =>
blanchet@41388
   386
               if top_level andalso length args = n then (name, [])
blanchet@41388
   387
               else (fname, ty_args)
blanchet@41388
   388
             | NONE =>
blanchet@41388
   389
               case strip_prefix_and_unascii const_prefix s of
blanchet@41388
   390
                 NONE => (name, ty_args)
blanchet@41388
   391
               | SOME s'' =>
blanchet@41388
   392
                 let
blanchet@41388
   393
                   val s'' = invert_const s''
blanchet@41388
   394
                   val ty_args =
blanchet@41388
   395
                     if needs_type_args thy type_sys s'' then ty_args else []
blanchet@41388
   396
                  in
blanchet@41388
   397
                    if top_level then
blanchet@41388
   398
                      case s of
blanchet@41388
   399
                        "c_False" => (("$false", s'), [])
blanchet@41388
   400
                      | "c_True" => (("$true", s'), [])
blanchet@41388
   401
                      | _ => (name, ty_args)
blanchet@41388
   402
                    else
blanchet@41388
   403
                      (name, ty_args)
blanchet@41388
   404
                  end)
blanchet@38506
   405
          | CombVar (name, _) => (name, [])
blanchet@38506
   406
          | CombApp _ => raise Fail "impossible \"CombApp\""
blanchet@41386
   407
        val t =
blanchet@41386
   408
          ATerm (x, map fo_term_for_combtyp ty_args @ map (aux false) args)
blanchet@41386
   409
        val ty = combtyp_of u
blanchet@38506
   410
    in
blanchet@41386
   411
      t |> (if should_tag_with_type ctxt type_sys ty then
blanchet@41386
   412
              tag_with_type (fo_term_for_combtyp ty)
blanchet@41382
   413
            else
blanchet@41382
   414
              I)
blanchet@38506
   415
    end
blanchet@38506
   416
  in aux true end
blanchet@38506
   417
blanchet@41386
   418
fun formula_for_combformula ctxt type_sys =
blanchet@38506
   419
  let
blanchet@38506
   420
    fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi)
blanchet@38506
   421
      | aux (AConn (c, phis)) = AConn (c, map aux phis)
blanchet@41386
   422
      | aux (AAtom tm) = AAtom (fo_term_for_combterm ctxt type_sys tm)
blanchet@38506
   423
  in aux end
blanchet@38506
   424
blanchet@41386
   425
fun formula_for_fact ctxt type_sys
blanchet@40445
   426
                     ({combformula, ctypes_sorts, ...} : translated_formula) =
blanchet@38506
   427
  mk_ahorn (map (formula_for_fo_literal o fo_literal_for_type_literal)
blanchet@41385
   428
                (atp_type_literals_for_types type_sys ctypes_sorts))
blanchet@41386
   429
           (formula_for_combformula ctxt type_sys combformula)
blanchet@38506
   430
blanchet@41386
   431
fun problem_line_for_fact ctxt prefix type_sys (formula as {name, kind, ...}) =
blanchet@41386
   432
  Fof (prefix ^ ascii_of name, kind, formula_for_fact ctxt type_sys formula)
blanchet@38506
   433
blanchet@38506
   434
fun problem_line_for_class_rel_clause (ClassRelClause {name, subclass,
blanchet@38506
   435
                                                       superclass, ...}) =
blanchet@38506
   436
  let val ty_arg = ATerm (("T", "T"), []) in
blanchet@38506
   437
    Fof (class_rel_clause_prefix ^ ascii_of name, Axiom,
blanchet@38506
   438
         AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
blanchet@38506
   439
                           AAtom (ATerm (superclass, [ty_arg]))]))
blanchet@38506
   440
  end
blanchet@38506
   441
blanchet@38506
   442
fun fo_literal_for_arity_literal (TConsLit (c, t, args)) =
blanchet@38506
   443
    (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
blanchet@38506
   444
  | fo_literal_for_arity_literal (TVarLit (c, sort)) =
blanchet@38506
   445
    (false, ATerm (c, [ATerm (sort, [])]))
blanchet@38506
   446
blanchet@38506
   447
fun problem_line_for_arity_clause (ArityClause {name, conclLit, premLits,
blanchet@38506
   448
                                                ...}) =
blanchet@38506
   449
  Fof (arity_clause_prefix ^ ascii_of name, Axiom,
blanchet@38506
   450
       mk_ahorn (map (formula_for_fo_literal o apfst not
blanchet@38506
   451
                      o fo_literal_for_arity_literal) premLits)
blanchet@38506
   452
                (formula_for_fo_literal
blanchet@38506
   453
                     (fo_literal_for_arity_literal conclLit)))
blanchet@38506
   454
blanchet@41386
   455
fun problem_line_for_conjecture ctxt type_sys
blanchet@40358
   456
        ({name, kind, combformula, ...} : translated_formula) =
blanchet@38506
   457
  Fof (conjecture_prefix ^ name, kind,
blanchet@41386
   458
       formula_for_combformula ctxt type_sys combformula)
blanchet@38506
   459
blanchet@41385
   460
fun free_type_literals_for_conjecture type_sys
blanchet@40358
   461
        ({ctypes_sorts, ...} : translated_formula) =
blanchet@41385
   462
  ctypes_sorts |> atp_type_literals_for_types type_sys
blanchet@41385
   463
               |> map fo_literal_for_type_literal
blanchet@38506
   464
blanchet@40156
   465
fun problem_line_for_free_type j lit =
blanchet@40156
   466
  Fof (tfree_prefix ^ string_of_int j, Hypothesis, formula_for_fo_literal lit)
blanchet@41385
   467
fun problem_lines_for_free_types type_sys conjectures =
blanchet@38506
   468
  let
blanchet@41385
   469
    val litss = map (free_type_literals_for_conjecture type_sys) conjectures
blanchet@38506
   470
    val lits = fold (union (op =)) litss []
blanchet@40156
   471
  in map2 problem_line_for_free_type (0 upto length lits - 1) lits end
blanchet@38506
   472
blanchet@38506
   473
(** "hBOOL" and "hAPP" **)
blanchet@38506
   474
blanchet@38506
   475
type const_info = {min_arity: int, max_arity: int, sub_level: bool}
blanchet@38506
   476
blanchet@38506
   477
fun consider_term top_level (ATerm ((s, _), ts)) =
blanchet@39692
   478
  (if is_atp_variable s then
blanchet@38506
   479
     I
blanchet@38506
   480
   else
blanchet@38506
   481
     let val n = length ts in
blanchet@38506
   482
       Symtab.map_default
blanchet@38506
   483
           (s, {min_arity = n, max_arity = 0, sub_level = false})
blanchet@38506
   484
           (fn {min_arity, max_arity, sub_level} =>
blanchet@38506
   485
               {min_arity = Int.min (n, min_arity),
blanchet@38506
   486
                max_arity = Int.max (n, max_arity),
blanchet@38506
   487
                sub_level = sub_level orelse not top_level})
blanchet@38506
   488
     end)
blanchet@41386
   489
  #> fold (consider_term (top_level andalso s = type_tag_name)) ts
blanchet@38506
   490
fun consider_formula (AQuant (_, _, phi)) = consider_formula phi
blanchet@38506
   491
  | consider_formula (AConn (_, phis)) = fold consider_formula phis
blanchet@38506
   492
  | consider_formula (AAtom tm) = consider_term true tm
blanchet@38506
   493
blanchet@38506
   494
fun consider_problem_line (Fof (_, _, phi)) = consider_formula phi
blanchet@38506
   495
fun consider_problem problem = fold (fold consider_problem_line o snd) problem
blanchet@38506
   496
blanchet@41388
   497
(* needed for helper facts if the problem otherwise does not involve equality *)
blanchet@41388
   498
val equal_entry = ("equal", {min_arity = 2, max_arity = 2, sub_level = false})
blanchet@41388
   499
blanchet@38506
   500
fun const_table_for_problem explicit_apply problem =
blanchet@41388
   501
  if explicit_apply then
blanchet@41388
   502
    NONE
blanchet@41388
   503
  else
blanchet@41388
   504
    SOME (Symtab.empty |> Symtab.update equal_entry |> consider_problem problem)
blanchet@38506
   505
blanchet@41382
   506
fun min_arity_of thy type_sys NONE s =
blanchet@41386
   507
    (if s = "equal" orelse s = type_tag_name orelse
blanchet@38506
   508
        String.isPrefix type_const_prefix s orelse
blanchet@38506
   509
        String.isPrefix class_prefix s then
blanchet@38506
   510
       16383 (* large number *)
blanchet@38987
   511
     else case strip_prefix_and_unascii const_prefix s of
blanchet@41384
   512
       SOME s' => num_atp_type_args thy type_sys (invert_const s')
blanchet@38506
   513
     | NONE => 0)
blanchet@38506
   514
  | min_arity_of _ _ (SOME the_const_tab) s =
blanchet@38506
   515
    case Symtab.lookup the_const_tab s of
blanchet@38506
   516
      SOME ({min_arity, ...} : const_info) => min_arity
blanchet@38506
   517
    | NONE => 0
blanchet@38506
   518
blanchet@38506
   519
fun full_type_of (ATerm ((s, _), [ty, _])) =
blanchet@41386
   520
    if s = type_tag_name then SOME ty else NONE
blanchet@41386
   521
  | full_type_of _ = NONE
blanchet@38506
   522
blanchet@38506
   523
fun list_hAPP_rev _ t1 [] = t1
blanchet@38506
   524
  | list_hAPP_rev NONE t1 (t2 :: ts2) =
blanchet@38506
   525
    ATerm (`I "hAPP", [list_hAPP_rev NONE t1 ts2, t2])
blanchet@38506
   526
  | list_hAPP_rev (SOME ty) t1 (t2 :: ts2) =
blanchet@41386
   527
    case full_type_of t2 of
blanchet@41386
   528
      SOME ty2 =>
blanchet@41386
   529
      let val ty' = ATerm (`make_fixed_type_const @{type_name fun},
blanchet@41386
   530
                           [ty2, ty]) in
blanchet@41386
   531
        ATerm (`I "hAPP",
blanchet@41386
   532
               [tag_with_type ty' (list_hAPP_rev (SOME ty') t1 ts2), t2])
blanchet@41386
   533
      end
blanchet@41386
   534
    | NONE => list_hAPP_rev NONE t1 (t2 :: ts2)
blanchet@38506
   535
blanchet@41382
   536
fun repair_applications_in_term thy type_sys const_tab =
blanchet@38506
   537
  let
blanchet@38506
   538
    fun aux opt_ty (ATerm (name as (s, _), ts)) =
blanchet@41386
   539
      if s = type_tag_name then
blanchet@38506
   540
        case ts of
blanchet@38506
   541
          [t1, t2] => ATerm (name, [aux NONE t1, aux (SOME t1) t2])
blanchet@41386
   542
        | _ => raise Fail "malformed type tag"
blanchet@38506
   543
      else
blanchet@38506
   544
        let
blanchet@38506
   545
          val ts = map (aux NONE) ts
blanchet@41382
   546
          val (ts1, ts2) = chop (min_arity_of thy type_sys const_tab s) ts
blanchet@38506
   547
        in list_hAPP_rev opt_ty (ATerm (name, ts1)) (rev ts2) end
blanchet@38506
   548
  in aux NONE end
blanchet@38506
   549
blanchet@38506
   550
fun boolify t = ATerm (`I "hBOOL", [t])
blanchet@38506
   551
blanchet@38506
   552
(* True if the constant ever appears outside of the top-level position in
blanchet@38506
   553
   literals, or if it appears with different arities (e.g., because of different
blanchet@38506
   554
   type instantiations). If false, the constant always receives all of its
blanchet@38506
   555
   arguments and is used as a predicate. *)
blanchet@38506
   556
fun is_predicate NONE s =
blanchet@38812
   557
    s = "equal" orelse s = "$false" orelse s = "$true" orelse
blanchet@38812
   558
    String.isPrefix type_const_prefix s orelse String.isPrefix class_prefix s
blanchet@38506
   559
  | is_predicate (SOME the_const_tab) s =
blanchet@38506
   560
    case Symtab.lookup the_const_tab s of
blanchet@38506
   561
      SOME {min_arity, max_arity, sub_level} =>
blanchet@38506
   562
      not sub_level andalso min_arity = max_arity
blanchet@38506
   563
    | NONE => false
blanchet@38506
   564
blanchet@41388
   565
fun repair_predicates_in_term pred_const_tab (t as ATerm ((s, _), ts)) =
blanchet@41386
   566
  if s = type_tag_name then
blanchet@38506
   567
    case ts of
blanchet@38506
   568
      [_, t' as ATerm ((s', _), _)] =>
blanchet@41388
   569
      if is_predicate pred_const_tab s' then t' else boolify t
blanchet@41386
   570
    | _ => raise Fail "malformed type tag"
blanchet@38506
   571
  else
blanchet@41388
   572
    t |> not (is_predicate pred_const_tab s) ? boolify
blanchet@38506
   573
blanchet@38506
   574
fun close_universally phi =
blanchet@38506
   575
  let
blanchet@38506
   576
    fun term_vars bounds (ATerm (name as (s, _), tms)) =
blanchet@39692
   577
        (is_atp_variable s andalso not (member (op =) bounds name))
blanchet@38506
   578
          ? insert (op =) name
blanchet@38506
   579
        #> fold (term_vars bounds) tms
blanchet@38917
   580
    fun formula_vars bounds (AQuant (_, xs, phi)) =
blanchet@38506
   581
        formula_vars (xs @ bounds) phi
blanchet@38506
   582
      | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis
blanchet@38506
   583
      | formula_vars bounds (AAtom tm) = term_vars bounds tm
blanchet@38506
   584
  in
blanchet@38506
   585
    case formula_vars [] phi [] of [] => phi | xs => AQuant (AForall, xs, phi)
blanchet@38506
   586
  end
blanchet@38506
   587
blanchet@41382
   588
fun repair_formula thy explicit_forall type_sys const_tab =
blanchet@38506
   589
  let
blanchet@41388
   590
    val pred_const_tab = case type_sys of Tags _ => NONE | _ => const_tab
blanchet@38506
   591
    fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi)
blanchet@38506
   592
      | aux (AConn (c, phis)) = AConn (c, map aux phis)
blanchet@38506
   593
      | aux (AAtom tm) =
blanchet@41382
   594
        AAtom (tm |> repair_applications_in_term thy type_sys const_tab
blanchet@41388
   595
                  |> repair_predicates_in_term pred_const_tab)
blanchet@38506
   596
  in aux #> explicit_forall ? close_universally end
blanchet@38506
   597
blanchet@41382
   598
fun repair_problem_line thy explicit_forall type_sys const_tab
blanchet@38506
   599
                        (Fof (ident, kind, phi)) =
blanchet@41382
   600
  Fof (ident, kind, repair_formula thy explicit_forall type_sys const_tab phi)
blanchet@41388
   601
fun repair_problem thy = map o apsnd o map ooo repair_problem_line thy
blanchet@38506
   602
blanchet@41388
   603
fun dest_Fof (Fof z) = z
blanchet@38506
   604
blanchet@41382
   605
fun prepare_atp_problem ctxt readable_names explicit_forall type_sys
blanchet@40445
   606
                        explicit_apply hyp_ts concl_t facts =
blanchet@38506
   607
  let
blanchet@38506
   608
    val thy = ProofContext.theory_of ctxt
blanchet@41388
   609
    val (fact_names, (conjectures, facts, class_rel_clauses, arity_clauses)) =
blanchet@41382
   610
      translate_formulas ctxt type_sys hyp_ts concl_t facts
blanchet@41386
   611
    val fact_lines = map (problem_line_for_fact ctxt fact_prefix type_sys) facts
blanchet@38506
   612
    val conjecture_lines =
blanchet@41386
   613
      map (problem_line_for_conjecture ctxt type_sys) conjectures
blanchet@41385
   614
    val tfree_lines = problem_lines_for_free_types type_sys conjectures
blanchet@38506
   615
    val class_rel_lines =
blanchet@38506
   616
      map problem_line_for_class_rel_clause class_rel_clauses
blanchet@38506
   617
    val arity_lines = map problem_line_for_arity_clause arity_clauses
blanchet@38506
   618
    (* Reordering these might or might not confuse the proof reconstruction
blanchet@38506
   619
       code or the SPASS Flotter hack. *)
blanchet@38506
   620
    val problem =
blanchet@40445
   621
      [("Relevant facts", fact_lines),
blanchet@38506
   622
       ("Class relationships", class_rel_lines),
blanchet@38506
   623
       ("Arity declarations", arity_lines),
blanchet@41388
   624
       ("Helper facts", []),
blanchet@38506
   625
       ("Conjectures", conjecture_lines),
blanchet@38506
   626
       ("Type variables", tfree_lines)]
blanchet@41388
   627
    val const_tab = const_table_for_problem explicit_apply problem
blanchet@41388
   628
    val problem =
blanchet@41388
   629
      problem |> repair_problem thy explicit_forall type_sys const_tab
blanchet@41388
   630
    val helper_facts =
blanchet@41388
   631
      get_helper_facts ctxt type_sys (maps (map (#3 o dest_Fof) o snd) problem)
blanchet@41388
   632
    val helper_lines =
blanchet@41388
   633
      helper_facts
blanchet@41388
   634
      |> map (problem_line_for_fact ctxt helper_prefix type_sys
blanchet@41388
   635
              #> repair_problem_line thy explicit_forall type_sys const_tab)
blanchet@41388
   636
    val (problem, pool) =
blanchet@41388
   637
      problem |> AList.update (op =) ("Helper facts", helper_lines)
blanchet@41388
   638
              |> nice_atp_problem readable_names
blanchet@38506
   639
    val conjecture_offset =
blanchet@40445
   640
      length fact_lines + length class_rel_lines + length arity_lines
blanchet@38506
   641
      + length helper_lines
blanchet@38506
   642
  in
blanchet@38506
   643
    (problem,
blanchet@38506
   644
     case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
blanchet@40445
   645
     conjecture_offset, fact_names)
blanchet@38506
   646
  end
blanchet@38506
   647
blanchet@38506
   648
end;