src/HOL/TPTP/atp_export.ML
author blanchet
Thu, 28 Jul 2011 11:49:03 +0200
changeset 44870 04fd92795458
parent 44867 4d1270ddf042
child 44959 3693baa6befb
permissions -rw-r--r--
clean up temporary directory hack
blanchet@44064
     1
(*  Title:      HOL/ex/atp_export.ML
blanchet@43473
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@43473
     3
    Copyright   2011
blanchet@43473
     4
blanchet@43473
     5
Export Isabelle theories as first-order TPTP inferences, exploiting
blanchet@43473
     6
Sledgehammer's translation.
blanchet@43473
     7
*)
blanchet@43473
     8
blanchet@44064
     9
signature ATP_EXPORT =
blanchet@43473
    10
sig
blanchet@44365
    11
  val theorems_mentioned_in_proof_term :
blanchet@44365
    12
    string list option -> thm -> string list
blanchet@43473
    13
  val generate_tptp_graph_file_for_theory :
blanchet@43473
    14
    Proof.context -> theory -> string -> unit
blanchet@43473
    15
  val generate_tptp_inference_file_for_theory :
blanchet@44263
    16
    Proof.context -> theory -> string -> string -> unit
blanchet@43473
    17
end;
blanchet@43473
    18
blanchet@44064
    19
structure ATP_Export : ATP_EXPORT =
blanchet@43473
    20
struct
blanchet@43473
    21
blanchet@44350
    22
open ATP_Problem
blanchet@44350
    23
open ATP_Translate
blanchet@44428
    24
open ATP_Proof
blanchet@44428
    25
open ATP_Systems
blanchet@43473
    26
blanchet@44350
    27
val fact_name_of = prefix fact_prefix o ascii_of
blanchet@43473
    28
blanchet@43473
    29
fun facts_of thy =
wenzelm@44156
    30
  Sledgehammer_Filter.all_facts (Proof_Context.init_global thy) Symtab.empty
blanchet@44217
    31
                                true [] []
blanchet@44438
    32
  |> filter (curry (op =) @{typ bool} o fastype_of
blanchet@44438
    33
             o Object_Logic.atomize_term thy o prop_of o snd)
blanchet@43473
    34
blanchet@44329
    35
(* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few
blanchet@44329
    36
   fixes that seem to be missing over there; or maybe the two code portions are
blanchet@44329
    37
   not doing the same? *)
blanchet@44369
    38
fun fold_body_thms thm_name all_names f =
blanchet@43473
    39
  let
blanchet@44329
    40
    fun app n (PBody {thms, ...}) =
blanchet@44350
    41
      thms |> fold (fn (_, (name, prop, body)) => fn x =>
blanchet@44350
    42
        let
blanchet@44350
    43
          val body' = Future.join body
blanchet@44350
    44
          val n' =
blanchet@44350
    45
            n + (if name = "" orelse
blanchet@44350
    46
                    (is_some all_names andalso
blanchet@44369
    47
                     not (member (op =) (the all_names) name)) orelse
blanchet@44369
    48
                    (* uncommon case where the proved theorem occurs twice
blanchet@44369
    49
                       (e.g., "Transitive_Closure.trancl_into_trancl") *)
blanchet@44369
    50
                    n = 1 andalso name = thm_name then
blanchet@44350
    51
                   0
blanchet@44350
    52
                 else
blanchet@44350
    53
                   1)
blanchet@44350
    54
          val x' = x |> n' <= 1 ? app n' body'
blanchet@44350
    55
        in (x' |> n = 1 ? f (name, prop, body')) end)
blanchet@44350
    56
  in fold (app 0) end
blanchet@43473
    57
blanchet@44350
    58
fun theorems_mentioned_in_proof_term all_names thm =
blanchet@43473
    59
  let
blanchet@43473
    60
    fun collect (s, _, _) = if s <> "" then insert (op =) s else I
blanchet@43473
    61
    val names =
blanchet@44369
    62
      [] |> fold_body_thms (Thm.get_name_hint thm) all_names collect
blanchet@44369
    63
                           [Thm.proof_body_of thm]
blanchet@44350
    64
         |> map fact_name_of
blanchet@43473
    65
  in names end
blanchet@43473
    66
blanchet@43473
    67
fun interesting_const_names ctxt =
wenzelm@44156
    68
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43473
    69
    Sledgehammer_Filter.const_names_in_fact thy
blanchet@44428
    70
        (Sledgehammer_Provers.is_built_in_const_for_prover ctxt eN)
blanchet@43473
    71
  end
