| author | wenzelm | 
| Fri, 20 Jan 2023 19:52:52 +0100 | |
| changeset 77028 | f5896dea6fce | 
| parent 76551 | c7996b073524 | 
| child 77055 | f56800b8b085 | 
| permissions | -rw-r--r-- | 
| 73815 | 1 | /* Title: Pure/System/components.scala | 
| 69395 | 2 | Author: Makarius | 
| 3 | ||
| 4 | Isabelle system components. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 10 | import java.io.{File => JFile}
 | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 11 | |
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 12 | |
| 75393 | 13 | object Components {
 | 
| 69413 | 14 | /* archive name */ | 
| 15 | ||
| 75393 | 16 |   object Archive {
 | 
| 69413 | 17 | val suffix: String = ".tar.gz" | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 18 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 19 | def apply(name: String): String = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 20 |       if (name == "") error("Bad component name: " + quote(name))
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 21 | else name + suffix | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 22 | |
| 75393 | 23 |     def unapply(archive: String): Option[String] = {
 | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 24 |       for {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 25 | name0 <- Library.try_unsuffix(suffix, archive) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 26 | name <- proper_string(name0) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 27 | } yield name | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 28 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 29 | |
| 69413 | 30 | def get_name(archive: String): String = | 
| 31 | unapply(archive) getOrElse | |
| 32 |         error("Bad component archive name (expecting .tar.gz): " + quote(archive))
 | |
| 33 | } | |
| 34 | ||
| 35 | ||
| 69395 | 36 | /* component collections */ | 
| 37 | ||
| 73240 
3e963d68d394
more robust ISABELLE_COMPONENT_REPOSITORY: use current value of managing process to avoid its fluctuation in ancient history;
 wenzelm parents: 
73173diff
changeset | 38 | def default_component_repository: String = | 
| 
3e963d68d394
more robust ISABELLE_COMPONENT_REPOSITORY: use current value of managing process to avoid its fluctuation in ancient history;
 wenzelm parents: 
73173diff
changeset | 39 |     Isabelle_System.getenv("ISABELLE_COMPONENT_REPOSITORY")
 | 
| 
3e963d68d394
more robust ISABELLE_COMPONENT_REPOSITORY: use current value of managing process to avoid its fluctuation in ancient history;
 wenzelm parents: 
73173diff
changeset | 40 | |
| 71601 | 41 |   val default_components_base: Path = Path.explode("$ISABELLE_COMPONENTS_BASE")
 | 
| 69434 | 42 | |
| 69395 | 43 |   def admin(dir: Path): Path = dir + Path.explode("Admin/components")
 | 
| 44 | ||
| 45 | def contrib(dir: Path = Path.current, name: String = ""): Path = | |
| 46 |     dir + Path.explode("contrib") + Path.explode(name)
 | |
| 47 | ||
| 75393 | 48 |   def unpack(dir: Path, archive: Path, progress: Progress = new Progress): String = {
 | 
| 69413 | 49 | val name = Archive.get_name(archive.file_name) | 
| 50 |     progress.echo("Unpacking " + name)
 | |
| 76540 
83de6e9ae983
clarified signature: prefer Scala functions instead of shell scripts;
 wenzelm parents: 
76519diff
changeset | 51 | Isabelle_System.extract(archive, dir) | 
| 69413 | 52 | name | 
| 53 | } | |
| 54 | ||
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 55 | def resolve(base_dir: Path, names: List[String], | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 56 | target_dir: Option[Path] = None, | 
| 70102 | 57 | copy_dir: Option[Path] = None, | 
| 75393 | 58 | progress: Progress = new Progress | 
| 59 |   ): Unit = {
 | |
| 72375 | 60 | Isabelle_System.make_directory(base_dir) | 
| 69398 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 61 |     for (name <- names) {
 | 
| 69413 | 62 | val archive_name = Archive(name) | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 63 | val archive = base_dir + Path.explode(archive_name) | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 64 |       if (!archive.is_file) {
 | 
| 73240 
3e963d68d394
more robust ISABELLE_COMPONENT_REPOSITORY: use current value of managing process to avoid its fluctuation in ancient history;
 wenzelm parents: 
73173diff
changeset | 65 | val remote = Components.default_component_repository + "/" + archive_name | 
| 73566 | 66 | Isabelle_System.download_file(remote, archive, progress = progress) | 
| 69398 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 67 | } | 
| 70102 | 68 |       for (dir <- copy_dir) {
 | 
| 72375 | 69 | Isabelle_System.make_directory(dir) | 
| 73317 | 70 | Isabelle_System.copy_file(archive, dir) | 
| 70102 | 71 | } | 
| 69413 | 72 | unpack(target_dir getOrElse base_dir, archive, progress = progress) | 
| 69398 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 73 | } | 
| 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 74 | } | 
| 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 75 | |
| 73637 | 76 | private val platforms_family: Map[Platform.Family.Value, Set[String]] = | 
| 77 | Map( | |
| 78 |       Platform.Family.linux_arm -> Set("arm64-linux", "arm64_32-linux"),
 | |
| 79 |       Platform.Family.linux -> Set("x86_64-linux", "x86_64_32-linux"),
 | |
| 80 | Platform.Family.macos -> | |
| 81 |         Set("arm64-darwin", "arm64_32-darwin", "x86_64-darwin", "x86_64_32-darwin"),
 | |
| 82 | Platform.Family.windows -> | |
| 83 |         Set("x86_64-cygwin", "x86_64-windows", "x86_64_32-windows", "x86-windows"))
 | |
| 84 | ||
| 85 | private val platforms_all: Set[String] = | |
| 86 |     Set("x86-linux", "x86-cygwin") ++ platforms_family.iterator.flatMap(_._2)
 | |
| 87 | ||
| 75393 | 88 |   def purge(dir: Path, platform: Platform.Family.Value): Unit = {
 | 
| 73637 | 89 | val purge_set = platforms_all -- platforms_family(platform) | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 90 | |
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 91 | File.find_files(dir.file, | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 92 | (file: JFile) => file.isDirectory && purge_set(file.getName), | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 93 | include_dirs = true).foreach(Isabelle_System.rm_tree) | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 94 | } | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 95 | |
| 69395 | 96 | |
| 73815 | 97 | /* component directories */ | 
| 98 | ||
| 99 | def directories(): List[Path] = | |
| 100 |     Path.split(Isabelle_System.getenv_strict("ISABELLE_COMPONENTS"))
 | |
| 101 | ||
| 102 | ||
| 69395 | 103 | /* component directory content */ | 
| 104 | ||
| 76518 | 105 |   object Directory {
 | 
| 106 | def apply(path: Path): Directory = new Directory(path.absolute) | |
| 107 | } | |
| 108 | ||
| 109 |   class Directory private(val path: Path) {
 | |
| 110 | override def toString: String = path.toString | |
| 111 | ||
| 112 |     def etc: Path = path + Path.basic("etc")
 | |
| 113 |     def src: Path = path + Path.basic("src")
 | |
| 114 |     def lib: Path = path + Path.basic("lib")
 | |
| 115 |     def settings: Path = etc + Path.basic("settings")
 | |
| 116 |     def components: Path = etc + Path.basic("components")
 | |
| 117 |     def build_props: Path = etc + Path.basic("build.props")
 | |
| 118 |     def README: Path = path + Path.basic("README")
 | |
| 119 |     def LICENSE: Path = path + Path.basic("LICENSE")
 | |
| 120 | ||
| 76547 | 121 |     def create(progress: Progress = new Progress): Directory = {
 | 
| 122 |       progress.echo("Creating component directory " + path)
 | |
| 123 | Isabelle_System.new_directory(path) | |
| 124 | Isabelle_System.make_directory(etc) | |
| 125 | this | |
| 126 | } | |
| 127 | ||
| 76550 | 128 | def ok: Boolean = settings.is_file || components.is_file || Sessions.is_session_dir(path) | 
| 129 | ||
| 130 | def check: Directory = | |
| 131 |       if (!path.is_dir) error("Bad component directory: " + path)
 | |
| 132 |       else if (!ok) {
 | |
| 133 |         error("Malformed component directory: " + path +
 | |
| 134 | "\n (missing \"etc/settings\" or \"etc/components\" or \"ROOT\" or \"ROOTS\")") | |
| 135 | } | |
| 136 | else this | |
| 69395 | 137 | |
| 76519 | 138 | def read_components(): List[String] = | 
| 139 | split_lines(File.read(components)).filter(_.nonEmpty) | |
| 140 | def write_components(lines: List[String]): Unit = | |
| 141 | File.write(components, terminate_lines(lines)) | |
| 76548 | 142 | |
| 143 | def write_settings(text: String): Unit = | |
| 144 | File.write(settings, "# -*- shell-script -*- :mode=shellscript:\n" + text) | |
| 76519 | 145 | } | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 146 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 147 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 148 | /* component repository content */ | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 149 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 150 |   val components_sha1: Path = Path.explode("~~/Admin/components/components.sha1")
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 151 | |
| 75393 | 152 |   sealed case class SHA1_Digest(digest: SHA1.Digest, name: String) {
 | 
| 75350 
93943e7e38a4
prefer Isabelle shasum over the old command-line tool with its extra marker character;
 wenzelm parents: 
75349diff
changeset | 153 | override def toString: String = digest.shasum(name) | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 154 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 155 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 156 | def read_components_sha1(lines: List[String] = Nil): List[SHA1_Digest] = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 157 | (proper_list(lines) getOrElse split_lines(File.read(components_sha1))).flatMap(line => | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 158 |       Word.explode(line) match {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 159 | case Nil => None | 
| 75349 | 160 | case List(sha1, name) => Some(SHA1_Digest(SHA1.fake_digest(sha1), name)) | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 161 |         case _ => error("Bad components.sha1 entry: " + quote(line))
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 162 | }) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 163 | |
| 71601 | 164 | def write_components_sha1(entries: List[SHA1_Digest]): Unit = | 
| 75349 | 165 |     File.write(components_sha1, entries.sortBy(_.name).mkString("", "\n", "\n"))
 | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 166 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 167 | |
| 73172 | 168 | /** manage user components **/ | 
| 169 | ||
| 75349 | 170 |   val components_path: Path = Path.explode("$ISABELLE_HOME_USER/etc/components")
 | 
| 73172 | 171 | |
| 172 | def read_components(): List[String] = | |
| 173 | if (components_path.is_file) Library.trim_split_lines(File.read(components_path)) | |
| 174 | else Nil | |
| 175 | ||
| 75393 | 176 |   def write_components(lines: List[String]): Unit = {
 | 
| 73172 | 177 | Isabelle_System.make_directory(components_path.dir) | 
| 178 | File.write(components_path, Library.terminate_lines(lines)) | |
| 179 | } | |
| 180 | ||
| 75393 | 181 |   def update_components(add: Boolean, path0: Path, progress: Progress = new Progress): Unit = {
 | 
| 73172 | 182 | val path = path0.expand.absolute | 
| 76551 
c7996b073524
clarified check: allow to remove bad directories;
 wenzelm parents: 
76550diff
changeset | 183 | if (add) Directory(path).check | 
| 73172 | 184 | |
| 185 | val lines1 = read_components() | |
| 186 | val lines2 = | |
| 187 | lines1.filter(line => | |
| 188 |         line.isEmpty || line.startsWith("#") || !File.eq(Path.explode(line), path))
 | |
| 189 | val lines3 = if (add) lines2 ::: List(path.implode) else lines2 | |
| 190 | ||
| 191 | if (lines1 != lines3) write_components(lines3) | |
| 192 | ||
| 193 | val prefix = if (lines1 == lines3) "Unchanged" else if (add) "Added" else "Removed" | |
| 194 | progress.echo(prefix + " component " + path) | |
| 195 | } | |
| 196 | ||
| 197 | ||
| 198 | /* main entry point */ | |
| 199 | ||
| 75393 | 200 |   def main(args: Array[String]): Unit = {
 | 
| 73172 | 201 |     Command_Line.tool {
 | 
| 202 |       for (arg <- args) {
 | |
| 203 | val add = | |
| 204 |           if (arg.startsWith("+")) true
 | |
| 205 |           else if (arg.startsWith("-")) false
 | |
| 206 |           else error("Bad argument: " + quote(arg))
 | |
| 207 | val path = Path.explode(arg.substring(1)) | |
| 208 | update_components(add, path, progress = new Console_Progress) | |
| 209 | } | |
| 210 | } | |
| 211 | } | |
| 212 | ||
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 213 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 214 | /** build and publish components **/ | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 215 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 216 | def build_components( | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 217 | options: Options, | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 218 | components: List[Path], | 
| 71726 
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
 wenzelm parents: 
71601diff
changeset | 219 | progress: Progress = new Progress, | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 220 | publish: Boolean = false, | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 221 | force: Boolean = false, | 
| 75393 | 222 | update_components_sha1: Boolean = false | 
| 223 |   ): Unit = {
 | |
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 224 | val archives: List[Path] = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 225 |       for (path <- components) yield {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 226 |         path.file_name match {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 227 | case Archive(_) => path | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 228 | case name => | 
| 76550 | 229 | Directory(path).check | 
| 230 | val component_path = path.expand | |
| 231 | val archive_dir = component_path.dir | |
| 232 | val archive_name = Archive(name) | |
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 233 | |
| 76550 | 234 | val archive = archive_dir + Path.explode(archive_name) | 
| 235 |             if (archive.is_file && !force) {
 | |
| 236 |               error("Component archive already exists: " + archive)
 | |
| 237 | } | |
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 238 | |
| 76550 | 239 |             progress.echo("Packaging " + archive_name)
 | 
| 240 |             Isabelle_System.gnutar("-czf " + File.bash_path(archive) + " " + Bash.string(name),
 | |
| 241 | dir = archive_dir).check | |
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 242 | |
| 76550 | 243 | archive | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 244 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 245 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 246 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 247 |     if ((publish && archives.nonEmpty) || update_components_sha1) {
 | 
| 76169 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 248 |       val server = options.string("isabelle_components_server")
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 249 |       if (server.isEmpty) error("Undefined option isabelle_components_server")
 | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 250 | |
| 76169 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 251 |       using(SSH.open_session(options, server)) { ssh =>
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 252 |         val components_dir = Path.explode(options.string("isabelle_components_dir"))
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 253 |         val contrib_dir = Path.explode(options.string("isabelle_components_contrib_dir"))
 | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 254 | |
| 76169 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 255 |         for (dir <- List(components_dir, contrib_dir) if !ssh.is_dir(dir)) {
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 256 |           error("Bad remote directory: " + dir)
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 257 | } | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 258 | |
| 76169 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 259 |         if (publish) {
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 260 |           for (archive <- archives) {
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 261 | val archive_name = archive.file_name | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 262 | val name = Archive.get_name(archive_name) | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 263 | val remote_component = components_dir + archive.base | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 264 | val remote_contrib = contrib_dir + Path.explode(name) | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 265 | |
| 76169 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 266 | // component archive | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 267 |             if (ssh.is_file(remote_component) && !force) {
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 268 |               error("Remote component archive already exists: " + remote_component)
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 269 | } | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 270 |             progress.echo("Uploading " + archive_name)
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 271 | ssh.write_file(remote_component, archive) | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 272 | |
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 273 | // contrib directory | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 274 | val is_standard_component = | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 275 |               Isabelle_System.with_tmp_dir("component") { tmp_dir =>
 | 
| 76540 
83de6e9ae983
clarified signature: prefer Scala functions instead of shell scripts;
 wenzelm parents: 
76519diff
changeset | 276 | Isabelle_System.extract(archive, tmp_dir) | 
| 76550 | 277 | Directory(tmp_dir + Path.explode(name)).ok | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 278 | } | 
| 76169 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 279 |             if (is_standard_component) {
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 280 |               if (ssh.is_dir(remote_contrib)) {
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 281 | if (force) ssh.rm_tree(remote_contrib) | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 282 |                 else error("Remote component directory already exists: " + remote_contrib)
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 283 | } | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 284 |               progress.echo("Unpacking remote " + archive_name)
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 285 |               ssh.execute("tar -C " + ssh.bash_path(contrib_dir) + " -xzf " +
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 286 | ssh.bash_path(remote_component)).check | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 287 | } | 
| 76169 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 288 |             else {
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 289 |               progress.echo_warning("No unpacking of non-standard component: " + archive_name)
 | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 290 | } | 
| 75394 | 291 | } | 
| 76169 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 292 | } | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 293 | |
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 294 | // remote SHA1 digests | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 295 |         if (update_components_sha1) {
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 296 | val lines = | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 297 |             for {
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 298 | entry <- ssh.read_dir(components_dir) | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 299 | if ssh.is_file(components_dir + Path.basic(entry)) && | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 300 | entry.endsWith(Archive.suffix) | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 301 | } | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 302 |             yield {
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 303 |               progress.echo("Digesting remote " + entry)
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 304 |               ssh.execute("cd " + ssh.bash_path(components_dir) +
 | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 305 | "; sha1sum " + Bash.string(entry)).check.out | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 306 | } | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 307 | write_components_sha1(read_components_sha1(lines)) | 
| 
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
 wenzelm parents: 
76122diff
changeset | 308 | } | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 309 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 310 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 311 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 312 | // local SHA1 digests | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 313 |     {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 314 | val new_entries = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 315 | for (archive <- archives) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 316 |         yield {
 | 
| 75349 | 317 | val name = archive.file_name | 
| 318 |           progress.echo("Digesting local " + name)
 | |
| 319 | SHA1_Digest(SHA1.digest(archive), name) | |
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 320 | } | 
| 75349 | 321 | val new_names = new_entries.map(_.name).toSet | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 322 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 323 | write_components_sha1( | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 324 | new_entries ::: | 
| 75349 | 325 | read_components_sha1().filterNot(entry => new_names.contains(entry.name))) | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 326 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 327 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 328 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 329 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 330 | /* Isabelle tool wrapper */ | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 331 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 332 | private val relevant_options = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 333 |     List("isabelle_components_server", "isabelle_components_dir", "isabelle_components_contrib_dir")
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 334 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 335 | val isabelle_tool = | 
| 72763 | 336 |     Isabelle_Tool("build_components", "build and publish Isabelle components",
 | 
| 75394 | 337 | Scala_Project.here, | 
| 338 |       { args =>
 | |
| 339 | var publish = false | |
| 340 | var update_components_sha1 = false | |
| 341 | var force = false | |
| 342 | var options = Options.init() | |
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 343 | |
| 75394 | 344 | def show_options: String = | 
| 75844 
7d27944d7141
clarified signature: avoid public representation;
 wenzelm parents: 
75394diff
changeset | 345 | cat_lines(relevant_options.flatMap(options.get).map(_.print)) | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 346 | |
| 75394 | 347 |         val getopts = Getopts("""
 | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 348 | Usage: isabelle build_components [OPTIONS] ARCHIVES... DIRS... | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 349 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 350 | Options are: | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 351 | -P publish on SSH server (see options below) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 352 | -f force: overwrite existing component archives and directories | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 353 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 354 | -u update all SHA1 keys in Isabelle repository Admin/components | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 355 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 356 | Build and publish Isabelle components as .tar.gz archives on SSH server, | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 357 | depending on system options: | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 358 | |
| 73736 | 359 | """ + Library.indent_lines(2, show_options) + "\n", | 
| 75394 | 360 | "P" -> (_ => publish = true), | 
| 361 | "f" -> (_ => force = true), | |
| 362 | "o:" -> (arg => options = options + arg), | |
| 363 | "u" -> (_ => update_components_sha1 = true)) | |
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 364 | |
| 75394 | 365 | val more_args = getopts(args) | 
| 366 | if (more_args.isEmpty && !update_components_sha1) getopts.usage() | |
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 367 | |
| 75394 | 368 | val progress = new Console_Progress | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 369 | |
| 75394 | 370 | build_components(options, more_args.map(Path.explode), progress = progress, | 
| 371 | publish = publish, force = force, update_components_sha1 = update_components_sha1) | |
| 372 | }) | |
| 69395 | 373 | } |