133 def extra_components: List[Component] = build_config.extra_components |
133 def extra_components: List[Component] = build_config.extra_components |
134 |
134 |
135 def build_cluster = build_config.build_cluster |
135 def build_cluster = build_config.build_cluster |
136 def build_hosts: List[Build_Cluster.Host] = |
136 def build_hosts: List[Build_Cluster.Host] = |
137 Build_Cluster.Host.parse(Registry.global, hosts_spec) |
137 Build_Cluster.Host.parse(Registry.global, hosts_spec) |
|
138 |
|
139 def >(t: Task): Boolean = |
|
140 priority.ordinal > t.priority.ordinal || |
|
141 (priority == t.priority && submit_date.time < t.submit_date.time) |
138 } |
142 } |
139 |
143 |
140 sealed case class Job( |
144 sealed case class Job( |
141 uuid: UUID.T, |
145 uuid: UUID.T, |
142 kind: String, |
146 kind: String, |
1245 (if (running.nonEmpty) render_job(running.head) |
1253 (if (running.nonEmpty) render_job(running.head) |
1246 else if (finished.nonEmpty) render_result(finished.head) |
1254 else if (finished.nonEmpty) render_result(finished.head) |
1247 else Nil)) |
1255 else Nil)) |
1248 } |
1256 } |
1249 |
1257 |
1250 text("Queue: " + state.pending.size + " tasks waiting") ::: |
1258 val running = state.running.values.toList |
1251 section("Builds") :: par(text("Total: " + state.num_builds + " builds")) :: |
1259 val idle = (build_hosts.map(_.hostname).toSet -- running.flatMap(_.hostnames).toSet).toList |
|
1260 |
|
1261 def render_rows(hostnames: List[String], body: XML.Body): XML.Body = |
|
1262 hostnames match { |
|
1263 case Nil => Nil |
|
1264 case hostname :: Nil => List(tr(List(td(text(hostname)), td(body)))) |
|
1265 case hostname :: hostnames1 => |
|
1266 tr(List(td(text(hostname)), td.apply(rowspan(hostnames.length), body))) :: |
|
1267 hostnames1.map(hostname => tr(List(td(text(hostname))))) |
|
1268 } |
|
1269 |
|
1270 def render_job(job: Job): XML.Body = |
|
1271 render_rows(job.hostnames, |
|
1272 page_link(Page.BUILD, job.name, Markup.Name(job.name)) :: summary(job.start_date)) |
|
1273 |
|
1274 par(text("Queue: " + state.pending.size + " tasks waiting")) :: |
|
1275 table(tr(List(th(text("Host")), th(text("Activity")))) :: |
|
1276 running.sortBy(_.name).flatMap(render_job) ::: |
|
1277 idle.sorted.map(List(_)).flatMap(render_rows(_, text("idle")))) :: |
|
1278 section("Builds") :: |
|
1279 par(text("Total: " + state.num_builds + " builds")) :: |
1252 state.kinds.sorted.map(render_kind) |
1280 state.kinds.sorted.map(render_kind) |
1253 } |
1281 } |
1254 |
1282 |
1255 def render_overview(kind: String, state: State): XML.Body = |
1283 def render_overview(kind: String, state: State): XML.Body = |
1256 render_page("Overview: " + kind + " job ") { |
1284 render_page("Overview: " + kind + " job ") { |
1282 |
1310 |
1283 if (!components.map(_.name).exists(hg_info.toSet)) text("Components: " + s) |
1311 if (!components.map(_.name).exists(hg_info.toSet)) text("Components: " + s) |
1284 else text("Components: ") :+ page_link(Page.DIFF, s, Markup.Name(build.name)) |
1312 else text("Components: ") :+ page_link(Page.DIFF, s, Markup.Name(build.name)) |
1285 } |
1313 } |
1286 |
1314 |
|
1315 def waiting_for(host: Build_Cluster.Host): XML.Body = |
|
1316 build_hosts.find(_.hostname == host.hostname) match { |
|
1317 case None => break ::: text(quote(host.hostname) + " is not a build host") |
|
1318 case Some(host) => |
|
1319 val active = state.running.values.filter(_.hostnames.contains(host.hostname)) |
|
1320 if (active.isEmpty) Nil |
|
1321 else break ::: |
|
1322 text(host.hostname + " is busy with " + active.map(_.name).mkString(" and ")) |
|
1323 } |
|
1324 |
|
1325 def waiting(task: Task): XML.Body = { |
|
1326 val num_before = state.pending.values.count(_ > task) |
|
1327 |
|
1328 if (num_before > 0) text("Waiting for " + num_before + " tasks to complete") |
|
1329 else Exn.capture(task.build_hosts) match { |
|
1330 case Exn.Res(hosts) => text("Hosts not ready:") ::: hosts.flatMap(waiting_for) |
|
1331 case _ => text("Unkown host spec") |
|
1332 } |
|
1333 } |
|
1334 |
|
1335 def started(user: Option[String], date: Date): String = |
|
1336 "Started" + if_proper(user, " by " + user.get) + " on " + Build_Log.print_date(date) |
|
1337 |
1287 build match { |
1338 build match { |
1288 case task: Task => |
1339 case task: Task => |
1289 par(text("Task from " + Build_Log.print_date(task.submit_date) + ". ")) :: |
1340 par(text("Task from " + Build_Log.print_date(task.submit_date) + ". ")) :: |
1290 par(text(task.components.mkString(", "))) :: |
1341 par(text("Components: " + task.components.mkString(", "))) :: |
|
1342 par(List(bold(waiting(task)))) :: |
1291 render_cancel(task.uuid) |
1343 render_cancel(task.uuid) |
1292 |
1344 |
1293 case job: Job => |
1345 case job: Job => |
1294 val report_data = cache.lookup(store.report(job.kind, job.id)) |
1346 val report_data = cache.lookup(store.report(job.kind, job.id)) |
1295 |
1347 |
1296 par(text("Start: " + Build_Log.print_date(job.start_date))) :: |
1348 par(text(started(job.user, job.start_date))) :: |
1297 par( |
1349 par( |
1298 if (job.cancelled) text("Cancelling ...") |
1350 if (job.cancelled) text("Cancelling ...") |
1299 else text("Running ...") ::: render_cancel(job.uuid)) :: |
1351 else text("Running ...") ::: render_cancel(job.uuid)) :: |
1300 par(render_rev(job.components, report_data)) :: |
1352 par(render_rev(job.components, report_data)) :: |
1301 par(List(source(report_data.build_log))) :: Nil |
1353 par(List(source(report_data.build_log))) :: Nil |
1302 |
1354 |
1303 case result: Result => |
1355 case result: Result => |
1304 val report_data = cache.lookup(store.report(result.kind, result.id)) |
1356 val report_data = cache.lookup(store.report(result.kind, result.id)) |
1305 |
1357 |
1306 par(text("Start: " + Build_Log.print_date(result.start_date) + |
1358 par(text(started(result.user, result.start_date) + |
1307 if_proper(result.end_date, |
1359 if_proper(result.end_date, |
1308 ", took " + (result.end_date.get - result.start_date).message_hms))) :: |
1360 ", took " + (result.end_date.get - result.start_date).message_hms))) :: |
1309 par(text("Status: " + result.status)) :: |
1361 par(text("Status: " + result.status)) :: |
1310 par(render_rev(result.components, report_data)) :: |
1362 par(render_rev(result.components, report_data)) :: |
1311 par(List(source(report_data.build_log))) :: Nil |
1363 par(List(source(report_data.build_log))) :: Nil |
1602 |
1654 |
1603 val processes = List( |
1655 val processes = List( |
1604 new Runner(store, build_hosts, isabelle_repository, sync_dirs, progress), |
1656 new Runner(store, build_hosts, isabelle_repository, sync_dirs, progress), |
1605 new Poller(ci_jobs, store, isabelle_repository, sync_dirs, progress), |
1657 new Poller(ci_jobs, store, isabelle_repository, sync_dirs, progress), |
1606 new Timer(ci_jobs, store, progress), |
1658 new Timer(ci_jobs, store, progress), |
1607 new Web_Server(port, store, progress)) |
1659 new Web_Server(port, store, build_hosts, progress)) |
1608 |
1660 |
1609 val threads = processes.map(Isabelle_Thread.create(_)) |
1661 val threads = processes.map(Isabelle_Thread.create(_)) |
1610 POSIX_Interrupt.handler { |
1662 POSIX_Interrupt.handler { |
1611 progress.stop() |
1663 progress.stop() |
1612 processes.foreach(_.interrupt()) |
1664 processes.foreach(_.interrupt()) |