blanchet@43473
    72
blanchet@43473
    73
fun generate_tptp_graph_file_for_theory ctxt thy file_name =
blanchet@43473
    74
  let
blanchet@43473
    75
    val path = file_name |> Path.explode
blanchet@43473
    76
    val _ = File.write path ""
blanchet@43473
    77
    val axioms = Theory.all_axioms_of thy |> map fst
blanchet@43473
    78
    fun do_thm thm =
blanchet@43473
    79
      let
blanchet@43473
    80
        val name = Thm.get_name_hint thm
blanchet@43473
    81
        val s =
blanchet@43473
    82
          "[" ^ Thm.legacy_get_kind thm ^ "] " ^
blanchet@43473
    83
          (if member (op =) axioms name then "A" else "T") ^ " " ^
blanchet@44350
    84
          prefix fact_prefix (ascii_of name) ^ ": " ^
blanchet@44350
    85
          commas (theorems_mentioned_in_proof_term NONE thm) ^ "; " ^
blanchet@44350
    86
          commas (map (prefix const_prefix o ascii_of)
blanchet@43473
    87
                      (interesting_const_names ctxt (Thm.prop_of thm))) ^ " \n"
blanchet@43473
    88
      in File.append path s end
blanchet@44086
    89
    val thms = facts_of thy |> map snd
blanchet@43473
    90
    val _ = map do_thm thms
blanchet@43473
    91
  in () end
blanchet@43473
    92
blanchet@43473
    93
fun inference_term [] = NONE
blanchet@43473
    94
  | inference_term ss =
blanchet@44350
    95
    ATerm ("inference",
blanchet@44350
    96
           [ATerm ("isabelle", []),
blanchet@44350
    97
            ATerm (tptp_empty_list, []),
blanchet@44350
    98
            ATerm (tptp_empty_list, map (fn s => ATerm (s, [])) ss)])
blanchet@43473
    99
    |> SOME
blanchet@43473
   100
fun inference infers ident =
blanchet@43473
   101
  these (AList.lookup (op =) infers ident) |> inference_term
blanchet@43473
   102
fun add_inferences_to_problem_line infers
blanchet@44350
   103
                                   (Formula (ident, Axiom, phi, NONE, NONE)) =
blanchet@44350
   104
    Formula (ident, Lemma, phi, inference infers ident, NONE)
blanchet@43473
   105
  | add_inferences_to_problem_line _ line = line
blanchet@44867
   106
fun add_inferences_to_problem infers =
blanchet@44867
   107
  map (apsnd (map (add_inferences_to_problem_line infers)))
blanchet@43473
   108
blanchet@44350
   109
fun ident_of_problem_line (Decl (ident, _, _)) = ident
blanchet@44350
   110
  | ident_of_problem_line (Formula (ident, _, _, _, _)) = ident
blanchet@44350
   111
blanchet@44428
   112
fun run_some_atp ctxt problem =
blanchet@44428
   113
  let
blanchet@44428
   114
    val thy = Proof_Context.theory_of ctxt
blanchet@44870
   115
    val prob_file = File.tmp_path (Path.explode "prob.tptp")
blanchet@44428
   116
    val {exec, arguments, proof_delims, known_failures, ...} =
blanchet@44428
   117
      get_atp thy spassN
blanchet@44428
   118
    val _ = problem |> tptp_lines_for_atp_problem FOF
blanchet@44428
   119
                    |> File.write_list prob_file
blanchet@44428
   120
    val command =
blanchet@44428
   121
      File.shell_path (Path.explode (getenv (fst exec) ^ "/" ^ snd exec)) ^
blanchet@44428
   122
      " " ^ arguments ctxt false "" (seconds 1.0) (K []) ^ " " ^
blanchet@44428
   123
      File.shell_path prob_file
blanchet@44428
   124
  in
wenzelm@44721
   125
    TimeLimit.timeLimit (seconds 0.3) Isabelle_System.bash_output command
blanchet@44428
   126
    |> fst
blanchet@44428
   127
    |> extract_tstplike_proof_and_outcome false true proof_delims
blanchet@44428
   128
                                          known_failures
blanchet@44428
   129
    |> snd
blanchet@44428
   130
  end
blanchet@44428
   131
  handle TimeLimit.TimeOut => SOME TimedOut
blanchet@44428
   132
blanchet@44428
   133
val likely_tautology_prefixes =
blanchet@44431
   134
  [@{theory HOL}, @{theory Meson}, @{theory ATP}, @{theory Metis}]
