--- a/src/HOL/Tools/SMT/smt_solver.ML Mon Nov 08 12:13:44 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_solver.ML Mon Nov 08 12:13:51 2010 +0100
@@ -108,7 +108,11 @@
fun run ctxt cmd args input =
(case C.certificates_of ctxt of
- NONE => Cache_IO.run (make_cmd (choose cmd) args) input
+ NONE =>
+ let
+ val {output, redirected_output, ...} =
+ Cache_IO.run (make_cmd (choose cmd) args) input
+ in (redirected_output, output) end
| SOME certs =>
(case Cache_IO.lookup certs input of
(NONE, key) =>
--- a/src/Tools/cache_io.ML Mon Nov 08 12:13:44 2010 +0100
+++ b/src/Tools/cache_io.ML Mon Nov 08 12:13:51 2010 +0100
@@ -6,10 +6,13 @@
signature CACHE_IO =
sig
+ (*IO wrapper*)
val with_tmp_file: string -> (Path.T -> 'a) -> 'a
val with_tmp_dir: string -> (Path.T -> 'a) -> 'a
- val run: (Path.T -> Path.T -> string) -> string -> string list * string list
+ val run: (Path.T -> Path.T -> string) -> string ->
+ {output: string list, redirected_output: string list, return_code: int}
+ (*cache*)
type cache
val make: Path.T -> cache
val cache_path_of: cache -> Path.T
@@ -23,6 +26,8 @@
structure Cache_IO : CACHE_IO =
struct
+(* IO wrapper *)
+
val cache_io_prefix = "cache-io-"
fun with_tmp_file name f =
@@ -45,21 +50,20 @@
with_tmp_file cache_io_prefix (fn out_path =>
let
val _ = File.write in_path str
- val (out2, _) = bash_output (make_cmd in_path out_path)
+ val (out2, rc) = bash_output (make_cmd in_path out_path)
val out1 = the_default [] (try (rev o File.fold_lines cons out_path) [])
- in (out1, split_lines out2) end))
+ in {output=split_lines out2, redirected_output=out1, return_code=rc} end))
+(* cache *)
abstype cache = Cache of {
path: Path.T,
table: (int * (int * int * int) Symtab.table) Synchronized.var }
with
-
fun cache_path_of (Cache {path, ...}) = path
-
fun load cache_path =
let
fun err () = error ("Cache IO: corrupted cache file: " ^
@@ -87,7 +91,6 @@
let val table = if File.exists path then load path else (1, Symtab.empty)
in Cache {path=path, table=Synchronized.var (Path.implode path) table} end
-
fun load_cached_result cache_path (p, len1, len2) =
let
fun load line (i, xsp) =
@@ -97,7 +100,6 @@
else (i, xsp)
in pairself rev (snd (File.fold_lines load cache_path (1, ([], [])))) end
-
fun lookup (Cache {path=cache_path, table}) str =
let val key = SHA1.digest str
in
@@ -106,10 +108,10 @@
| SOME pos => (SOME (load_cached_result cache_path pos), key))
end
-
fun run_and_cache (Cache {path=cache_path, table}) key make_cmd str =
let
- val res as (out, err) = run make_cmd str
+ val {output=err, redirected_output=out, ...} = run make_cmd str
+ val res = (out, err)
val (l1, l2) = pairself length res
val header = key ^ " " ^ string_of_int l1 ^ " " ^ string_of_int l2
val lines = map (suffix "\n") (header :: out @ err)
@@ -121,11 +123,11 @@
in (p+l1+l2+1, Symtab.update (key, (p+1, l1, l2)) tab) end)
in res end
-
fun run_cached cache make_cmd str =
(case lookup cache str of
(NONE, key) => run_and_cache cache key make_cmd str
| (SOME output, _) => output)
end
+
end