src/HOL/TPTP/atp_theory_export.ML
author blanchet
Tue, 26 Jun 2012 11:14:39 +0200
changeset 49145 defbcdc60fd6
parent 49144 933d43c31689
child 49146 1016664b8feb
permissions -rw-r--r--
tuning
blanchet@47149
     1
(*  Title:      HOL/TPTP/atp_theory_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@47149
     9
signature ATP_THEORY_EXPORT =
blanchet@43473
    10
sig
blanchet@46176
    11
  type atp_format = ATP_Problem.atp_format
blanchet@46176
    12
blanchet@44365
    13
  val theorems_mentioned_in_proof_term :
blanchet@44365
    14
    string list option -> thm -> string list
blanchet@43473
    15
  val generate_tptp_graph_file_for_theory :
blanchet@43473
    16
    Proof.context -> theory -> string -> unit
blanchet@43473
    17
  val generate_tptp_inference_file_for_theory :
blanchet@46176
    18
    Proof.context -> theory -> atp_format -> string -> string -> unit
blanchet@43473
    19
end;
blanchet@43473
    20
blanchet@47149
    21
structure ATP_Theory_Export : ATP_THEORY_EXPORT =
blanchet@43473
    22
struct
blanchet@43473
    23
blanchet@44350
    24
open ATP_Problem
blanchet@44428
    25
open ATP_Proof
blanchet@47148
    26
open ATP_Problem_Generate
blanchet@44428
    27
open ATP_Systems
blanchet@43473
    28
blanchet@44350
    29
val fact_name_of = prefix fact_prefix o ascii_of
blanchet@43473
    30
blanchet@43473
    31
fun facts_of thy =
blanchet@47601
    32
  let val ctxt = Proof_Context.init_global thy in
blanchet@47601
    33
    Sledgehammer_Filter.all_facts ctxt false
blanchet@47601
    34
        Symtab.empty true [] []
blanchet@47601
    35
        (Sledgehammer_Filter.clasimpset_rule_table_of ctxt)
blanchet@47601
    36
    |> filter (curry (op =) @{typ bool} o fastype_of
blanchet@47601
    37
               o Object_Logic.atomize_term thy o prop_of o snd)
blanchet@47601
    38
  end
blanchet@43473
    39
blanchet@44329
    40
(* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few
blanchet@44329
    41
   fixes that seem to be missing over there; or maybe the two code portions are
blanchet@44329
    42
   not doing the same? *)
blanchet@44369
    43
fun fold_body_thms thm_name all_names f =
blanchet@43473
    44
  let
blanchet@44329
    45
    fun app n (PBody {thms, ...}) =
blanchet@44350
    46
      thms |> fold (fn (_, (name, prop, body)) => fn x =>
blanchet@44350
    47
        let
blanchet@44350
    48
          val body' = Future.join body
blanchet@44350
    49
          val n' =
blanchet@44350
    50
            n + (if name = "" orelse
blanchet@44350
    51
                    (is_some all_names andalso
blanchet@44369
    52
                     not (member (op =) (the all_names) name)) orelse
blanchet@44369
    53
                    (* uncommon case where the proved theorem occurs twice
blanchet@44369
    54
                       (e.g., "Transitive_Closure.trancl_into_trancl") *)
blanchet@44369
    55
                    n = 1 andalso name = thm_name then
blanchet@44350
    56
                   0
blanchet@44350
    57
                 else
blanchet@44350
    58
                   1)
blanchet@44350
    59
          val x' = x |> n' <= 1 ? app n' body'
blanchet@44350
    60
        in (x' |> n = 1 ? f (name, prop, body')) end)
blanchet@44350
    61
  in fold (app 0) end
blanchet@43473
    62
blanchet@44350
    63
fun theorems_mentioned_in_proof_term all_names thm =
blanchet@43473
    64
  let
blanchet@43473
    65
    fun collect (s, _, _) = if s <> "" then insert (op =) s else I
blanchet@43473
    66
    val names =
blanchet@44369
    67
      [] |> fold_body_thms (Thm.get_name_hint thm) all_names collect
blanchet@44369
    68
                           [Thm.proof_body_of thm]
blanchet@44350
    69
         |> map fact_name_of
blanchet@43473
    70
  in names end
blanchet@43473
    71
blanchet@43473
    72
fun interesting_const_names ctxt =
wenzelm@44156
    73
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43473
    74
    Sledgehammer_Filter.const_names_in_fact thy
blanchet@44428
    75
        (Sledgehammer_Provers.is_built_in_const_for_prover ctxt eN)
blanchet@43473
    76
  end
blanchet@43473
    77
blanchet@43473
    78
fun generate_tptp_graph_file_for_theory ctxt thy file_name =
blanchet@43473
    79
  let
blanchet@43473
    80
    val path = file_name |> Path.explode
blanchet@43473
    81
    val _ = File.write path ""
blanchet@43473
    82
    val axioms = Theory.all_axioms_of thy |> map fst
blanchet@43473
    83
    fun do_thm thm =
blanchet@43473
    84
      let
