| author | wenzelm |
| Mon, 05 Sep 2016 11:43:37 +0200 | |
| changeset 63789 | af28929ff219 |
| parent 63708 | 77323d963ca4 |
| child 63805 | c272680df665 |
| permissions | -rw-r--r-- |
| 63686 | 1 |
/* Title: Pure/Tools/build_stats.scala |
2 |
Author: Makarius |
|
3 |
||
4 |
Statistics from session build output. |
|
5 |
*/ |
|
6 |
||
7 |
package isabelle |
|
8 |
||
9 |
||
10 |
import scala.collection.mutable |
|
11 |
import scala.util.matching.Regex |
|
12 |
||
13 |
||
14 |
object Build_Stats |
|
15 |
{
|
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
16 |
/* parse build output */ |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
17 |
|
|
63702
fed1d4dab990
cpu time is optional (see Timing.message_resources);
wenzelm
parents:
63701
diff
changeset
|
18 |
private val Session_Finished1 = |
| 63686 | 19 |
new Regex("""^Finished (\S+) \((\d+):(\d+):(\d+) elapsed time, (\d+):(\d+):(\d+) cpu time.*$""")
|
|
63702
fed1d4dab990
cpu time is optional (see Timing.message_resources);
wenzelm
parents:
63701
diff
changeset
|
20 |
private val Session_Finished2 = |
|
fed1d4dab990
cpu time is optional (see Timing.message_resources);
wenzelm
parents:
63701
diff
changeset
|
21 |
new Regex("""^Finished (\S+) \((\d+):(\d+):(\d+) elapsed time.*$""")
|
| 63686 | 22 |
private val Session_Timing = |
23 |
new Regex("""^Timing (\S+) \((\d) threads, (\d+\.\d+)s elapsed time, (\d+\.\d+)s cpu time, (\d+\.\d+)s GC time.*$""")
|
|
24 |
||
25 |
private object ML_Option |
|
26 |
{
|
|
27 |
def unapply(s: String): Option[(String, String)] = |
|
28 |
s.indexOf('=') match {
|
|
29 |
case -1 => None |
|
30 |
case i => |
|
31 |
val a = s.substring(0, i) |
|
32 |
Library.try_unquote(s.substring(i + 1)) match {
|
|
33 |
case Some(b) if Build.ml_options.contains(a) => Some((a, b)) |
|
34 |
case _ => None |
|
35 |
} |
|
36 |
} |
|
37 |
} |
|
38 |
||
39 |
def parse(text: String): Build_Stats = |
|
40 |
{
|
|
41 |
import Properties.Value |
|
42 |
||
43 |
val ml_options = new mutable.ListBuffer[(String, String)] |
|
44 |
var finished = Map.empty[String, Timing] |
|
45 |
var timing = Map.empty[String, Timing] |
|
46 |
var threads = Map.empty[String, Int] |
|
47 |
||
48 |
for (line <- split_lines(text)) {
|
|
49 |
line match {
|
|
|
63702
fed1d4dab990
cpu time is optional (see Timing.message_resources);
wenzelm
parents:
63701
diff
changeset
|
50 |
case Session_Finished1(name, |
| 63686 | 51 |
Value.Int(e1), Value.Int(e2), Value.Int(e3), |
52 |
Value.Int(c1), Value.Int(c2), Value.Int(c3)) => |
|
| 63700 | 53 |
val elapsed = Time.hms(e1, e2, e3) |
54 |
val cpu = Time.hms(c1, c2, c3) |
|
| 63686 | 55 |
finished += (name -> Timing(elapsed, cpu, Time.zero)) |
|
63702
fed1d4dab990
cpu time is optional (see Timing.message_resources);
wenzelm
parents:
63701
diff
changeset
|
56 |
case Session_Finished2(name, |
|
fed1d4dab990
cpu time is optional (see Timing.message_resources);
wenzelm
parents:
63701
diff
changeset
|
57 |
Value.Int(e1), Value.Int(e2), Value.Int(e3)) => |
|
fed1d4dab990
cpu time is optional (see Timing.message_resources);
wenzelm
parents:
63701
diff
changeset
|
58 |
val elapsed = Time.hms(e1, e2, e3) |
|
fed1d4dab990
cpu time is optional (see Timing.message_resources);
wenzelm
parents:
63701
diff
changeset
|
59 |
finished += (name -> Timing(elapsed, Time.zero, Time.zero)) |
| 63686 | 60 |
case Session_Timing(name, |
61 |
Value.Int(t), Value.Double(e), Value.Double(c), Value.Double(g)) => |
|
62 |
val elapsed = Time.seconds(e) |
|
63 |
val cpu = Time.seconds(c) |
|
64 |
val gc = Time.seconds(g) |
|
65 |
timing += (name -> Timing(elapsed, cpu, gc)) |
|
66 |
threads += (name -> t) |
|
67 |
case ML_Option(option) => ml_options += option |
|
68 |
case _ => |
|
69 |
} |
|
70 |
} |
|
71 |
||
72 |
Build_Stats(ml_options.toList, finished, timing, threads) |
|
73 |
} |
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
74 |
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
75 |
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
76 |
/* presentation */ |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
77 |
|
| 63700 | 78 |
private val default_history_length = 100 |
79 |
private val default_size = (800, 600) |
|
80 |
private val default_only_sessions = Set.empty[String] |
|
81 |
private val default_elapsed_threshold = Time.zero |
|
| 63706 | 82 |
private val default_ml_timing: Option[Boolean] = None |
| 63700 | 83 |
|
| 63703 | 84 |
def present_job(job: String, dir: Path, |
| 63700 | 85 |
history_length: Int = default_history_length, |
86 |
size: (Int, Int) = default_size, |
|
87 |
only_sessions: Set[String] = default_only_sessions, |
|
| 63706 | 88 |
elapsed_threshold: Time = default_elapsed_threshold, |
89 |
ml_timing: Option[Boolean] = default_ml_timing): List[String] = |
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
90 |
{
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
91 |
val build_infos = CI_API.build_job_builds(job).sortBy(_.timestamp).reverse.take(history_length) |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
92 |
if (build_infos.isEmpty) error("No build infos for job " + quote(job))
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
93 |
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
94 |
val all_build_stats = |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
95 |
Par_List.map((info: CI_API.Build_Info) => |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
96 |
(info.timestamp / 1000, parse(Url.read(info.output))), build_infos) |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
97 |
val all_sessions = |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
98 |
(Set.empty[String] /: all_build_stats)( |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
99 |
{ case (s, (_, stats)) => s ++ stats.sessions })
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
100 |
|
| 63700 | 101 |
def check_threshold(stats: Build_Stats, session: String): Boolean = |
102 |
stats.finished.get(session) match {
|
|
103 |
case Some(t) => t.elapsed >= elapsed_threshold |
|
104 |
case None => false |
|
105 |
} |
|
106 |
||
| 63703 | 107 |
val sessions = |
108 |
for {
|
|
109 |
session <- (if (only_sessions.isEmpty) all_sessions else all_sessions & only_sessions) |
|
110 |
if all_build_stats.filter({ case (_, stats) => check_threshold(stats, session) }).length >= 3
|
|
111 |
} yield session |
|
112 |
||
113 |
Isabelle_System.mkdirs(dir) |
|
114 |
for (session <- sessions) {
|
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
115 |
Isabelle_System.with_tmp_file(session, "png") { data_file =>
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
116 |
Isabelle_System.with_tmp_file(session, "gnuplot") { plot_file =>
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
117 |
val data = |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
118 |
for { (t, stats) <- all_build_stats if stats.finished.isDefinedAt(session) }
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
119 |
yield {
|
| 63708 | 120 |
val finished = stats.finished.getOrElse(session, Timing.zero) |
121 |
val timing = stats.timing.getOrElse(session, Timing.zero) |
|
| 63706 | 122 |
List(t.toString, finished.elapsed.minutes, finished.cpu.minutes, |
123 |
timing.elapsed.minutes, timing.cpu.minutes, timing.gc.minutes).mkString(" ")
|
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
124 |
} |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
125 |
File.write(data_file, cat_lines(data)) |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
126 |
|
| 63706 | 127 |
val plots1 = |
128 |
List( |
|
|
63707
b7aab1a6cf0d
clarified presentation order, according to typical amounts;
wenzelm
parents:
63706
diff
changeset
|
129 |
""" using 1:3 smooth sbezier title "cpu time (smooth)" """, |
|
b7aab1a6cf0d
clarified presentation order, according to typical amounts;
wenzelm
parents:
63706
diff
changeset
|
130 |
""" using 1:3 smooth csplines title "cpu time" """, |
| 63706 | 131 |
""" using 1:2 smooth sbezier title "elapsed time (smooth)" """, |
|
63707
b7aab1a6cf0d
clarified presentation order, according to typical amounts;
wenzelm
parents:
63706
diff
changeset
|
132 |
""" using 1:2 smooth csplines title "elapsed time" """) |
| 63706 | 133 |
val plots2 = |
134 |
List( |
|
|
63707
b7aab1a6cf0d
clarified presentation order, according to typical amounts;
wenzelm
parents:
63706
diff
changeset
|
135 |
""" using 1:5 smooth sbezier title "ML cpu time (smooth)" """, |
|
b7aab1a6cf0d
clarified presentation order, according to typical amounts;
wenzelm
parents:
63706
diff
changeset
|
136 |
""" using 1:5 smooth csplines title "ML cpu time" """, |
| 63706 | 137 |
""" using 1:4 smooth sbezier title "ML elapsed time (smooth)" """, |
138 |
""" using 1:4 smooth csplines title "ML elapsed time" """, |
|
139 |
""" using 1:6 smooth sbezier title "ML gc time (smooth)" """, |
|
140 |
""" using 1:6 smooth csplines title "ML gc time" """) |
|
141 |
val plots = |
|
142 |
ml_timing match {
|
|
143 |
case None => plots1 |
|
144 |
case Some(false) => plots1 ::: plots2 |
|
145 |
case Some(true) => plots2 |
|
146 |
} |
|
147 |
||
148 |
val data_file_name = File.standard_path(data_file.getAbsolutePath) |
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
149 |
File.write(plot_file, """ |
| 63700 | 150 |
set terminal png size """ + size._1 + "," + size._2 + """ |
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
151 |
set output """ + quote(File.standard_path(dir + Path.basic(session + ".png"))) + """ |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
152 |
set xdata time |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
153 |
set timefmt "%s" |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
154 |
set format x "%d-%b" |
| 63701 | 155 |
set xlabel """ + quote(session) + """ noenhanced |
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
156 |
set key left top |
| 63706 | 157 |
plot [] [0:] """ + plots.map(s => quote(data_file_name) + " " + s).mkString(", ") + "\n")
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
158 |
val result = Isabelle_System.bash("\"$ISABELLE_GNUPLOT\" " + File.bash_path(plot_file))
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
159 |
if (result.rc != 0) {
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
160 |
Output.error_message("Session " + session + ": gnuplot error")
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
161 |
result.print |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
162 |
} |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
163 |
} |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
164 |
} |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
165 |
} |
| 63703 | 166 |
|
167 |
sessions.toList.sorted |
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
168 |
} |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
169 |
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
170 |
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
171 |
/* Isabelle tool wrapper */ |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
172 |
|
| 63703 | 173 |
private val html_header = """<!DOCTYPE HTML PUBLIC "-//IETF//DTD HTML//EN"> |
174 |
<html> |
|
175 |
<head><title>Performance statistics from session build output</title></head> |
|
176 |
<body> |
|
177 |
""" |
|
178 |
private val html_footer = """ |
|
179 |
</body> |
|
180 |
</html> |
|
181 |
""" |
|
182 |
||
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
183 |
val isabelle_tool = |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
184 |
Isabelle_Tool("build_stats", "present statistics from session build output", args =>
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
185 |
{
|
| 63703 | 186 |
var target_dir = Path.explode("stats")
|
| 63706 | 187 |
var ml_timing = default_ml_timing |
| 63700 | 188 |
var only_sessions = default_only_sessions |
189 |
var elapsed_threshold = default_elapsed_threshold |
|
190 |
var history_length = default_history_length |
|
191 |
var size = default_size |
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
192 |
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
193 |
val getopts = Getopts("""
|
| 63700 | 194 |
Usage: isabelle build_stats [OPTIONS] [JOBS ...] |
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
195 |
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
196 |
Options are: |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
197 |
-D DIR target directory (default "stats") |
| 63706 | 198 |
-M only ML timing |
| 63700 | 199 |
-S SESSIONS only given SESSIONS (comma separated) |
200 |
-T THRESHOLD only sessions with elapsed time >= THRESHOLD (minutes) |
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
201 |
-l LENGTH length of history (default 100) |
| 63706 | 202 |
-m include ML timing |
| 63700 | 203 |
-s WxH size of PNG image (default 800x600) |
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
204 |
|
| 63700 | 205 |
Present statistics from session build output of the given JOBS, from Jenkins |
206 |
continuous build service specified as URL via ISABELLE_JENKINS_ROOT. |
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
207 |
""", |
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
208 |
"D:" -> (arg => target_dir = Path.explode(arg)), |
| 63706 | 209 |
"M" -> (_ => ml_timing = Some(true)), |
| 63700 | 210 |
"S:" -> (arg => only_sessions = space_explode(',', arg).toSet),
|
211 |
"T:" -> (arg => elapsed_threshold = Time.minutes(Properties.Value.Double.parse(arg))), |
|
212 |
"l:" -> (arg => history_length = Properties.Value.Int.parse(arg)), |
|
| 63706 | 213 |
"m" -> (_ => ml_timing = Some(false)), |
| 63700 | 214 |
"s:" -> (arg => |
215 |
space_explode('x', arg).map(Properties.Value.Int.parse(_)) match {
|
|
216 |
case List(w, h) if w > 0 && h > 0 => size = (w, h) |
|
217 |
case _ => error("Error bad PNG image size: " + quote(arg))
|
|
218 |
})) |
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
219 |
|
| 63700 | 220 |
val jobs = getopts(args) |
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
221 |
val all_jobs = CI_API.build_jobs() |
| 63700 | 222 |
val bad_jobs = jobs.filterNot(all_jobs.contains(_)).sorted |
223 |
||
224 |
if (jobs.isEmpty) |
|
225 |
error("No build jobs given. Available jobs: " + commas(all_jobs))
|
|
226 |
||
227 |
if (bad_jobs.nonEmpty) |
|
228 |
error("Unknown build jobs: " + commas(bad_jobs) +
|
|
229 |
"\nAvailable jobs: " + commas(all_jobs.sorted)) |
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
230 |
|
|
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
231 |
for (job <- jobs) {
|
| 63703 | 232 |
val dir = target_dir + Path.basic(job) |
233 |
Output.writeln(dir.implode) |
|
| 63706 | 234 |
val sessions = |
235 |
present_job(job, dir, history_length, size, only_sessions, elapsed_threshold, ml_timing) |
|
| 63703 | 236 |
File.write(dir + Path.basic("index.html"),
|
237 |
html_header + "\n<h1>" + HTML.output(job) + "</h1>\n" + |
|
238 |
cat_lines( |
|
239 |
sessions.map(session => |
|
240 |
"""<br/><img src=""" + quote(HTML.output(session + ".png")) + """><br/>""")) + |
|
241 |
"\n" + html_footer) |
|
|
63688
cc57255bf6ae
gnuplot presentation similar to former isatest-statistics;
wenzelm
parents:
63686
diff
changeset
|
242 |
} |
| 63703 | 243 |
|
244 |
File.write(target_dir + Path.basic("index.html"),
|
|
245 |
html_header + "\n<ul>\n" + |
|
246 |
cat_lines( |
|
247 |
jobs.map(job => """<li> <a href=""" + quote(HTML.output(job + "/index.html")) + """>""" + |
|
248 |
HTML.output(job) + """</a> </li>""")) + |
|
249 |
"\n</ul>\n" + html_footer) |
|
250 |
}) |
|
| 63686 | 251 |
} |
252 |
||
253 |
sealed case class Build_Stats( |
|
254 |
ml_options: List[(String, String)], |
|
255 |
finished: Map[String, Timing], |
|
256 |
timing: Map[String, Timing], |
|
257 |
threads: Map[String, Int]) |
|
258 |
{
|
|
259 |
val sessions: Set[String] = finished.keySet ++ timing.keySet |
|
260 |
||
261 |
override def toString: String = |
|
262 |
sessions.toList.sorted.mkString("Build_Stats(", ", ", ")")
|
|
263 |
} |