src/HOL/TPTP/atp_problem_import.ML
changeset 47714 d6683fe037b1
parent 47670 24babc4b1925
child 47715 04400144c6fc
equal deleted inserted replaced
47713:bd0683000a0f 47714:d6683fe037b1
    20 open ATP_Util
    20 open ATP_Util
    21 open ATP_Problem
    21 open ATP_Problem
    22 open ATP_Proof
    22 open ATP_Proof
    23 
    23 
    24 
    24 
    25 (** Crude TPTP CNF and FOF parsing (obsolescent) **)
       
    26 
       
    27 exception SYNTAX of string
       
    28 exception THF of unit
       
    29 
       
    30 val tptp_explode = raw_explode o strip_spaces_except_between_idents
       
    31 
       
    32 fun parse_file_path (c :: ss) =
       
    33     if c = "'" orelse c = "\"" then
       
    34       ss |> chop_while (curry (op <>) c) |>> implode ||> tl
       
    35     else
       
    36       raise SYNTAX "invalid file path"
       
    37   | parse_file_path [] = raise SYNTAX "invalid file path"
       
    38 
       
    39 fun parse_include x =
       
    40   let
       
    41     val (file_name, rest) =
       
    42       (Scan.this_string "include" |-- $$ "(" |-- parse_file_path --| $$ ")"
       
    43        --| $$ ".") x
       
    44     val path = file_name |> Path.explode
       
    45     val path =
       
    46       path |> not (Path.is_absolute path) ? Path.append (Path.explode "$TPTP")
       
    47   in ((), (path |> File.read |> tptp_explode) @ rest) end
       
    48 
       
    49 val parse_cnf_or_fof =
       
    50   (Scan.this_string "cnf" || Scan.this_string "fof") |-- $$ "(" |--
       
    51   Scan.many (not_equal ",") |-- $$ "," |--
       
    52   (Scan.this_string "axiom" || Scan.this_string "definition"
       
    53    || Scan.this_string "theorem" || Scan.this_string "lemma"
       
    54    || Scan.this_string "hypothesis" || Scan.this_string "conjecture"
       
    55    || Scan.this_string "negated_conjecture") --| $$ "," -- parse_formula
       
    56       --| $$ ")" --| $$ "."
       
    57     >> (fn ("conjecture", phi) => (true, AConn (ANot, [phi]))
       
    58          | (_, phi) => (false, phi))
       
    59   || Scan.this_string "thf" >> (fn _ => raise THF ())
       
    60 
       
    61 val parse_problem =
       
    62   Scan.repeat parse_include
       
    63   |-- Scan.repeat (parse_cnf_or_fof --| Scan.repeat parse_include)
       
    64 
       
    65 val parse_tptp_problem =
       
    66   Scan.finite Symbol.stopper
       
    67       (Scan.error (!! (fn _ => raise SYNTAX "malformed TPTP input")
       
    68                   parse_problem))
       
    69   o tptp_explode
       
    70 
       
    71 val iota_T = @{typ iota}
       
    72 val quant_T = @{typ "(iota => bool) => bool"}
       
    73 
       
    74 fun is_variable s = Char.isUpper (String.sub (s, 0))
       
    75 
       
    76 fun hol_term_from_fo_term res_T (ATerm (x, us)) =
       
    77   let val ts = map (hol_term_from_fo_term iota_T) us in
       
    78     list_comb ((case x of
       
    79                   "$true" => @{const_name True}
       
    80                 | "$false" => @{const_name False}
       
    81                 | "=" => @{const_name HOL.eq}
       
    82                 | "equal" => @{const_name HOL.eq}
       
    83                 | _ => x, map fastype_of ts ---> res_T)
       
    84                |> (if is_variable x then Free else Const), ts)
       
    85   end
       
    86 
       
    87 fun hol_prop_from_formula phi =
       
    88   case phi of
       
    89     AQuant (_, [], phi') => hol_prop_from_formula phi'
       
    90   | AQuant (q, (x, _) :: xs, phi') =>
       
    91     Const (case q of AForall => @{const_name All} | AExists => @{const_name Ex},
       
    92            quant_T)
       
    93     $ lambda (Free (x, iota_T)) (hol_prop_from_formula (AQuant (q, xs, phi')))
       
    94   | AConn (ANot, [u']) => HOLogic.mk_not (hol_prop_from_formula u')
       
    95   | AConn (c, [u1, u2]) =>
       
    96     pairself hol_prop_from_formula (u1, u2)
       
    97     |> (case c of
       
    98           AAnd => HOLogic.mk_conj
       
    99         | AOr => HOLogic.mk_disj
       
   100         | AImplies => HOLogic.mk_imp
       
   101         | AIff => HOLogic.mk_eq
       
   102         | ANot => raise Fail "binary \"ANot\"")
       
   103   | AConn _ => raise Fail "malformed AConn"
       
   104   | AAtom u => hol_term_from_fo_term @{typ bool} u
       
   105 
       
   106 fun mk_all x t = Const (@{const_name All}, quant_T) $ lambda x t
       
   107 
       
   108 fun close_hol_prop t = fold (mk_all o Free) (Term.add_frees t []) t
       
   109 
       
   110 
       
   111 (** TPTP parsing based on ML-Yacc-generated "TPTP_Parser" **)
    25 (** TPTP parsing based on ML-Yacc-generated "TPTP_Parser" **)
   112 
    26 
   113 fun mk_meta_not P = Logic.mk_implies (P, @{prop False})
    27 fun mk_meta_not P = Logic.mk_implies (P, @{prop False})
   114 
    28 
   115 (* cf. "close_form" in "refute.ML" *)
    29 (* cf. "close_form" in "refute.ML" *)
   127     val path =
    41     val path =
   128       Path.explode file_name
    42       Path.explode file_name
   129       |> (fn path =>
    43       |> (fn path =>
   130              path |> not (Path.is_absolute path)
    44              path |> not (Path.is_absolute path)
   131                      ? Path.append (Path.explode "$PWD"))
    45                      ? Path.append (Path.explode "$PWD"))
       
    46     val ((_, _, problem), thy) =
       
    47       TPTP_Interpret.interpret_file true (Path.explode "$TPTP") path [] [] thy
   132   in
    48   in
   133     (case parse_tptp_problem (File.read path) of
    49     (exists (fn (_, role, _, _) => role = TPTP_Syntax.Role_Conjecture) problem,
   134        (_, s :: ss) =>
    50      map get_tptp_formula problem, thy)
   135        raise SYNTAX ("cannot parse " ^ quote (implode (s :: ss)))
       
   136      | (problem, []) =>
       
   137        (exists fst problem,
       
   138         map (HOLogic.mk_Trueprop o close_hol_prop o hol_prop_from_formula o snd)
       
   139             problem))
       
   140     handle THF () =>
       
   141     let
       
   142       val problem =
       
   143         TPTP_Interpret.interpret_file true (Path.explode "$TPTP") path [] [] thy
       
   144         |> fst |> #3
       
   145     in
       
   146       (exists (fn (_, role, _, _) => role = TPTP_Syntax.Role_Conjecture) problem,
       
   147        map get_tptp_formula problem)
       
   148     end
       
   149   end
    51   end
   150 
    52 
   151 (** Isabelle (combination of provers) **)
    53 (** Isabelle (combination of provers) **)
   152 
    54 
   153 fun isabelle_tptp_file timeout file_name = ()
    55 fun isabelle_tptp_file timeout file_name = ()
   155 
    57 
   156 (** Nitpick (alias Nitrox) **)
    58 (** Nitpick (alias Nitrox) **)
   157 
    59 
   158 fun nitpick_tptp_file timeout file_name =
    60 fun nitpick_tptp_file timeout file_name =
   159   let
    61   let
   160     val (falsify, ts) = read_tptp_file @{theory} file_name
    62     val (falsify, ts, thy) = read_tptp_file @{theory} file_name
   161     val state = Proof.init @{context}
    63     val ctxt = Proof_Context.init_global thy
       
    64     val state = Proof.init ctxt
   162     val params =
    65     val params =
   163       [("card", "1\<emdash>100"),
    66       [("card", "1\<emdash>100"),
   164        ("box", "false"),
    67        ("box", "false"),
   165        ("sat_solver", "smart"),
    68        ("sat_solver", "smart"),
   166        ("max_threads", "1"),
    69        ("max_threads", "1"),
   171        (* ("overlord", "true"), *)
    74        (* ("overlord", "true"), *)
   172        ("show_consts", "true"),
    75        ("show_consts", "true"),
   173        ("format", "1000"),
    76        ("format", "1000"),
   174        ("max_potential", "0"),
    77        ("max_potential", "0"),
   175        ("timeout", string_of_int timeout)]
    78        ("timeout", string_of_int timeout)]
   176       |> Nitpick_Isar.default_params @{theory}
    79       |> Nitpick_Isar.default_params thy
   177     val i = 1
    80     val i = 1
   178     val n = 1
    81     val n = 1
   179     val step = 0
    82     val step = 0
   180     val subst = []
    83     val subst = []
   181   in
    84   in
   194      "Unknown")
    97      "Unknown")
   195   |> writeln
    98   |> writeln
   196 
    99 
   197 fun refute_tptp_file timeout file_name =
   100 fun refute_tptp_file timeout file_name =
   198   let
   101   let
   199     val (falsify, ts) = read_tptp_file @{theory} file_name
   102     val (falsify, ts, thy) = read_tptp_file @{theory} file_name
       
   103     val ctxt = Proof_Context.init_global thy
   200     val params =
   104     val params =
   201       [("maxtime", string_of_int timeout)]
   105       [("maxtime", string_of_int timeout),
       
   106        ("maxvars", "100000")]
   202   in
   107   in
   203     Refute.refute_term @{context} params ts @{prop False}
   108     Refute.refute_term ctxt params ts @{prop False}
   204     |> print_szs_from_outcome falsify
   109     |> print_szs_from_outcome falsify
   205   end
   110   end
   206 
   111 
   207 
   112 
   208 (** Sledgehammer **)
   113 (** Sledgehammer **)