| author | wenzelm |
| Wed, 02 Jan 2013 21:55:57 +0100 | |
| changeset 50691 | 20beafe66748 |
| parent 50690 | 03c4d75e8e32 |
| child 50697 | 82e9178e6a98 |
| permissions | -rw-r--r-- |
| 50685 | 1 |
/* Title: Pure/ML/ml_statistics.ML |
2 |
Author: Makarius |
|
3 |
||
4 |
ML runtime statistics. |
|
5 |
*/ |
|
6 |
||
7 |
package isabelle |
|
8 |
||
9 |
||
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
10 |
import scala.collection.immutable.{SortedSet, SortedMap}
|
| 50691 | 11 |
import scala.swing.{Frame, Component}
|
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
12 |
|
| 50689 | 13 |
import org.jfree.data.xy.{XYSeries, XYSeriesCollection}
|
14 |
import org.jfree.chart.{JFreeChart, ChartPanel, ChartFactory}
|
|
15 |
import org.jfree.chart.plot.PlotOrientation |
|
16 |
||
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
17 |
|
| 50685 | 18 |
object ML_Statistics |
19 |
{
|
|
| 50690 | 20 |
/* content interpretation */ |
21 |
||
22 |
final case class Entry(time: Double, data: Map[String, Double]) |
|
23 |
||
24 |
def apply(stats: List[Properties.T]): ML_Statistics = new ML_Statistics(stats) |
|
25 |
def apply(log: Path): ML_Statistics = apply(read_log(log)) |
|
26 |
||
27 |
||
28 |
/* standard fields */ |
|
29 |
||
30 |
val GC_fields = ("GCs", List("partial_GCs", "full_GCs"))
|
|
31 |
||
32 |
val heap_fields = |
|
33 |
("Heap", List("size_heap", "size_allocation", "size_allocation_free",
|
|
34 |
"size_heap_free_last_full_GC", "size_heap_free_last_GC")) |
|
35 |
||
36 |
val threads_fields = |
|
37 |
("Threads", List("threads_total", "threads_in_ML", "threads_wait_condvar",
|
|
38 |
"threads_wait_IO", "threads_wait_mutex", "threads_wait_signal")) |
|
39 |
||
40 |
val time_fields = |
|
41 |
("Time", List("time_GC_system", "time_GC_user", "time_non_GC_system", "time_non_GC_user"))
|
|
42 |
||
43 |
val tasks_fields = |
|
44 |
("Future tasks", List("tasks_ready", "tasks_pending", "tasks_running", "tasks_passive"))
|
|
45 |
||
46 |
val workers_fields = |
|
47 |
("Worker threads", List("workers_total", "workers_active", "workers_waiting"))
|
|
48 |
||
49 |
val standard_fields = |
|
50 |
List(GC_fields, heap_fields, threads_fields, time_fields, tasks_fields, workers_fields) |
|
51 |
||
52 |
||
| 50685 | 53 |
/* read properties from build log */ |
54 |
||
55 |
private val line_prefix = "\fML_statistics = " |
|
56 |
||
57 |
private val syntax = Outer_Syntax.empty + "," + "(" + ")" + "[" + "]"
|
|
58 |
||
59 |
private object Parser extends Parse.Parser |
|
60 |
{
|
|
61 |
private def stat: Parser[(String, String)] = |
|
62 |
keyword("(") ~ string ~ keyword(",") ~ string ~ keyword(")") ^^
|
|
63 |
{ case _ ~ x ~ _ ~ y ~ _ => (x, y) }
|
|
64 |
private def stats: Parser[Properties.T] = |
|
65 |
keyword("[") ~> repsep(stat, keyword(",")) <~ keyword("]")
|
|
66 |
||
67 |
def parse_stats(s: String): Properties.T = |
|
68 |
{
|
|
69 |
parse_all(stats, Token.reader(syntax.scan(s))) match {
|
|
70 |
case Success(result, _) => result |
|
71 |
case bad => error(bad.toString) |
|
72 |
} |
|
73 |
} |
|
74 |
} |
|
75 |
||
76 |
def read_log(log: Path): List[Properties.T] = |
|
77 |
for {
|
|
78 |
line <- split_lines(File.read_gzip(log)) |
|
79 |
if line.startsWith(line_prefix) |
|
80 |
stats = line.substring(line_prefix.length) |
|
81 |
} yield Parser.parse_stats(stats) |
|
82 |
} |
|
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
83 |
|
|
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
84 |
final class ML_Statistics private(val stats: List[Properties.T]) |
|
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
85 |
{
|
| 50690 | 86 |
val Now = new Properties.Double("now")
|
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
87 |
|
| 50690 | 88 |
require(!stats.isEmpty && stats.forall(props => Now.unapply(props).isDefined)) |
89 |
||
90 |
val time_start = Now.unapply(stats.head).get |
|
91 |
val duration = Now.unapply(stats.last).get - time_start |
|
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
92 |
|
| 50689 | 93 |
val fields: Set[String] = |
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
94 |
SortedSet.empty[String] ++ |
| 50690 | 95 |
(for (props <- stats.iterator; (x, _) <- props.iterator if x != Now.name) |
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
96 |
yield x) |
|
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
97 |
|
|
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
98 |
val content: List[ML_Statistics.Entry] = |
|
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
99 |
stats.map(props => {
|
| 50690 | 100 |
val time = Now.unapply(props).get - time_start |
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
101 |
require(time >= 0.0) |
|
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
102 |
val data = |
|
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
103 |
SortedMap.empty[String, Double] ++ |
| 50690 | 104 |
(for ((x, y) <- props.iterator if x != Now.name) |
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
105 |
yield (x, java.lang.Double.parseDouble(y))) |
|
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
106 |
ML_Statistics.Entry(time, data) |
|
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
107 |
}) |
| 50689 | 108 |
|
109 |
||
110 |
/* charts */ |
|
111 |
||
| 50690 | 112 |
def chart(title: String, selected_fields: Iterable[String]): JFreeChart = |
| 50689 | 113 |
{
|
114 |
val data = new XYSeriesCollection |
|
115 |
||
116 |
for {
|
|
| 50690 | 117 |
field <- selected_fields.iterator |
| 50689 | 118 |
series = new XYSeries(field) |
119 |
} {
|
|
120 |
content.foreach(entry => series.add(entry.time, entry.data(field))) |
|
121 |
data.addSeries(series) |
|
122 |
} |
|
123 |
||
124 |
ChartFactory.createXYLineChart(title, "time", "value", data, |
|
125 |
PlotOrientation.VERTICAL, true, true, true) |
|
126 |
} |
|
127 |
||
| 50690 | 128 |
def chart_panel(title: String, selected_fields: Iterable[String]): ChartPanel = |
129 |
new ChartPanel(chart(title, selected_fields)) |
|
| 50691 | 130 |
|
131 |
def standard_frames: Unit = |
|
132 |
for ((title, selected_fields) <- ML_Statistics.standard_fields) {
|
|
133 |
val c = chart(title, selected_fields) |
|
134 |
Swing_Thread.later {
|
|
135 |
new Frame { contents = Component.wrap(new ChartPanel(c)); visible = true }
|
|
136 |
} |
|
137 |
} |
|
|
50688
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
138 |
} |
|
f02864682307
some support for ML statistics content interpretation;
wenzelm
parents:
50685
diff
changeset
|
139 |