| author | Fabian Huch <huch@in.tum.de> | 
| Wed, 20 Mar 2024 14:56:16 +0100 | |
| changeset 79946 | 05e034a54924 | 
| parent 79651 | 155bb0ae4ae2 | 
| permissions | -rw-r--r-- | 
| 78939 | 1 | /* Title: Pure/System/registry.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Hierarchic system configuration in TOML notation. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 10 | object Registry {
 | |
| 78950 | 11 | /* registry files */ | 
| 12 | ||
| 78951 | 13 | def load(files: Iterable[Path]): Registry = new Registry(TOML.parse_files(files)) | 
| 14 | ||
| 15 | def global_files(): List[Path] = | |
| 78939 | 16 |     Path.split_permissive_files(Isabelle_System.getenv("ISABELLE_REGISTRY"))
 | 
| 17 | ||
| 78951 | 18 | lazy val global: Registry = load(global_files()) | 
| 78942 
409442cb7814
support interpreted/typed entries via Registry.Category and Registry.Table;
 wenzelm parents: 
78940diff
changeset | 19 | |
| 78950 | 20 | |
| 21 | /* interpreted entries */ | |
| 22 | ||
| 78962 | 23 |   trait Category {
 | 
| 24 | def prefix: String | |
| 78942 
409442cb7814
support interpreted/typed entries via Registry.Category and Registry.Table;
 wenzelm parents: 
78940diff
changeset | 25 |     override def toString: String = "Registry.Category(" + quote(prefix) + ")"
 | 
| 78961 | 26 | def qualify(name: String): String = Long_Name.qualify(prefix, name) | 
| 27 | def try_unqualify(name: String): Option[String] = Long_Name.try_unqualify(prefix, name) | |
| 28 | ||
| 78965 | 29 | type Value | 
| 78962 | 30 |     def bad_value(name: String): Nothing = error("Bad registry entry " + quote(qualify(name)))
 | 
| 78965 | 31 | def default_value(registry: Registry, name: String): Value | 
| 32 | def value(registry: Registry, t: TOML.T, name: String): Value | |
| 78963 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 33 | |
| 78965 | 34 |     def get(registry: Registry, name: String): Value = {
 | 
| 78963 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 35 |       registry.root.any.get(prefix) match {
 | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 36 | case None => default_value(registry, name) | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 37 | case Some(t: TOML.Table) => | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 38 |           t.any.get(name) match {
 | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 39 | case None => default_value(registry, name) | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 40 | case Some(u) => value(registry, u, name) | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 41 | } | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 42 | case Some(_) => bad_value(name) | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 43 | } | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 44 | } | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 45 | } | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 46 | |
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 47 |   trait Strict extends Category {
 | 
| 78965 | 48 | override def default_value(registry: Registry, name: String): Value = bad_value(name) | 
| 78942 
409442cb7814
support interpreted/typed entries via Registry.Category and Registry.Table;
 wenzelm parents: 
78940diff
changeset | 49 | } | 
| 
409442cb7814
support interpreted/typed entries via Registry.Category and Registry.Table;
 wenzelm parents: 
78940diff
changeset | 50 | |
| 78962 | 51 |   trait Table extends Category {
 | 
| 78965 | 52 | def table_value(registry: Registry, table: TOML.Table, name: String): Value | 
| 53 | override def default_value(registry: Registry, name: String): Value = | |
| 78963 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 54 | table_value(registry, TOML.Table.empty, name) | 
| 78965 | 55 | override def value(registry: Registry, t: TOML.T, name: String): Value = | 
| 78942 
409442cb7814
support interpreted/typed entries via Registry.Category and Registry.Table;
 wenzelm parents: 
78940diff
changeset | 56 |       t match {
 | 
| 78963 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 57 | case table: TOML.Table => table_value(registry, table, name) | 
| 78960 | 58 | case _ => bad_value(name) | 
| 78942 
409442cb7814
support interpreted/typed entries via Registry.Category and Registry.Table;
 wenzelm parents: 
78940diff
changeset | 59 | } | 
| 
409442cb7814
support interpreted/typed entries via Registry.Category and Registry.Table;
 wenzelm parents: 
78940diff
changeset | 60 | } | 
| 78945 
73162a487f94
build cluster host specifications are based on registry entries (table prefix "host");
 wenzelm parents: 
