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