src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML
author blanchet
Wed, 18 Aug 2010 12:03:44 +0200
changeset 38743 54727b44e277
parent 38721 dafcd0d19b11
child 38812 b03f8fe043ec
permissions -rw-r--r--
handle bound name conflicts gracefully in FOF translation
blanchet@38506
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_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@38506
     6
Translation of HOL to FOL.
blanchet@38506
     7
*)
blanchet@38506
     8
blanchet@38506
     9
signature SLEDGEHAMMER_TRANSLATE =
blanchet@38506
    10
sig
blanchet@38506
    11
  type 'a problem = 'a ATP_Problem.problem
blanchet@38506
    12
blanchet@38506
    13
  val axiom_prefix : string
blanchet@38506
    14
  val conjecture_prefix : string
blanchet@38506
    15
  val helper_prefix : string
blanchet@38506
    16
  val class_rel_clause_prefix : string
blanchet@38506
    17
  val arity_clause_prefix : string
blanchet@38506
    18
  val tfrees_name : string
blanchet@38506
    19
  val prepare_problem :
blanchet@38506
    20
    Proof.context -> bool -> bool -> bool -> bool -> term list -> term
blanchet@38506
    21
    -> (string * thm) list
blanchet@38506
    22
    -> string problem * string Symtab.table * int * string Vector.vector
blanchet@38506
    23
end;
blanchet@38506
    24
blanchet@38506
    25
structure Sledgehammer_Translate : SLEDGEHAMMER_TRANSLATE =
blanchet@38506
    26
struct
blanchet@38506
    27
blanchet@38506
    28
open ATP_Problem
blanchet@38506
    29
open Metis_Clauses
blanchet@38506
    30
open Sledgehammer_Util
blanchet@38506
    31
blanchet@38506
    32
val axiom_prefix = "ax_"
blanchet@38506
    33
val conjecture_prefix = "conj_"
blanchet@38506
    34
val helper_prefix = "help_"
blanchet@38506
    35
val class_rel_clause_prefix = "clrel_";
blanchet@38506
    36
val arity_clause_prefix = "arity_"
blanchet@38506
    37
val tfrees_name = "tfrees"
blanchet@38506
    38
blanchet@38506
    39
(* Freshness almost guaranteed! *)
blanchet@38506
    40
val sledgehammer_weak_prefix = "Sledgehammer:"
blanchet@38506
    41
blanchet@38506
    42
datatype fol_formula =
blanchet@38506
    43
  FOLFormula of {name: string,
blanchet@38506
    44
                 kind: kind,
blanchet@38506
    45
                 combformula: (name, combterm) formula,
blanchet@38506
    46
                 ctypes_sorts: typ list}
blanchet@38506
    47
blanchet@38506
    48
fun mk_anot phi = AConn (ANot, [phi])
blanchet@38506
    49
fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2])
blanchet@38506
    50
fun mk_ahorn [] phi = phi
blanchet@38506
    51
  | mk_ahorn (phi :: phis) psi =
blanchet@38506
    52
    AConn (AImplies, [fold (mk_aconn AAnd) phis phi, psi])
blanchet@38506
    53
blanchet@38506
    54
fun combformula_for_prop thy =
blanchet@38506
    55
  let
blanchet@38506
    56
    val do_term = combterm_from_term thy
blanchet@38506
    57
    fun do_quant bs q s T t' =
blanchet@38743
    58
      let val s = Name.variant (map fst bs) s in
blanchet@38743
    59
        do_formula ((s, T) :: bs) t'
blanchet@38743
    60
        #>> (fn phi => AQuant (q, [`make_bound_var s], phi))
blanchet@38743
    61
      end
blanchet@38506
    62
    and do_conn bs c t1 t2 =
blanchet@38506
    63
      do_formula bs t1 ##>> do_formula bs t2
blanchet@38506
    64
      #>> (fn (phi1, phi2) => AConn (c, [phi1, phi2]))
blanchet@38506
    65
    and do_formula bs t =
blanchet@38506
    66
      case t of
blanchet@38506
    67
        @{const Not} $ t1 =>
blanchet@38506
    68
        do_formula bs t1 #>> (fn phi => AConn (ANot, [phi]))
