better compliance with TPTP SZS standard
authorblanchet
Fri, 02 Oct 2015 21:06:32 +0200
changeset 61310 9a50ea544fd3
parent 61309 a2548e708f03
child 61311 150aa3015c47
better compliance with TPTP SZS standard
src/HOL/TPTP/atp_problem_import.ML
src/HOL/Tools/Nitpick/nitpick.ML
--- a/src/HOL/TPTP/atp_problem_import.ML	Fri Oct 02 20:28:56 2015 +0200
+++ b/src/HOL/TPTP/atp_problem_import.ML	Fri Oct 02 21:06:32 2015 +0200
@@ -116,7 +116,7 @@
       (if s = "genuine" then
          if falsify then "CounterSatisfiable" else "Satisfiable"
        else
-         "Unknown")
+         "GaveUp")
       |> writeln
     val (conjs, (defs, nondefs), ctxt) = read_tptp_file thy snd file_name
     val params =
@@ -268,7 +268,7 @@
     (if success then
        if null conjs then "Unsatisfiable" else "Theorem"
      else
-       "Unknown"))
+       "GaveUp"))
 
 fun sledgehammer_tptp_file thy timeout file_name =
   let
--- a/src/HOL/Tools/Nitpick/nitpick.ML	Fri Oct 02 20:28:56 2015 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick.ML	Fri Oct 02 21:06:32 2015 +0200
@@ -636,7 +636,7 @@
                (if genuine then
                   if falsify then "CounterSatisfiable" else "Satisfiable"
                 else
-                  "Unknown") ^ "\n" ^
+                  "GaveUp") ^ "\n" ^
                "% SZS output start FiniteModel")
         val (reconstructed_model, codatatypes_ok) =
           reconstruct_hol_model