src/HOL/ex/atp_export.ML
author blanchet
Mon, 27 Jun 2011 13:52:47 +0200
changeset 44428 a818d5a34cca
parent 44370 9ca694caa61b
child 44431 b342cd125533
permissions -rw-r--r--
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
     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 
    33 (* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few
    34    fixes that seem to be missing over there; or maybe the two code portions are
    35    not doing the same? *)
    36 fun fold_body_thms thm_name all_names f =
    37   let
    38     fun app n (PBody {thms, ...}) =
    39       thms |> fold (fn (_, (name, prop, body)) => fn x =>
    40         let
    41           val body' = Future.join body
    42           val n' =
    43             n + (if name = "" orelse
    44                     (is_some all_names andalso
    45                      not (member (op =) (the all_names) name)) orelse
    46                     (* uncommon case where the proved theorem occurs twice
    47                        (e.g., "Transitive_Closure.trancl_into_trancl") *)
    48                     n = 1 andalso name = thm_name then
    49                    0
    50                  else
    51                    1)
    52           val x' = x |> n' <= 1 ? app n' body'
    53         in (x' |> n = 1 ? f (name, prop, body')) end)
    54   in fold (app 0) end
    55 
    56 fun theorems_mentioned_in_proof_term all_names thm =
    57   let
    58     fun collect (s, _, _) = if s <> "" then insert (op =) s else I
    59     val names =
    60       [] |> fold_body_thms (Thm.get_name_hint thm) all_names collect
    61                            [Thm.proof_body_of thm]
    62          |> map fact_name_of
    63   in names end
    64 
    65 fun interesting_const_names ctxt =
    66   let val thy = Proof_Context.theory_of ctxt in
    67     Sledgehammer_Filter.const_names_in_fact thy
    68         (Sledgehammer_Provers.is_built_in_const_for_prover ctxt eN)
    69   end
    70 
    71 fun generate_tptp_graph_file_for_theory ctxt thy file_name =
    72   let
    73     val path = file_name |> Path.explode
    74     val _ = File.write path ""
    75     val axioms = Theory.all_axioms_of thy |> map fst
    76     fun do_thm thm =
    77       let
    78         val name = Thm.get_name_hint thm
    79         val s =
    80           "[" ^ Thm.legacy_get_kind thm ^ "] " ^
    81           (if member (op =) axioms name then "A" else "T") ^ " " ^
    82           prefix fact_prefix (ascii_of name) ^ ": " ^
    83           commas (theorems_mentioned_in_proof_term NONE thm) ^ "; " ^
    84           commas (map (prefix const_prefix o ascii_of)
    85                       (interesting_const_names ctxt (Thm.prop_of thm))) ^ " \n"
    86       in File.append path s end
    87     val thms = facts_of thy |> map snd
    88     val _ = map do_thm thms
    89   in () end
    90 
    91 fun inference_term [] = NONE
    92   | inference_term ss =
    93     ATerm ("inference",
    94            [ATerm ("isabelle", []),
    95             ATerm (tptp_empty_list, []),
    96             ATerm (tptp_empty_list, map (fn s => ATerm (s, [])) ss)])
    97     |> SOME
    98 fun inference infers ident =
    99   these (AList.lookup (op =) infers ident) |> inference_term
   100 fun add_inferences_to_problem_line infers
   101                                    (Formula (ident, Axiom, phi, NONE, NONE)) =
   102     Formula (ident, Lemma, phi, inference infers ident, NONE)
   103   | add_inferences_to_problem_line _ line = line
   104 val add_inferences_to_problem =
   105   map o apsnd o map o add_inferences_to_problem_line
   106 
   107 fun ident_of_problem_line (Decl (ident, _, _)) = ident
   108   | ident_of_problem_line (Formula (ident, _, _, _, _)) = ident
   109 
   110 fun run_some_atp ctxt problem =
   111   let
   112     val thy = Proof_Context.theory_of ctxt
   113     val prob_file = Path.explode "/tmp/prob.tptp"
   114     val {exec, arguments, proof_delims, known_failures, ...} =
   115       get_atp thy spassN
   116     val _ = problem |> tptp_lines_for_atp_problem FOF
   117                     |> File.write_list prob_file
   118     val command =
   119       File.shell_path (Path.explode (getenv (fst exec) ^ "/" ^ snd exec)) ^
   120       " " ^ arguments ctxt false "" (seconds 1.0) (K []) ^ " " ^
   121       File.shell_path prob_file
   122   in
   123     TimeLimit.timeLimit (seconds 0.15) bash_output command
   124     |> fst
   125     |> extract_tstplike_proof_and_outcome false true proof_delims
   126                                           known_failures
   127     |> snd
   128   end
   129   handle TimeLimit.TimeOut => SOME TimedOut
   130 
   131 val likely_tautology_prefixes =
   132   [@{theory HOL}, @{theory Meson}, @{theory Metis}]
   133   |> map (fact_name_of o Context.theory_name)
   134 
   135 fun is_problem_line_tautology ctxt (Formula (ident, _, phi, _, _)) =
   136     exists (fn prefix => String.isPrefix prefix ident)
   137            likely_tautology_prefixes andalso
   138     is_none (run_some_atp ctxt
   139                  [(factsN, [Formula (ident, Conjecture, phi, NONE, NONE)])])
   140   | is_problem_line_tautology _ _ = false
   141 
   142 structure String_Graph = Graph(type key = string val ord = string_ord);
   143 
   144 fun order_facts ord = sort (ord o pairself ident_of_problem_line)
   145 fun order_problem_facts _ [] = []
   146   | order_problem_facts ord ((heading, lines) :: problem) =
   147     if heading = factsN then (heading, order_facts ord lines) :: problem
   148     else (heading, lines) :: order_problem_facts ord problem
   149 
   150 fun generate_tptp_inference_file_for_theory ctxt thy type_sys file_name =
   151   let
   152     val format = FOF
   153     val type_sys = type_sys |> type_sys_from_string
   154     val path = file_name |> Path.explode
   155     val _ = File.write path ""
   156     val facts0 = facts_of thy
   157     val facts =
   158       facts0 |> map (fn ((_, loc), th) =>
   159                         ((Thm.get_name_hint th, loc), prop_of th))
   160     val (atp_problem, _, _, _, _, _, _) =
   161       prepare_atp_problem ctxt format Axiom Axiom type_sys true false true []
   162                           @{prop False} facts
   163     val atp_problem =
   164       atp_problem
   165       |> map (apsnd (filter_out (is_problem_line_tautology ctxt)))
   166     val all_names = facts0 |> map (Thm.get_name_hint o snd)
   167     val infers =
   168       facts0 |> map (fn (_, th) =>
   169                         (fact_name_of (Thm.get_name_hint th),
   170                          theorems_mentioned_in_proof_term (SOME all_names) th))
   171 
   172     val all_atp_problem_names =
   173       atp_problem |> maps (map ident_of_problem_line o snd)
   174     val infers =
   175       infers |> filter (member (op =) all_atp_problem_names o fst)
   176              |> map (apsnd (filter (member (op =) all_atp_problem_names)))
   177     val ordered_names =
   178       String_Graph.empty
   179       |> fold (String_Graph.new_node o rpair ()) all_atp_problem_names
   180       |> fold (fn (to, froms) =>
   181                   fold (fn from => String_Graph.add_edge (from, to)) froms)
   182               infers
   183       |> String_Graph.topological_order
   184     val order_tab =
   185       Symtab.empty
   186       |> fold (Symtab.insert (op =))
   187               (ordered_names ~~ (1 upto length ordered_names))
   188     val name_ord = int_ord o pairself (the o Symtab.lookup order_tab)
   189     val atp_problem =
   190       atp_problem |> add_inferences_to_problem infers
   191                   |> order_problem_facts name_ord
   192     val ss = tptp_lines_for_atp_problem FOF atp_problem
   193     val _ = app (File.append path) ss
   194   in () end
   195 
   196 end;