turned SMT counterexamples into verbose messages (they had been swallowed before, following the state of smt_trace -- which is off by default), because they might be useful for the user
--- a/src/HOL/Tools/SMT/smt_failure.ML Fri Nov 12 15:56:08 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_failure.ML Fri Nov 12 15:56:10 2010 +0100
@@ -26,7 +26,9 @@
fun string_of_failure ctxt (Counterexample (real, ex)) =
let
- val msg = (if real then "C" else "Potential c") ^ "ounterexample found"
+ val msg =
+ if real then "Counterexample found (possibly spurious)"
+ else "Potential counterexample found"
in
if null ex then msg
else Pretty.string_of (Pretty.big_list (msg ^ ":")
--- a/src/HOL/Tools/SMT/smt_solver.ML Fri Nov 12 15:56:08 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_solver.ML Fri Nov 12 15:56:10 2010 +0100
@@ -342,8 +342,12 @@
val str_of = SMT_Failure.string_of_failure ctxt'
fun safe_solve irules =
if pass_exns then SOME (solve irules)
- else (SOME (solve irules) handle SMT_Failure.SMT fail =>
- (C.trace_msg ctxt' (prefix tag o str_of) fail; NONE))
+ else (SOME (solve irules)
+ handle
+ SMT_Failure.SMT (fail as SMT_Failure.Counterexample _) =>
+ (C.verbose_msg ctxt' (prefix tag o str_of) fail; NONE)
+ | SMT_Failure.SMT fail =>
+ (C.trace_msg ctxt' (prefix tag o str_of) fail; NONE))
in
safe_solve (map (pair ~1) (rules @ prems))
|> (fn SOME thm => Tactic.rtac thm 1 | _ => Tactical.no_tac)