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