blanchet@43473
    85
        val name = Thm.get_name_hint thm
blanchet@43473
    86
        val s =
blanchet@43473
    87
          "[" ^ Thm.legacy_get_kind thm ^ "] " ^
blanchet@43473
    88
          (if member (op =) axioms name then "A" else "T") ^ " " ^
blanchet@44350
    89
          prefix fact_prefix (ascii_of name) ^ ": " ^
blanchet@44350
    90
          commas (theorems_mentioned_in_proof_term NONE thm) ^ "; " ^
blanchet@44350
    91
          commas (map (prefix const_prefix o ascii_of)
blanchet@43473
    92
                      (interesting_const_names ctxt (Thm.prop_of thm))) ^ " \n"
blanchet@43473
    93
      in File.append path s end
blanchet@44086
    94
    val thms = facts_of thy |> map snd
blanchet@43473
    95
    val _ = map do_thm thms
blanchet@43473
    96
  in () end
blanchet@43473
    97
blanchet@43473
    98
fun inference_term [] = NONE
blanchet@43473
    99
  | inference_term ss =
blanchet@44350
   100
    ATerm ("inference",
blanchet@44350
   101
           [ATerm ("isabelle", []),
blanchet@44350
   102
            ATerm (tptp_empty_list, []),
blanchet@44350
   103
            ATerm (tptp_empty_list, map (fn s => ATerm (s, [])) ss)])
blanchet@43473
   104
    |> SOME
blanchet@43473
   105
fun inference infers ident =
blanchet@43473
   106
  these (AList.lookup (op =) infers ident) |> inference_term
blanchet@43473
   107
fun add_inferences_to_problem_line infers
blanchet@47234
   108
                                   (Formula (ident, Axiom, phi, NONE, tms)) =
blanchet@47234
   109
    Formula (ident, Lemma, phi, inference infers ident, tms)
blanchet@43473
   110
  | add_inferences_to_problem_line _ line = line
blanchet@44867
   111
fun add_inferences_to_problem infers =
blanchet@44867
   112
  map (apsnd (map (add_inferences_to_problem_line infers)))
blanchet@43473
   113
blanchet@44350
   114
fun ident_of_problem_line (Decl (ident, _, _)) = ident
blanchet@44350
   115
  | ident_of_problem_line (Formula (ident, _, _, _, _)) = ident
blanchet@44350
   116
blanchet@46176
   117
fun run_some_atp ctxt format problem =
blanchet@44428
   118
  let
blanchet@44428
   119
    val thy = Proof_Context.theory_of ctxt
blanchet@47925
   120
    val prob_file = File.tmp_path (Path.explode "prob")
blanchet@49144
   121
    val atp = if format = DFG then spassN else eN
blanchet@48469
   122
    val {exec, arguments, proof_delims, known_failures, ...} =
blanchet@48469
   123
      get_atp thy atp ()
blanchet@47909
   124
    val ord = effective_term_order ctxt atp
blanchet@47909
   125
    val _ = problem |> lines_for_atp_problem format ord (K [])
blanchet@47270
   126
                    |> File.write_list prob_file
blanchet@44428
   127
    val command =
blanchet@47925
   128
      File.shell_path (Path.explode (getenv (hd (fst exec)) ^ "/" ^ snd exec)) ^
blanchet@47909
   129
      " " ^ arguments ctxt false "" (seconds 1.0) (ord, K [], K []) ^ " " ^
blanchet@44428
   130
      File.shell_path prob_file
blanchet@44428
   131
  in
wenzelm@44721
   132
    TimeLimit.timeLimit (seconds 0.3) Isabelle_System.bash_output command
blanchet@44428
   133
    |> fst
blanchet@47909
   134
    |> extract_tstplike_proof_and_outcome false true proof_delims known_failures
blanchet@44428
   135
    |> snd
blanchet@44428
   136
  end
blanchet@44428
   137
  handle TimeLimit.TimeOut => SOME TimedOut
blanchet@44428
   138
blanchet@44428
   139
val likely_tautology_prefixes =
blanchet@44431
   140
  [@{theory HOL}, @{theory Meson}, @{theory ATP}, @{theory Metis}]
blanchet@44428
   141
  |> map (fact_name_of o Context.theory_name)
blanchet@44428
   142
blanchet@46176
   143
fun is_problem_line_tautology ctxt format (Formula (ident, _, phi, _, _)) =
blanchet@44428
   144
    exists (fn prefix => String.isPrefix prefix ident)
blanchet@44428
   145
           likely_tautology_prefixes andalso
blanchet@46176
   146
    is_none (run_some_atp ctxt format
blanchet@47234
   147
                 [(factsN, [Formula (ident, Conjecture, phi, NONE, [])])])
blanchet@46176
   148
  | is_problem_line_tautology _ _ _ = false
blanchet@44428
   149
blanchet@44370
   150
fun order_facts ord = sort (ord o pairself ident_of_problem_line)
blanchet@44370
   151
fun order_problem_facts _ [] = []
blanchet@44370
   152
  | order_problem_facts ord ((heading, lines) :: problem) =
