| author | nipkow | 
| Thu, 04 Oct 2018 11:18:39 +0200 | |
| changeset 69118 | 12dce58bcd3f | 
| parent 67817 | 93faefc25fe7 | 
| child 69693 | 06153e2e0cdb | 
| 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 | ||
| 10 | object AFP | |
| 11 | {
 | |
| 66854 | 12 | val repos_source = "https://bitbucket.org/isa-afp/afp-devel" | 
| 13 | ||
| 66824 | 14 |   def init(options: Options, base_dir: Path = Path.explode("$AFP_BASE")): AFP =
 | 
| 15 | new AFP(options, base_dir) | |
| 66820 | 16 | |
| 66821 | 17 | sealed case class Entry(name: String, sessions: List[String]) | 
| 66820 | 18 | } | 
| 19 | ||
| 66824 | 20 | class AFP private(options: Options, val base_dir: Path) | 
| 66820 | 21 | {
 | 
| 66821 | 22 | override def toString: String = base_dir.expand.toString | 
| 23 | ||
| 66820 | 24 |   val main_dir: Path = base_dir + Path.explode("thys")
 | 
| 25 | ||
| 66821 | 26 | |
| 27 | /* entries and sessions */ | |
| 28 | ||
| 66820 | 29 | val entries: List[AFP.Entry] = | 
| 66821 | 30 |     (for (name <- Sessions.parse_roots(main_dir + Sessions.ROOTS)) yield {
 | 
| 66820 | 31 | val sessions = | 
| 32 | Sessions.parse_root_entries(main_dir + Path.explode(name) + Sessions.ROOT).map(_.name) | |
| 33 | AFP.Entry(name, sessions) | |
| 34 | }).sortBy(_.name) | |
| 35 | ||
| 66821 | 36 | val sessions: List[String] = entries.flatMap(_.sessions) | 
| 66824 | 37 | |
| 67052 | 38 | val sessions_structure: Sessions.Structure = | 
| 67026 | 39 | Sessions.load_structure(options, dirs = List(main_dir)). | 
| 67025 
961285f581e6
clarifified selection: always wrt. build_graph structure;
 wenzelm parents: 
66861diff
changeset | 40 | selection(Sessions.Selection(sessions = sessions.toList)) | 
| 66821 | 41 | |
| 42 | ||
| 66824 | 43 | /* dependency graph */ | 
| 66821 | 44 | |
| 66824 | 45 | private def sessions_deps(entry: AFP.Entry): List[String] = | 
| 46 | entry.sessions.flatMap(sessions_structure.imports_graph.imm_preds(_)).distinct.sorted | |
| 47 | ||
| 66852 
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
 wenzelm parents: 
66832diff
changeset | 48 | lazy val entries_graph: Graph[String, Unit] = | 
| 66821 | 49 |   {
 | 
| 50 | val session_entries = | |
| 66831 
29ea2b900a05
tuned: each session has at most one defining entry;
 wenzelm parents: 
66824diff
changeset | 51 |       (Map.empty[String, String] /: entries) {
 | 
| 
29ea2b900a05
tuned: each session has at most one defining entry;
 wenzelm parents: 
66824diff
changeset | 52 |         case (m1, e) => (m1 /: e.sessions) { case (m2, s) => m2 + (s -> e.name) }
 | 
| 66821 | 53 | } | 
| 66831 
29ea2b900a05
tuned: each session has at most one defining entry;
 wenzelm parents: 
66824diff
changeset | 54 |     (Graph.empty[String, Unit] /: entries) { case (g, entry) =>
 | 
| 
29ea2b900a05
tuned: each session has at most one defining entry;
 wenzelm parents: 
66824diff
changeset | 55 | val e1 = entry.name | 
| 66832 | 56 |       (g.default_node(e1, ()) /: sessions_deps(entry)) { case (g1, s) =>
 | 
| 57 |         (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: 
66832diff
changeset | 58 |           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: 
66832diff
changeset | 59 |           catch {
 | 
| 
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
 wenzelm parents: 
66832diff
changeset | 60 | case exn: Graph.Cycles[_] => | 
| 
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
 wenzelm parents: 
66832diff
changeset | 61 | error(cat_lines(exn.cycles.map(cycle => | 
| 
d20a668b394e
entries_graph requires acyclic graph, but lazy val allows forming the AFP object nonetheless;
 wenzelm parents: 
66832diff
changeset | 62 |                 "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: 
66832diff
changeset | 63 | " due to session " + quote(s)))) | 
| 66832 | 64 | } | 
| 66821 | 65 | } | 
| 66 | } | |
| 66824 | 67 | } | 
| 66821 | 68 | } | 
| 69 | ||
| 66824 | 70 | def entries_graph_display: Graph_Display.Graph = | 
| 71 | Graph_Display.make_graph(entries_graph) | |
| 66823 | 72 | |
| 66824 | 73 | def entries_json_text: String = | 
| 74 |     (for (entry <- entries.iterator) yield {
 | |
| 75 | val distrib_deps = sessions_deps(entry).filterNot(sessions.contains(_)) | |
| 76 | val afp_deps = entries_graph.imm_preds(entry.name).toList | |
| 77 | """ | |
| 66821 | 78 |  {""" + JSON.Format(entry.name) + """:
 | 
| 79 |   {"distrib_deps": """ + JSON.Format(distrib_deps) + """,
 | |
| 80 | "afp_deps": """ + JSON.Format(afp_deps) + """ | |
| 81 | } | |
| 82 | }""" | |
| 66824 | 83 |     }).mkString("[", ", ", "\n]\n")
 | 
| 66861 
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
 wenzelm parents: 
66854diff
changeset | 84 | |
| 
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
 wenzelm parents: 
66854diff
changeset | 85 | |
| 
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
 wenzelm parents: 
66854diff
changeset | 86 | /* partition sessions */ | 
| 
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
 wenzelm parents: 
66854diff
changeset | 87 | |
| 67817 | 88 |   val force_partition1: List[String] = List("Category3", "HOL-ODE")
 | 
| 67776 | 89 | |
| 66861 
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
 wenzelm parents: 
66854diff
changeset | 90 | 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 | 91 |     n match {
 | 
| 
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
 wenzelm parents: 
66854diff
changeset | 92 | case 0 => Nil | 
| 
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
 wenzelm parents: 
66854diff
changeset | 93 | case 1 | 2 => | 
| 
f6676691ef8a
partition AFP sessions according to structure, which happens to cut it roughly into equal parts;
 wenzelm parents: 
66854diff
changeset | 94 | val graph = sessions_structure.build_graph.restrict(sessions.toSet) | 
| 67817 | 95 | val force_part1 = | 
| 96 | graph.all_preds(graph.all_succs(force_partition1.filter(graph.defined(_)))).toSet | |
| 97 | val (part1, part2) = graph.keys.partition(a => force_part1(a) || graph.is_isolated(a)) | |
| 98 | 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 | 99 |       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 | 100 | } | 
| 66820 | 101 | } |