blanchet@38506
    69
      | Const (@{const_name All}, _) $ Abs (s, T, t') =>
blanchet@38506
    70
        do_quant bs AForall s T t'
blanchet@38506
    71
      | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
blanchet@38506
    72
        do_quant bs AExists s T t'
blanchet@38506
    73
      | @{const "op &"} $ t1 $ t2 => do_conn bs AAnd t1 t2
blanchet@38506
    74
      | @{const "op |"} $ t1 $ t2 => do_conn bs AOr t1 t2
blanchet@38506
    75
      | @{const "op -->"} $ t1 $ t2 => do_conn bs AImplies t1 t2
blanchet@38506
    76
      | Const (@{const_name "op ="}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
blanchet@38506
    77
        do_conn bs AIff t1 t2
blanchet@38506
    78
      | _ => (fn ts => do_term bs (Envir.eta_contract t)
blanchet@38506
    79
                       |>> AAtom ||> union (op =) ts)
blanchet@38506
    80
  in do_formula [] end
blanchet@38506
    81
blanchet@38506
    82
(* Converts an elim-rule into an equivalent theorem that does not have the
blanchet@38506
    83
   predicate variable. Leaves other theorems unchanged. We simply instantiate
blanchet@38506
    84
   the conclusion variable to False. (Cf. "transform_elim_term" in
blanchet@38506
    85
   "ATP_Systems".) *)
blanchet@38506
    86
fun transform_elim_term t =
blanchet@38506
    87
  case Logic.strip_imp_concl t of
blanchet@38506
    88
    @{const Trueprop} $ Var (z, @{typ bool}) =>
blanchet@38506
    89
    subst_Vars [(z, @{const False})] t
blanchet@38506
    90
  | Var (z, @{typ prop}) => subst_Vars [(z, @{prop False})] t
blanchet@38506
    91
  | _ => t
blanchet@38506
    92
blanchet@38506
    93
fun presimplify_term thy =
blanchet@38506
    94
  Skip_Proof.make_thm thy
blanchet@38506
    95
  #> Meson.presimplify
blanchet@38506
    96
  #> prop_of
blanchet@38506
    97
blanchet@38506
    98
fun concealed_bound_name j = sledgehammer_weak_prefix ^ Int.toString j
blanchet@38506
    99
fun conceal_bounds Ts t =
blanchet@38506
   100
  subst_bounds (map (Free o apfst concealed_bound_name)
blanchet@38506
   101
                    (0 upto length Ts - 1 ~~ Ts), t)
blanchet@38506
   102
fun reveal_bounds Ts =
blanchet@38506
   103
  subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
blanchet@38506
   104
                    (0 upto length Ts - 1 ~~ Ts))
blanchet@38506
   105
blanchet@38506
   106
fun introduce_combinators_in_term ctxt kind t =
blanchet@38716
   107
  let val thy = ProofContext.theory_of ctxt in
blanchet@38716
   108
    if Meson.is_fol_term thy t then
blanchet@38716
   109
      t
blanchet@38716
   110
    else
blanchet@38716
   111
      let
blanchet@38716
   112
        fun aux Ts t =
blanchet@38716
   113
          case t of
blanchet@38716
   114
            @{const Not} $ t1 => @{const Not} $ aux Ts t1
blanchet@38716
   115
          | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
blanchet@38716
   116
            t0 $ Abs (s, T, aux (T :: Ts) t')
blanchet@38716
   117
          | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
blanchet@38716
   118
            t0 $ Abs (s, T, aux (T :: Ts) t')
blanchet@38716
   119
          | (t0 as @{const "op &"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
blanchet@38716
   120
          | (t0 as @{const "op |"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
blanchet@38716
   121
          | (t0 as @{const "op -->"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
blanchet@38716
   122
          | (t0 as Const (@{const_name "op ="}, Type (_, [@{typ bool}, _])))
blanchet@38716
   123
              $ t1 $ t2 =>
blanchet@38716
   124
            t0 $ aux Ts t1 $ aux Ts t2
blanchet@38716
   125
          | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
blanchet@38716
   126
                   t
blanchet@38716
   127
                 else
blanchet@38716
   128
                   t |> conceal_bounds Ts
blanchet@38716
   129
                     |> Envir.eta_contract
blanchet@38716
   130
                     |> cterm_of thy
blanchet@38716
   131
                     |> Clausifier.introduce_combinators_in_cterm
blanchet@38716
   132
                     |> prop_of |> Logic.dest_equals |> snd
blanchet@38716
   133
                     |> reveal_bounds Ts
blanchet@38716
   134
        val ([t], ctxt') = Variable.import_terms true [t] ctxt
blanchet@38716
   135
      in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
blanchet@38716
   136
      handle THM _ =>
blanchet@38716
   137
             (* A type variable of sort "{}" will make abstraction fail. *)
blanchet@38716
   138
             case kind of
blanchet@38716
   139
               Axiom => HOLogic.true_const
blanchet@38716
   140
             | Conjecture => HOLogic.false_const
blanchet@38716
   141
  end
blanchet@38506
   142
blanchet@38506
   143
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
blanchet@38506
   144
   same in Sledgehammer to prevent the discovery of unreplable proofs. *)
blanchet@38506
   145
fun freeze_term t =
blanchet@38506
   146
  let
blanchet@38506
   147
    fun aux (t $ u) = aux t $ aux u
blanchet@38506
   148
      | aux (Abs (s, T, t)) = Abs (s, T, aux t)
blanchet@38506
   149
      | aux (Var ((s, i), T)) =
blanchet@38506
   150
        Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
blanchet@38506
   151
      | aux t = t
blanchet@38506
   152
  in t |> exists_subterm is_Var t ? aux end
blanchet@38506
   153
blanchet@38506
   154
(* making axiom and conjecture formulas *)
blanchet@38506
   155
fun make_formula ctxt presimp (name, kind, t) =
blanchet@38506
   156
  let
blanchet@38506
   157
    val thy = ProofContext.theory_of ctxt
blanchet@38506
   158
    val t = t |> transform_elim_term
blanchet@38506
   159
              |> Object_Logic.atomize_term thy
blanchet@38506
   160
    val t = t |> fastype_of t = HOLogic.boolT ? HOLogic.mk_Trueprop
blanchet@38506
   161
              |> extensionalize_term
blanchet@38506
   162
              |> presimp ? presimplify_term thy
blanchet@38506
   163
              |> perhaps (try (HOLogic.dest_Trueprop))
blanchet@38506
   164
              |> introduce_combinators_in_term ctxt kind
blanchet@38506
   165
              |> kind = Conjecture ? freeze_term
blanchet@38506
   166
    val (combformula, ctypes_sorts) = combformula_for_prop thy t []
blanchet@38506
   167
  in
blanchet@38506
   168
    FOLFormula {name = name, combformula = combformula, kind = kind,
blanchet@38506
   169
                ctypes_sorts = ctypes_sorts}
blanchet@38506
   170
  end
blanchet@38506
   171
blanchet@38506
   172
fun make_axiom ctxt presimp (name, th) =
blanchet@38506
   173
  (name, make_formula ctxt presimp (name, Axiom, prop_of th))
blanchet@38506
   174
fun make_conjectures ctxt ts =
blanchet@38506
   175
  map2 (fn j => fn t => make_formula ctxt true (Int.toString j, Conjecture, t))
blanchet@38506
   176
       (0 upto length ts - 1) ts
blanchet@38506
   177
blanchet@38506
   178
(** Helper facts **)
blanchet@38506
   179
blanchet@38506
   180
fun count_combterm (CombConst ((s, _), _, _)) =
blanchet@38506
   181
    Symtab.map_entry s (Integer.add 1)
blanchet@38506
   182
  | count_combterm (CombVar _) = I
blanchet@38506
   183
  | count_combterm (CombApp (t1, t2)) = fold count_combterm [t1, t2]
blanchet@38506
   184
fun count_combformula (AQuant (_, _, phi)) = count_combformula phi
blanchet@38506
   185
  | count_combformula (AConn (_, phis)) = fold count_combformula phis
blanchet@38506
   186
  | count_combformula (AAtom tm) = count_combterm tm
blanchet@38506
   187
fun count_fol_formula (FOLFormula {combformula, ...}) =
blanchet@38506
   188
  count_combformula combformula
blanchet@38506
   189
blanchet@38506
   190
val optional_helpers =
blanchet@38506
   191
  [(["c_COMBI", "c_COMBK"], @{thms COMBI_def COMBK_def}),
blanchet@38506
   192
   (["c_COMBB", "c_COMBC"], @{thms COMBB_def COMBC_def}),
blanchet@38506
   193
   (["c_COMBS"], @{thms COMBS_def})]
blanchet@38506
   194
val optional_typed_helpers =
blanchet@38506
   195
  [(["c_True", "c_False"], @{thms True_or_False}),
blanchet@38506
   196
   (["c_If"], @{thms if_True if_False True_or_False})]
blanchet@38506
   197
val mandatory_helpers = @{thms fequal_imp_equal equal_imp_fequal}
blanchet@38506
   198
blanchet@38506
   199
val init_counters =
blanchet@38506
   200
  Symtab.make (maps (maps (map (rpair 0) o fst))
blanchet@38506
   201
                    [optional_helpers, optional_typed_helpers])
blanchet@38506
   202
blanchet@38506
   203
fun get_helper_facts ctxt is_FO full_types conjectures axioms =
blanchet@38506
   204
  let
blanchet@38506
   205
    val ct = fold (fold count_fol_formula) [conjectures, axioms] init_counters
blanchet@38506
   206
    fun is_needed c = the (Symtab.lookup ct c) > 0
blanchet@38506
   207
  in
blanchet@38506
   208
    (optional_helpers
blanchet@38506
   209
     |> full_types ? append optional_typed_helpers
blanchet@38506
   210
     |> maps (fn (ss, ths) =>
blanchet@38506
   211
                 if exists is_needed ss then map (`Thm.get_name_hint) ths
blanchet@38506
   212
                 else [])) @
blanchet@38506
   213
    (if is_FO then [] else map (`Thm.get_name_hint) mandatory_helpers)
blanchet@38506
   214
    |> map (snd o make_axiom ctxt false)
blanchet@38506
   215
  end
blanchet@38506
   216
blanchet@38506
   217
fun meta_not t = @{const "==>"} $ t $ @{prop False}
blanchet@38506
   218
blanchet@38506
   219
fun prepare_formulas ctxt full_types hyp_ts concl_t axioms =
blanchet@38506
   220
  let
blanchet@38506
   221
    val thy = ProofContext.theory_of ctxt
blanchet@38506
   222
    val axiom_ts = map (prop_of o snd) axioms
blanchet@38506
   223
    val hyp_ts =
blanchet@38506
   224
      if null hyp_ts then
blanchet@38506
   225
        []
blanchet@38506
   226
      else
blanchet@38506
   227
        (* Remove existing axioms from the conjecture, as this can dramatically
blanchet@38506
   228
           boost an ATP's performance (for some reason). *)
blanchet@38506
   229
        let
blanchet@38506
   230
          val axiom_table = fold (Termtab.update o rpair ()) axiom_ts
blanchet@38506
   231
                                 Termtab.empty
blanchet@38506
   232
        in hyp_ts |> filter_out (Termtab.defined axiom_table) end
blanchet@38506
   233
    val goal_t = Logic.list_implies (hyp_ts, concl_t)
blanchet@38506
   234
    val is_FO = Meson.is_fol_term thy goal_t
blanchet@38506
   235
    val subs = tfree_classes_of_terms [goal_t]
blanchet@38506
   236
    val supers = tvar_classes_of_terms axiom_ts
blanchet@38506
   237
    val tycons = type_consts_of_terms thy (goal_t :: axiom_ts)
blanchet@38506
   238
    (* TFrees in the conjecture; TVars in the axioms *)
blanchet@38506
   239
    val conjectures = map meta_not hyp_ts @ [concl_t] |> make_conjectures ctxt
blanchet@38506
   240
    val (axiom_names, axioms) =
blanchet@38506
   241
      ListPair.unzip (map (make_axiom ctxt true) axioms)
blanchet@38506
   242
    val helper_facts = get_helper_facts ctxt is_FO full_types conjectures axioms
blanchet@38506
   243
    val (supers', arity_clauses) = make_arity_clauses thy tycons supers
blanchet@38506
   244
    val class_rel_clauses = make_class_rel_clauses thy subs supers'
blanchet@38506
   245
  in
blanchet@38506
   246
    (Vector.fromList axiom_names,
blanchet@38506
   247
     (conjectures, axioms, helper_facts, class_rel_clauses, arity_clauses))
blanchet@38506
   248
  end
blanchet@38506
   249
blanchet@38506
   250
fun wrap_type ty t = ATerm ((type_wrapper_name, type_wrapper_name), [ty, t])
blanchet@38506
   251
blanchet@38506
   252
fun fo_term_for_combtyp (CombTVar name) = ATerm (name, [])
blanchet@38506
   253
  | fo_term_for_combtyp (CombTFree name) = ATerm (name, [])
blanchet@38506
   254
  | fo_term_for_combtyp (CombType (name, tys)) =
blanchet@38506
   255
    ATerm (name, map fo_term_for_combtyp tys)
blanchet@38506
   256
blanchet@38506
   257
fun fo_literal_for_type_literal (TyLitVar (class, name)) =
blanchet@38506
   258
    (true, ATerm (class, [ATerm (name, [])]))
blanchet@38506
   259
  | fo_literal_for_type_literal (TyLitFree (class, name)) =
blanchet@38506
   260
    (true, ATerm (class, [ATerm (name, [])]))
blanchet@38506
   261
blanchet@38506
   262
fun formula_for_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
blanchet@38506
   263
blanchet@38506
   264
fun fo_term_for_combterm full_types =
blanchet@38506
   265
  let
blanchet@38506
   266
    fun aux top_level u =
blanchet@38506
   267
      let
blanchet@38506
   268
        val (head, args) = strip_combterm_comb u
blanchet@38506
   269
        val (x, ty_args) =
blanchet@38506
   270
          case head of
blanchet@38506
   271
            CombConst (name as (s, s'), _, ty_args) =>
blanchet@38721
   272
            let val ty_args = if full_types then [] else ty_args in
blanchet@38721
   273
              if s = "equal" then
blanchet@38721
   274
                if top_level andalso length args = 2 then (name, [])
blanchet@38721
   275
                else (("c_fequal", @{const_name fequal}), ty_args)
blanchet@38721
   276
              else if top_level then
blanchet@38721
   277
                case s of
blanchet@38721
   278
                  "c_False" => (("$false", s'), [])
blanchet@38721
   279
                | "c_True" => (("$true", s'), [])
blanchet@38721
   280
                | _ => (name, ty_args)
blanchet@38721
   281
              else
blanchet@38721
   282
                (name, ty_args)
blanchet@38721
   283
            end
blanchet@38506
   284
          | CombVar (name, _) => (name, [])
blanchet@38506
   285
          | CombApp _ => raise Fail "impossible \"CombApp\""
blanchet@38506
   286
        val t = ATerm (x, map fo_term_for_combtyp ty_args @
blanchet@38506
   287
                          map (aux false) args)
blanchet@38506
   288
    in
blanchet@38506
   289
      if full_types then wrap_type (fo_term_for_combtyp (combtyp_of u)) t else t
blanchet@38506
   290
    end
blanchet@38506
   291
  in aux true end
blanchet@38506
   292
blanchet@38506
   293
fun formula_for_combformula full_types =
blanchet@38506
   294
  let
blanchet@38506
   295
    fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi)
blanchet@38506
   296
      | aux (AConn (c, phis)) = AConn (c, map aux phis)
blanchet@38506
   297
      | aux (AAtom tm) = AAtom (fo_term_for_combterm full_types tm)
blanchet@38506
   298
  in aux end
blanchet@38506
   299
blanchet@38506
   300
fun formula_for_axiom full_types (FOLFormula {combformula, ctypes_sorts, ...}) =
blanchet@38506
   301
  mk_ahorn (map (formula_for_fo_literal o fo_literal_for_type_literal)
blanchet@38506
   302
                (type_literals_for_types ctypes_sorts))
blanchet@38506
   303
           (formula_for_combformula full_types combformula)
blanchet@38506
   304
blanchet@38506
   305
fun problem_line_for_fact prefix full_types
blanchet@38506
   306
                          (formula as FOLFormula {name, kind, ...}) =
blanchet@38506
   307
  Fof (prefix ^ ascii_of name, kind, formula_for_axiom full_types formula)
blanchet@38506
   308
blanchet@38506
   309
fun problem_line_for_class_rel_clause (ClassRelClause {name, subclass,
blanchet@38506
   310
                                                       superclass, ...}) =
blanchet@38506
   311
  let val ty_arg = ATerm (("T", "T"), []) in
blanchet@38506
   312
    Fof (class_rel_clause_prefix ^ ascii_of name, Axiom,
blanchet@38506
   313
         AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
blanchet@38506
   314
                           AAtom (ATerm (superclass, [ty_arg]))]))
blanchet@38506
   315
  end
blanchet@38506
   316
blanchet@38506
   317
fun fo_literal_for_arity_literal (TConsLit (c, t, args)) =
blanchet@38506
   318
    (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
blanchet@38506
   319
  | fo_literal_for_arity_literal (TVarLit (c, sort)) =
blanchet@38506
   320
    (false, ATerm (c, [ATerm (sort, [])]))
blanchet@38506
   321
blanchet@38506
   322
fun problem_line_for_arity_clause (ArityClause {name, conclLit, premLits,
blanchet@38506
   323
                                                ...}) =
blanchet@38506
   324
  Fof (arity_clause_prefix ^ ascii_of name, Axiom,
blanchet@38506
   325
       mk_ahorn (map (formula_for_fo_literal o apfst not
blanchet@38506
   326
                      o fo_literal_for_arity_literal) premLits)
blanchet@38506
   327
                (formula_for_fo_literal
blanchet@38506
   328
                     (fo_literal_for_arity_literal conclLit)))
blanchet@38506
   329
blanchet@38506
   330
fun problem_line_for_conjecture full_types
blanchet@38506
   331
                                (FOLFormula {name, kind, combformula, ...}) =
blanchet@38506
   332
  Fof (conjecture_prefix ^ name, kind,
blanchet@38506
   333
       formula_for_combformula full_types combformula)
blanchet@38506
   334
blanchet@38506
   335
fun free_type_literals_for_conjecture (FOLFormula {ctypes_sorts, ...}) =
blanchet@38506
   336
  map fo_literal_for_type_literal (type_literals_for_types ctypes_sorts)
blanchet@38506
   337
blanchet@38506
   338
fun problem_line_for_free_type lit =
blanchet@38506
   339
  Fof (tfrees_name, Conjecture, mk_anot (formula_for_fo_literal lit))
blanchet@38506
   340
fun problem_lines_for_free_types conjectures =
blanchet@38506
   341
  let
blanchet@38506
   342
    val litss = map free_type_literals_for_conjecture conjectures
blanchet@38506
   343
    val lits = fold (union (op =)) litss []
blanchet@38506
   344
  in map problem_line_for_free_type lits end
blanchet@38506
   345
blanchet@38506
   346
(** "hBOOL" and "hAPP" **)
blanchet@38506
   347
blanchet@38506
   348
type const_info = {min_arity: int, max_arity: int, sub_level: bool}
blanchet@38506
   349
blanchet@38506
   350
fun consider_term top_level (ATerm ((s, _), ts)) =
blanchet@38506
   351
  (if is_tptp_variable s then
blanchet@38506
   352
     I
blanchet@38506
   353
   else
blanchet@38506
   354
     let val n = length ts in
blanchet@38506
   355
       Symtab.map_default
blanchet@38506
   356
           (s, {min_arity = n, max_arity = 0, sub_level = false})
blanchet@38506
   357
           (fn {min_arity, max_arity, sub_level} =>
blanchet@38506
   358
               {min_arity = Int.min (n, min_arity),
blanchet@38506
   359
                max_arity = Int.max (n, max_arity),
blanchet@38506
   360
                sub_level = sub_level orelse not top_level})
blanchet@38506
   361
     end)
blanchet@38506
   362
  #> fold (consider_term (top_level andalso s = type_wrapper_name)) ts
blanchet@38506
   363
fun consider_formula (AQuant (_, _, phi)) = consider_formula phi
blanchet@38506
   364
  | consider_formula (AConn (_, phis)) = fold consider_formula phis
blanchet@38506
   365
  | consider_formula (AAtom tm) = consider_term true tm
blanchet@38506
   366
blanchet@38506
   367
fun consider_problem_line (Fof (_, _, phi)) = consider_formula phi
blanchet@38506
   368
fun consider_problem problem = fold (fold consider_problem_line o snd) problem
blanchet@38506
   369
blanchet@38506
   370
fun const_table_for_problem explicit_apply problem =
blanchet@38506
   371
  if explicit_apply then NONE
blanchet@38506
   372
  else SOME (Symtab.empty |> consider_problem problem)
blanchet@38506
   373
blanchet@38506
   374
fun min_arity_of thy full_types NONE s =
blanchet@38506
   375
    (if s = "equal" orelse s = type_wrapper_name orelse
blanchet@38506
   376
        String.isPrefix type_const_prefix s orelse
blanchet@38506
   377
        String.isPrefix class_prefix s then
blanchet@38506
   378
       16383 (* large number *)
blanchet@38506
   379
     else if full_types then
blanchet@38506
   380
       0
blanchet@38506
   381
     else case strip_prefix_and_undo_ascii const_prefix s of
blanchet@38506
   382
       SOME s' => num_type_args thy (invert_const s')
blanchet@38506
   383
     | NONE => 0)
blanchet@38506
   384
  | min_arity_of _ _ (SOME the_const_tab) s =
blanchet@38506
   385
    case Symtab.lookup the_const_tab s of
blanchet@38506
   386
      SOME ({min_arity, ...} : const_info) => min_arity
blanchet@38506
   387
    | NONE => 0
blanchet@38506
   388
blanchet@38506
   389
fun full_type_of (ATerm ((s, _), [ty, _])) =
blanchet@38506
   390
    if s = type_wrapper_name then ty else raise Fail "expected type wrapper"
blanchet@38506
   391
  | full_type_of _ = raise Fail "expected type wrapper"
blanchet@38506
   392
blanchet@38506
   393
fun list_hAPP_rev _ t1 [] = t1
blanchet@38506
   394
  | list_hAPP_rev NONE t1 (t2 :: ts2) =
blanchet@38506
   395
    ATerm (`I "hAPP", [list_hAPP_rev NONE t1 ts2, t2])
blanchet@38506
   396
  | list_hAPP_rev (SOME ty) t1 (t2 :: ts2) =
blanchet@38506
   397
    let val ty' = ATerm (`make_fixed_type_const @{type_name fun},
blanchet@38506
   398
                         [full_type_of t2, ty]) in
blanchet@38506
   399
      ATerm (`I "hAPP", [wrap_type ty' (list_hAPP_rev (SOME ty') t1 ts2), t2])
blanchet@38506
   400
    end
blanchet@38506
   401
blanchet@38506
   402
fun repair_applications_in_term thy full_types const_tab =
blanchet@38506
   403
  let
blanchet@38506
   404
    fun aux opt_ty (ATerm (name as (s, _), ts)) =
blanchet@38506
   405
      if s = type_wrapper_name then
blanchet@38506
   406
        case ts of
blanchet@38506
   407
          [t1, t2] => ATerm (name, [aux NONE t1, aux (SOME t1) t2])
blanchet@38506
   408
        | _ => raise Fail "malformed type wrapper"
blanchet@38506
   409
      else
blanchet@38506
   410
        let
blanchet@38506
   411
          val ts = map (aux NONE) ts
blanchet@38506
   412
          val (ts1, ts2) = chop (min_arity_of thy full_types const_tab s) ts
blanchet@38506
   413
        in list_hAPP_rev opt_ty (ATerm (name, ts1)) (rev ts2) end
blanchet@38506
   414
  in aux NONE end
blanchet@38506
   415
blanchet@38506
   416
fun boolify t = ATerm (`I "hBOOL", [t])
blanchet@38506
   417
blanchet@38506
   418
(* True if the constant ever appears outside of the top-level position in
blanchet@38506
   419
   literals, or if it appears with different arities (e.g., because of different
blanchet@38506
   420
   type instantiations). If false, the constant always receives all of its
blanchet@38506
   421
   arguments and is used as a predicate. *)
blanchet@38506
   422
fun is_predicate NONE s =
blanchet@38506
   423
    s = "equal" orelse String.isPrefix type_const_prefix s orelse
blanchet@38506
   424
    String.isPrefix class_prefix s
blanchet@38506
   425
  | is_predicate (SOME the_const_tab) s =
blanchet@38506
   426
    case Symtab.lookup the_const_tab s of
blanchet@38506
   427
      SOME {min_arity, max_arity, sub_level} =>
blanchet@38506
   428
      not sub_level andalso min_arity = max_arity
blanchet@38506
   429
    | NONE => false
blanchet@38506
   430
blanchet@38506
   431
fun repair_predicates_in_term const_tab (t as ATerm ((s, _), ts)) =
blanchet@38506
   432
  if s = type_wrapper_name then
blanchet@38506
   433
    case ts of
blanchet@38506
   434
      [_, t' as ATerm ((s', _), _)] =>
blanchet@38506
   435
      if is_predicate const_tab s' then t' else boolify t
blanchet@38506
   436
    | _ => raise Fail "malformed type wrapper"
blanchet@38506
   437
  else
blanchet@38506
   438
    t |> not (is_predicate const_tab s) ? boolify
blanchet@38506
   439
blanchet@38506
   440
fun close_universally phi =
blanchet@38506
   441
  let
blanchet@38506
   442
    fun term_vars bounds (ATerm (name as (s, _), tms)) =
blanchet@38506
   443
        (is_tptp_variable s andalso not (member (op =) bounds name))
blanchet@38506
   444
          ? insert (op =) name
blanchet@38506
   445
        #> fold (term_vars bounds) tms
blanchet@38506
   446
    fun formula_vars bounds (AQuant (q, xs, phi)) =
blanchet@38506
   447
        formula_vars (xs @ bounds) phi
blanchet@38506
   448
      | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis
blanchet@38506
   449
      | formula_vars bounds (AAtom tm) = term_vars bounds tm
blanchet@38506
   450
  in
blanchet@38506
   451
    case formula_vars [] phi [] of [] => phi | xs => AQuant (AForall, xs, phi)
blanchet@38506
   452
  end
blanchet@38506
   453
blanchet@38506
   454
fun repair_formula thy explicit_forall full_types const_tab =
blanchet@38506
   455
  let
blanchet@38506
   456
    fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi)
blanchet@38506
   457
      | aux (AConn (c, phis)) = AConn (c, map aux phis)
blanchet@38506
   458
      | aux (AAtom tm) =
blanchet@38506
   459
        AAtom (tm |> repair_applications_in_term thy full_types const_tab
blanchet@38506
   460
                  |> repair_predicates_in_term const_tab)
blanchet@38506
   461
  in aux #> explicit_forall ? close_universally end
blanchet@38506
   462
blanchet@38506
   463
fun repair_problem_line thy explicit_forall full_types const_tab
blanchet@38506
   464
                        (Fof (ident, kind, phi)) =
blanchet@38506
   465
  Fof (ident, kind, repair_formula thy explicit_forall full_types const_tab phi)
blanchet@38506
   466
fun repair_problem_with_const_table thy =
blanchet@38506
   467
  map o apsnd o map ooo repair_problem_line thy
blanchet@38506
   468
blanchet@38506
   469
fun repair_problem thy explicit_forall full_types explicit_apply problem =
blanchet@38506
   470
  repair_problem_with_const_table thy explicit_forall full_types
blanchet@38506
   471
      (const_table_for_problem explicit_apply problem) problem
blanchet@38506
   472
blanchet@38506
   473
fun prepare_problem ctxt readable_names explicit_forall full_types
blanchet@38506
   474
                    explicit_apply hyp_ts concl_t axiom_ts =
blanchet@38506
   475
  let
blanchet@38506
   476
    val thy = ProofContext.theory_of ctxt
blanchet@38506
   477
    val (axiom_names, (conjectures, axioms, helper_facts, class_rel_clauses,
blanchet@38506
   478
                       arity_clauses)) =
blanchet@38506
   479
      prepare_formulas ctxt full_types hyp_ts concl_t axiom_ts
blanchet@38506
   480
    val axiom_lines = map (problem_line_for_fact axiom_prefix full_types) axioms
blanchet@38506
   481
    val helper_lines =
blanchet@38506
   482
      map (problem_line_for_fact helper_prefix full_types) helper_facts
blanchet@38506
   483
    val conjecture_lines =
blanchet@38506
   484
      map (problem_line_for_conjecture full_types) conjectures
blanchet@38506
   485
    val tfree_lines = problem_lines_for_free_types conjectures
blanchet@38506
   486
    val class_rel_lines =
blanchet@38506
   487
      map problem_line_for_class_rel_clause class_rel_clauses
blanchet@38506
   488
    val arity_lines = map problem_line_for_arity_clause arity_clauses
blanchet@38506
   489
    (* Reordering these might or might not confuse the proof reconstruction
blanchet@38506
   490
       code or the SPASS Flotter hack. *)
blanchet@38506
   491
    val problem =
blanchet@38506
   492
      [("Relevant facts", axiom_lines),
blanchet@38506
   493
       ("Class relationships", class_rel_lines),
blanchet@38506
   494
       ("Arity declarations", arity_lines),
blanchet@38506
   495
       ("Helper facts", helper_lines),
blanchet@38506
   496
       ("Conjectures", conjecture_lines),
blanchet@38506
   497
       ("Type variables", tfree_lines)]
blanchet@38506
   498
      |> repair_problem thy explicit_forall full_types explicit_apply
blanchet@38506
   499
    val (problem, pool) = nice_tptp_problem readable_names problem
blanchet@38506
   500
    val conjecture_offset =
blanchet@38506
   501
      length axiom_lines + length class_rel_lines + length arity_lines
blanchet@38506
   502
      + length helper_lines
blanchet@38506
   503
  in
blanchet@38506
   504
    (problem,
blanchet@38506
   505
     case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
blanchet@38506
   506
     conjecture_offset, axiom_names)
blanchet@38506
   507
  end
blanchet@38506
   508
blanchet@38506
   509
end;