| author | wenzelm | 
| Wed, 25 Nov 2020 12:57:45 +0100 | |
| changeset 72701 | 1c42ac589fa0 | 
| parent 72700 | c6981f55e60d | 
| child 72711 | 9e89c2e15d36 | 
| permissions | -rw-r--r-- | 
| 72662 | 1 | /* Title: Pure/Tools/build_job.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Build job running prover process, with rudimentary PIDE session. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 10 | import scala.collection.mutable | |
| 11 | ||
| 12 | ||
| 13 | class Build_Job(progress: Progress, | |
| 14 | session_name: String, | |
| 15 | val info: Sessions.Info, | |
| 16 | deps: Sessions.Deps, | |
| 17 | store: Sessions.Store, | |
| 18 | do_store: Boolean, | |
| 19 | verbose: Boolean, | |
| 20 | val numa_node: Option[Int], | |
| 21 | command_timings0: List[Properties.T]) | |
| 22 | {
 | |
| 23 | val options: Options = NUMA.policy_options(info.options, numa_node) | |
| 24 | ||
| 25 | private val sessions_structure = deps.sessions_structure | |
| 26 | ||
| 27 | private val future_result: Future[Process_Result] = | |
| 28 |     Future.thread("build", uninterruptible = true) {
 | |
| 29 |       val parent = info.parent.getOrElse("")
 | |
| 30 | val base = deps(parent) | |
| 31 | ||
| 32 | val env = | |
| 33 | Isabelle_System.settings() + | |
| 34 |           ("ISABELLE_ML_DEBUGGER" -> options.bool("ML_debugger").toString)
 | |
| 35 | ||
| 36 | val is_pure = Sessions.is_pure(session_name) | |
| 37 | ||
| 38 | val use_prelude = if (is_pure) Thy_Header.ml_roots.map(_._1) else Nil | |
| 39 | ||
| 40 | val eval_store = | |
| 41 |         if (do_store) {
 | |
| 42 |           (if (info.theories.nonEmpty) List("ML_Heap.share_common_data ()") else Nil) :::
 | |
| 43 |           List("ML_Heap.save_child " +
 | |
| 44 | ML_Syntax.print_string_bytes(File.platform_path(store.output_heap(session_name)))) | |
| 45 | } | |
| 46 | else Nil | |
| 47 | ||
| 48 | val resources = new Resources(sessions_structure, base, command_timings = command_timings0) | |
| 49 | val session = | |
| 50 |         new Session(options, resources) {
 | |
| 51 | override val xml_cache: XML.Cache = store.xml_cache | |
| 52 | override val xz_cache: XZ.Cache = store.xz_cache | |
| 53 | } | |
| 54 | ||
| 55 | object Build_Session_Errors | |
| 56 |       {
 | |
| 57 | private val promise: Promise[List[String]] = Future.promise | |
| 58 | ||
| 59 | def result: Exn.Result[List[String]] = promise.join_result | |
| 60 | def cancel: Unit = promise.cancel | |
| 61 | def apply(errs: List[String]) | |
| 62 |         {
 | |
| 63 |           try { promise.fulfill(errs) }
 | |
| 64 |           catch { case _: IllegalStateException => }
 | |
| 65 | } | |
| 66 | } | |
| 67 | ||
| 68 | val export_consumer = | |
| 69 | Export.consumer(store.open_database(session_name, output = true), cache = store.xz_cache) | |
| 70 | ||
| 71 | val stdout = new StringBuilder(1000) | |
| 72 | val stderr = new StringBuilder(1000) | |
| 73 | val messages = new mutable.ListBuffer[XML.Elem] | |
| 74 | val command_timings = new mutable.ListBuffer[Properties.T] | |
| 75 | val theory_timings = new mutable.ListBuffer[Properties.T] | |
| 76 | val session_timings = new mutable.ListBuffer[Properties.T] | |
| 77 | val runtime_statistics = new mutable.ListBuffer[Properties.T] | |
| 78 | val task_statistics = new mutable.ListBuffer[Properties.T] | |
| 79 | ||
| 80 | def fun( | |
| 81 | name: String, | |
| 82 | acc: mutable.ListBuffer[Properties.T], | |
| 83 | unapply: Properties.T => Option[Properties.T]): (String, Session.Protocol_Function) = | |
| 84 |       {
 | |
| 85 | name -> ((msg: Prover.Protocol_Output) => | |
| 86 |           unapply(msg.properties) match {
 | |
| 87 | case Some(props) => acc += props; true | |
| 88 | case _ => false | |
| 89 | }) | |
| 90 | } | |
| 91 | ||
| 92 | session.init_protocol_handler(new Session.Protocol_Handler | |
| 93 |         {
 | |
| 94 |           override def exit() { Build_Session_Errors.cancel }
 | |
| 95 | ||
| 96 | private def build_session_finished(msg: Prover.Protocol_Output): Boolean = | |
| 97 |           {
 | |
| 98 | val (rc, errors) = | |
| 99 |               try {
 | |
| 100 | val (rc, errs) = | |
| 101 |                 {
 | |
| 102 | import XML.Decode._ | |
| 103 | pair(int, list(x => x))(Symbol.decode_yxml(msg.text)) | |
| 104 | } | |
| 105 | val errors = | |
| 106 |                   for (err <- errs) yield {
 | |
| 107 | val prt = Protocol_Message.expose_no_reports(err) | |
| 108 | Pretty.string_of(prt, metric = Symbol.Metric) | |
| 109 | } | |
| 110 | (rc, errors) | |
| 111 | } | |
| 112 |               catch { case ERROR(err) => (2, List(err)) }
 | |
| 113 | ||
| 114 |             session.protocol_command("Prover.stop", rc.toString)
 | |
| 115 | Build_Session_Errors(errors) | |
| 116 | true | |
| 117 | } | |
| 118 | ||
| 119 | private def loading_theory(msg: Prover.Protocol_Output): Boolean = | |
| 120 |             msg.properties match {
 | |
| 72692 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 121 | case Markup.Loading_Theory(Markup.Name(name)) => | 
| 72662 | 122 | progress.theory(Progress.Theory(name, session = session_name)) | 
| 72692 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 123 | false | 
| 72662 | 124 | case _ => false | 
| 125 | } | |
| 126 | ||
| 127 | private def export(msg: Prover.Protocol_Output): Boolean = | |
| 128 |             msg.properties match {
 | |
| 129 | case Protocol.Export(args) => | |
| 130 | export_consumer(session_name, args, msg.bytes) | |
| 131 | true | |
| 132 | case _ => false | |
| 133 | } | |
| 134 | ||
| 135 | private def command_timing(props: Properties.T): Option[Properties.T] = | |
| 136 |             for {
 | |
| 137 | props1 <- Markup.Command_Timing.unapply(props) | |
| 138 | elapsed <- Markup.Elapsed.unapply(props1) | |
| 139 | elapsed_time = Time.seconds(elapsed) | |
| 140 |               if elapsed_time.is_relevant && elapsed_time >= options.seconds("command_timing_threshold")
 | |
| 141 | } yield props1.filter(p => Markup.command_timing_properties(p._1)) | |
| 142 | ||
| 143 | override val functions = | |
| 144 | List( | |
| 145 | Markup.Build_Session_Finished.name -> build_session_finished, | |
| 146 | Markup.Loading_Theory.name -> loading_theory, | |
| 147 | Markup.EXPORT -> export, | |
| 148 | fun(Markup.Command_Timing.name, command_timings, command_timing), | |
| 149 | fun(Markup.Theory_Timing.name, theory_timings, Markup.Theory_Timing.unapply), | |
| 150 | fun(Markup.Session_Timing.name, session_timings, Markup.Session_Timing.unapply), | |
| 151 | fun(Markup.Task_Statistics.name, task_statistics, Markup.Task_Statistics.unapply)) | |
| 152 | }) | |
| 153 | ||
| 154 |       session.runtime_statistics += Session.Consumer("ML_statistics")
 | |
| 155 |         {
 | |
| 156 | case Session.Runtime_Statistics(props) => runtime_statistics += props | |
| 157 | } | |
| 158 | ||
| 72692 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 159 |       session.finished_theories += Session.Consumer[Command.State]("finished_theories")
 | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 160 |         {
 | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 161 | case st => | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 162 | val command = st.command | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 163 | val theory_name = command.node_name.theory | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 164 | val args = Protocol.Export.Args(theory_name = theory_name, name = Export.MARKUP) | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 165 | val xml = | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 166 | st.markups(Command.Markup_Index.markup) | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 167 | .to_XML(command.range, command.source, Markup.Elements.full) | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 168 | export_consumer(session_name, args, Bytes(YXML.string_of_body(xml))) | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 169 | } | 
| 
22aeec526ffd
support for PIDE markup in batch build (inactive due to pide_reports=false);
 wenzelm parents: 
72683diff
changeset | 170 | |
| 72662 | 171 |       session.all_messages += Session.Consumer[Any]("build_session_output")
 | 
| 172 |         {
 | |
| 173 | case msg: Prover.Output => | |
| 174 | val message = msg.message | |
| 175 |             if (msg.is_stdout) {
 | |
| 176 | stdout ++= Symbol.encode(XML.content(message)) | |
| 177 | } | |
| 178 |             else if (msg.is_stderr) {
 | |
| 179 | stderr ++= Symbol.encode(XML.content(message)) | |
| 180 | } | |
| 181 |             else if (Protocol.is_exported(message)) {
 | |
| 182 | messages += message | |
| 183 | } | |
| 184 |             else if (msg.is_exit) {
 | |
| 185 | val err = | |
| 186 | "Prover terminated" + | |
| 187 |                   (msg.properties match {
 | |
| 188 | case Markup.Process_Result(result) => ": " + result.print_rc | |
| 189 | case _ => "" | |
| 190 | }) | |
| 191 | Build_Session_Errors(List(err)) | |
| 192 | } | |
| 193 | case _ => | |
| 194 | } | |
| 195 | ||
| 196 |       val eval_main = Command_Line.ML_tool("Isabelle_Process.init_build ()" :: eval_store)
 | |
| 197 | ||
| 198 | val process = | |
| 199 | Isabelle_Process(session, options, sessions_structure, store, | |
| 200 | logic = parent, raw_ml_system = is_pure, | |
| 201 | use_prelude = use_prelude, eval_main = eval_main, | |
| 202 | cwd = info.dir.file, env = env) | |
| 203 | ||
| 204 | val build_errors = | |
| 205 |         Isabelle_Thread.interrupt_handler(_ => process.terminate) {
 | |
| 206 |           Exn.capture { process.await_startup } match {
 | |
| 207 | case Exn.Res(_) => | |
| 208 | val resources_yxml = resources.init_session_yxml | |
| 209 | val args_yxml = | |
| 210 | YXML.string_of_body( | |
| 211 |                   {
 | |
| 212 | import XML.Encode._ | |
| 213 | pair(string, list(pair(Options.encode, list(pair(string, properties)))))( | |
| 214 | (session_name, info.theories)) | |
| 215 | }) | |
| 216 |               session.protocol_command("build_session", resources_yxml, args_yxml)
 | |
| 217 | Build_Session_Errors.result | |
| 218 | case Exn.Exn(exn) => Exn.Res(List(Exn.message(exn))) | |
| 219 | } | |
| 220 | } | |
| 221 | ||
| 222 | val process_result = | |
| 223 |         Isabelle_Thread.interrupt_handler(_ => process.terminate) { process.await_shutdown }
 | |
| 224 | ||
| 72697 | 225 | session.stop() | 
| 226 | ||
| 72662 | 227 | val export_errors = | 
| 228 | export_consumer.shutdown(close = true).map(Output.error_message_text) | |
| 229 | ||
| 72699 | 230 | val (document_output, document_errors) = | 
| 72662 | 231 |         try {
 | 
| 72675 | 232 | if (build_errors.isInstanceOf[Exn.Res[_]] && process_result.ok && info.documents.nonEmpty) | 
| 233 |           {
 | |
| 72662 | 234 | val documents = | 
| 72683 | 235 | using(store.open_database_context(deps.sessions_structure))(db_context => | 
| 236 | Presentation.build_documents(session_name, deps, db_context, | |
| 237 | output_sources = info.document_output, | |
| 238 | output_pdf = info.document_output, | |
| 72699 | 239 | progress = progress, | 
| 240 | verbose = verbose)) | |
| 72673 
8ff7a0e394f9
clarified PDF/HTML presentation, based on pdf blobs from session database (e.g. from earlier builds);
 wenzelm parents: 
72663diff
changeset | 241 | using(store.open_database(session_name, output = true))(db => | 
| 72701 | 242 | documents.foreach(_.write(db, session_name))) | 
| 72699 | 243 | (documents.flatMap(_.log_lines), Nil) | 
| 72662 | 244 | } | 
| 72699 | 245 | (Nil, Nil) | 
| 72662 | 246 | } | 
| 72699 | 247 |         catch { case Exn.Interrupt.ERROR(msg) => (Nil, List(msg)) }
 | 
| 72662 | 248 | |
| 249 | val result = | |
| 250 |       {
 | |
| 251 | val more_output = | |
| 252 | Library.trim_line(stdout.toString) :: | |
| 253 | messages.toList.map(message => | |
| 254 | Symbol.encode(Protocol.message_text(List(message), metric = Symbol.Metric))) ::: | |
| 255 | command_timings.toList.map(Protocol.Command_Timing_Marker.apply) ::: | |
| 256 | theory_timings.toList.map(Protocol.Theory_Timing_Marker.apply) ::: | |
| 257 | session_timings.toList.map(Protocol.Session_Timing_Marker.apply) ::: | |
| 258 | runtime_statistics.toList.map(Protocol.ML_Statistics_Marker.apply) ::: | |
| 259 | task_statistics.toList.map(Protocol.Task_Statistics_Marker.apply) ::: | |
| 72699 | 260 | document_output | 
| 72662 | 261 | |
| 262 | val more_errors = | |
| 263 | Library.trim_line(stderr.toString) :: export_errors ::: document_errors | |
| 264 | ||
| 265 | process_result.output(more_output).errors(more_errors) | |
| 266 | } | |
| 267 | ||
| 268 |       build_errors match {
 | |
| 269 | case Exn.Res(build_errs) => | |
| 270 | val errs = build_errs ::: document_errors | |
| 271 | if (errs.isEmpty) result | |
| 272 |           else {
 | |
| 273 | result.error_rc.output( | |
| 274 | errs.flatMap(s => split_lines(Output.error_message_text(s))) ::: | |
| 275 | errs.map(Protocol.Error_Message_Marker.apply)) | |
| 276 | } | |
| 277 | case Exn.Exn(Exn.Interrupt()) => | |
| 278 | if (result.ok) result.copy(rc = Exn.Interrupt.return_code) else result | |
| 279 | case Exn.Exn(exn) => throw exn | |
| 280 | } | |
| 281 | } | |
| 282 | ||
| 283 | def terminate: Unit = future_result.cancel | |
| 284 | def is_finished: Boolean = future_result.is_finished | |
| 285 | ||
| 286 | private val timeout_request: Option[Event_Timer.Request] = | |
| 287 |   {
 | |
| 288 | if (info.timeout > Time.zero) | |
| 289 |       Some(Event_Timer.request(Time.now() + info.timeout) { terminate })
 | |
| 290 | else None | |
| 291 | } | |
| 292 | ||
| 293 | def join: (Process_Result, Option[String]) = | |
| 294 |   {
 | |
| 295 | val result1 = future_result.join | |
| 296 | ||
| 297 | val was_timeout = | |
| 298 |       timeout_request match {
 | |
| 299 | case None => false | |
| 300 | case Some(request) => !request.cancel | |
| 301 | } | |
| 302 | ||
| 303 | val result2 = | |
| 304 |       if (result1.interrupted) {
 | |
| 305 |         if (was_timeout) result1.error(Output.error_message_text("Timeout")).was_timeout
 | |
| 306 |         else result1.error(Output.error_message_text("Interrupt"))
 | |
| 307 | } | |
| 308 | else result1 | |
| 309 | ||
| 310 | val heap_digest = | |
| 311 | if (result2.ok && do_store && store.output_heap(session_name).is_file) | |
| 312 | Some(Sessions.write_heap_digest(store.output_heap(session_name))) | |
| 313 | else None | |
| 314 | ||
| 315 | (result2, heap_digest) | |
| 316 | } | |
| 317 | } |