| author | wenzelm | 
| Fri, 18 Feb 2011 17:05:19 +0100 | |
| changeset 41780 | 7eb9eac392b6 | 
| parent 41307 | bb8468ae414e | 
| child 41945 | 8e32c3992cb3 | 
| permissions | -rw-r--r-- | 
| 35942 
667fd8553cd5
use internal SHA1 digest implementation for generating hash keys
 boehmes parents: 
35941diff
changeset | 1 | (* Title: Tools/cache_io.ML | 
| 35151 | 2 | Author: Sascha Boehme, TU Muenchen | 
| 3 | ||
| 4 | Cache for output of external processes. | |
| 5 | *) | |
| 6 | ||
| 7 | signature CACHE_IO = | |
| 8 | sig | |
| 40425 
c9b5e0fcee31
return the process return code along with the process outputs
 boehmes parents: 
37740diff
changeset | 9 | (*IO wrapper*) | 
| 40538 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 10 |   type result = {
 | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 11 | output: string list, | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 12 | redirected_output: string list, | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 13 | return_code: int} | 
| 40578 
2b098a549450
keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
 boehmes parents: 
40538diff
changeset | 14 | val raw_run: (Path.T -> Path.T -> string) -> string -> Path.T -> Path.T -> | 
| 
2b098a549450
keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
 boehmes parents: 
40538diff
changeset | 15 | result | 
| 40538 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 16 | val run: (Path.T -> Path.T -> string) -> string -> result | 
| 35151 | 17 | |
| 40425 
c9b5e0fcee31
return the process return code along with the process outputs
 boehmes parents: 
37740diff
changeset | 18 | (*cache*) | 
| 35151 | 19 | type cache | 
| 20 | val make: Path.T -> cache | |
| 21 | val cache_path_of: cache -> Path.T | |
| 40538 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 22 | val lookup: cache -> string -> result option * string | 
| 36086 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 23 | val run_and_cache: cache -> string -> (Path.T -> Path.T -> string) -> | 
| 40538 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 24 | string -> result | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 25 | val run_cached: cache -> (Path.T -> Path.T -> string) -> string -> result | 
| 35151 | 26 | end | 
| 27 | ||
| 28 | structure Cache_IO : CACHE_IO = | |
| 29 | struct | |
| 30 | ||
| 40425 
c9b5e0fcee31
return the process return code along with the process outputs
 boehmes parents: 
37740diff
changeset | 31 | (* IO wrapper *) | 
| 
c9b5e0fcee31
return the process return code along with the process outputs
 boehmes parents: 
37740diff
changeset | 32 | |
| 36086 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 33 | val cache_io_prefix = "cache-io-" | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 34 | |
| 40538 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 35 | type result = {
 | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 36 | output: string list, | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 37 | redirected_output: string list, | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 38 | return_code: int} | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 39 | |
| 40578 
2b098a549450
keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
 boehmes parents: 
40538diff
changeset | 40 | fun raw_run make_cmd str in_path out_path = | 
| 
2b098a549450
keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
 boehmes parents: 
40538diff
changeset | 41 | let | 
| 
2b098a549450
keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
 boehmes parents: 
40538diff
changeset | 42 | val _ = File.write in_path str | 
| 
2b098a549450
keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
 boehmes parents: 
40538diff
changeset | 43 | val (out2, rc) = bash_output (make_cmd in_path out_path) | 
| 
2b098a549450
keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
 boehmes parents: 
40538diff
changeset | 44 | val out1 = the_default [] (try (rev o File.fold_lines cons out_path) []) | 
| 
2b098a549450
keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
 boehmes parents: 
40538diff
changeset | 45 |   in {output=split_lines out2, redirected_output=out1, return_code=rc} end
 | 
| 
2b098a549450
keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
 boehmes parents: 
40538diff
changeset | 46 | |
| 36086 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 47 | fun run make_cmd str = | 
| 41307 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 wenzelm parents: 
40743diff
changeset | 48 | Isabelle_System.with_tmp_file cache_io_prefix (fn in_path => | 
| 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 wenzelm parents: 
40743diff
changeset | 49 | Isabelle_System.with_tmp_file cache_io_prefix (fn out_path => | 
| 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 wenzelm parents: 
40743diff
changeset | 50 | raw_run make_cmd str in_path out_path)) | 
| 35151 | 51 | |
| 52 | ||
| 40425 
c9b5e0fcee31
return the process return code along with the process outputs
 boehmes parents: 
