src/HOL/TPTP/atp_theory_export.ML
author blanchet
Wed, 11 Jul 2012 21:43:19 +0200
changeset 49266 6cdcfbddc077
parent 49265 1065c307fafe
child 49314 5e5c6616f0fe
permissions -rw-r--r--
moved most of MaSh exporter code to Sledgehammer
blanchet@47149
     1
(*  Title:      HOL/TPTP/atp_theory_export.ML
blanchet@43473
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@43473
     3
    Copyright   2011
blanchet@43473
     4
blanchet@49249
     5
Export Isabelle theories as first-order TPTP inferences.
blanchet@43473
     6
*)
blanchet@43473
     7
blanchet@47149
     8
signature ATP_THEORY_EXPORT =
blanchet@43473
     9
sig
blanchet@46176
    10
  type atp_format = ATP_Problem.atp_format
blanchet@46176
    11
blanchet@49249
    12
  val generate_atp_inference_file_for_theory :
blanchet@46176
    13
    Proof.context -> theory -> atp_format -> string -> string -> unit
blanchet@43473
    14
end;
blanchet@43473
    15
blanchet@49249
    16
structure ATP_Theory_Export : ATP_THEORY_EXPORT =
blanchet@43473
    17
struct
blanchet@43473
    18
blanchet@44350
    19
open ATP_Problem
blanchet@44428
    20
open ATP_Proof
blanchet@47148
    21
open ATP_Problem_Generate
blanchet@44428
    22
open ATP_Systems
blanchet@43473
    23
blanchet@49249
    24
val fact_name_of = prefix fact_prefix o ascii_of
blanchet@43473
    25
blanchet@43473
    26
fun inference_term [] = NONE
blanchet@43473
    27
  | inference_term ss =
blanchet@49147
    28
    ATerm (("inference", []),
blanchet@49147
    29
           [ATerm (("isabelle", []), []),
blanchet@49147
    30
            ATerm ((tptp_empty_list, []), []),
blanchet@49147
    31
            ATerm ((tptp_empty_list, []),
blanchet@49147
    32
            map (fn s => ATerm ((s, []), [])) ss)])
blanchet@43473
    33
    |> SOME
blanchet@43473
    34
fun inference infers ident =
blanchet@43473
    35
  these (AList.lookup (op =) infers ident) |> inference_term
blanchet@43473
    36
fun add_inferences_to_problem_line infers
blanchet@47234
    37
                                   (Formula (ident, Axiom, phi, NONE, tms)) =
blanchet@47234
    38
    Formula (ident, Lemma, phi, inference infers ident, tms)
blanchet@43473
    39
  | add_inferences_to_problem_line _ line = line
blanchet@44867
    40
fun add_inferences_to_problem infers =
blanchet@44867
    41
  map (apsnd (map (add_inferences_to_problem_line infers)))
blanchet@43473
    42
blanchet@49157
    43
fun ident_of_problem_line (Class_Decl (ident, _, _)) = ident
blanchet@49157
    44
  | ident_of_problem_line (Type_Decl (ident, _, _)) = ident
blanchet@49152
    45
  | ident_of_problem_line (Sym_Decl (ident, _, _)) = ident
blanchet@49157
    46
  | ident_of_problem_line (Class_Memb (ident, _, _, _)) = ident
blanchet@44350
    47
  | ident_of_problem_line (Formula (ident, _, _, _, _)) = ident
blanchet@44350
    48
blanchet@46176
    49
fun run_some_atp ctxt format problem =
blanchet@44428
    50
  let
blanchet@44428
    51
    val thy = Proof_Context.theory_of ctxt
blanchet@47925
    52
    val prob_file = File.tmp_path (Path.explode "prob")
blanchet@49146
    53
    val atp = case format of DFG _ => spassN | _ => eN
blanchet@48469
    54
    val {exec, arguments, proof_delims, known_failures, ...} =
blanchet@48469
    55
      get_atp thy atp ()
blanchet@47909
    56
    val ord = effective_term_order ctxt atp
blanchet@47909
    57
    val _ = problem |> lines_for_atp_problem format ord (K [])
blanchet@47270
    58
                    |> File.write_list prob_file
blanchet@49228
    59
    val path = getenv (List.last (fst exec)) ^ "/" ^ snd exec
blanchet@44428
    60
    val command =
blanchet@49228
    61
      File.shell_path (Path.explode path) ^
blanchet@47909
    62
      " " ^ arguments ctxt false "" (seconds 1.0) (ord, K [], K []) ^ " " ^
blanchet@44428
    63
      File.shell_path prob_file
blanchet@44428
    64
  in
wenzelm@44721
    65
    TimeLimit.timeLimit (seconds 0.3) Isabelle_System.bash_output command
blanchet@44428
    66
    |> fst
blanchet@47909
    67
    |> extract_tstplike_proof_and_outcome false true proof_delims known_failures
blanchet@44428
    68
    |> snd
blanchet@44428
    69
  end
blanchet@44428
    70
  handle TimeLimit.TimeOut => SOME TimedOut
blanchet@44428
    71
blanchet@49232
    72
val tautology_prefixes =
blanchet@44431
    73
  [@{theory HOL}, @{theory Meson}, @{theory ATP}, @{theory Metis}]
blanchet@44428
    74
  |> map (fact_name_of o Context.theory_name)
blanchet@44428
    75
blanchet@46176
    76
fun is_problem_line_tautology ctxt format (Formula (ident, _, phi, _, _)) =
blanchet@44428
    77
    exists (fn prefix => String.isPrefix prefix ident)
blanchet@49232
    78
           tautology_prefixes andalso
