src/HOL/ex/atp_export.ML
author blanchet
Tue, 07 Jun 2011 07:04:53 +0200
changeset 44064 c9e87dc92d9e
parent 44063 src/HOL/ex/tptp_export.ML@d90151a666cc
child 44065 97906dfd39b7
permissions -rw-r--r--
renamed example theory to "ATP_Export", for consistency with its underlying "ATP_" modules
blanchet@44064
     1
(*  Title:      HOL/ex/atp_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@44064
     9
signature ATP_EXPORT =
blanchet@43473
    10
sig
blanchet@43473
    11
  val generate_tptp_graph_file_for_theory :
blanchet@43473
    12
    Proof.context -> theory -> string -> unit
blanchet@43473
    13
  val generate_tptp_inference_file_for_theory :
blanchet@43473
    14
    Proof.context -> theory -> bool -> string -> unit
blanchet@43473
    15
end;
blanchet@43473
    16
blanchet@44064
    17
structure ATP_Export : ATP_EXPORT =
blanchet@43473
    18
struct
blanchet@43473
    19
blanchet@43929
    20
val ascii_of = ATP_Translate.ascii_of
blanchet@43473
    21
blanchet@43929
    22
val fact_name_of = prefix ATP_Translate.fact_prefix o ascii_of
blanchet@43473
    23
blanchet@43473
    24
fun facts_of thy =
blanchet@43509
    25
  Sledgehammer_Filter.all_facts (ProofContext.init_global thy) Symtab.empty
blanchet@43785
    26
                                true (K true) [] []
blanchet@43473
    27
blanchet@43473
    28
fun fold_body_thms f =
blanchet@43473
    29
  let
blanchet@43473
    30
    fun app n (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) => fn (x, seen) =>
blanchet@43473
    31
      if Inttab.defined seen i then (x, seen)
blanchet@43473
    32
      else
blanchet@43473
    33
        let
blanchet@43473
    34
          val body' = Future.join body;
blanchet@43473
    35
          val (x', seen') = app (n + (if name = "" then 0 else 1)) body' (x, Inttab.update (i, ()) seen);
blanchet@43473
    36
        in (x' |> n = 1 ? f (name, prop, body'), seen') end);
blanchet@43473
    37
  in fn bodies => fn x => #1 (fold (app 0) bodies (x, Inttab.empty)) end;
blanchet@43473
    38
blanchet@43473
    39
fun theorems_mentioned_in_proof_term thm =
blanchet@43473
    40
  let
blanchet@43473
    41
    fun collect (s, _, _) = if s <> "" then insert (op =) s else I
blanchet@43473
    42
    val names =
blanchet@43473
    43
      [] |> fold_body_thms collect [Thm.proof_body_of thm] |> map fact_name_of
blanchet@43473
    44
  in names end
blanchet@43473
    45
blanchet@43473
    46
fun interesting_const_names ctxt =
blanchet@43473
    47
  let val thy = ProofContext.theory_of ctxt in
blanchet@43473
    48
    Sledgehammer_Filter.const_names_in_fact thy
blanchet@43473
    49
        (Sledgehammer_Provers.is_built_in_const_for_prover ctxt ATP_Systems.eN)
blanchet@43473
    50
  end
blanchet@43473
    51
blanchet@43473
    52
fun generate_tptp_graph_file_for_theory ctxt thy file_name =
blanchet@43473
    53
  let
blanchet@43473
    54
    val path = file_name |> Path.explode
blanchet@43473
    55
    val _ = File.write path ""
blanchet@43473
    56
    val axioms = Theory.all_axioms_of thy |> map fst
blanchet@43473
    57
    fun do_thm thm =
blanchet@43473
    58
      let
blanchet@43473
    59
        val name = Thm.get_name_hint thm
blanchet@43473
    60
        val s =
blanchet@43473
    61
          "[" ^ Thm.legacy_get_kind thm ^ "] " ^
blanchet@43473
    62
          (if member (op =) axioms name then "A" else "T") ^ " " ^
blanchet@43929
    63
          prefix ATP_Translate.fact_prefix (ascii_of name) ^ ": " ^
blanchet@43473
    64
          commas (theorems_mentioned_in_proof_term thm) ^ "; " ^
blanchet@43929
    65
          commas (map (prefix ATP_Translate.const_prefix o ascii_of)
blanchet@43473
    66
                      (interesting_const_names ctxt (Thm.prop_of thm))) ^ " \n"
blanchet@43473
    67
      in File.append path s end
blanchet@43473
    68
    val thms = facts_of thy |> map (snd o snd)
blanchet@43473
    69
    val _ = map do_thm thms
blanchet@43473
    70
  in () end
blanchet@43473
    71
blanchet@43473
    72
fun inference_term [] = NONE
blanchet@43473
    73
  | inference_term ss =
blanchet@43473
    74
    ATP_Problem.ATerm ("inference",
blanchet@43473
    75
           [ATP_Problem.ATerm ("isabelle", []),
blanchet@43473
    76
            ATP_Problem.ATerm ("[]", []),
blanchet@43473
    77
            ATP_Problem.ATerm ("[]",
blanchet@43473
    78
                map (fn s => ATP_Problem.ATerm (s, [])) ss)])
blanchet@43473
    79
    |> SOME
blanchet@43473
    80
fun inference infers ident =
blanchet@43473
    81
  these (AList.lookup (op =) infers ident) |> inference_term
blanchet@43473
    82
fun add_inferences_to_problem_line infers
blanchet@43473
    83
        (ATP_Problem.Formula (ident, ATP_Problem.Axiom, phi, NONE, NONE)) =
blanchet@43473
    84
    ATP_Problem.Formula (ident, ATP_Problem.Lemma, phi, inference infers ident,
blanchet@43473
    85
                         NONE)
blanchet@43473
    86
  | add_inferences_to_problem_line _ line = line
blanchet@43473
    87
val add_inferences_to_problem =
blanchet@43473
    88
  map o apsnd o map o add_inferences_to_problem_line
blanchet@43473
    89
blanchet@43473
    90
fun generate_tptp_inference_file_for_theory ctxt thy full_types file_name =
blanchet@43473
    91
  let
blanchet@43803
    92
    val format = ATP_Problem.FOF
blanchet@43835
    93
    val type_sys =
blanchet@43929
    94
      ATP_Translate.Preds
blanchet@43929
    95
          (ATP_Translate.Polymorphic,
blanchet@43929
    96
           if full_types then ATP_Translate.All_Types
blanchet@43929
    97
           else ATP_Translate.Const_Arg_Types,
blanchet@43969
    98
           ATP_Translate.Heavyweight)
blanchet@43473
    99
    val path = file_name |> Path.explode
blanchet@43473
   100
    val _ = File.write path ""
blanchet@43473
   101
    val facts0 = facts_of thy
blanchet@43473
   102
    val facts =
blanchet@44055
   103
      facts0 |> map (fn ((_, loc), (_, th)) =>
blanchet@44063
   104
                        ((Thm.get_name_hint th, loc), prop_of th))
blanchet@43905
   105
    val explicit_apply = NONE
blanchet@43750
   106
    val (atp_problem, _, _, _, _, _, _) =
blanchet@43929
   107
      ATP_Translate.prepare_atp_problem ctxt format
blanchet@43951
   108
          ATP_Problem.Axiom ATP_Problem.Axiom type_sys explicit_apply false true
blanchet@43951
   109
          [] @{prop False} facts
blanchet@43473
   110
    val infers =
blanchet@43473
   111
      facts0 |> map (fn (_, (_, th)) =>
blanchet@43473
   112
                        (fact_name_of (Thm.get_name_hint th),
blanchet@43473
   113
                         theorems_mentioned_in_proof_term th))
blanchet@43473
   114
    val infers =
blanchet@43473
   115
      infers |> map (apsnd (filter (member (op = o apsnd fst) infers)))
blanchet@43473
   116
    val atp_problem = atp_problem |> add_inferences_to_problem infers
blanchet@43473
   117
    val ss =
blanchet@43778
   118
      ATP_Problem.tptp_strings_for_atp_problem ATP_Problem.FOF atp_problem
blanchet@43473
   119
    val _ = app (File.append path) ss
blanchet@43473
   120
  in () end
blanchet@43473
   121
blanchet@43473
   122
end;