37740diff
changeset | 53 | (* cache *) | 
| 35151 | 54 | |
| 55 | abstype cache = Cache of {
 | |
| 56 | path: Path.T, | |
| 57 | table: (int * (int * int * int) Symtab.table) Synchronized.var } | |
| 58 | with | |
| 59 | ||
| 60 | fun cache_path_of (Cache {path, ...}) = path
 | |
| 61 | ||
| 62 | fun load cache_path = | |
| 63 | let | |
| 64 |     fun err () = error ("Cache IO: corrupted cache file: " ^
 | |
| 65 | File.shell_path cache_path) | |
| 66 | ||
| 67 | fun int_of_string s = | |
| 40627 
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
 wenzelm parents: 
40578diff
changeset | 68 | (case read_int (raw_explode s) of | 
| 35151 | 69 | (i, []) => i | 
| 70 | | _ => err ()) | |
| 71 | ||
| 72 | fun split line = | |
| 73 | (case space_explode " " line of | |
| 74 | [key, len1, len2] => (key, int_of_string len1, int_of_string len2) | |
| 75 | | _ => err ()) | |
| 76 | ||
| 77 | fun parse line ((i, l), tab) = | |
| 78 | if i = l | |
| 79 | then | |
| 80 | let val (key, l1, l2) = split line | |
| 81 | in ((i+1, l+l1+l2+1), Symtab.update (key, (i+1, l1, l2)) tab) end | |
| 82 | else ((i+1, l), tab) | |
| 83 | in apfst fst (File.fold_lines parse cache_path ((1, 1), Symtab.empty)) end | |
| 84 | ||
| 85 | fun make path = | |
| 86 | let val table = if File.exists path then load path else (1, Symtab.empty) | |
| 87 |   in Cache {path=path, table=Synchronized.var (Path.implode path) table} end
 | |
| 88 | ||
| 89 | fun load_cached_result cache_path (p, len1, len2) = | |
| 90 | let | |
| 91 | fun load line (i, xsp) = | |
| 92 | if i < p then (i+1, xsp) | |
| 93 | else if i < p + len1 then (i+1, apfst (cons line) xsp) | |
| 94 | else if i < p + len2 then (i+1, apsnd (cons line) xsp) | |
| 95 | else (i, xsp) | |
| 40538 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 96 | val (out, err) = | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 97 | pairself rev (snd (File.fold_lines load cache_path (1, ([], [])))) | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 98 |   in {output=err, redirected_output=out, return_code=0} end
 | 
| 35151 | 99 | |
| 36086 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 100 | fun lookup (Cache {path=cache_path, table}) str =
 | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 101 | let val key = SHA1.digest str | 
| 35151 | 102 | in | 
| 103 | (case Symtab.lookup (snd (Synchronized.value table)) key of | |
| 36086 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 104 | NONE => (NONE, key) | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 105 | | SOME pos => (SOME (load_cached_result cache_path pos), key)) | 
| 35151 | 106 | end | 
| 107 | ||
| 36086 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 108 | fun run_and_cache (Cache {path=cache_path, table}) key make_cmd str =
 | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 109 | let | 
| 40538 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 110 |     val {output=err, redirected_output=out, return_code} = run make_cmd str
 | 
| 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 111 | val (l1, l2) = pairself length (out, err) | 
| 36086 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 112 | val header = key ^ " " ^ string_of_int l1 ^ " " ^ string_of_int l2 | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 113 | val lines = map (suffix "\n") (header :: out @ err) | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 114 | |
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 115 | val _ = Synchronized.change table (fn (p, tab) => | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 116 | if Symtab.defined tab key then (p, tab) | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 117 | else | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 118 | let val _ = File.append_list cache_path lines | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 119 | in (p+l1+l2+1, Symtab.update (key, (p+1, l1, l2)) tab) end) | 
| 40538 
b8482ff0bc92
check the return code of the SMT solver and raise an exception if the prover failed
 boehmes parents: 
40425diff
changeset | 120 |   in {output=err, redirected_output=out, return_code=return_code} end
 | 
| 36086 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 121 | |
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 122 | fun run_cached cache make_cmd str = | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 123 | (case lookup cache str of | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 124 | (NONE, key) => run_and_cache cache key make_cmd str | 
| 
8e5454761f26
simplified Cache_IO interface (input is just a string and not already stored in a file)
 boehmes parents: 
35942diff
changeset | 125 | | (SOME output, _) => output) | 
| 35151 | 126 | |
| 127 | end | |
| 40425 
c9b5e0fcee31
return the process return code along with the process outputs
 boehmes parents: 
37740diff
changeset | 128 | |
| 35151 | 129 | end |