--- a/src/HOL/Tools/ATP/atp_systems.ML Thu Sep 09 14:47:06 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_systems.ML Thu Sep 09 16:06:11 2010 +0200
@@ -10,7 +10,7 @@
datatype failure =
Unprovable | IncompleteUnprovable | CantConnect | TimedOut |
OutOfResources | SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl |
- MalformedInput | MalformedOutput | UnknownError
+ MalformedInput | MalformedOutput | Interrupted | UnknownError
type prover_config =
{exec: string * string,
@@ -42,7 +42,7 @@
datatype failure =
Unprovable | IncompleteUnprovable | CantConnect | TimedOut | OutOfResources |
SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl | MalformedInput |
- MalformedOutput | UnknownError
+ MalformedOutput | Interrupted | UnknownError
type prover_config =
{exec: string * string,
@@ -85,6 +85,7 @@
"The ATP problem is malformed. Please report this to the Isabelle \
\developers."
| string_for_failure MalformedOutput = "The ATP output is malformed."
+ | string_for_failure Interrupted = "The ATP was interrupted."
| string_for_failure UnknownError = "An unknown ATP error occurred."
fun known_failure_in_output output =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML Thu Sep 09 14:47:06 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML Thu Sep 09 16:06:11 2010 +0200
@@ -162,7 +162,10 @@
output =
case known_failure_in_output output known_failures of
NONE => (case extract_proof proof_delims output of
- "" => ("", SOME MalformedOutput)
+ "" => ("", SOME (if res_code = 0 andalso output = "" then
+ Interrupted
+ else
+ UnknownError))
| proof => if res_code = 0 then (proof, NONE)
else ("", SOME UnknownError))
| SOME failure =>
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Thu Sep 09 14:47:06 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Thu Sep 09 16:06:11 2010 +0200
@@ -31,6 +31,7 @@
fun string_for_failure Unprovable = "Unprovable."
| string_for_failure TimedOut = "Timed out."
+ | string_for_failure Interrupted = "Interrupted."
| string_for_failure _ = "Unknown error."
fun n_theorems names =