| author | haftmann | 
| Wed, 10 Aug 2022 18:26:22 +0000 | |
| changeset 75800 | a21debbc7074 | 
| parent 75440 | 39011d0d2128 | 
| child 77115 | 97a425ecf96d | 
| 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 | ||
| 75393 | 33 |   sealed case class Memory_Status(heap_size: Long, heap_free: Long, gc_percent: Int) {
 | 
| 72138 | 34 | def heap_used: Long = (heap_size - heap_free) max 0 | 
| 35 | def heap_used_fraction: Double = | |
| 36 | if (heap_size == 0) 0.0 else heap_used.toDouble / heap_size | |
| 37 | def gc_progress: Option[Double] = | |
| 38 | if (1 <= gc_percent && gc_percent <= 100) Some((gc_percent - 1) * 0.01) else None | |
| 39 | } | |
| 40 | ||
| 75393 | 41 |   def memory_status(props: Properties.T): Memory_Status = {
 | 
| 72138 | 42 | val heap_size = Heap_Size.get(props) | 
| 43 | val heap_free = Heap_Free.get(props) | |
| 44 | val gc_percent = GC_Percent.get(props) | |
| 45 | Memory_Status(heap_size, heap_free, gc_percent) | |
| 46 | } | |
| 47 | ||
| 48 | ||
| 72035 | 49 | /* monitor process */ | 
| 50 | ||
| 51 | def monitor(pid: Long, | |
| 72119 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 52 | stats_dir: String = "", | 
| 72035 | 53 | delay: Time = Time.seconds(0.5), | 
| 75393 | 54 | consume: Properties.T => Unit = Console.println | 
| 55 |   ): Unit = {
 | |
| 56 |     def progress_stdout(line: String): Unit = {
 | |
| 73716 
00ef0f401a29
more uniform use of Properties.Eq.unapply, with slightly changed semantics in boundary cases;
 wenzelm parents: 
73604diff
changeset | 57 |       val props = Library.space_explode(',', line).flatMap(Properties.Eq.unapply)
 | 
| 72035 | 58 | if (props.nonEmpty) consume(props) | 
| 59 | } | |
| 60 | ||
| 72119 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 61 | val env_prefix = | 
| 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 62 | if (stats_dir.isEmpty) "" else "export POLYSTATSDIR=" + Bash.string(stats_dir) + "\n" | 
| 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 63 | |
| 73604 | 64 | 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 | 65 |         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 | 66 | ML_Syntax.print_double(delay.seconds)), | 
| 73522 | 67 | cwd = Path.ISABELLE_HOME.file) | 
| 72035 | 68 | .result(progress_stdout = progress_stdout, strict = false).check | 
| 69 | } | |
| 70 | ||
| 71 | ||
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 72 | /* protocol handler */ | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 73 | |
| 75393 | 74 |   class Handler extends Session.Protocol_Handler {
 | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 75 | private var session: Session = null | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 76 | private var monitoring: Future[Unit] = Future.value(()) | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 77 | |
| 75380 | 78 |     override def init(session: Session): Unit = synchronized {
 | 
| 72213 | 79 | this.session = session | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 80 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 81 | |
| 75380 | 82 |     override def exit(): Unit = synchronized {
 | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 83 | session = null | 
| 73367 | 84 | monitoring.cancel() | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 85 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 86 | |
| 75380 | 87 |     private def consume(props: Properties.T): Unit = synchronized {
 | 
| 72136 
98dca728fc9c
removed pointless option "ML_statistics": always enabled;
 wenzelm parents: 
72119diff
changeset | 88 |       if (session != null) {
 | 
| 73031 
f93f0597f4fb
clarified signature: absorb XZ.Cache into XML.Cache;
 wenzelm parents: 
72250diff
changeset | 89 | val props1 = (session.cache.props(props ::: Java_Statistics.jvm_statistics())) | 
| 72149 | 90 | 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 | 91 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 92 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 93 | |
| 75380 | 94 |     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 | 95 |       msg.properties match {
 | 
| 72119 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 96 | 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 | 97 | monitoring = | 
| 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 98 |             Future.thread("ML_statistics") {
 | 
| 
d115d50a19c0
provide POLYSTATSDIR to keep $HOME/.polyml clean (requires Poly/ML 52881757b127, otherwise ignored);
 wenzelm parents: 
72117diff
changeset | 99 | 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 | 100 | } | 
| 72112 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 101 | true | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 102 | case _ => false | 
| 
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 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 105 | |
| 75440 | 106 | override val functions: Session.Protocol_Functions = | 
| 107 | 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 | 108 | } | 
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 109 | |
| 
3546dd4ade74
ML statistics via external process: allows monitoring RTS while ML program sleeps;
 wenzelm parents: 
72044diff
changeset | 110 | |
| 69832 | 111 | /* memory fields (mega bytes) */ | 
| 112 | ||
| 113 | def mem_print(x: Long): Option[String] = | |
| 114 | if (x == 0L) None else Some(x.toString + " M") | |
| 115 | ||
| 116 | def mem_scale(x: Long): Long = x / 1024 / 1024 | |
| 50690 | 117 | |
| 69832 | 118 | def mem_field_scale(name: String, x: Double): Double = | 
| 119 | if (heap_fields._2.contains(name) || program_fields._2.contains(name)) | |
| 120 | mem_scale(x.toLong).toDouble | |
| 121 | else x | |
| 122 | ||
| 123 | val CODE_SIZE = "size_code" | |
| 124 | val STACK_SIZE = "size_stacks" | |
| 65858 | 125 | val HEAP_SIZE = "size_heap" | 
| 126 | ||
| 65866 | 127 | |
| 128 | /* standard fields */ | |
| 129 | ||
| 65864 | 130 | type Fields = (String, List[String]) | 
| 65051 | 131 | |
| 132 | val tasks_fields: Fields = | |
| 60610 
f52b4b0c10c4
improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
 wenzelm parents: 
57869diff
changeset | 133 |     ("Future tasks",
 | 
| 68129 | 134 |       List("tasks_ready", "tasks_pending", "tasks_running", "tasks_passive",
 | 
| 135 | "tasks_urgent", "tasks_total")) | |
| 57869 | 136 | |
| 65051 | 137 | val workers_fields: Fields = | 
| 57869 | 138 |     ("Worker threads", List("workers_total", "workers_active", "workers_waiting"))
 | 
| 139 | ||
| 65051 | 140 | val GC_fields: Fields = | 
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 141 |     ("GCs", List("partial_GCs", "full_GCs", "share_passes"))
 | 
| 50690 | 142 | |
| 65051 | 143 | val heap_fields: Fields = | 
| 65858 | 144 |     ("Heap", List(HEAP_SIZE, "size_allocation", "size_allocation_free",
 | 
| 50690 | 145 | "size_heap_free_last_full_GC", "size_heap_free_last_GC")) | 
| 146 | ||
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 147 | val program_fields: Fields = | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 148 |     ("Program", List("size_code", "size_stacks"))
 | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 149 | |
| 65051 | 150 | val threads_fields: Fields = | 
| 50690 | 151 |     ("Threads", List("threads_total", "threads_in_ML", "threads_wait_condvar",
 | 
| 152 | "threads_wait_IO", "threads_wait_mutex", "threads_wait_signal")) | |
| 153 | ||
| 65051 | 154 | val time_fields: Fields = | 
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 155 |     ("Time", List("time_elapsed", "time_elapsed_GC", "time_CPU", "time_GC"))
 | 
| 51432 | 156 | |
| 65051 | 157 | val speed_fields: Fields = | 
| 158 |     ("Speed", List("speed_CPU", "speed_GC"))
 | |
| 50690 | 159 | |
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 160 |   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 | 161 | |
| 72149 | 162 | val java_heap_fields: Fields = | 
| 163 |     ("Java heap", List("java_heap_size", "java_heap_used"))
 | |
| 164 | ||
| 165 | val java_thread_fields: Fields = | |
| 166 |     ("Java threads", List("java_threads_total", "java_workers_total", "java_workers_active"))
 | |
| 167 | ||
| 65053 | 168 | |
| 72143 | 169 | val main_fields: List[Fields] = | 
| 170 | List(heap_fields, tasks_fields, workers_fields) | |
| 171 | ||
| 72149 | 172 | val other_fields: List[Fields] = | 
| 173 | List(threads_fields, GC_fields, program_fields, time_fields, speed_fields, | |
| 174 | java_heap_fields, java_thread_fields) | |
| 175 | ||
| 176 | val all_fields: List[Fields] = main_fields ::: other_fields | |
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 177 | |
| 
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 | /* content interpretation */ | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 180 | |
| 75393 | 181 |   final case class Entry(time: Double, data: Map[String, Double]) {
 | 
| 65858 | 182 | def get(field: String): Double = data.getOrElse(field, 0.0) | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 183 | } | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 184 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 185 | val empty: ML_Statistics = apply(Nil) | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 186 | |
| 75393 | 187 | def apply( | 
| 188 | ml_statistics0: List[Properties.T], | |
| 189 | heading: String = "", | |
| 190 | domain: String => Boolean = _ => true | |
| 191 |   ): ML_Statistics = {
 | |
| 73120 
c3589f2dff31
more informative errors: simplify diagnosis of spurious failures reported by users;
 wenzelm parents: 
73031diff
changeset | 192 | 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 | 193 | |
| 72224 | 194 | val ml_statistics = ml_statistics0.sortBy(now) | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 195 | 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 | 196 | 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 | 197 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 198 | val fields = | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 199 | SortedSet.empty[String] ++ | 
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 200 |         (for {
 | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 201 | props <- ml_statistics.iterator | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 202 | (x, _) <- props.iterator | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 203 | if x != Now.name && domain(x) } yield x) | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 204 | |
| 75393 | 205 |     val content = {
 | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 206 | var last_edge = Map.empty[String, (Double, Double, Double)] | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 207 | val result = new mutable.ListBuffer[ML_Statistics.Entry] | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 208 |       for (props <- ml_statistics) {
 | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 209 | val time = now(props) - time_start | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 210 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 211 | // rising edges -- relative speed | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 212 | val speeds = | 
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 213 |           (for {
 | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 214 | (key, value) <- props.iterator | 
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 215 | key1 <- time_speed.get(key) | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 216 | if domain(key1) | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 217 |           } yield {
 | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 218 | 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 | 219 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 220 | val x1 = time | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 221 | val y1 = java.lang.Double.parseDouble(value) | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 222 | 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 | 223 | |
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 224 |             if (y1 > y0) {
 | 
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 225 | last_edge += (key -> (x1, y1, s1)) | 
| 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 226 | (key1, s1.toString) | 
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 227 | } | 
| 69822 
8c587dd44f51
updated to polyml-5.8-20190220 (pre-release of Poly/ML 5.8);
 wenzelm parents: 
68129diff
changeset | 228 | else (key1, s0.toString) | 
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 229 | }).toList | 
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 230 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 231 | val data = | 
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 232 | SortedMap.empty[String, Double] ++ | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 233 |             (for {
 | 
| 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 234 | (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 | 235 | 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 | 236 | z = java.lang.Double.parseDouble(y) if z != 0.0 | 
| 69832 | 237 |             } yield { (x.intern, mem_field_scale(x, z)) })
 | 
| 65866 | 238 | |
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 239 | result += ML_Statistics.Entry(time, data) | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 240 | } | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 241 | result.toList | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 242 | } | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 243 | |
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 244 | new ML_Statistics(heading, fields, content, time_start, duration) | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 245 | } | 
| 50685 | 246 | } | 
| 50688 
f02864682307
some support for ML statistics content interpretation;
 wenzelm parents: 
50685diff
changeset | 247 | |
| 65855 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 248 | final class ML_Statistics private( | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 249 | val heading: String, | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 250 | val fields: Set[String], | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 251 | val content: List[ML_Statistics.Entry], | 
| 
33b3e76114f8
store processed content instead of somewhat bulky properties;
 wenzelm parents: 
65854diff
changeset | 252 | val time_start: Double, | 
| 75393 | 253 | val duration: Double | 
| 254 | ) {
 | |
| 74853 | 255 | override def toString: String = | 
| 256 | if (content.isEmpty) "ML_Statistics.empty" | |
| 257 | else "ML_Statistics(length = " + content.length + ", fields = " + fields.size + ")" | |
| 258 | ||
| 259 | ||
| 65858 | 260 | /* content */ | 
| 261 | ||
| 262 | def maximum(field: String): Double = | |
| 73359 | 263 |     content.foldLeft(0.0) { case (m, e) => m max e.get(field) }
 | 
| 65858 | 264 | |
| 75393 | 265 |   def average(field: String): Double = {
 | 
| 65858 | 266 | @tailrec def sum(t0: Double, list: List[ML_Statistics.Entry], acc: Double): Double = | 
| 267 |       list match {
 | |
| 268 | case Nil => acc | |
| 269 | case e :: es => | |
| 270 | val t = e.time | |
| 271 | sum(t, es, (t - t0) * e.get(field) + acc) | |
| 272 | } | |
| 273 |     content match {
 | |
| 274 | case Nil => 0.0 | |
| 275 | case List(e) => e.get(field) | |
| 276 | case e :: es => sum(e.time, es, 0.0) / duration | |
| 277 | } | |
| 278 | } | |
| 279 | ||
| 50689 | 280 | |
| 281 | /* charts */ | |
| 282 | ||
| 75393 | 283 |   def update_data(data: XYSeriesCollection, selected_fields: List[String]): Unit = {
 | 
| 74852 | 284 | data.removeAllSeries() | 
| 69832 | 285 |     for (field <- selected_fields) {
 | 
| 286 | val series = new XYSeries(field) | |
| 67760 
553d9ad7d679
more compact ML_Statistics, to make build_status work with less than 2GB heap;
 wenzelm parents: 
65866diff
changeset | 287 | content.foreach(entry => series.add(entry.time, entry.get(field))) | 
| 50689 | 288 | data.addSeries(series) | 
| 289 | } | |
| 50697 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 290 | } | 
| 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 291 | |
| 75393 | 292 |   def chart(title: String, selected_fields: List[String]): JFreeChart = {
 | 
| 50697 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 293 | val data = new XYSeriesCollection | 
| 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 294 | update_data(data, selected_fields) | 
| 50689 | 295 | |
| 296 | ChartFactory.createXYLineChart(title, "time", "value", data, | |
| 297 | PlotOrientation.VERTICAL, true, true, true) | |
| 298 | } | |
| 299 | ||
| 65051 | 300 | def chart(fields: ML_Statistics.Fields): JFreeChart = | 
| 301 | chart(fields._1, fields._2) | |
| 50691 | 302 | |
| 65053 | 303 | def show_frames(fields: List[ML_Statistics.Fields] = ML_Statistics.main_fields): Unit = | 
| 71601 | 304 | fields.map(chart).foreach(c => | 
| 57612 
990ffb84489b
clarified module name: facilitate alternative GUI frameworks;
 wenzelm parents: 
53189diff
changeset | 305 |       GUI_Thread.later {
 | 
| 50854 | 306 |         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 | 307 | iconImage = GUI.isabelle_image() | 
| 65851 | 308 | title = heading | 
| 50854 | 309 | contents = Component.wrap(new ChartPanel(c)) | 
| 310 | visible = true | |
| 311 | } | |
| 50697 
82e9178e6a98
improved Monitor_Dockable, based on ML_Statistics operations;
 wenzelm parents: 
50691diff
changeset | 312 | }) | 
| 50688 
f02864682307
some support for ML statistics content interpretation;
 wenzelm parents: 
50685diff
changeset | 313 | } |