| author | wenzelm | 
| Tue, 10 Dec 2024 21:06:04 +0100 | |
| changeset 81572 | 693a95492008 | 
| parent 80224 | db92e0b6a11a | 
| child 82466 | d5ef492dd673 | 
| permissions | -rw-r--r-- | 
| 65477 | 1 | /* Title: Pure/ML/ml_statistics.scala | 
| 50685 | 2 | Author: Makarius | 
| 3 | ||
| 4 | ML runtime statistics. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 65858 | 10 | import scala.annotation.tailrec | 
| 51432 | 11 | import scala.collection.mutable | 
| 50688 
f02864682307
some support for ML statistics content interpretation;
 wenzelm parents: 
50685diff
changeset | 12 | import scala.collection.immutable.{SortedSet, SortedMap}
 | 
| 50691 | 13 | import scala.swing.{Frame, Component}
 | 
| 50688 
f02864682307
some support for ML statistics content interpretation;
 wenzelm parents: 
50685diff
changeset | 14 | |
| 50689 | 15 | import org.jfree.data.xy.{XYSeries, XYSeriesCollection}
 | 
| 16 | import org.jfree.chart.{JFreeChart, ChartPanel, ChartFactory}
 | |
| 17 | import org.jfree.chart.plot.PlotOrientation | |
| 18 | ||
| 50688 
f02864682307
some support for ML statistics content interpretation;
 wenzelm parents: 
50685diff
changeset | 19 | |
| 75393 | 20 | object ML_Statistics {
 | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 21 | /* properties */ | 
| 50690 | 22 | |
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 23 |   val Now = new Properties.Double("now")
 | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 24 | def now(props: Properties.T): Double = Now.unapply(props).get | 
| 50697 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 25 | |
| 50690 | 26 | |
| 72138 | 27 | /* memory status */ | 
| 28 | ||
| 29 |   val Heap_Size = new Properties.Long("size_heap")
 | |
| 30 |   val Heap_Free = new Properties.Long("size_heap_free_last_GC")
 | |
| 31 |   val GC_Percent = new Properties.Int("GC_percent")
 | |
