| author | wenzelm |
| Fri, 04 Oct 2019 16:25:45 +0200 | |
| changeset 70785 | edaeb8feb4d0 |
| parent 70100 | d9ea307aac2a |
| child 70854 | 85c2cbd35632 |
| permissions | -rw-r--r-- |
| 66820 | 1 |
/* Title: Pure/Admin/afp.scala |
2 |
Author: Makarius |
|
3 |
||
4 |
Administrative support for the Archive of Formal Proofs. |
|
5 |
*/ |
|
6 |
||
7 |
package isabelle |
|
8 |
||
9 |
||
| 69980 | 10 |
import java.time.LocalDate |
| 69974 | 11 |
import scala.collection.immutable.SortedMap |
12 |
||
13 |
||
| 66820 | 14 |
object AFP |
15 |
{
|
|
|
69787
60b5a4731695
clarified URL -- avoid odd certificate problem with api.media.atlassian.com;
wenzelm
parents:
69693
diff
changeset
|
16 |
val repos_source = "https://isabelle.sketis.net/repos/afp-devel" |
| 66854 | 17 |
|
| 69693 | 18 |
val groups: Map[String, String] = |
19 |
Map("large" -> "full 64-bit memory model or word arithmetic required",
|
|
20 |
"slow" -> "CPU time much higher than 60min (on mid-range hardware)", |
|
21 |
"very_slow" -> "elapsed time of many hours (on high-end hardware)") |
|
22 |
||
23 |
def groups_bulky: List[String] = List("large", "slow")
|
|
24 |
||
| 66824 | 25 |
def init(options: Options, base_dir: Path = Path.explode("$AFP_BASE")): AFP =
|
26 |
new AFP(options, base_dir) |
|
| 66820 | 27 |
|
| 69973 | 28 |
|
29 |
/* entries */ |
|
30 |
||
|
69977
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
31 |
def parse_date(s: String): Date = |
|
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
32 |
{
|
|
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
33 |
val t = Date.Formatter.pattern("uuuu-MM-dd").parse(s)
|
| 69980 | 34 |
Date(LocalDate.from(t).atStartOfDay(Date.timezone_berlin)) |
|
69977
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
35 |
} |
|
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
36 |
|
| 69995 | 37 |
def trim_mail(s: String): String = s.replaceAll("<[^>]*>", "").trim
|
38 |
||
| 69975 | 39 |
sealed case class Entry(name: String, metadata: Properties.T, sessions: List[String]) |
| 69976 | 40 |
{
|
|
69977
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
41 |
def get(prop: String): Option[String] = Properties.get(metadata, prop) |
|
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
42 |
def get_string(prop: String): String = get(prop).getOrElse("")
|
|
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
43 |
def get_strings(prop: String): List[String] = |
| 70100 | 44 |
space_explode(',', get_string(prop)).map(_.trim).filter(_.nonEmpty)
|
|
69977
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
45 |
|
|
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
46 |
def title: String = get_string("title")
|
|
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
47 |
def authors: List[String] = get_strings("author")
|
| 69981 | 48 |
def date: Date = |
49 |
parse_date(get("date").getOrElse(error("Missing date for entry " + quote(name))))
|
|
|
69977
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
50 |
def topics: List[String] = get_strings("topic")
|
| 69981 | 51 |
def `abstract`: String = get_string("abstract").trim
|
|
69977
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
52 |
def maintainers: List[String] = get_strings("notify")
|
|
3c166df11085
clarified signature: explicitly typed interfaces;
wenzelm
parents:
69976
diff
changeset
|
53 |
def contributors: List[String] = get_strings("contributors")
|
| 69981 | 54 |
def license: String = get("license").getOrElse("BSD")
|
| 69982 | 55 |
|
56 |
def rdf_meta_data: Properties.T = |
|
57 |
RDF.meta_data( |
|
58 |
proper_string(title).map(Markup.META_TITLE -> _).toList ::: |
|
59 |
authors.map(Markup.META_CREATOR -> _) ::: |
|
60 |
contributors.map(Markup.META_CONTRIBUTOR -> _) ::: |
|
61 |
List(Markup.META_DATE -> RDF.date_format(date)) ::: |
|
62 |
List(Markup.META_LICENSE -> license) ::: |
|
63 |
proper_string(`abstract`).map(Markup.META_DESCRIPTION -> _).toList) |
|
| 69976 | 64 |
} |
| 66820 | 65 |
} |
66 |
||
| 66824 | 67 |
class AFP private(options: Options, val base_dir: Path) |
| 66820 | 68 |
{
|
| 66821 | 69 |
override def toString: String = base_dir.expand.toString |
70 |
||
| 66820 | 71 |
val main_dir: Path = base_dir + Path.explode("thys")
|
72 |
||
| 66821 | 73 |
|
| 69975 | 74 |
/* metadata */ |
75 |
||
76 |
private val entry_metadata: Map[String, Properties.T] = |
|
77 |
{
|
|
78 |
val metadata_file = base_dir + Path.explode("metadata/metadata")
|
|
79 |
||
80 |
var result = Map.empty[String, Properties.T] |
|
81 |
var section = "" |
|
82 |
var props = List.empty[Properties.Entry] |
|
83 |
||
84 |
val Section = """^\[(\S+)\]\s*$""".r |
|
85 |
val Property = """^(\S+)\s*=(.*)$""".r |
|
86 |
val Extra_Line = """^\s+(.*)$""".r |
|
87 |
val Blank_Line = """^\s*$""".r |
|
88 |
||
89 |
def flush() |
|
90 |
{
|
|
|
69978
4ecdd3eaec04
proper treatment of empty extra lines (amending 98a440cfbb2b);
wenzelm
parents:
69977
diff
changeset
|
91 |
if (section != "") result += (section -> props.reverse.filter(p => p._2.nonEmpty)) |
| 69975 | 92 |
section = "" |
93 |
props = Nil |
|
94 |
} |
|
95 |
||
96 |
for ((line, i) <- split_lines(File.read(metadata_file)).zipWithIndex) |
|
97 |
{
|
|
98 |
def err(msg: String): Nothing = |
|
99 |
error(msg + Position.here(Position.Line_File(i + 1, metadata_file.expand.implode))) |
|
100 |
||
101 |
line match {
|
|
102 |
case Section(name) => flush(); section = name |
|
103 |
case Property(a, b) => |
|
104 |
if (section == "") err("Property without a section")
|
|
|
69978
4ecdd3eaec04
proper treatment of empty extra lines (amending 98a440cfbb2b);
wenzelm
parents:
69977
diff
changeset
|
105 |
props = (a -> b.trim) :: props |
| 69975 | 106 |
case Extra_Line(line) => |
107 |
props match {
|
|
108 |
case Nil => err("Extra line without a property")
|
|
109 |
case (a, b) :: rest => props = (a, b + "\n" + line.trim) :: rest |
|
110 |
} |
|
111 |
case Blank_Line() => |
|
112 |
case _ => err("Bad input")
|
|
113 |
} |
|
114 |
} |
|
115 |
||
116 |
flush() |
|
117 |
result |
|
118 |
} |
|
119 |
||
120 |
||
| 69979 | 121 |
/* entries */ |
| 66821 | 122 |
|
| 69974 | 123 |
val entries_map: SortedMap[String, AFP.Entry] = |
| 69973 | 124 |
{
|
125 |
val entries = |
|
| 69974 | 126 |
for (name <- Sessions.parse_roots(main_dir + Sessions.ROOTS)) yield {
|
| 69973 | 127 |
val metadata = |
128 |
entry_metadata.getOrElse(name, error("Entry without metadata: " + quote(name)))
|
|
129 |
val sessions = |
|
130 |
Sessions.parse_root_entries(main_dir + Path.explode(name) + Sessions.ROOT).map(_.name) |
|
131 |
AFP.Entry(name, metadata, sessions) |
|
| 69974 | 132 |
} |
| 69973 | 133 |
|
| 69974 | 134 |
val entries_map = |
135 |
(SortedMap.empty[String, AFP.Entry] /: entries)({ case (m, e) => m + (e.name -> e) })
|
|
136 |
||
| 69973 | 137 |
val extra_metadata = |
| 69974 | 138 |
(for ((name, _) <- entry_metadata.iterator if !entries_map.isDefinedAt(name)) yield name). |
139 |
toList.sorted |
|
140 |
if (extra_metadata.nonEmpty) |
|
141 |
error("Meta data without entry: " + commas_quote(extra_metadata))
|
|
| 69973 | 142 |
|
| 69974 | 143 |
entries_map |
| 69973 | 144 |
} |
| 66820 | 145 |
|
| 69974 | 146 |
val entries: List[AFP.Entry] = entries_map.toList.map(_._2) |
| 69979 | 147 |
|
148 |
||
149 |
/* sessions */ |
|
150 |
||
151 |
val sessions_map: SortedMap[String, AFP.Entry] = |
|
152 |
(SortedMap.empty[String, AFP.Entry] /: entries)( |
|
153 |
{ case (m1, e) => (m1 /: e.sessions)({ case (m2, s) => m2 + (s -> e) }) })
|
|
154 |
||
| 66821 | 155 |
val sessions: List[String] = entries.flatMap(_.sessions) |
| 66824 | 156 |
|
| 67052 | 157 |
val sessions_structure: Sessions.Structure = |
| 67026 | 158 |
Sessions.load_structure(options, dirs = List(main_dir)). |
|
67025
961285f581e6
clarifified selection: always wrt. build_graph structure;
wenzelm
parents:
66861
diff
changeset
|
159 |
selection(Sessions.Selection(sessions = sessions.toList)) |
| 66821 | 160 |
|
161 |
||
| 66824 | 162 |
/* dependency graph */ |
| 66821 | 163 |
|
| 66824 | 164 |
private def sessions_deps(entry: AFP.Entry): List[String] = |
165 |
entry.sessions.flatMap(sessions_structure.imports_graph.imm_preds(_)).distinct.sorted |
|
166 |
||
|
66852
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
wenzelm
parents:
66832
diff
changeset
|
167 |
lazy val entries_graph: Graph[String, Unit] = |
| 66821 | 168 |
{
|
169 |
val session_entries = |
|
|
66831
29ea2b900a05
tuned: each session has at most one defining entry;
wenzelm
parents:
66824
diff
changeset
|
170 |
(Map.empty[String, String] /: entries) {
|
|
29ea2b900a05
tuned: each session has at most one defining entry;
wenzelm
parents:
66824
diff
changeset
|
171 |
case (m1, e) => (m1 /: e.sessions) { case (m2, s) => m2 + (s -> e.name) }
|
| 66821 | 172 |
} |
|
66831
29ea2b900a05
tuned: each session has at most one defining entry;
wenzelm
parents:
66824
diff
changeset
|
173 |
(Graph.empty[String, Unit] /: entries) { case (g, entry) =>
|
|
29ea2b900a05
tuned: each session has at most one defining entry;
wenzelm
parents:
66824
diff
changeset
|
174 |
val e1 = entry.name |
| 66832 | 175 |
(g.default_node(e1, ()) /: sessions_deps(entry)) { case (g1, s) =>
|
176 |
(g1 /: session_entries.get(s).filterNot(_ == e1)) { case (g2, e2) =>
|
|
|
66852
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
wenzelm
parents:
66832
diff
changeset
|
177 |
try { g2.default_node(e2, ()).add_edge_acyclic(e2, e1) }
|
|
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
wenzelm
parents:
66832
diff
changeset
|
178 |
catch {
|
|
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
wenzelm
parents:
66832
diff
changeset
|
179 |
case exn: Graph.Cycles[_] => |
|
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
wenzelm
parents:
66832
diff
changeset
|
180 |
error(cat_lines(exn.cycles.map(cycle => |
|
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
wenzelm
parents:
66832
diff
changeset
|
181 |
"Cyclic dependency of " + cycle.map(c => quote(c.toString)).mkString(" via ") +
|
|
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
wenzelm
parents:
66832
diff
changeset
|
182 |
" due to session " + quote(s)))) |
| 66832 | 183 |
} |
| 66821 | 184 |
} |
185 |
} |
|
| 66824 | 186 |
} |
| 66821 | 187 |
} |
188 |
||
| 66824 | 189 |
def entries_graph_display: Graph_Display.Graph = |
190 |
Graph_Display.make_graph(entries_graph) |
|
| 66823 | 191 |
|
| 66824 | 192 |
def entries_json_text: String = |
193 |
(for (entry <- entries.iterator) yield {
|
|
194 |
val distrib_deps = sessions_deps(entry).filterNot(sessions.contains(_)) |
|
195 |
val afp_deps = entries_graph.imm_preds(entry.name).toList |
|
196 |
""" |
|
| 66821 | 197 |
{""" + JSON.Format(entry.name) + """:
|
198 |
{"distrib_deps": """ + JSON.Format(distrib_deps) + """,
|
|
199 |
"afp_deps": """ + JSON.Format(afp_deps) + """ |
|
200 |
} |
|
201 |
}""" |
|
| 66824 | 202 |
}).mkString("[", ", ", "\n]\n")
|
|
66861
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
203 |
|
|
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
204 |
|
|
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
205 |
/* partition sessions */ |
|
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
206 |
|
| 67817 | 207 |
val force_partition1: List[String] = List("Category3", "HOL-ODE")
|
| 67776 | 208 |
|
|
66861
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
209 |
def partition(n: Int): List[String] = |
|
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
210 |
n match {
|
|
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
211 |
case 0 => Nil |
|
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
212 |
case 1 | 2 => |
|
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
213 |
val graph = sessions_structure.build_graph.restrict(sessions.toSet) |
| 67817 | 214 |
val force_part1 = |
215 |
graph.all_preds(graph.all_succs(force_partition1.filter(graph.defined(_)))).toSet |
|
216 |
val (part1, part2) = graph.keys.partition(a => force_part1(a) || graph.is_isolated(a)) |
|
217 |
if (n == 1) part1 else part2 |
|
|
66861
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
218 |
case _ => error("Bad AFP partition: " + n + " (should be 0, 1, 2)")
|
|
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
wenzelm
parents:
66854
diff
changeset
|
219 |
} |
| 66820 | 220 |
} |