src/HOL/Tools/Nitpick/nitrox.ML
author blanchet
Fri, 27 May 2011 10:30:08 +0200
changeset 43861 abb5d1f907e4
parent 43802 f30ae82cb62e
child 43863 7d3ce43d9464
permissions -rw-r--r--
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet@42919
     1
(*  Title:      HOL/Tools/Nitpick/nitrox.ML
blanchet@42919
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@42919
     3
    Copyright   2010, 2011
blanchet@42919
     4
blanchet@43800
     5
Finite model generation for TPTP first-order formulas (FOF and CNF) via Nitpick.
blanchet@42919
     6
*)
blanchet@42919
     7
blanchet@42919
     8
signature NITROX =
blanchet@42919
     9
sig
blanchet@42919
    10
  val pick_nits_in_fof_file : string -> string
blanchet@42919
    11
end;
blanchet@42919
    12
blanchet@42919
    13
structure Nitrox : NITROX =
blanchet@42919
    14
struct
blanchet@42919
    15
blanchet@43801
    16
open ATP_Problem
blanchet@43802
    17
open ATP_Proof
blanchet@42919
    18
blanchet@42919
    19
exception SYNTAX of string
blanchet@42919
    20
blanchet@42919
    21
val parse_keyword = Scan.this_string
blanchet@42919
    22
blanchet@43802
    23
fun parse_file_path (c :: ss) =
blanchet@43802
    24
    if c = "'" orelse c = "\"" then
blanchet@43802
    25
      ss |> chop_while (curry (op <>) c) |>> implode ||> tl
blanchet@43802
    26
    else
blanchet@43802
    27
      raise SYNTAX "invalid file path"
blanchet@43802
    28
  | parse_file_path [] = raise SYNTAX "invalid file path"
blanchet@42919
    29
blanchet@42919
    30
fun parse_include x =
blanchet@42919
    31
  let
blanchet@42919
    32
    val (file_name, rest) =
blanchet@42919
    33
      (parse_keyword "include" |-- $$ "(" |-- parse_file_path --| $$ ")"
blanchet@42919
    34
       --| $$ ".") x
blanchet@42919
    35
  in
blanchet@43802
    36
    ((), (file_name |> Path.explode |> File.read
blanchet@43802
    37
                    |> strip_spaces true (K true)
blanchet@43802
    38
                    |> raw_explode) @ rest)
blanchet@42919
    39
  end
blanchet@42919
    40
blanchet@42919
    41
val parse_fof_or_cnf =
blanchet@42919
    42
  (parse_keyword "fof" || parse_keyword "cnf") |-- $$ "(" |--
blanchet@42919
    43
  Scan.many (not_equal ",") |-- $$ "," |--
blanchet@42919
    44
  (parse_keyword "axiom" || parse_keyword "definition"
blanchet@42919
    45
   || parse_keyword "theorem" || parse_keyword "lemma"
blanchet@42919
    46
   || parse_keyword "hypothesis" || parse_keyword "conjecture"
blanchet@42919
    47
   || parse_keyword "negated_conjecture") --| $$ "," -- parse_formula
blanchet@42919
    48
      --| $$ ")" --| $$ "."
blanchet@42919
    49
  >> (fn ("conjecture", phi) => AConn (ANot, [phi]) | (_, phi) => phi)
blanchet@42919
    50
blanchet@42919
    51
val parse_problem =
blanchet@42919
    52
  Scan.repeat parse_include
blanchet@42919
    53
  |-- Scan.repeat (parse_fof_or_cnf --| Scan.repeat parse_include)
blanchet@42919
    54
blanchet@42919
    55
val parse_tptp_problem =
blanchet@42919
    56
  Scan.finite Symbol.stopper
blanchet@42919
    57
      (Scan.error (!! (fn _ => raise SYNTAX "malformed TPTP input")
blanchet@42919
    58
                  parse_problem))
blanchet@43802
    59
  o raw_explode o strip_spaces true (K true)
blanchet@42919
    60
blanchet@42919
    61
val boolT = @{typ bool}
blanchet@42919
    62
val iotaT = @{typ iota}
blanchet@42919
    63
val quantT = (iotaT --> boolT) --> boolT
blanchet@42919
    64
blanchet@42919
    65
fun is_variable s = Char.isUpper (String.sub (s, 0))
blanchet@42919
    66
blanchet@42919
    67
fun hol_term_from_fo_term res_T (ATerm (x, us)) =
blanchet@42919
    68
  let val ts = map (hol_term_from_fo_term iotaT) us in
