author | wenzelm |
Wed, 21 Feb 2024 19:59:35 +0100 | |
changeset 79684 | 0554a32a6ef4 |
parent 79682 | 1fa1b32b0379 |
child 79685 | 45af93b0370a |
permissions | -rw-r--r-- |
79502 | 1 |
/* Title: Pure/Build/store.scala |
78178 | 2 |
Author: Makarius |
3 |
||
4 |
Persistent store for session content: within file-system and/or SQL database. |
|
5 |
*/ |
|
6 |
||
7 |
package isabelle |
|
8 |
||
9 |
||
10 |
import java.sql.SQLException |
|
11 |
||
12 |
||
13 |
object Store { |
|
79676
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
14 |
def apply( |
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
15 |
options: Options, |
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
16 |
build_cluster: Boolean = false, |
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
17 |
cache: Term.Cache = Term.Cache.make() |
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
18 |
): Store = new Store(options, build_cluster, cache) |
78178 | 19 |
|
20 |
||
79684 | 21 |
/* file names */ |
22 |
||
23 |
def heap(name: String): Path = Path.basic(name) |
|
24 |
def log(name: String): Path = Path.basic("log") + Path.basic(name) |
|
25 |
def log_db(name: String): Path = log(name).db |
|
26 |
def log_gz(name: String): Path = log(name).gz |
|
27 |
||
28 |
||
79662 | 29 |
/* session */ |
30 |
||
79674 | 31 |
final class Session private[Store]( |
32 |
val name: String, |
|
33 |
val heap: Option[Path], |
|
34 |
val log_db: Option[Path], |
|
35 |
dirs: List[Path] |
|
36 |
) { |
|
79662 | 37 |
def defined: Boolean = heap.isDefined || log_db.isDefined |
38 |
||
79674 | 39 |
def the_heap: Path = |
40 |
heap getOrElse |
|
41 |
error("Missing heap image for session " + quote(name) + " -- expected in:\n" + |
|
42 |
cat_lines(dirs.map(dir => " " + File.standard_path(dir)))) |
|
43 |
||
79663
4a299bdb5d61
clarified signature: more comprehensive operations;
wenzelm
parents:
79662
diff
changeset
|
44 |
def heap_digest(): Option[SHA1.Digest] = |
4a299bdb5d61
clarified signature: more comprehensive operations;
wenzelm
parents:
79662
diff
changeset
|
45 |
heap.flatMap(ML_Heap.read_file_digest) |
4a299bdb5d61
clarified signature: more comprehensive operations;
wenzelm
parents:
79662
diff
changeset
|
46 |
|
79662 | 47 |
override def toString: String = name |
48 |
} |
|
49 |
||
50 |
||
51 |
||
78179 | 52 |
/* session build info */ |
53 |
||
54 |
sealed case class Build_Info( |
|
55 |
sources: SHA1.Shasum, |
|
56 |
input_heaps: SHA1.Shasum, |
|
57 |
output_heap: SHA1.Shasum, |
|
58 |
return_code: Int, |
|
59 |
uuid: String |
|
60 |
) { |
|
61 |
def ok: Boolean = return_code == 0 |
|
62 |
} |
|
63 |
||
64 |
||
65 |
/* session sources */ |
|
78178 | 66 |
|
67 |
sealed case class Source_File( |
|
68 |
name: String, |
|
69 |
digest: SHA1.Digest, |
|
70 |
compressed: Boolean, |
|
71 |
body: Bytes, |
|
72 |
cache: Compress.Cache |
|
73 |
) { |
|
74 |
override def toString: String = name |
|
75 |
||
76 |
def bytes: Bytes = if (compressed) body.uncompress(cache = cache) else body |
|
77 |
} |
|
78 |
||
79 |
object Sources { |
|
80 |
def load(session_base: Sessions.Base, cache: Compress.Cache = Compress.Cache.none): Sources = |
|
81 |
new Sources( |
|
82 |
session_base.session_sources.foldLeft(Map.empty) { |
|
83 |
case (sources, (path, digest)) => |
|
84 |
def err(): Nothing = error("Incoherent digest for source file: " + path) |
|
85 |
val name = File.symbolic_path(path) |
|
86 |
sources.get(name) match { |
|
87 |
case Some(source_file) => |
|
88 |
if (source_file.digest == digest) sources else err() |
|
89 |
case None => |
|
90 |
val bytes = Bytes.read(path) |
|
91 |
if (bytes.sha1_digest == digest) { |
|
92 |
val (compressed, body) = |
|
93 |
bytes.maybe_compress(Compress.Options_Zstd(), cache = cache) |
|
94 |
val file = Source_File(name, digest, compressed, body, cache) |
|
95 |
sources + (name -> file) |
|
96 |
} |
|
97 |
else err() |
|
98 |
} |
|
99 |
}) |
|
100 |
} |
|
101 |
||
102 |
class Sources private(rep: Map[String, Source_File]) extends Iterable[Source_File] { |
|
103 |
override def toString: String = rep.values.toList.sortBy(_.name).mkString("Sources(", ", ", ")") |
|
104 |
override def iterator: Iterator[Source_File] = rep.valuesIterator |
|
105 |
||
106 |
def get(name: String): Option[Source_File] = rep.get(name) |
|
107 |
def apply(name: String): Source_File = |
|
108 |
get(name).getOrElse(error("Missing session sources entry " + quote(name))) |
|
109 |
} |
|
110 |
||
111 |
||
78179 | 112 |
/* SQL data model */ |
78178 | 113 |
|
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
114 |
object private_data extends SQL.Data() { |
78260
0a7f7abbe4f0
more robust transaction_lock: avoid overlapping data spaces;
wenzelm
parents:
78227
diff
changeset
|
115 |
override lazy val tables = SQL.Tables(Session_Info.table, Sources.table) |
78187
2df0f3604a67
clarified signature: more explicit class SQL.Data;
wenzelm
parents:
78186
diff
changeset
|
116 |
|
78179 | 117 |
object Session_Info { |
118 |
val session_name = SQL.Column.string("session_name").make_primary_key |
|
119 |
||
120 |
// Build_Log.Session_Info |
|
121 |
val session_timing = SQL.Column.bytes("session_timing") |
|
122 |
val command_timings = SQL.Column.bytes("command_timings") |
|
123 |
val theory_timings = SQL.Column.bytes("theory_timings") |
|
124 |
val ml_statistics = SQL.Column.bytes("ml_statistics") |
|
125 |
val task_statistics = SQL.Column.bytes("task_statistics") |
|
126 |
val errors = SQL.Column.bytes("errors") |
|
127 |
val build_log_columns = |
|
128 |
List(session_name, session_timing, command_timings, theory_timings, |
|
129 |
ml_statistics, task_statistics, errors) |
|
78178 | 130 |
|
78179 | 131 |
// Build_Info |
132 |
val sources = SQL.Column.string("sources") |
|
133 |
val input_heaps = SQL.Column.string("input_heaps") |
|
134 |
val output_heap = SQL.Column.string("output_heap") |
|
135 |
val return_code = SQL.Column.int("return_code") |
|
136 |
val uuid = SQL.Column.string("uuid") |
|
137 |
val build_columns = List(sources, input_heaps, output_heap, return_code, uuid) |
|
138 |
||
139 |
val table = SQL.Table("isabelle_session_info", build_log_columns ::: build_columns) |
|
140 |
} |
|
141 |
||
142 |
object Sources { |
|
143 |
val session_name = SQL.Column.string("session_name").make_primary_key |
|
144 |
val name = SQL.Column.string("name").make_primary_key |
|
145 |
val digest = SQL.Column.string("digest") |
|
146 |
val compressed = SQL.Column.bool("compressed") |
|
147 |
val body = SQL.Column.bytes("body") |
|
78178 | 148 |
|
78179 | 149 |
val table = |
150 |
SQL.Table("isabelle_sources", List(session_name, name, digest, compressed, body)) |
|
151 |
||
152 |
def where_equal(session_name: String, name: String = ""): SQL.Source = |
|
153 |
SQL.where_and( |
|
154 |
Sources.session_name.equal(session_name), |
|
155 |
if_proper(name, Sources.name.equal(name))) |
|
156 |
} |
|
157 |
||
78265 | 158 |
def read_bytes(db: SQL.Database, name: String, column: SQL.Column): Bytes = |
159 |
db.execute_query_statementO[Bytes]( |
|
160 |
Session_Info.table.select(List(column), sql = Session_Info.session_name.where_equal(name)), |
|
161 |
res => res.bytes(column) |
|
162 |
).getOrElse(Bytes.empty) |
|
163 |
||
164 |
def read_properties( |
|
165 |
db: SQL.Database, name: String, column: SQL.Column, cache: Term.Cache |
|
166 |
): List[Properties.T] = Properties.uncompress(read_bytes(db, name, column), cache = cache) |
|
167 |
||
168 |
def read_session_timing(db: SQL.Database, name: String, cache: Term.Cache): Properties.T = |
|
169 |
Properties.decode(read_bytes(db, name, Session_Info.session_timing), cache = cache) |
|
170 |
||
171 |
def read_command_timings(db: SQL.Database, name: String): Bytes = |
|
172 |
read_bytes(db, name, Session_Info.command_timings) |
|
173 |
||
174 |
def read_theory_timings(db: SQL.Database, name: String, cache: Term.Cache): List[Properties.T] = |
|
175 |
read_properties(db, name, Session_Info.theory_timings, cache) |
|
176 |
||
177 |
def read_ml_statistics(db: SQL.Database, name: String, cache: Term.Cache): List[Properties.T] = |
|
178 |
read_properties(db, name, Session_Info.ml_statistics, cache) |
|
179 |
||
180 |
def read_task_statistics(db: SQL.Database, name: String, cache: Term.Cache): List[Properties.T] = |
|
181 |
read_properties(db, name, Session_Info.task_statistics, cache) |
|
182 |
||
183 |
def read_errors(db: SQL.Database, name: String, cache: Term.Cache): List[String] = |
|
184 |
Build_Log.uncompress_errors(read_bytes(db, name, Session_Info.errors), cache = cache) |
|
185 |
||
78377 | 186 |
def read_build(db: SQL.Database, name: String): Option[Store.Build_Info] = |
187 |
db.execute_query_statementO[Store.Build_Info]( |
|
188 |
Session_Info.table.select(sql = Session_Info.session_name.where_equal(name)), |
|
189 |
{ res => |
|
190 |
val uuid = |
|
191 |
try { Option(res.string(Session_Info.uuid)).getOrElse("") } |
|
192 |
catch { case _: SQLException => "" } |
|
193 |
Store.Build_Info( |
|
194 |
SHA1.fake_shasum(res.string(Session_Info.sources)), |
|
195 |
SHA1.fake_shasum(res.string(Session_Info.input_heaps)), |
|
196 |
SHA1.fake_shasum(res.string(Session_Info.output_heap)), |
|
197 |
res.int(Session_Info.return_code), |
|
198 |
uuid) |
|
199 |
}) |
|
78265 | 200 |
|
79682
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
201 |
def read_build_uuid(db: SQL.Database, name: String): String = |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
202 |
db.execute_query_statementO[String]( |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
203 |
Session_Info.table.select(List(Session_Info.uuid), |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
204 |
sql = Session_Info.session_name.where_equal(name)), |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
205 |
{ res => |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
206 |
try { Option(res.string(Session_Info.uuid)).getOrElse("") } |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
207 |
catch { case _: SQLException => "" } |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
208 |
}).getOrElse("") |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
209 |
|
78181 | 210 |
def write_session_info( |
211 |
db: SQL.Database, |
|
212 |
cache: Compress.Cache, |
|
213 |
session_name: String, |
|
214 |
build_log: Build_Log.Session_Info, |
|
215 |
build: Build_Info |
|
216 |
): Unit = { |
|
78262 | 217 |
db.execute_statement(Session_Info.table.insert(), body = |
78181 | 218 |
{ stmt => |
219 |
stmt.string(1) = session_name |
|
220 |
stmt.bytes(2) = Properties.encode(build_log.session_timing) |
|
221 |
stmt.bytes(3) = Properties.compress(build_log.command_timings, cache = cache) |
|
222 |
stmt.bytes(4) = Properties.compress(build_log.theory_timings, cache = cache) |
|
223 |
stmt.bytes(5) = Properties.compress(build_log.ml_statistics, cache = cache) |
|
224 |
stmt.bytes(6) = Properties.compress(build_log.task_statistics, cache = cache) |
|
225 |
stmt.bytes(7) = Build_Log.compress_errors(build_log.errors, cache = cache) |
|
226 |
stmt.string(8) = build.sources.toString |
|
227 |
stmt.string(9) = build.input_heaps.toString |
|
228 |
stmt.string(10) = build.output_heap.toString |
|
229 |
stmt.int(11) = build.return_code |
|
230 |
stmt.string(12) = build.uuid |
|
231 |
}) |
|
232 |
} |
|
233 |
||
78555 | 234 |
def write_sources( |
235 |
db: SQL.Database, |
|
236 |
session_name: String, |
|
237 |
source_files: Iterable[Source_File] |
|
238 |
): Unit = { |
|
239 |
db.execute_batch_statement(Sources.table.insert(), batch = |
|
240 |
for (source_file <- source_files) yield { (stmt: SQL.Statement) => |
|
241 |
stmt.string(1) = session_name |
|
242 |
stmt.string(2) = source_file.name |
|
243 |
stmt.string(3) = source_file.digest.toString |
|
244 |
stmt.bool(4) = source_file.compressed |
|
245 |
stmt.bytes(5) = source_file.body |
|
246 |
}) |
|
247 |
} |
|
78178 | 248 |
|
78179 | 249 |
def read_sources( |
250 |
db: SQL.Database, |
|
251 |
session_name: String, |
|
78265 | 252 |
name: String, |
253 |
cache: Compress.Cache |
|
78179 | 254 |
): List[Source_File] = { |
255 |
db.execute_query_statement( |
|
256 |
Sources.table.select( |
|
257 |
sql = Sources.where_equal(session_name, name = name) + SQL.order_by(List(Sources.name))), |
|
258 |
List.from[Source_File], |
|
259 |
{ res => |
|
260 |
val res_name = res.string(Sources.name) |
|
261 |
val digest = SHA1.fake_digest(res.string(Sources.digest)) |
|
262 |
val compressed = res.bool(Sources.compressed) |
|
263 |
val body = res.bytes(Sources.body) |
|
78180 | 264 |
Source_File(res_name, digest, compressed, body, cache) |
78179 | 265 |
} |
266 |
) |
|
267 |
} |
|
78178 | 268 |
} |
79682
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
269 |
|
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
270 |
def read_build_uuid(path: Path, session: String): String = |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
271 |
try { using(SQLite.open_database(path))(private_data.read_build_uuid(_, session)) } |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
272 |
catch { case _: SQLException => "" } |
78178 | 273 |
} |
274 |
||
79676
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
275 |
class Store private( |
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
276 |
val options: Options, |
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
277 |
val build_cluster: Boolean, |
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
278 |
val cache: Term.Cache |
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
279 |
) { |
78178 | 280 |
store => |
281 |
||
282 |
override def toString: String = "Store(output_dir = " + output_dir.absolute + ")" |
|
283 |
||
284 |
||
285 |
/* directories */ |
|
286 |
||
287 |
val system_output_dir: Path = Path.explode("$ISABELLE_HEAPS_SYSTEM/$ML_IDENTIFIER") |
|
288 |
val user_output_dir: Path = Path.explode("$ISABELLE_HEAPS/$ML_IDENTIFIER") |
|
289 |
||
290 |
def system_heaps: Boolean = options.bool("system_heaps") |
|
291 |
||
292 |
val output_dir: Path = |
|
293 |
if (system_heaps) system_output_dir else user_output_dir |
|
294 |
||
295 |
val input_dirs: List[Path] = |
|
296 |
if (system_heaps) List(system_output_dir) |
|
297 |
else List(user_output_dir, system_output_dir) |
|
298 |
||
299 |
def presentation_dir: Path = |
|
300 |
if (system_heaps) Path.explode("$ISABELLE_BROWSER_INFO_SYSTEM") |
|
301 |
else Path.explode("$ISABELLE_BROWSER_INFO") |
|
302 |
||
303 |
||
304 |
/* file names */ |
|
305 |
||
79684 | 306 |
def output_heap(name: String): Path = output_dir + Store.heap(name) |
307 |
def output_log(name: String): Path = output_dir + Store.log(name) |
|
308 |
def output_log_db(name: String): Path = output_dir + Store.log_db(name) |
|
309 |
def output_log_gz(name: String): Path = output_dir + Store.log_gz(name) |
|
78178 | 310 |
|
311 |
||
79662 | 312 |
/* session */ |
313 |
||
79663
4a299bdb5d61
clarified signature: more comprehensive operations;
wenzelm
parents:
79662
diff
changeset
|
314 |
def get_session(name: String): Store.Session = { |
79684 | 315 |
val heap = input_dirs.view.map(_ + Store.heap(name)).find(_.is_file) |
316 |
val log_db = input_dirs.view.map(_ + Store.log_db(name)).find(_.is_file) |
|
79674 | 317 |
new Store.Session(name, heap, log_db, input_dirs) |
79663
4a299bdb5d61
clarified signature: more comprehensive operations;
wenzelm
parents:
79662
diff
changeset
|
318 |
} |
79662 | 319 |
|
79682
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
320 |
def output_session(name: String, store_heap: Boolean = false): Store.Session = { |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
321 |
val heap = if (store_heap) Some(output_heap(name)) else None |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
322 |
val log_db = if (!build_database_server) Some(output_log_db(name)) else None |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
323 |
new Store.Session(name, heap, log_db, List(output_dir)) |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
324 |
} |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
325 |
|
79662 | 326 |
|
78178 | 327 |
/* heap */ |
328 |
||
78212 | 329 |
def heap_shasum(database_server: Option[SQL.Database], name: String): SHA1.Shasum = { |
78510
8f45302a9ff0
more thorough ML_Heap.restore: include ancestors; prefer simultaneous ML_Heap.get_entries: just one database access for heap hierarchy;
wenzelm
parents:
78400
diff
changeset
|
330 |
def get_database: Option[SHA1.Digest] = { |
8f45302a9ff0
more thorough ML_Heap.restore: include ancestors; prefer simultaneous ML_Heap.get_entries: just one database access for heap hierarchy;
wenzelm
parents:
78400
diff
changeset
|
331 |
for { |
8f45302a9ff0
more thorough ML_Heap.restore: include ancestors; prefer simultaneous ML_Heap.get_entries: just one database access for heap hierarchy;
wenzelm
parents:
78400
diff
changeset
|
332 |
db <- database_server |
79677 | 333 |
digest <- ML_Heap.read_digests(db, List(name)).valuesIterator.nextOption() |
78510
8f45302a9ff0
more thorough ML_Heap.restore: include ancestors; prefer simultaneous ML_Heap.get_entries: just one database access for heap hierarchy;
wenzelm
parents:
78400
diff
changeset
|
334 |
} yield digest |
8f45302a9ff0
more thorough ML_Heap.restore: include ancestors; prefer simultaneous ML_Heap.get_entries: just one database access for heap hierarchy;
wenzelm
parents:
78400
diff
changeset
|
335 |
} |
8f45302a9ff0
more thorough ML_Heap.restore: include ancestors; prefer simultaneous ML_Heap.get_entries: just one database access for heap hierarchy;
wenzelm
parents:
78400
diff
changeset
|
336 |
|
79663
4a299bdb5d61
clarified signature: more comprehensive operations;
wenzelm
parents:
79662
diff
changeset
|
337 |
get_database orElse get_session(name).heap_digest() match { |
78196
140a6f2e3728
restore heaps from database, which takes precedence over file-system;
wenzelm
parents:
78190
diff
changeset
|
338 |
case Some(digest) => SHA1.shasum(digest, name) |
140a6f2e3728
restore heaps from database, which takes precedence over file-system;
wenzelm
parents:
78190
diff
changeset
|
339 |
case None => SHA1.no_shasum |
140a6f2e3728
restore heaps from database, which takes precedence over file-system;
wenzelm
parents:
78190
diff
changeset
|
340 |
} |
140a6f2e3728
restore heaps from database, which takes precedence over file-system;
wenzelm
parents:
78190
diff
changeset
|
341 |
} |
140a6f2e3728
restore heaps from database, which takes precedence over file-system;
wenzelm
parents:
78190
diff
changeset
|
342 |
|
78178 | 343 |
|
344 |
/* databases for build process and session content */ |
|
345 |
||
346 |
def build_database_server: Boolean = options.bool("build_database_server") |
|
78511 | 347 |
def build_database: Boolean = options.bool("build_database") |
78178 | 348 |
|
78366 | 349 |
def open_server(): SSH.Server = |
350 |
PostgreSQL.open_server(options, |
|
351 |
host = options.string("build_database_host"), |
|
352 |
port = options.int("build_database_port"), |
|
353 |
ssh_host = options.string("build_database_ssh_host"), |
|
354 |
ssh_port = options.int("build_database_ssh_port"), |
|
355 |
ssh_user = options.string("build_database_ssh_user")) |
|
356 |
||
78347 | 357 |
def open_database_server(server: SSH.Server = SSH.no_server): PostgreSQL.Database = |
358 |
PostgreSQL.open_database_server(options, server = server, |
|
78178 | 359 |
user = options.string("build_database_user"), |
360 |
password = options.string("build_database_password"), |
|
361 |
database = options.string("build_database_name"), |
|
362 |
host = options.string("build_database_host"), |
|
363 |
port = options.int("build_database_port"), |
|
78347 | 364 |
ssh_host = options.string("build_database_ssh_host"), |
365 |
ssh_port = options.int("build_database_ssh_port"), |
|
78863
f627ab8c276c
discontinued pointless option (reverting 63d55ba90a9f): performance tuning works better via SQL.Database.execute_batch_statement;
wenzelm
parents:
78555
diff
changeset
|
366 |
ssh_user = options.string("build_database_ssh_user")) |
78178 | 367 |
|
79682
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
368 |
def maybe_open_database_server( |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
369 |
server: SSH.Server = SSH.no_server, |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
370 |
guard: Boolean = build_database_server |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
371 |
): Option[SQL.Database] = { |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
372 |
if (guard) Some(open_database_server(server = server)) else None |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
373 |
} |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
374 |
|
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
375 |
def maybe_open_heaps_database( |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
376 |
database_server: Option[SQL.Database], |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
377 |
server: SSH.Server = SSH.no_server |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
378 |
): Option[SQL.Database] = { |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
379 |
if (database_server.isDefined) None |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
380 |
else store.maybe_open_database_server(server = server, guard = build_cluster) |
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
381 |
} |
78205
a40ae2df39ad
clarified database for heaps: do not depend on build_database_test;
wenzelm
parents:
78198
diff
changeset
|
382 |
|
78372 | 383 |
def open_build_database(path: Path, server: SSH.Server = SSH.no_server): SQL.Database = |
79676
0cac7e3634d0
more explicit build_cluster flag to guard open_build_database server;
wenzelm
parents:
79674
diff
changeset
|
384 |
if (build_database_server || build_cluster) open_database_server(server = server) |
78178 | 385 |
else SQLite.open_database(path, restrict = true) |
386 |
||
78223 | 387 |
def maybe_open_build_database( |
78372 | 388 |
path: Path = Path.explode("$ISABELLE_HOME_USER/build.db"), |
389 |
server: SSH.Server = SSH.no_server |
|
390 |
): Option[SQL.Database] = { |
|
78511 | 391 |
if (build_database) Some(open_build_database(path, server = server)) else None |
78372 | 392 |
} |
78184 | 393 |
|
78178 | 394 |
def try_open_database( |
395 |
name: String, |
|
396 |
output: Boolean = false, |
|
78372 | 397 |
server: SSH.Server = SSH.no_server, |
78367 | 398 |
server_mode: Boolean = build_database_server |
78178 | 399 |
): Option[SQL.Database] = { |
400 |
def check(db: SQL.Database): Option[SQL.Database] = |
|
401 |
if (output || session_info_exists(db)) Some(db) else { db.close(); None } |
|
402 |
||
78372 | 403 |
if (server_mode) check(open_database_server(server = server)) |
79661
2a9d8c74eb3c
clarified signature: emphasize physical db files;
wenzelm
parents:
79502
diff
changeset
|
404 |
else if (output) Some(SQLite.open_database(output_log_db(name))) |
78178 | 405 |
else { |
406 |
(for { |
|
407 |
dir <- input_dirs.view |
|
79684 | 408 |
path = dir + Store.log_db(name) if path.is_file |
78178 | 409 |
db <- check(SQLite.open_database(path)) |
410 |
} yield db).headOption |
|
411 |
} |
|
412 |
} |
|
413 |
||
414 |
def error_database(name: String): Nothing = |
|
415 |
error("Missing build database for session " + quote(name)) |
|
416 |
||
78372 | 417 |
def open_database( |
418 |
name: String, |
|
419 |
output: Boolean = false, |
|
420 |
server: SSH.Server = SSH.no_server |
|
421 |
): SQL.Database = { |
|
422 |
try_open_database(name, output = output, server = server) getOrElse error_database(name) |
|
423 |
} |
|
78178 | 424 |
|
78213
fd0430a7b7a4
avoid repeated open_database_server: synchronized transaction_lock;
wenzelm
parents:
78212
diff
changeset
|
425 |
def clean_output( |
fd0430a7b7a4
avoid repeated open_database_server: synchronized transaction_lock;
wenzelm
parents:
78212
diff
changeset
|
426 |
database_server: Option[SQL.Database], |
fd0430a7b7a4
avoid repeated open_database_server: synchronized transaction_lock;
wenzelm
parents:
78212
diff
changeset
|
427 |
name: String, |
78217 | 428 |
session_init: Boolean = false |
78213
fd0430a7b7a4
avoid repeated open_database_server: synchronized transaction_lock;
wenzelm
parents:
78212
diff
changeset
|
429 |
): Option[Boolean] = { |
78178 | 430 |
val relevant_db = |
78213
fd0430a7b7a4
avoid repeated open_database_server: synchronized transaction_lock;
wenzelm
parents:
78212
diff
changeset
|
431 |
database_server match { |
79682
1fa1b32b0379
build local log_db, with store/restore via optional database server;
wenzelm
parents:
79677
diff
changeset
|
432 |
case Some(db) => clean_session_info(db, name) |
78227
1ba48d402005
proper session_init *after* deleting db files (amending af6c493b0441);
wenzelm
parents:
78223
diff
changeset
|
433 |
case None => false |
78213
fd0430a7b7a4
avoid repeated open_database_server: synchronized transaction_lock;
wenzelm
parents:
78212
diff
changeset
|
434 |
} |
78178 | 435 |
|
436 |
val del = |
|
437 |
for { |
|
438 |
dir <- |
|
439 |
(if (system_heaps) List(user_output_dir, system_output_dir) else List(user_output_dir)) |
|
79684 | 440 |
file <- List(Store.heap(name), Store.log_db(name), Store.log(name), Store.log_gz(name)) |
78178 | 441 |
path = dir + file if path.is_file |
442 |
} yield path.file.delete |
|
443 |
||
78227
1ba48d402005
proper session_init *after* deleting db files (amending af6c493b0441);
wenzelm
parents:
78223
diff
changeset
|
444 |
if (database_server.isEmpty && session_init) { |
1ba48d402005
proper session_init *after* deleting db files (amending af6c493b0441);
wenzelm
parents:
78223
diff
changeset
|
445 |
using(open_database(name, output = true))(clean_session_info(_, name)) |
1ba48d402005
proper session_init *after* deleting db files (amending af6c493b0441);
wenzelm
parents:
78223
diff
changeset
|
446 |
} |
1ba48d402005
proper session_init *after* deleting db files (amending af6c493b0441);
wenzelm
parents:
78223
diff
changeset
|
447 |
|
78185 | 448 |
if (relevant_db || del.nonEmpty) Some(del.forall(identity)) else None |
78178 | 449 |
} |
450 |
||
451 |
def check_output( |
|
78374 | 452 |
database_server: Option[SQL.Database], |
78178 | 453 |
name: String, |
454 |
session_options: Options, |
|
455 |
sources_shasum: SHA1.Shasum, |
|
456 |
input_shasum: SHA1.Shasum, |
|
457 |
fresh_build: Boolean, |
|
458 |
store_heap: Boolean |
|
459 |
): (Boolean, SHA1.Shasum) = { |
|
78374 | 460 |
def no_check: (Boolean, SHA1.Shasum) = (false, SHA1.no_shasum) |
461 |
||
462 |
def check(db: SQL.Database): (Boolean, SHA1.Shasum) = |
|
463 |
read_build(db, name) match { |
|
464 |
case Some(build) => |
|
465 |
val output_shasum = heap_shasum(if (db.is_postgresql) Some(db) else None, name) |
|
466 |
val current = |
|
467 |
!fresh_build && |
|
468 |
build.ok && |
|
469 |
Sessions.eq_sources(session_options, build.sources, sources_shasum) && |
|
470 |
build.input_heaps == input_shasum && |
|
471 |
build.output_heap == output_shasum && |
|
472 |
!(store_heap && output_shasum.is_empty) |
|
473 |
(current, output_shasum) |
|
474 |
case None => no_check |
|
475 |
} |
|
476 |
||
477 |
database_server match { |
|
478 |
case Some(db) => if (session_info_exists(db)) check(db) else no_check |
|
479 |
case None => using_option(try_open_database(name))(check) getOrElse no_check |
|
78178 | 480 |
} |
481 |
} |
|
482 |
||
483 |
||
484 |
/* session info */ |
|
485 |
||
78375
234f2ff9afe6
clarified signature: more specific exists_table --- avoid retrieving full list beforehand;
wenzelm
parents:
78374
diff
changeset
|
486 |
def session_info_exists(db: SQL.Database): Boolean = |
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
487 |
Store.private_data.tables.forall(db.exists_table) |
78181 | 488 |
|
489 |
def session_info_defined(db: SQL.Database, name: String): Boolean = |
|
490 |
db.execute_query_statementB( |
|
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
491 |
Store.private_data.Session_Info.table.select(List(Store.private_data.Session_Info.session_name), |
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
492 |
sql = Store.private_data.Session_Info.session_name.where_equal(name))) |
78181 | 493 |
|
78260
0a7f7abbe4f0
more robust transaction_lock: avoid overlapping data spaces;
wenzelm
parents:
78227
diff
changeset
|
494 |
def clean_session_info(db: SQL.Database, name: String): Boolean = { |
0a7f7abbe4f0
more robust transaction_lock: avoid overlapping data spaces;
wenzelm
parents:
78227
diff
changeset
|
495 |
Export.clean_session(db, name) |
0a7f7abbe4f0
more robust transaction_lock: avoid overlapping data spaces;
wenzelm
parents:
78227
diff
changeset
|
496 |
Document_Build.clean_session(db, name) |
0a7f7abbe4f0
more robust transaction_lock: avoid overlapping data spaces;
wenzelm
parents:
78227
diff
changeset
|
497 |
|
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
498 |
Store.private_data.transaction_lock(db, create = true, label = "Store.clean_session_info") { |
78178 | 499 |
val already_defined = session_info_defined(db, name) |
500 |
||
501 |
db.execute_statement( |
|
78555 | 502 |
SQL.multi( |
503 |
Store.private_data.Session_Info.table.delete( |
|
504 |
sql = Store.private_data.Session_Info.session_name.where_equal(name)), |
|
505 |
Store.private_data.Sources.table.delete( |
|
506 |
sql = Store.private_data.Sources.where_equal(name)))) |
|
78178 | 507 |
|
508 |
already_defined |
|
509 |
} |
|
78260
0a7f7abbe4f0
more robust transaction_lock: avoid overlapping data spaces;
wenzelm
parents:
78227
diff
changeset
|
510 |
} |
78178 | 511 |
|
512 |
def write_session_info( |
|
513 |
db: SQL.Database, |
|
514 |
session_name: String, |
|
515 |
sources: Store.Sources, |
|
516 |
build_log: Build_Log.Session_Info, |
|
517 |
build: Store.Build_Info |
|
518 |
): Unit = { |
|
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
519 |
Store.private_data.transaction_lock(db, label = "Store.write_session_info") { |
78555 | 520 |
for (source_files <- sources.iterator.toList.grouped(200)) { |
521 |
Store.private_data.write_sources(db, session_name, source_files) |
|
522 |
} |
|
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
523 |
Store.private_data.write_session_info(db, cache.compress, session_name, build_log, build) |
78178 | 524 |
} |
525 |
} |
|
526 |
||
78265 | 527 |
def read_session_timing(db: SQL.Database, session: String): Properties.T = |
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
528 |
Store.private_data.transaction_lock(db, label = "Store.read_session_timing") { |
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
529 |
Store.private_data.read_session_timing(db, session, cache) |
78356 | 530 |
} |
78178 | 531 |
|
78265 | 532 |
def read_command_timings(db: SQL.Database, session: String): Bytes = |
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
533 |
Store.private_data.transaction_lock(db, label = "Store.read_command_timings") { |
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
534 |
Store.private_data.read_command_timings(db, session) |
78356 | 535 |
} |
78178 | 536 |
|
78265 | 537 |
def read_theory_timings(db: SQL.Database, session: String): List[Properties.T] = |
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
538 |
Store.private_data.transaction_lock(db, label = "Store.read_theory_timings") { |
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
539 |
Store.private_data.read_theory_timings(db, session, cache) |
78356 | 540 |
} |
78178 | 541 |
|
78265 | 542 |
def read_ml_statistics(db: SQL.Database, session: String): List[Properties.T] = |
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
543 |
Store.private_data.transaction_lock(db, label = "Store.read_ml_statistics") { |
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
544 |
Store.private_data.read_ml_statistics(db, session, cache) |
78356 | 545 |
} |
78178 | 546 |
|
78265 | 547 |
def read_task_statistics(db: SQL.Database, session: String): List[Properties.T] = |
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
548 |
Store.private_data.transaction_lock(db, label = "Store.read_task_statistics") { |
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
549 |
Store.private_data.read_task_statistics(db, session, cache) |
78356 | 550 |
} |
78265 | 551 |
|
552 |
def read_theories(db: SQL.Database, session: String): List[String] = |
|
553 |
read_theory_timings(db, session).flatMap(Markup.Name.unapply) |
|
554 |
||
555 |
def read_errors(db: SQL.Database, session: String): List[String] = |
|
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
556 |
Store.private_data.transaction_lock(db, label = "Store.read_errors") { |
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
557 |
Store.private_data.read_errors(db, session, cache) |
78356 | 558 |
} |
78265 | 559 |
|
560 |
def read_build(db: SQL.Database, session: String): Option[Store.Build_Info] = |
|
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
561 |
Store.private_data.transaction_lock(db, label = "Store.read_build") { |
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
562 |
if (session_info_exists(db)) Store.private_data.read_build(db, session) else None |
78356 | 563 |
} |
78178 | 564 |
|
78179 | 565 |
def read_sources(db: SQL.Database, session: String, name: String = ""): List[Store.Source_File] = |
78396
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
566 |
Store.private_data.transaction_lock(db, label = "Store.read_sources") { |
7853d9072d1b
renamed object Data to private_data, to emphasize its intended scope (but it is publicly accessible in the database);
wenzelm
parents:
78377
diff
changeset
|
567 |
Store.private_data.read_sources(db, session, name, cache.compress) |
78356 | 568 |
} |
78178 | 569 |
} |