blanchet@44428
   135
  |> map (fact_name_of o Context.theory_name)
blanchet@44428
   136
blanchet@44428
   137
fun is_problem_line_tautology ctxt (Formula (ident, _, phi, _, _)) =
blanchet@44428
   138
    exists (fn prefix => String.isPrefix prefix ident)
blanchet@44428
   139
           likely_tautology_prefixes andalso
blanchet@44428
   140
    is_none (run_some_atp ctxt
blanchet@44428
   141
                 [(factsN, [Formula (ident, Conjecture, phi, NONE, NONE)])])
blanchet@44428
   142
  | is_problem_line_tautology _ _ = false
blanchet@44428
   143
blanchet@44370
   144
structure String_Graph = Graph(type key = string val ord = string_ord);
blanchet@44370
   145
blanchet@44370
   146
fun order_facts ord = sort (ord o pairself ident_of_problem_line)
blanchet@44370
   147
fun order_problem_facts _ [] = []
blanchet@44370
   148
  | order_problem_facts ord ((heading, lines) :: problem) =
blanchet@44370
   149
    if heading = factsN then (heading, order_facts ord lines) :: problem
blanchet@44370
   150
    else (heading, lines) :: order_problem_facts ord problem
blanchet@44370
   151
blanchet@44493
   152
fun generate_tptp_inference_file_for_theory ctxt thy type_enc file_name =
blanchet@43473
   153
  let
blanchet@44350
   154
    val format = FOF
blanchet@44493
   155
    val type_enc = type_enc |> type_enc_from_string
blanchet@43473
   156
    val path = file_name |> Path.explode
blanchet@43473
   157
    val _ = File.write path ""
blanchet@44438
   158
    val facts = facts_of thy
blanchet@43750
   159
    val (atp_problem, _, _, _, _, _, _) =
blanchet@44438
   160
      facts
blanchet@44438
   161
      |> map (fn ((_, loc), th) => ((Thm.get_name_hint th, loc), prop_of th))
blanchet@44692
   162
      |> prepare_atp_problem ctxt format Axiom Axiom type_enc true true
blanchet@44734
   163
                             (rpair [] o map (introduce_combinators ctxt))
blanchet@44734
   164
                             false true [] @{prop False}
blanchet@44428
   165
    val atp_problem =
blanchet@44428
   166
      atp_problem
blanchet@44428
   167
      |> map (apsnd (filter_out (is_problem_line_tautology ctxt)))
blanchet@44438
   168
    val all_names = facts |> map (Thm.get_name_hint o snd)
blanchet@43473
   169
    val infers =
blanchet@44438
   170
      facts |> map (fn (_, th) =>
blanchet@44438
   171
                       (fact_name_of (Thm.get_name_hint th),
blanchet@44438
   172
                        theorems_mentioned_in_proof_term (SOME all_names) th))
blanchet@44350
   173
    val all_atp_problem_names =
blanchet@44350
   174
      atp_problem |> maps (map ident_of_problem_line o snd)
blanchet@43473
   175
    val infers =
blanchet@44370
   176
      infers |> filter (member (op =) all_atp_problem_names o fst)
blanchet@44370
   177
             |> map (apsnd (filter (member (op =) all_atp_problem_names)))
blanchet@44370
   178
    val ordered_names =
blanchet@44370
   179
      String_Graph.empty
blanchet@44370
   180
      |> fold (String_Graph.new_node o rpair ()) all_atp_problem_names
blanchet@44370
   181
      |> fold (fn (to, froms) =>
blanchet@44428
   182
                  fold (fn from => String_Graph.add_edge (from, to)) froms)
blanchet@44428
   183
              infers
blanchet@44370
   184
      |> String_Graph.topological_order
blanchet@44370
   185
    val order_tab =
blanchet@44370
   186
      Symtab.empty
blanchet@44370
   187
      |> fold (Symtab.insert (op =))
blanchet@44370
   188
              (ordered_names ~~ (1 upto length ordered_names))
blanchet@44370
   189
    val name_ord = int_ord o pairself (the o Symtab.lookup order_tab)
blanchet@44370
   190
    val atp_problem =
blanchet@44370
   191
      atp_problem |> add_inferences_to_problem infers
blanchet@44370
   192
                  |> order_problem_facts name_ord
blanchet@44350
   193
    val ss = tptp_lines_for_atp_problem FOF atp_problem
blanchet@43473
   194
    val _ = app (File.append path) ss
blanchet@43473
   195
  in () end
blanchet@43473
   196
blanchet@43473
   197
end;