blanchet@42919
    69
    list_comb ((case x of
blanchet@42919
    70
                  "$true" => @{const_name True}
blanchet@42919
    71
                | "$false" => @{const_name False}
blanchet@42919
    72
                | "=" => @{const_name HOL.eq}
blanchet@42919
    73
                | _ => x, map fastype_of ts ---> res_T)
blanchet@42919
    74
               |> (if is_variable x then Free else Const), ts)
blanchet@42919
    75
  end
blanchet@42919
    76
blanchet@42919
    77
fun hol_prop_from_formula phi =
blanchet@42919
    78
  case phi of
blanchet@42919
    79
    AQuant (_, [], phi') => hol_prop_from_formula phi'
blanchet@43801
    80
  | AQuant (q, (x, _) :: xs, phi') =>
blanchet@42919
    81
    Const (case q of AForall => @{const_name All} | AExists => @{const_name Ex},
blanchet@42919
    82
           quantT)
blanchet@42919
    83
    $ lambda (Free (x, iotaT)) (hol_prop_from_formula (AQuant (q, xs, phi')))
blanchet@42919
    84
  | AConn (ANot, [u']) => HOLogic.mk_not (hol_prop_from_formula u')
blanchet@42919
    85
  | AConn (c, [u1, u2]) =>
blanchet@42919
    86
    pairself hol_prop_from_formula (u1, u2)
blanchet@42919
    87
    |> (case c of
blanchet@42919
    88
          AAnd => HOLogic.mk_conj
blanchet@42919
    89
        | AOr => HOLogic.mk_disj
blanchet@42919
    90
        | AImplies => HOLogic.mk_imp
blanchet@42919
    91
        | AIf => HOLogic.mk_imp o swap
blanchet@42919
    92
        | AIff => HOLogic.mk_eq
blanchet@42919
    93
        | ANotIff => HOLogic.mk_not o HOLogic.mk_eq
blanchet@42919
    94
        | ANot => raise Fail "binary \"ANot\"")
blanchet@42919
    95
  | AConn _ => raise Fail "malformed AConn"
blanchet@43801
    96
  | AAtom u => hol_term_from_fo_term boolT u
blanchet@42919
    97
blanchet@42919
    98
fun mk_all x t = Const (@{const_name All}, quantT) $ lambda x t
blanchet@42919
    99
blanchet@42919
   100
fun close_hol_prop t = fold (mk_all o Free) (Term.add_frees t []) t
blanchet@42919
   101
blanchet@42919
   102
fun pick_nits_in_fof_file file_name =
blanchet@42919
   103
  case parse_tptp_problem (File.read (Path.explode file_name)) of
blanchet@42919
   104
    (_, s :: ss) => raise SYNTAX ("cannot parse " ^ quote (implode (s :: ss)))
blanchet@42919
   105
  | (phis, []) =>
blanchet@42919
   106
    let
blanchet@42919
   107
      val ts = map (HOLogic.mk_Trueprop o close_hol_prop
blanchet@42919
   108
                    o hol_prop_from_formula) phis
blanchet@42919
   109
(*
blanchet@42919
   110
      val _ = warning (PolyML.makestring phis)
blanchet@42919
   111
      val _ = app (warning o Syntax.string_of_term @{context}) ts
blanchet@42919
   112
*)
blanchet@42919
   113
      val state = Proof.init @{context}
blanchet@42919
   114
      val params =
blanchet@43800
   115
        [("card", "1\<emdash>8"),
blanchet@42919
   116
         ("box", "false"),
blanchet@42919
   117
         ("sat_solver", "smart"),
blanchet@42919
   118
         ("max_threads", "1"),
blanchet@42919
   119
         ("batch_size", "10"),
blanchet@42919
   120
         (* ("debug", "true"), *)
blanchet@42919
   121
         ("verbose", "true"),
blanchet@42919
   122
         (* ("overlord", "true"), *)
blanchet@42919
   123
         ("show_consts", "true"),
blanchet@42919
   124
         ("format", "1000"),
blanchet@42919
   125
         ("max_potential", "0"),
blanchet@42919
   126
         (* ("timeout", "240 s"), *)
blanchet@43861
   127
         ("expect", Nitpick.genuineN)]
blanchet@42919
   128
        |> Nitpick_Isar.default_params @{theory}
blanchet@42919
   129
      val auto = false
blanchet@42919
   130
      val i = 1
blanchet@42919
   131
      val n = 1
blanchet@42919
   132
      val step = 0
blanchet@42919
   133
      val subst = []
blanchet@42919
   134
    in
blanchet@42919
   135
      Nitpick.pick_nits_in_term state params auto i n step subst ts
blanchet@42919
   136
                                @{prop False} |> fst
blanchet@42919
   137
    end
blanchet@42919
   138
blanchet@42919
   139
end;