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