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