| author | wenzelm | 
| Fri, 24 Sep 2021 22:23:26 +0200 | |
| changeset 74362 | 0135a0c77b64 | 
| parent 73815 | 43882e34c038 | 
| child 75310 | 42baf7ffa088 | 
| 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 | |
| 69395 | 13 | object Components | 
| 14 | {
 | |
| 69413 | 15 | /* archive name */ | 
| 16 | ||
| 17 | object Archive | |
| 18 |   {
 | |
| 19 | 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 | 20 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 21 | 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 | 22 |       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 | 23 | else name + suffix | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 24 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 25 | def unapply(archive: String): Option[String] = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 26 |     {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 27 |       for {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 28 | 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 | 29 | name <- proper_string(name0) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 30 | } yield name | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 31 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 32 | |
| 69413 | 33 | def get_name(archive: String): String = | 
| 34 | unapply(archive) getOrElse | |
| 35 |         error("Bad component archive name (expecting .tar.gz): " + quote(archive))
 | |
| 36 | } | |
| 37 | ||
| 38 | ||
| 69395 | 39 | /* component collections */ | 
| 40 | ||
| 73240 
3e963d68d394
more robust ISABELLE_COMPONENT_REPOSITORY: use current value of managing process to avoid its fluctuation in ancient history;
 wenzelm parents: 
73173diff
changeset | 41 | 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 | 42 |     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 | 43 | |
| 71601 | 44 |   val default_components_base: Path = Path.explode("$ISABELLE_COMPONENTS_BASE")
 | 
| 69434 | 45 | |
| 69395 | 46 |   def admin(dir: Path): Path = dir + Path.explode("Admin/components")
 | 
| 47 | ||
| 48 | def contrib(dir: Path = Path.current, name: String = ""): Path = | |
| 49 |     dir + Path.explode("contrib") + Path.explode(name)
 | |
| 50 | ||
| 71726 
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
 wenzelm parents: 
71601diff
changeset | 51 | def unpack(dir: Path, archive: Path, progress: Progress = new Progress): String = | 
| 69413 | 52 |   {
 | 
| 53 | val name = Archive.get_name(archive.file_name) | |
| 54 |     progress.echo("Unpacking " + name)
 | |
| 69425 | 55 |     Isabelle_System.gnutar("-xzf " + File.bash_path(archive), dir = dir).check
 | 
| 69413 | 56 | name | 
| 57 | } | |
| 58 | ||
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 59 | 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 | 60 | target_dir: Option[Path] = None, | 
| 70102 | 61 | copy_dir: Option[Path] = None, | 
| 73340 | 62 | progress: Progress = new Progress): Unit = | 
| 69398 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 63 |   {
 | 
| 72375 | 64 | Isabelle_System.make_directory(base_dir) | 
| 69398 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 65 |     for (name <- names) {
 | 
| 69413 | 66 | val archive_name = Archive(name) | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 67 | 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 | 68 |       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 | 69 | val remote = Components.default_component_repository + "/" + archive_name | 
| 73566 | 70 | Isabelle_System.download_file(remote, archive, progress = progress) | 
| 69398 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 71 | } | 
| 70102 | 72 |       for (dir <- copy_dir) {
 | 
| 72375 | 73 | Isabelle_System.make_directory(dir) | 
| 73317 | 74 | Isabelle_System.copy_file(archive, dir) | 
| 70102 | 75 | } | 
| 69413 | 76 | unpack(target_dir getOrElse base_dir, archive, progress = progress) | 
| 69398 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 77 | } | 
| 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 78 | } | 
| 
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
 wenzelm parents: 
69395diff
changeset | 79 | |
| 73637 | 80 | private val platforms_family: Map[Platform.Family.Value, Set[String]] = | 
| 81 | Map( | |
| 82 |       Platform.Family.linux_arm -> Set("arm64-linux", "arm64_32-linux"),
 | |
| 83 |       Platform.Family.linux -> Set("x86_64-linux", "x86_64_32-linux"),
 | |
| 84 | Platform.Family.macos -> | |
| 85 |         Set("arm64-darwin", "arm64_32-darwin", "x86_64-darwin", "x86_64_32-darwin"),
 | |
| 86 | Platform.Family.windows -> | |
| 87 |         Set("x86_64-cygwin", "x86_64-windows", "x86_64_32-windows", "x86-windows"))
 | |