blanchet@46176
    79
    is_none (run_some_atp ctxt format
blanchet@47234
    80
                 [(factsN, [Formula (ident, Conjecture, phi, NONE, [])])])
blanchet@46176
    81
  | is_problem_line_tautology _ _ _ = false
blanchet@44428
    82
blanchet@44370
    83
fun order_facts ord = sort (ord o pairself ident_of_problem_line)
blanchet@44370
    84
fun order_problem_facts _ [] = []
blanchet@44370
    85
  | order_problem_facts ord ((heading, lines) :: problem) =
blanchet@44370
    86
    if heading = factsN then (heading, order_facts ord lines) :: problem
blanchet@44370
    87
    else (heading, lines) :: order_problem_facts ord problem
blanchet@44370
    88
blanchet@46176
    89
(* A fairly random selection of types used for monomorphizing. *)
blanchet@46176
    90
val ground_types =
blanchet@46176
    91
  [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}, @{typ bool},
blanchet@46176
    92
   @{typ unit}]
blanchet@46176
    93
blanchet@46176
    94
fun ground_type_for_tvar _ [] tvar =
blanchet@46176
    95
    raise TYPE ("ground_type_for_sorts", [TVar tvar], [])
blanchet@46176
    96
  | ground_type_for_tvar thy (T :: Ts) tvar =
blanchet@46176
    97
    if can (Sign.typ_match thy (TVar tvar, T)) Vartab.empty then T
blanchet@46176
    98
    else ground_type_for_tvar thy Ts tvar
blanchet@46176
    99
blanchet@46176
   100
fun monomorphize_term ctxt t =
blanchet@46176
   101
  let val thy = Proof_Context.theory_of ctxt in
blanchet@46176
   102
    t |> map_types (map_type_tvar (ground_type_for_tvar thy ground_types))
blanchet@46176
   103
    handle TYPE _ => @{prop True}
blanchet@46176
   104
  end
blanchet@46176
   105
blanchet@49249
   106
fun generate_atp_inference_file_for_theory ctxt thy format type_enc file_name =
blanchet@43473
   107
  let
blanchet@47129
   108
    val type_enc = type_enc |> type_enc_from_string Strict
blanchet@46176
   109
                            |> adjust_type_enc format
blanchet@49146
   110
    val mono = not (is_type_enc_polymorphic type_enc)
blanchet@43473
   111
    val path = file_name |> Path.explode
blanchet@43473
   112
    val _ = File.write path ""
blanchet@49266
   113
    val facts = Sledgehammer_Fact.all_facts_of thy
blanchet@46422
   114
    val atp_problem =
blanchet@44438
   115
      facts
blanchet@46176
   116
      |> map (fn ((_, loc), th) =>
blanchet@46176
   117
                 ((Thm.get_name_hint th, loc),
blanchet@46176
   118
                   th |> prop_of |> mono ? monomorphize_term ctxt))
blanchet@48961
   119
      |> prepare_atp_problem ctxt format Axiom type_enc Exporter combsN false
blanchet@48961
   120
                             false true [] @{prop False}
blanchet@46422
   121
      |> #1
blanchet@44428
   122
    val atp_problem =
blanchet@44428
   123
      atp_problem
blanchet@46176
   124
      |> map (apsnd (filter_out (is_problem_line_tautology ctxt format)))
blanchet@49232
   125
    val ths = facts |> map snd
blanchet@49244
   126
    val all_names = ths |> map Thm.get_name_hint
blanchet@43473
   127
    val infers =
blanchet@44438
   128
      facts |> map (fn (_, th) =>
blanchet@44438
   129
                       (fact_name_of (Thm.get_name_hint th),
blanchet@49266
   130
                        th |> Sledgehammer_Util.thms_in_proof (SOME all_names)
blanchet@49249
   131
                           |> map fact_name_of))
blanchet@44350
   132
    val all_atp_problem_names =
blanchet@44350
   133
      atp_problem |> maps (map ident_of_problem_line o snd)
blanchet@43473
   134
    val infers =
blanchet@44370
   135
      infers |> filter (member (op =) all_atp_problem_names o fst)
blanchet@44370
   136
             |> map (apsnd (filter (member (op =) all_atp_problem_names)))
blanchet@44370
   137
    val ordered_names =
blanchet@44370
   138
      String_Graph.empty
blanchet@44370
   139
      |> fold (String_Graph.new_node o rpair ()) all_atp_problem_names
blanchet@44370
   140
      |> fold (fn (to, froms) =>
blanchet@44428
   141
                  fold (fn from => String_Graph.add_edge (from, to)) froms)
blanchet@44428
   142
              infers
blanchet@44370
   143
      |> String_Graph.topological_order
blanchet@44370
   144
    val order_tab =
blanchet@44370
   145
      Symtab.empty
blanchet@44370
   146
      |> fold (Symtab.insert (op =))
blanchet@44370
   147
              (ordered_names ~~ (1 upto length ordered_names))
blanchet@44370
   148
    val name_ord = int_ord o pairself (the o Symtab.lookup order_tab)
blanchet@44370
   149
    val atp_problem =
blanchet@46176
   150
      atp_problem
blanchet@49146
   151
      |> (case format of DFG _ => I | _ => add_inferences_to_problem infers)
blanchet@46176
   152
      |> order_problem_facts name_ord
blanchet@47909
   153
    val ord = effective_term_order ctxt eN (* dummy *)
blanchet@47909
   154
    val ss = lines_for_atp_problem format ord (K []) atp_problem
blanchet@43473
   155
    val _ = app (File.append path) ss
blanchet@43473
   156
  in () end
blanchet@43473
   157
blanchet@43473
   158
end;