78942diff
changeset | 61 | |
| 78950 | 62 | |
| 63 | /* build cluster resources */ | |
| 64 | ||
| 78963 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 65 |   trait Host extends Table {
 | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 66 | def prefix = "host" | 
| 78965 | 67 | type Value = List[Options.Spec] | 
| 78962 | 68 | |
| 78945 
73162a487f94
build cluster host specifications are based on registry entries (table prefix "host");
 wenzelm parents: 
78942diff
changeset | 69 | def options_spec(a: TOML.Key, b: TOML.T): Option[Options.Spec] = | 
| 78947 | 70 | TOML.Scalar.unapply(b).map(Options.Spec.eq(a, _)) | 
| 78945 
73162a487f94
build cluster host specifications are based on registry entries (table prefix "host");
 wenzelm parents: 
78942diff
changeset | 71 | |
| 78965 | 72 | override def table_value(registry: Registry, t: TOML.Table, a: String): Value = | 
| 78945 
73162a487f94
build cluster host specifications are based on registry entries (table prefix "host");
 wenzelm parents: 
78942diff
changeset | 73 | for ((a, b) <- t.any.values; s <- options_spec(a, b)) yield s | 
| 
73162a487f94
build cluster host specifications are based on registry entries (table prefix "host");
 wenzelm parents: 
78942diff
changeset | 74 | } | 
| 78963 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 75 | object Host extends Host | 
| 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 76 | object Host_Strict extends Host with Strict | 
| 78964 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 77 |   object Host_Cluster extends Host with Strict { override def prefix = Cluster.prefix }
 | 
| 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 78 | |
| 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 79 |   object Cluster extends Table with Strict {
 | 
| 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 80 | def prefix = "cluster" | 
| 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 81 | def prefix_hosts = "hosts" | 
| 78965 | 82 | type Value = List[(String, List[Options.Spec])] | 
| 78964 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 83 | |
| 78965 | 84 |     override def table_value(registry: Registry, t: TOML.Table, a: String): Value = {
 | 
| 78964 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 85 | val hosts = | 
| 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 86 |         t.array.get(prefix_hosts) match {
 | 
| 79651 
155bb0ae4ae2
more robust: disallow empty clusters, so "isabelle build -H" really means cluster build;
 wenzelm parents: 
78965diff
changeset | 87 | case Some(arr) if arr.length > 0 => arr.string.values.map(_.rep) | 
| 
155bb0ae4ae2
more robust: disallow empty clusters, so "isabelle build -H" really means cluster build;
 wenzelm parents: 
78965diff
changeset | 88 | case _ => bad_value(Long_Name.qualify(a, prefix_hosts)) | 
| 78964 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 89 | } | 
| 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 90 | val cluster_specs = Host_Cluster.get(registry, a) | 
| 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 91 | hosts.map(h => (h, Host_Strict.get(registry, h) ::: cluster_specs)) | 
| 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 92 | } | 
| 
a2de1f6ff94e
support for "cluster" table with "hosts" array, and params/options as for "host" table;
 wenzelm parents: 
78963diff
changeset | 93 | } | 
| 78939 | 94 | } | 
| 95 | ||
| 78963 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 96 | class Registry private(val root: TOML.Table) {
 | 
| 78959 | 97 | override def toString: String = | 
| 78963 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 98 |     (for (a <- root.domain.toList.sorted.iterator) yield {
 | 
| 78959 | 99 | val size = | 
| 78963 
30360ee939f4
clarified signature: more operations, allow recursive get;
 wenzelm parents: 
78962diff
changeset | 100 |         root.any.get(a) match {
 | 
| 78959 | 101 |           case Some(t: TOML.Array) => "(" + t.length + ")"
 | 
| 102 |           case Some(t: TOML.Table) => "(" + t.domain.size + ")"
 | |
| 103 | case _ => "" | |
| 104 | } | |
| 105 | a + size | |
| 106 |     }).mkString("Registry(", ", ", ")")
 | |
| 78940 | 107 | } |