| 88 | ||
| 89 | private val platforms_all: Set[String] = | |
| 90 |     Set("x86-linux", "x86-cygwin") ++ platforms_family.iterator.flatMap(_._2)
 | |
| 91 | ||
| 73340 | 92 | def purge(dir: Path, platform: Platform.Family.Value): Unit = | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 93 |   {
 | 
| 73637 | 94 | 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 | 95 | |
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 96 | File.find_files(dir.file, | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 97 | (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 | 98 | 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 | 99 | } | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69398diff
changeset | 100 | |
| 69395 | 101 | |
| 73815 | 102 | /* component directories */ | 
| 103 | ||
| 104 | def directories(): List[Path] = | |
| 105 |     Path.split(Isabelle_System.getenv_strict("ISABELLE_COMPONENTS"))
 | |
| 106 | ||
| 107 | ||
| 69395 | 108 | /* component directory content */ | 
| 109 | ||
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 110 |   def settings(dir: Path = Path.current): Path = dir + Path.explode("etc/settings")
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 111 |   def components(dir: Path = Path.current): Path = dir + Path.explode("etc/components")
 | 
| 69395 | 112 | |
| 113 | def check_dir(dir: Path): Boolean = | |
| 114 | settings(dir).is_file || components(dir).is_file | |
| 115 | ||
| 116 | def read_components(dir: Path): List[String] = | |
| 117 | split_lines(File.read(components(dir))).filter(_.nonEmpty) | |
| 118 | ||
| 119 | def write_components(dir: Path, lines: List[String]): Unit = | |
| 120 | File.write(components(dir), terminate_lines(lines)) | |
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 121 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 122 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 123 | /* component repository content */ | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 124 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 125 |   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 | 126 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 127 | sealed case class SHA1_Digest(sha1: String, file_name: String) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 128 |   {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 129 | override def toString: String = sha1 + " " + file_name | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 130 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 131 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 132 | 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 | 133 | (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 | 134 |       Word.explode(line) match {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 135 | case Nil => None | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 136 | case List(sha1, name) => Some(SHA1_Digest(sha1, name)) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 137 |         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 | 138 | }) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 139 | |
| 71601 | 140 | def write_components_sha1(entries: List[SHA1_Digest]): Unit = | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 141 |     File.write(components_sha1, entries.sortBy(_.file_name).mkString("", "\n", "\n"))
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 142 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 143 | |
| 73172 | 144 | /** manage user components **/ | 
| 145 | ||
| 73173 | 146 |   val components_path = Path.explode("$ISABELLE_HOME_USER/etc/components")
 | 
| 73172 | 147 | |
| 148 | def read_components(): List[String] = | |
| 149 | if (components_path.is_file) Library.trim_split_lines(File.read(components_path)) | |
| 150 | else Nil | |
| 151 | ||
| 152 | def write_components(lines: List[String]): Unit = | |
| 153 |   {
 | |
| 154 | Isabelle_System.make_directory(components_path.dir) | |
| 155 | File.write(components_path, Library.terminate_lines(lines)) | |
| 156 | } | |
| 157 | ||
| 158 | def update_components(add: Boolean, path0: Path, progress: Progress = new Progress): Unit = | |
| 159 |   {
 | |
| 160 | val path = path0.expand.absolute | |
| 73814 
c8b4a4f69068
clarified check (refining fc828f64da5b): etc/settings or etc/components is not strictly required according to "init_component", and notable components only have session ROOTS (e.g. AFP/thys);
 wenzelm parents: 
73813diff
changeset | 161 |     if (!check_dir(path) && !Sessions.is_session_dir(path)) error("Bad component directory: " + path)
 | 
| 73172 | 162 | |
| 163 | val lines1 = read_components() | |
| 164 | val lines2 = | |
| 165 | lines1.filter(line => | |
| 166 |         line.isEmpty || line.startsWith("#") || !File.eq(Path.explode(line), path))
 | |
| 167 | val lines3 = if (add) lines2 ::: List(path.implode) else lines2 | |
| 168 | ||
| 169 | if (lines1 != lines3) write_components(lines3) | |
| 170 | ||
| 171 | val prefix = if (lines1 == lines3) "Unchanged" else if (add) "Added" else "Removed" | |
| 172 | progress.echo(prefix + " component " + path) | |
| 173 | } | |
| 174 | ||
| 175 | ||
| 176 | /* main entry point */ | |
| 177 | ||
| 178 | def main(args: Array[String]): Unit = | |
| 179 |   {
 | |
| 180 |     Command_Line.tool {
 | |
| 181 |       for (arg <- args) {
 | |
| 182 | val add = | |
| 183 |           if (arg.startsWith("+")) true
 | |
| 184 |           else if (arg.startsWith("-")) false
 | |
| 185 |           else error("Bad argument: " + quote(arg))
 | |
| 186 | val path = Path.explode(arg.substring(1)) | |
| 187 | update_components(add, path, progress = new Console_Progress) | |
| 188 | } | |
| 189 | } | |
| 190 | } | |
| 191 | ||
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 192 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 193 | /** build and publish components **/ | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 194 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 195 | def build_components( | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 196 | options: Options, | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 197 | components: List[Path], | 
| 71726 
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
 wenzelm parents: 
71601diff
changeset | 198 | 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 | 199 | publish: Boolean = false, | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 200 | force: Boolean = false, | 
| 73340 | 201 | update_components_sha1: Boolean = false): Unit = | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 202 |   {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 203 | val archives: List[Path] = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 204 |       for (path <- components) yield {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 205 |         path.file_name match {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 206 | case Archive(_) => path | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 207 | case name => | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 208 |             if (!path.is_dir) error("Bad component directory: " + path)
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 209 |             else if (!check_dir(path)) {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 210 |               error("Malformed component directory: " + path +
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 211 | "\n (requires " + settings() + " or " + Components.components() + ")") | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 212 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 213 |             else {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 214 | val component_path = path.expand | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 215 | val archive_dir = component_path.dir | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 216 | val archive_name = Archive(name) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 217 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 218 | val archive = archive_dir + Path.explode(archive_name) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 219 |               if (archive.is_file && !force) {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 220 |                 error("Component archive already exists: " + archive)
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 221 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 222 | |
| 69439 | 223 |               progress.echo("Packaging " + archive_name)
 | 
| 69430 | 224 |               Isabelle_System.gnutar("-czf " + File.bash_path(archive) + " " + Bash.string(name),
 | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 225 | dir = archive_dir).check | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 226 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 227 | archive | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 228 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 229 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 230 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 231 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 232 |     if ((publish && archives.nonEmpty) || update_components_sha1) {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 233 |       options.string("isabelle_components_server") match {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 234 | case SSH.Target(user, host) => | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 235 | using(SSH.open_session(options, host = host, user = user))(ssh => | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 236 |           {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 237 |             val components_dir = Path.explode(options.string("isabelle_components_dir"))
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 238 |             val contrib_dir = Path.explode(options.string("isabelle_components_contrib_dir"))
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 239 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 240 |             for (dir <- List(components_dir, contrib_dir) if !ssh.is_dir(dir)) {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 241 |               error("Bad remote directory: " + dir)
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 242 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 243 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 244 |             if (publish) {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 245 |               for (archive <- archives) {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 246 | val archive_name = archive.file_name | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 247 | val name = Archive.get_name(archive_name) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 248 | val remote_component = components_dir + archive.base | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 249 | val remote_contrib = contrib_dir + Path.explode(name) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 250 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 251 | // component archive | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 252 |                 if (ssh.is_file(remote_component) && !force) {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 253 |                   error("Remote component archive already exists: " + remote_component)
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 254 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 255 |                 progress.echo("Uploading " + archive_name)
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 256 | ssh.write_file(remote_component, archive) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 257 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 258 | // contrib directory | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 259 | val is_standard_component = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 260 |                   Isabelle_System.with_tmp_dir("component")(tmp_dir =>
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 261 |                   {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 262 |                     Isabelle_System.gnutar("-xzf " + File.bash_path(archive), dir = tmp_dir).check
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 263 | check_dir(tmp_dir + Path.explode(name)) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 264 | }) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 265 |                 if (is_standard_component) {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 266 |                   if (ssh.is_dir(remote_contrib)) {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 267 | if (force) ssh.rm_tree(remote_contrib) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 268 |                     else error("Remote component directory already exists: " + remote_contrib)
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 269 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 270 |                   progress.echo("Unpacking remote " + archive_name)
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 271 |                   ssh.execute("tar -C " + ssh.bash_path(contrib_dir) + " -xzf " +
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 272 | ssh.bash_path(remote_component)).check | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 273 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 274 |                 else {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 275 |                   progress.echo_warning("No unpacking of non-standard component: " + archive_name)
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 276 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 277 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 278 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 279 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 280 | // remote SHA1 digests | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 281 |             if (update_components_sha1) {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 282 | val lines = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 283 |                 for {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 284 | entry <- ssh.read_dir(components_dir) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 285 | if entry.is_file && entry.name.endsWith(Archive.suffix) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 286 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 287 |                 yield {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 288 |                   progress.echo("Digesting remote " + entry.name)
 | 
| 73277 
0110e2e2964c
clarified signature: always trim_line of Process_Result.out/err, uniformly in ML and Scala;
 wenzelm parents: 
73240diff
changeset | 289 |                   ssh.execute("cd " + ssh.bash_path(components_dir) +
 | 
| 
0110e2e2964c
clarified signature: always trim_line of Process_Result.out/err, uniformly in ML and Scala;
 wenzelm parents: 
73240diff
changeset | 290 | "; sha1sum " + Bash.string(entry.name)).check.out | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 291 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 292 | write_components_sha1(read_components_sha1(lines)) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 293 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 294 | }) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 295 |         case s => error("Bad isabelle_components_server: " + quote(s))
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 296 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 297 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 298 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 299 | // local SHA1 digests | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 300 |     {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 301 | val new_entries = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 302 | for (archive <- archives) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 303 |         yield {
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 304 | val file_name = archive.file_name | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 305 |           progress.echo("Digesting local " + file_name)
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 306 | val sha1 = SHA1.digest(archive).rep | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 307 | SHA1_Digest(sha1, file_name) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 308 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 309 | val new_names = new_entries.map(_.file_name).toSet | 
| 
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 | write_components_sha1( | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 312 | new_entries ::: | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 313 | read_components_sha1().filterNot(entry => new_names.contains(entry.file_name))) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 314 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 315 | } | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 316 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 317 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 318 | /* Isabelle tool wrapper */ | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 319 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 320 | private val relevant_options = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 321 |     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 | 322 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 323 | val isabelle_tool = | 
| 72763 | 324 |     Isabelle_Tool("build_components", "build and publish Isabelle components",
 | 
| 325 | Scala_Project.here, args => | |
| 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 | var publish = false | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 328 | var update_components_sha1 = false | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 329 | var force = false | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 330 | var options = Options.init() | 
| 
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 | def show_options: String = | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 333 | cat_lines(relevant_options.map(name => options.options(name).print)) | 
| 
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 getopts = Getopts("""
 | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 336 | 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 | 337 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 338 | Options are: | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 339 | -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 | 340 | -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 | 341 | -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 | 342 | -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 | 343 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 344 | 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 | 345 | depending on system options: | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 346 | |
| 73736 | 347 | """ + Library.indent_lines(2, show_options) + "\n", | 
| 69429 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 348 | "P" -> (_ => publish = true), | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 349 | "f" -> (_ => force = true), | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 350 | "o:" -> (arg => options = options + arg), | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 351 | "u" -> (_ => update_components_sha1 = true)) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 352 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 353 | val more_args = getopts(args) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 354 | if (more_args.isEmpty && !update_components_sha1) getopts.usage() | 
| 
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 | val progress = new Console_Progress | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 357 | |
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 358 | build_components(options, more_args.map(Path.explode), progress = progress, | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 359 | publish = publish, force = force, update_components_sha1 = update_components_sha1) | 
| 
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
 wenzelm parents: 
69426diff
changeset | 360 | }) | 
| 69395 | 361 | } |