| 32 | ||
| 77708 | 33 |   sealed case class Memory_Status(heap_size: Space, heap_free: Space, gc_percent: Int) {
 | 
| 34 | def heap_used: Space = heap_size.used(heap_free) | |
| 35 | def heap_used_fraction: Double = heap_size.used_fraction(heap_free) | |
| 72138 | 36 | def gc_progress: Option[Double] = | 
| 37 | if (1 <= gc_percent && gc_percent <= 100) Some((gc_percent - 1) * 0.01) else None | |
| 38 | } | |
| 39 | ||
| 75393 | 40 |   def memory_status(props: Properties.T): Memory_Status = {
 | 
| 77708 | 41 | val heap_size = Space.bytes(Heap_Size.get(props)) | 
| 42 | val heap_free = Space.bytes(Heap_Free.get(props)) | |
| 72138 | 43 | val gc_percent = GC_Percent.get(props) | 
| 44 | Memory_Status(heap_size, heap_free, gc_percent) | |
| 45 | } | |
| 46 | ||
| 47 | ||
| 72035 | 48 | /* monitor process */ | 
| 49 | ||
| 50 | def monitor(pid: Long, | |
| 72119 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 51 | stats_dir: String = "", | 
| 72035 | 52 | delay: Time = Time.seconds(0.5), | 
| 75393 | 53 | consume: Properties.T => Unit = Console.println | 
| 54 |   ): Unit = {
 | |
| 55 |     def progress_stdout(line: String): Unit = {
 | |
| 77218 | 56 |       val props = space_explode(',', line).flatMap(Properties.Eq.unapply)
 | 
| 72035 | 57 | if (props.nonEmpty) consume(props) | 
| 58 | } | |
| 59 | ||
| 72119 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 60 | val env_prefix = | 
| 77368 | 61 | if_proper(stats_dir, "export POLYSTATSDIR=" + Bash.string(stats_dir) + "\n") | 
| 72119 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 62 | |
| 73604 | 63 | Bash.process(env_prefix + "\"$POLYML_EXE\" -q --use src/Pure/ML/ml_statistics.ML --eval " + | 
| 72044 
efd169aed4dc
more robust: avoid potential problems with encoding of directory name;
 wenzelm parents: 
72037diff
changeset | 64 |         Bash.string("ML_Statistics.monitor " + ML_Syntax.print_long(pid) + " " +
 | 
| 
efd169aed4dc
more robust: avoid potential problems with encoding of directory name;
 wenzelm parents: 
72037diff
changeset | 65 | ML_Syntax.print_double(delay.seconds)), | 
| 80224 
db92e0b6a11a
clarified signature: prefer symbolic isabelle.Path over physical java.io.File;
 wenzelm parents: 
77710diff
changeset | 66 | cwd = Path.ISABELLE_HOME) | 
| 72035 | 67 | .result(progress_stdout = progress_stdout, strict = false).check | 
| 68 | } | |
| 69 | ||
| 70 | ||
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 71 | /* protocol handler */ | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 72 | |
| 75393 | 73 |   class Handler extends Session.Protocol_Handler {
 | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 74 | private var session: Session = null | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 75 | private var monitoring: Future[Unit] = Future.value(()) | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 76 | |
| 75380 | 77 |     override def init(session: Session): Unit = synchronized {
 | 
| 72213 | 78 | this.session = session | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 79 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 80 | |
| 75380 | 81 |     override def exit(): Unit = synchronized {
 | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 82 | session = null | 
| 73367 | 83 | monitoring.cancel() | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 84 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 85 | |
| 75380 | 86 |     private def consume(props: Properties.T): Unit = synchronized {
 | 
| 72136 
98dca728fc9c
removed pointless option "ML_statistics": always enabled;
 wenzelm parents: 
72119diff
changeset | 87 |       if (session != null) {
 | 
| 73031 
f93f0597f4fb
clarified signature: absorb XZ.Cache into XML.Cache;
 wenzelm parents: 
72250diff
changeset | 88 | val props1 = (session.cache.props(props ::: Java_Statistics.jvm_statistics())) | 
| 72149 | 89 | session.runtime_statistics.post(Session.Runtime_Statistics(props1)) | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 90 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 91 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 92 | |
| 75380 | 93 |     private def ml_statistics(msg: Prover.Protocol_Output): Boolean = synchronized {
 | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 94 |       msg.properties match {
 | 
| 72119 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 95 | case Markup.ML_Statistics(pid, stats_dir) => | 
| 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 96 | monitoring = | 
| 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 97 |             Future.thread("ML_statistics") {
 | 
| 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 98 | monitor(pid, stats_dir = stats_dir, consume = consume) | 
| 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 99 | } | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 100 | true | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 101 | case _ => false | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 102 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 103 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 104 | |
| 75440 | 105 | override val functions: Session.Protocol_Functions = | 
| 106 | List(Markup.ML_Statistics.name -> ml_statistics) | |
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 107 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 108 | |
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 109 | |
| 77121 | 110 | /* memory fields */ | 
| 50690 | 111 | |
| 69832 | 112 | val CODE_SIZE = "size_code" | 
| 113 | val STACK_SIZE = "size_stacks" | |
| 65858 | 114 | val HEAP_SIZE = "size_heap" | 
| 115 | ||
| 65866 | 116 | |
| 117 | /* standard fields */ | |
| 118 | ||
| 77710 | 119 |   sealed case class Fields(title: String, names: List[String], scale_MiB: Boolean = false) {
 | 
| 120 | def scale(y: Double): Double = if (scale_MiB) Space.B(y).MiB else y | |
| 121 | } | |
| 65051 | 122 | |
| 123 | val tasks_fields: Fields = | |
| 77117 | 124 |     Fields("Future tasks",
 | 
| 68129 | 125 |       List("tasks_ready", "tasks_pending", "tasks_running", "tasks_passive",
 | 
| 126 | "tasks_urgent", "tasks_total")) | |
| 57869 | 127 | |
| 65051 | 128 | val workers_fields: Fields = | 
| 77117 | 129 |     Fields("Worker threads", List("workers_total", "workers_active", "workers_waiting"))
 | 
| 57869 | 130 | |
| 65051 | 131 | val GC_fields: Fields = | 
| 77117 | 132 |     Fields("GCs", List("partial_GCs", "full_GCs", "share_passes"))
 | 
| 50690 | 133 | |
| 65051 | 134 | val heap_fields: Fields = | 
| 77117 | 135 |     Fields("Heap", List(HEAP_SIZE, "size_allocation", "size_allocation_free",
 | 
| 77710 | 136 | "size_heap_free_last_full_GC", "size_heap_free_last_GC"), scale_MiB = true) | 
| 50690 | 137 | |
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 138 | val program_fields: Fields = | 
| 77710 | 139 |     Fields("Program", List("size_code", "size_stacks"), scale_MiB = true)
 | 
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 140 | |
| 65051 | 141 | val threads_fields: Fields = | 
| 77117 | 142 |     Fields("Threads", List("threads_total", "threads_in_ML", "threads_wait_condvar",
 | 
| 50690 | 143 | "threads_wait_IO", "threads_wait_mutex", "threads_wait_signal")) | 
| 144 | ||
| 65051 | 145 | val time_fields: Fields = | 
| 77117 | 146 |     Fields("Time", List("time_elapsed", "time_elapsed_GC", "time_CPU", "time_GC"))
 | 
| 51432 | 147 | |
| 65051 | 148 | val speed_fields: Fields = | 
| 77117 | 149 |     Fields("Speed", List("speed_CPU", "speed_GC"))
 | 
| 50690 | 150 | |
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 151 |   private val time_speed = Map("time_CPU" -> "speed_CPU", "time_GC" -> "speed_GC")
 | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 152 | |
| 72149 | 153 | val java_heap_fields: Fields = | 
| 77117 | 154 |     Fields("Java heap", List("java_heap_size", "java_heap_used"))
 | 
| 72149 | 155 | |
| 156 | val java_thread_fields: Fields = | |
| 77117 | 157 |     Fields("Java threads", List("java_threads_total", "java_workers_total", "java_workers_active"))
 | 
| 72149 | 158 | |
| 65053 | 159 | |
| 72143 | 160 | val main_fields: List[Fields] = | 
| 161 | List(heap_fields, tasks_fields, workers_fields) | |
| 162 | ||
| 72149 | 163 | val other_fields: List[Fields] = | 
| 164 | List(threads_fields, GC_fields, program_fields, time_fields, speed_fields, | |
| 165 | java_heap_fields, java_thread_fields) | |
| 166 | ||
| 167 | val all_fields: List[Fields] = main_fields ::: other_fields | |
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 168 | |
| 77118 | 169 | def field_scale(x: String, y: Double): Double = | 
| 77710 | 170 |     all_fields.collectFirst({ case fields if fields.names.contains(x) => fields.scale(y) })
 | 
| 77118 | 171 | .getOrElse(y) | 
| 172 | ||
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 173 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 174 | /* content interpretation */ | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 175 | |
| 75393 | 176 |   final case class Entry(time: Double, data: Map[String, Double]) {
 | 
| 65858 | 177 | def get(field: String): Double = data.getOrElse(field, 0.0) | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 178 | } | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 179 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 180 | val empty: ML_Statistics = apply(Nil) | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 181 | |
| 75393 | 182 | def apply( | 
| 183 | ml_statistics0: List[Properties.T], | |
| 184 | heading: String = "", | |
| 185 | domain: String => Boolean = _ => true | |
| 186 |   ): ML_Statistics = {
 | |
| 73120 
c3589f2dff31
more informative errors: simplify diagnosis of spurious failures reported by users;
 wenzelm parents: 
73031diff
changeset | 187 | require(ml_statistics0.forall(props => Now.unapply(props).isDefined), "missing \"now\" field") | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 188 | |
| 72224 | 189 | val ml_statistics = ml_statistics0.sortBy(now) | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 190 | val time_start = if (ml_statistics.isEmpty) 0.0 else now(ml_statistics.head) | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 191 | val duration = if (ml_statistics.isEmpty) 0.0 else now(ml_statistics.last) - time_start | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 192 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 193 | val fields = | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 194 | SortedSet.empty[String] ++ | 
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 195 |         (for {
 | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 196 | props <- ml_statistics.iterator | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 197 | (x, _) <- props.iterator | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 198 | if x != Now.name && domain(x) } yield x) | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 199 | |
| 75393 | 200 |     val content = {
 | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 201 | var last_edge = Map.empty[String, (Double, Double, Double)] | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 202 | val result = new mutable.ListBuffer[ML_Statistics.Entry] | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 203 |       for (props <- ml_statistics) {
 | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 204 | val time = now(props) - time_start | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 205 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 206 | // rising edges -- relative speed | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 207 | val speeds = | 
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 208 |           (for {
 | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 209 | (key, value) <- props.iterator | 
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 210 | key1 <- time_speed.get(key) | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 211 | if domain(key1) | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 212 |           } yield {
 | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 213 | val (x0, y0, s0) = last_edge.getOrElse(key, (0.0, 0.0, 0.0)) | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 214 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 215 | val x1 = time | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 216 | val y1 = java.lang.Double.parseDouble(value) | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 217 | val s1 = if (x1 == x0) 0.0 else (y1 - y0) / (x1 - x0) | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 218 | |
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 219 |             if (y1 > y0) {
 | 
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 220 | last_edge += (key -> (x1, y1, s1)) | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 221 | (key1, s1.toString) | 
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 222 | } | 
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 223 | else (key1, s0.toString) | 
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 224 | }).toList | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 225 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 226 | val data = | 
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 227 | SortedMap.empty[String, Double] ++ | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 228 |             (for {
 | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 229 | (x, y) <- props.iterator ++ speeds.iterator | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 230 | if x != Now.name && domain(x) | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 231 | z = java.lang.Double.parseDouble(y) if z != 0.0 | 
| 77121 | 232 |             } yield { (x.intern, z) })
 | 
| 65866 | 233 | |
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 234 | result += ML_Statistics.Entry(time, data) | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 235 | } | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 236 | result.toList | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 237 | } | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 238 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 239 | new ML_Statistics(heading, fields, content, time_start, duration) | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 240 | } | 
| 50685 | 241 | } | 
| 50688 
f02864682307
some support for ML statistics content interpretation;
 wenzelm parents: 
50685diff
changeset | 242 | |
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 243 | final class ML_Statistics private( | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 244 | val heading: String, | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 245 | val fields: Set[String], | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 246 | val content: List[ML_Statistics.Entry], | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 247 | val time_start: Double, | 
| 75393 | 248 | val duration: Double | 
| 249 | ) {
 | |
| 74853 | 250 | override def toString: String = | 
| 251 | if (content.isEmpty) "ML_Statistics.empty" | |
| 252 | else "ML_Statistics(length = " + content.length + ", fields = " + fields.size + ")" | |
| 253 | ||
| 254 | ||
| 65858 | 255 | /* content */ | 
| 256 | ||
| 257 | def maximum(field: String): Double = | |
| 73359 | 258 |     content.foldLeft(0.0) { case (m, e) => m max e.get(field) }
 | 
| 65858 | 259 | |
| 75393 | 260 |   def average(field: String): Double = {
 | 
| 65858 | 261 | @tailrec def sum(t0: Double, list: List[ML_Statistics.Entry], acc: Double): Double = | 
| 262 |       list match {
 | |
| 263 | case Nil => acc | |
| 264 | case e :: es => | |
| 265 | val t = e.time | |
| 266 | sum(t, es, (t - t0) * e.get(field) + acc) | |
| 267 | } | |
| 268 |     content match {
 | |
| 269 | case Nil => 0.0 | |
| 270 | case List(e) => e.get(field) | |
| 271 | case e :: es => sum(e.time, es, 0.0) / duration | |
| 272 | } | |
| 273 | } | |
| 274 | ||
| 50689 | 275 | |
| 276 | /* charts */ | |
| 277 | ||
| 75393 | 278 |   def update_data(data: XYSeriesCollection, selected_fields: List[String]): Unit = {
 | 
| 74852 | 279 | data.removeAllSeries() | 
| 69832 | 280 |     for (field <- selected_fields) {
 | 
| 281 | val series = new XYSeries(field) | |
| 77121 | 282 | content.foreach(e => series.add(e.time, ML_Statistics.field_scale(field, e.get(field)))) | 
| 50689 | 283 | data.addSeries(series) | 
| 284 | } | |
| 50697 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 285 | } | 
| 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 286 | |
| 75393 | 287 |   def chart(title: String, selected_fields: List[String]): JFreeChart = {
 | 
| 50697 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 288 | val data = new XYSeriesCollection | 
| 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 289 | update_data(data, selected_fields) | 
| 50689 | 290 | |
| 291 | ChartFactory.createXYLineChart(title, "time", "value", data, | |
| 292 | PlotOrientation.VERTICAL, true, true, true) | |
| 293 | } | |
| 294 | ||
| 65051 | 295 | def chart(fields: ML_Statistics.Fields): JFreeChart = | 
| 77117 | 296 | chart(fields.title, fields.names) | 
| 50691 | 297 | |
| 65053 | 298 | def show_frames(fields: List[ML_Statistics.Fields] = ML_Statistics.main_fields): Unit = | 
| 71601 | 299 | fields.map(chart).foreach(c => | 
| 57612 
990ffb84489b
clarified module name: facilitate alternative GUI frameworks;
 wenzelm parents: 
53189diff
changeset | 300 |       GUI_Thread.later {
 | 
| 50854 | 301 |         new Frame {
 | 
| 51615 
072a7249e1ac
separate module "GUI", to keep this out of the way of generic Isabelle_System operations, notably for non-Isabelle/jEdit applications;
 wenzelm parents: 
51432diff
changeset | 302 | iconImage = GUI.isabelle_image() | 
| 65851 | 303 | title = heading | 
| 50854 | 304 | contents = Component.wrap(new ChartPanel(c)) | 
| 305 | visible = true | |
| 306 | } | |
| 50697 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 307 | }) | 
| 50688 
f02864682307
some support for ML statistics content interpretation;
 wenzelm parents: 
50685diff
changeset | 308 | } |