blanchet@44370
   153
    if heading = factsN then (heading, order_facts ord lines) :: problem
blanchet@44370
   154
    else (heading, lines) :: order_problem_facts ord problem
blanchet@44370
   155
blanchet@46176
   156
(* A fairly random selection of types used for monomorphizing. *)
blanchet@46176
   157
val ground_types =
blanchet@46176
   158
  [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}, @{typ bool},
blanchet@46176
   159
   @{typ unit}]
blanchet@46176
   160
blanchet@46176
   161
fun ground_type_for_tvar _ [] tvar =
blanchet@46176
   162
    raise TYPE ("ground_type_for_sorts", [TVar tvar], [])
blanchet@46176
   163
  | ground_type_for_tvar thy (T :: Ts) tvar =
blanchet@46176
   164
    if can (Sign.typ_match thy (TVar tvar, T)) Vartab.empty then T
blanchet@46176
   165
    else ground_type_for_tvar thy Ts tvar
blanchet@46176
   166
blanchet@46176
   167
fun monomorphize_term ctxt t =
blanchet@46176
   168
  let val thy = Proof_Context.theory_of ctxt in
blanchet@46176
   169
    t |> map_types (map_type_tvar (ground_type_for_tvar thy ground_types))
blanchet@46176
   170
    handle TYPE _ => @{prop True}
blanchet@46176
   171
  end
blanchet@46176
   172
blanchet@46176
   173
fun generate_tptp_inference_file_for_theory ctxt thy format type_enc file_name =
blanchet@43473
   174
  let
blanchet@47129
   175
    val type_enc = type_enc |> type_enc_from_string Strict
blanchet@46176
   176
                            |> adjust_type_enc format
blanchet@49145
   177
    val mono = polymorphism_of_type_enc type_enc <> Raw_Polymorphic
blanchet@43473
   178
    val path = file_name |> Path.explode
blanchet@43473
   179
    val _ = File.write path ""
blanchet@44438
   180
    val facts = facts_of thy
blanchet@46422
   181
    val atp_problem =
blanchet@44438
   182
      facts
blanchet@46176
   183
      |> map (fn ((_, loc), th) =>
blanchet@46176
   184
                 ((Thm.get_name_hint th, loc),
blanchet@46176
   185
                   th |> prop_of |> mono ? monomorphize_term ctxt))
blanchet@48961
   186
      |> prepare_atp_problem ctxt format Axiom type_enc Exporter combsN false
blanchet@48961
   187
                             false true [] @{prop False}
blanchet@46422
   188
      |> #1
blanchet@44428
   189
    val atp_problem =
blanchet@44428
   190
      atp_problem
blanchet@46176
   191
      |> map (apsnd (filter_out (is_problem_line_tautology ctxt format)))
blanchet@44438
   192
    val all_names = facts |> map (Thm.get_name_hint o snd)
blanchet@43473
   193
    val infers =
blanchet@44438
   194
      facts |> map (fn (_, th) =>
blanchet@44438
   195
                       (fact_name_of (Thm.get_name_hint th),
blanchet@44438
   196
                        theorems_mentioned_in_proof_term (SOME all_names) th))
blanchet@44350
   197
    val all_atp_problem_names =
blanchet@44350
   198
      atp_problem |> maps (map ident_of_problem_line o snd)
blanchet@43473
   199
    val infers =
blanchet@44370
   200
      infers |> filter (member (op =) all_atp_problem_names o fst)
blanchet@44370
   201
             |> map (apsnd (filter (member (op =) all_atp_problem_names)))
blanchet@44370
   202
    val ordered_names =
blanchet@44370
   203
      String_Graph.empty
blanchet@44370
   204
      |> fold (String_Graph.new_node o rpair ()) all_atp_problem_names
blanchet@44370
   205
      |> fold (fn (to, froms) =>
blanchet@44428
   206
                  fold (fn from => String_Graph.add_edge (from, to)) froms)
blanchet@44428
   207
              infers
blanchet@44370
   208
      |> String_Graph.topological_order
blanchet@44370
   209
    val order_tab =
blanchet@44370
   210
      Symtab.empty
blanchet@44370
   211
      |> fold (Symtab.insert (op =))
blanchet@44370
   212
              (ordered_names ~~ (1 upto length ordered_names))
blanchet@44370
   213
    val name_ord = int_ord o pairself (the o Symtab.lookup order_tab)
blanchet@44370
   214
    val atp_problem =
blanchet@46176
   215
      atp_problem
blanchet@49144
   216
      |> (format <> DFG ? add_inferences_to_problem infers)
blanchet@46176
   217
      |> order_problem_facts name_ord
blanchet@47909
   218
    val ord = effective_term_order ctxt eN (* dummy *)
blanchet@47909
   219
    val ss = lines_for_atp_problem format ord (K []) atp_problem
blanchet@43473
   220
    val _ = app (File.append path) ss
blanchet@43473
   221
  in () end
blanchet@43473
   222
blanchet@43473
   223
end;