src/HOL/TPTP/atp_problem_import.ML
author blanchet
Sun, 22 Apr 2012 14:16:46 +0200
changeset 47670 24babc4b1925
parent 47644 2d90e10f61f2
child 47714 d6683fe037b1
permissions -rw-r--r--
added timeout argument to TPTP tools
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/TPTP/atp_problem_import.ML
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     3
    Copyright   2012
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     4
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     5
Import TPTP problems as Isabelle terms or goals.
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     6
*)
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     7
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     8
signature ATP_PROBLEM_IMPORT =
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
     9
sig
47670
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
    10
  val isabelle_tptp_file : int -> string -> unit
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
    11
  val nitpick_tptp_file : int -> string -> unit
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
    12
  val refute_tptp_file : int -> string -> unit
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
    13
  val sledgehammer_tptp_file : int -> string -> unit
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
    14
  val translate_tptp_file : string -> string -> string -> unit
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    15
end;
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    16
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    17
structure ATP_Problem_Import : ATP_PROBLEM_IMPORT =
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    18
struct
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
    19
47643
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    20
open ATP_Util
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    21
open ATP_Problem
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    22
open ATP_Proof
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    23
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    24
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    25
(** Crude TPTP CNF and FOF parsing (obsolescent) **)
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    26
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    27
exception SYNTAX of string
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    28
exception THF of unit
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    29
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    30
val tptp_explode = raw_explode o strip_spaces_except_between_idents
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    31
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    32
fun parse_file_path (c :: ss) =
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    33
    if c = "'" orelse c = "\"" then
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    34
      ss |> chop_while (curry (op <>) c) |>> implode ||> tl
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    35
    else
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    36
      raise SYNTAX "invalid file path"
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    37
  | parse_file_path [] = raise SYNTAX "invalid file path"
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    38
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    39
fun parse_include x =
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    40
  let
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    41
    val (file_name, rest) =
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    42
      (Scan.this_string "include" |-- $$ "(" |-- parse_file_path --| $$ ")"
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    43
       --| $$ ".") x
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    44
    val path = file_name |> Path.explode
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    45
    val path =
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    46
      path |> not (Path.is_absolute path) ? Path.append (Path.explode "$TPTP")
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    47
  in ((), (path |> File.read |> tptp_explode) @ rest) end
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    48
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    49
val parse_cnf_or_fof =
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    50
  (Scan.this_string "cnf" || Scan.this_string "fof") |-- $$ "(" |--
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    51
  Scan.many (not_equal ",") |-- $$ "," |--
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    52
  (Scan.this_string "axiom" || Scan.this_string "definition"
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    53
   || Scan.this_string "theorem" || Scan.this_string "lemma"
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    54
   || Scan.this_string "hypothesis" || Scan.this_string "conjecture"
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    55
   || Scan.this_string "negated_conjecture") --| $$ "," -- parse_formula
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    56
      --| $$ ")" --| $$ "."
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    57
    >> (fn ("conjecture", phi) => (true, AConn (ANot, [phi]))
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    58
         | (_, phi) => (false, phi))
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    59
  || Scan.this_string "thf" >> (fn _ => raise THF ())
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    60
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    61
val parse_problem =
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    62
  Scan.repeat parse_include
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    63
  |-- Scan.repeat (parse_cnf_or_fof --| Scan.repeat parse_include)
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    64
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    65
val parse_tptp_problem =
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    66
  Scan.finite Symbol.stopper
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    67
      (Scan.error (!! (fn _ => raise SYNTAX "malformed TPTP input")
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    68
                  parse_problem))
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    69
  o tptp_explode
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    70
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    71
val iota_T = @{typ iota}
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    72
val quant_T = @{typ "(iota => bool) => bool"}
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    73
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    74
fun is_variable s = Char.isUpper (String.sub (s, 0))
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    75
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    76
fun hol_term_from_fo_term res_T (ATerm (x, us)) =
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    77
  let val ts = map (hol_term_from_fo_term iota_T) us in
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    78
    list_comb ((case x of
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    79
                  "$true" => @{const_name True}
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    80
                | "$false" => @{const_name False}
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    81
                | "=" => @{const_name HOL.eq}
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    82
                | "equal" => @{const_name HOL.eq}
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    83
                | _ => x, map fastype_of ts ---> res_T)
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    84
               |> (if is_variable x then Free else Const), ts)
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    85
  end
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    86
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    87
fun hol_prop_from_formula phi =
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    88
  case phi of
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    89
    AQuant (_, [], phi') => hol_prop_from_formula phi'
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    90
  | AQuant (q, (x, _) :: xs, phi') =>
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    91
    Const (case q of AForall => @{const_name All} | AExists => @{const_name Ex},
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    92
           quant_T)
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    93
    $ lambda (Free (x, iota_T)) (hol_prop_from_formula (AQuant (q, xs, phi')))
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    94
  | AConn (ANot, [u']) => HOLogic.mk_not (hol_prop_from_formula u')
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    95
  | AConn (c, [u1, u2]) =>
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    96
    pairself hol_prop_from_formula (u1, u2)
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    97
    |> (case c of
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    98
          AAnd => HOLogic.mk_conj
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
    99
        | AOr => HOLogic.mk_disj
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   100
        | AImplies => HOLogic.mk_imp
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   101
        | AIff => HOLogic.mk_eq
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   102
        | ANot => raise Fail "binary \"ANot\"")
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   103
  | AConn _ => raise Fail "malformed AConn"
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   104
  | AAtom u => hol_term_from_fo_term @{typ bool} u
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   105
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   106
fun mk_all x t = Const (@{const_name All}, quant_T) $ lambda x t
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   107
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   108
fun close_hol_prop t = fold (mk_all o Free) (Term.add_frees t []) t
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   109
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   110
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   111
(** TPTP parsing based on ML-Yacc-generated "TPTP_Parser" **)
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   112
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   113
fun mk_meta_not P = Logic.mk_implies (P, @{prop False})
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   114
47565
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
   115
(* cf. "close_form" in "refute.ML" *)
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
   116
fun close_form t =
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
   117
  fold (fn ((s, i), T) => fn t' =>
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
   118
           Logic.all_const T $ Abs (s, T, abstract_over (Var ((s, i), T), t')))
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
   119
       (Term.add_vars t []) t
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   120
47562
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   121
fun get_tptp_formula (_, role, P, _) =
47565
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
   122
  P |> Logic.varify_global |> close_form
762eb0dacaa6 remove old TPTP CNF/FOF parser; always use Nik's new parser
blanchet
parents: 47562
diff changeset
   123
    |> role = TPTP_Syntax.Role_Conjecture ? mk_meta_not
47557
32f35b3d9e42 started integrating Nik's parser into TPTP command-line tools
blanchet
parents: 46325
diff changeset
   124
32f35b3d9e42 started integrating Nik's parser into TPTP command-line tools
blanchet
parents: 46325
diff changeset
   125
fun read_tptp_file thy file_name =
47644
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
   126
  let
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
   127
    val path =
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
   128
      Path.explode file_name
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
   129
      |> (fn path =>
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
   130
             path |> not (Path.is_absolute path)
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
   131
                     ? Path.append (Path.explode "$PWD"))
2d90e10f61f2 prepend PWD to relative paths
blanchet
parents: 47643
diff changeset
   132
  in
47643
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   133
    (case parse_tptp_problem (File.read path) of
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   134
       (_, s :: ss) =>
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   135
       raise SYNTAX ("cannot parse " ^ quote (implode (s :: ss)))
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   136
     | (problem, []) =>
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   137
       (exists fst problem,
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   138
        map (HOLogic.mk_Trueprop o close_hol_prop o hol_prop_from_formula o snd)
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   139
            problem))
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   140
    handle THF () =>
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   141
    let
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   142
      val problem =
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   143
        TPTP_Interpret.interpret_file true (Path.explode "$TPTP") path [] [] thy
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   144
        |> fst |> #3
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   145
    in
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   146
      (exists (fn (_, role, _, _) => role = TPTP_Syntax.Role_Conjecture) problem,
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   147
       map get_tptp_formula problem)
e33c2be488fe reintroduced old FOF and CNF parsers, to work around TPTP_Parser failures
blanchet
parents: 47565
diff changeset
   148
    end
47557
32f35b3d9e42 started integrating Nik's parser into TPTP command-line tools
blanchet
parents: 46325
diff changeset
   149
  end
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   150
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   151
(** Isabelle (combination of provers) **)
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   152
47670
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
   153
fun isabelle_tptp_file timeout file_name = ()
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   154
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   155
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   156
(** Nitpick (alias Nitrox) **)
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   157
47670
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
   158
fun nitpick_tptp_file timeout file_name =
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   159
  let
47562
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   160
    val (falsify, ts) = read_tptp_file @{theory} file_name
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   161
    val state = Proof.init @{context}
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   162
    val params =
47557
32f35b3d9e42 started integrating Nik's parser into TPTP command-line tools
blanchet
parents: 46325
diff changeset
   163
      [("card", "1\<emdash>100"),
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   164
       ("box", "false"),
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   165
       ("sat_solver", "smart"),
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   166
       ("max_threads", "1"),
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   167
       ("batch_size", "10"),
47562
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   168
       ("falsify", if falsify then "true" else "false"),
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   169
       (* ("debug", "true"), *)
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   170
       ("verbose", "true"),
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   171
       (* ("overlord", "true"), *)
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   172
       ("show_consts", "true"),
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   173
       ("format", "1000"),
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   174
       ("max_potential", "0"),
47670
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
   175
       ("timeout", string_of_int timeout)]
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   176
      |> Nitpick_Isar.default_params @{theory}
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   177
    val i = 1
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   178
    val n = 1
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   179
    val step = 0
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   180
    val subst = []
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   181
  in
47562
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   182
    Nitpick.pick_nits_in_term state params Nitpick.TPTP i n step subst ts
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   183
        (if falsify then @{prop False} else @{prop True}); ()
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   184
  end
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   185
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   186
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   187
(** Refute **)
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   188
47562
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   189
fun print_szs_from_outcome falsify s =
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   190
  "% SZS status " ^
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   191
  (if s = "genuine" then
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   192
     if falsify then "CounterSatisfiable" else "Satisfiable"
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   193
   else
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   194
     "Unknown")
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   195
  |> writeln
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   196
47670
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
   197
fun refute_tptp_file timeout file_name =
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   198
  let
47562
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   199
    val (falsify, ts) = read_tptp_file @{theory} file_name
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   200
    val params =
47670
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
   201
      [("maxtime", string_of_int timeout)]
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   202
  in
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   203
    Refute.refute_term @{context} params ts @{prop False}
47562
a72239723ae8 tuned SZS status output
blanchet
parents: 47557
diff changeset
   204
    |> print_szs_from_outcome falsify
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   205
  end
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   206
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   207
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   208
(** Sledgehammer **)
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   209
47670
24babc4b1925 added timeout argument to TPTP tools
blanchet
parents: 47644
diff changeset
   210
fun sledgehammer_tptp_file timeout file_name = ()
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   211
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   212
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   213
(** Translator between TPTP(-like) file formats **)
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   214
46325
b170ab46513a implemented "tptp_refute" tool
blanchet
parents: 46324
diff changeset
   215
fun translate_tptp_file format in_file_name out_file_name = ()
46324
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   216
e4bccf5ec61e added problem importer
blanchet
parents:
diff changeset
   217
end;