| author | wenzelm |
| Wed, 01 Mar 2023 13:30:35 +0100 | |
| changeset 77438 | 0030eabbe6c3 |
| parent 77216 | ee7dc5151db5 |
| child 77509 | 3bc49507bae5 |
| 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:
69398
diff
changeset
|
10 |
import java.io.{File => JFile}
|
|
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
11 |
|
|
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
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:
69426
diff
changeset
|
18 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
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:
69426
diff
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:
69426
diff
changeset
|
21 |
else name + suffix |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
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:
69426
diff
changeset
|
24 |
for {
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
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:
69426
diff
changeset
|
26 |
name <- proper_string(name0) |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
27 |
} yield name |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
28 |
} |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
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 |
||
|
77097
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
36 |
/* platforms */ |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
37 |
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
38 |
private val family_platforms: Map[Platform.Family.Value, List[String]] = |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
39 |
Map( |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
40 |
Platform.Family.linux_arm -> List("arm64-linux", "arm64_32-linux"),
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
41 |
Platform.Family.linux -> List("x86_64-linux", "x86_64_32-linux"),
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
42 |
Platform.Family.macos -> |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
43 |
List("arm64-darwin", "arm64_32-darwin", "x86_64-darwin", "x86_64_32-darwin"),
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
44 |
Platform.Family.windows -> |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
45 |
List("x86_64-cygwin", "x86_64-windows", "x86_64_32-windows", "x86-windows"))
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
46 |
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
47 |
private val platform_names: Set[String] = |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
48 |
Set("x86-linux", "x86-cygwin") ++ family_platforms.iterator.flatMap(_._2)
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
49 |
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
50 |
def platform_purge(platforms: List[Platform.Family.Value]): String => Boolean = {
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
51 |
val preserve = |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
52 |
(for {
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
53 |
family <- platforms.iterator |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
54 |
platform <- family_platforms(family) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
55 |
} yield platform).toSet |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
56 |
(name: String) => platform_names(name) && !preserve(name) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
57 |
} |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
58 |
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
59 |
|
| 69395 | 60 |
/* component collections */ |
61 |
||
|
77128
f40c36ab154d
clarified names to emphasize suble differences in meaning;
wenzelm
parents:
77098
diff
changeset
|
62 |
def static_component_repository: String = |
|
73240
3e963d68d394
more robust ISABELLE_COMPONENT_REPOSITORY: use current value of managing process to avoid its fluctuation in ancient history;
wenzelm
parents:
73173
diff
changeset
|
63 |
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:
73173
diff
changeset
|
64 |
|
| 71601 | 65 |
val default_components_base: Path = Path.explode("$ISABELLE_COMPONENTS_BASE")
|
|
77128
f40c36ab154d
clarified names to emphasize suble differences in meaning;
wenzelm
parents:
77098
diff
changeset
|
66 |
val dynamic_components_base: String = "${ISABELLE_COMPONENTS_BASE:-$USER_HOME/.isabelle/contrib}"
|
| 69434 | 67 |
|
| 77055 | 68 |
val default_catalogs: List[String] = List("main")
|
69 |
val optional_catalogs: List[String] = List("main", "optional")
|
|
70 |
||
| 69395 | 71 |
def admin(dir: Path): Path = dir + Path.explode("Admin/components")
|
72 |
||
73 |
def contrib(dir: Path = Path.current, name: String = ""): Path = |
|
74 |
dir + Path.explode("contrib") + Path.explode(name)
|
|
75 |
||
| 77059 | 76 |
def unpack( |
77 |
dir: Path, |
|
78 |
archive: Path, |
|
79 |
ssh: SSH.System = SSH.Local, |
|
80 |
progress: Progress = new Progress |
|
81 |
): String = {
|
|
| 69413 | 82 |
val name = Archive.get_name(archive.file_name) |
|
77098
9d6118cdc0fd
tuned message, following "isabelle components -a";
wenzelm
parents:
77097
diff
changeset
|
83 |
progress.echo("Unpacking " + archive.base)
|
| 77059 | 84 |
ssh.execute( |
| 77066 | 85 |
"tar -C " + ssh.bash_path(dir) + " -x -z -f " + ssh.bash_path(archive), |
| 77059 | 86 |
progress_stdout = progress.echo, |
87 |
progress_stderr = progress.echo).check |
|
| 69413 | 88 |
name |
89 |
} |
|
90 |
||
|
77097
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
91 |
def clean( |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
92 |
component_dir: Path, |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
93 |
platforms: List[Platform.Family.Value] = Platform.Family.list, |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
94 |
ssh: SSH.System = SSH.Local, |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
95 |
progress: Progress = new Progress |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
96 |
): Unit = {
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
97 |
val purge = platform_purge(platforms) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
98 |
for {
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
99 |
name <- ssh.read_dir(component_dir) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
100 |
path = Path.basic(name) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
101 |
if purge(name) && ssh.is_dir(component_dir + path) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
102 |
} {
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
103 |
progress.echo("Removing " + (component_dir.base + path))
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
104 |
ssh.rm_tree(component_dir + path) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
105 |
} |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
106 |
} |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
107 |
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
108 |
def clean_base( |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
109 |
base_dir: Path, |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
110 |
platforms: List[Platform.Family.Value] = Platform.Family.list, |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
111 |
ssh: SSH.System = SSH.Local, |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
112 |
progress: Progress = new Progress |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
113 |
): Unit = {
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
114 |
for {
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
115 |
name <- ssh.read_dir(base_dir) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
116 |
dir = base_dir + Path.basic(name) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
117 |
if is_component_dir(dir) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
118 |
} clean(dir, platforms = platforms, ssh = ssh, progress = progress) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
119 |
} |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
120 |
|
| 77059 | 121 |
def resolve( |
122 |
base_dir: Path, |
|
| 77067 | 123 |
name: String, |
|
69401
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
124 |
target_dir: Option[Path] = None, |
| 70102 | 125 |
copy_dir: Option[Path] = None, |
|
77097
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
126 |
clean_platforms: Option[List[Platform.Family.Value]] = None, |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
127 |
clean_archives: Boolean = false, |
|
77128
f40c36ab154d
clarified names to emphasize suble differences in meaning;
wenzelm
parents:
77098
diff
changeset
|
128 |
component_repository: String = Components.static_component_repository, |
| 77059 | 129 |
ssh: SSH.System = SSH.Local, |
| 75393 | 130 |
progress: Progress = new Progress |
131 |
): Unit = {
|
|
| 77059 | 132 |
ssh.make_directory(base_dir) |
|
77097
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
133 |
|
| 77067 | 134 |
val archive_name = Archive(name) |
| 77068 | 135 |
val archive = base_dir + Path.basic(archive_name) |
| 77067 | 136 |
if (!ssh.is_file(archive)) {
|
137 |
val remote = Url.append_path(component_repository, archive_name) |
|
138 |
ssh.download_file(remote, archive, progress = progress) |
|
|
69398
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
wenzelm
parents:
69395
diff
changeset
|
139 |
} |
|
77097
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
140 |
|
| 77067 | 141 |
for (dir <- copy_dir) {
|
142 |
ssh.make_directory(dir) |
|
143 |
ssh.copy_file(archive, dir) |
|
144 |
} |
|
|
77097
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
145 |
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
146 |
val unpack_dir = target_dir getOrElse base_dir |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
147 |
unpack(unpack_dir, archive, ssh = ssh, progress = progress) |
|
69398
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
wenzelm
parents:
69395
diff
changeset
|
148 |
|
|
77097
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
149 |
if (clean_platforms.isDefined) {
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
150 |
clean(unpack_dir + Path.basic(name), |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
151 |
platforms = clean_platforms.get, ssh = ssh, progress = progress) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
152 |
} |
| 73637 | 153 |
|
|
77097
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
154 |
if (clean_archives) {
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
155 |
progress.echo("Removing " + quote(archive_name))
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
156 |
ssh.delete(archive) |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
157 |
} |
|
69401
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
158 |
} |
|
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
159 |
|
| 69395 | 160 |
|
| 73815 | 161 |
/* component directories */ |
162 |
||
163 |
def directories(): List[Path] = |
|
164 |
Path.split(Isabelle_System.getenv_strict("ISABELLE_COMPONENTS"))
|
|
165 |
||
|
77097
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
166 |
def is_component_dir(dir: Path, ssh: SSH.System = SSH.Local): Boolean = |
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
167 |
ssh.is_file(dir + Path.explode("etc/settings")) ||
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
168 |
ssh.is_file(dir + Path.explode("etc/components"))
|
|
023273cf2651
clean components more accurately: purge other platforms or archives;
wenzelm
parents:
77090
diff
changeset
|
169 |
|
| 73815 | 170 |
|
| 69395 | 171 |
/* component directory content */ |
172 |
||
| 76518 | 173 |
object Directory {
|
174 |
def apply(path: Path): Directory = new Directory(path.absolute) |
|
175 |
} |
|
176 |
||
177 |
class Directory private(val path: Path) {
|
|
178 |
override def toString: String = path.toString |
|
179 |
||
180 |
def etc: Path = path + Path.basic("etc")
|
|
181 |
def src: Path = path + Path.basic("src")
|
|
182 |
def lib: Path = path + Path.basic("lib")
|
|
183 |
def settings: Path = etc + Path.basic("settings")
|
|
184 |
def components: Path = etc + Path.basic("components")
|
|
185 |
def build_props: Path = etc + Path.basic("build.props")
|
|
186 |
def README: Path = path + Path.basic("README")
|
|
187 |
def LICENSE: Path = path + Path.basic("LICENSE")
|
|
188 |
||
| 76547 | 189 |
def create(progress: Progress = new Progress): Directory = {
|
190 |
progress.echo("Creating component directory " + path)
|
|
191 |
Isabelle_System.new_directory(path) |
|
192 |
Isabelle_System.make_directory(etc) |
|
193 |
this |
|
194 |
} |
|
195 |
||
| 76550 | 196 |
def ok: Boolean = settings.is_file || components.is_file || Sessions.is_session_dir(path) |
197 |
||
198 |
def check: Directory = |
|
199 |
if (!path.is_dir) error("Bad component directory: " + path)
|
|
200 |
else if (!ok) {
|
|
201 |
error("Malformed component directory: " + path +
|
|
202 |
"\n (missing \"etc/settings\" or \"etc/components\" or \"ROOT\" or \"ROOTS\")") |
|
203 |
} |
|
204 |
else this |
|
| 69395 | 205 |
|
| 76519 | 206 |
def read_components(): List[String] = |
207 |
split_lines(File.read(components)).filter(_.nonEmpty) |
|
208 |
def write_components(lines: List[String]): Unit = |
|
209 |
File.write(components, terminate_lines(lines)) |
|
| 76548 | 210 |
|
211 |
def write_settings(text: String): Unit = |
|
212 |
File.write(settings, "# -*- shell-script -*- :mode=shellscript:\n" + text) |
|
| 76519 | 213 |
} |
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
214 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
215 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
216 |
/* component repository content */ |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
217 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
218 |
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:
69426
diff
changeset
|
219 |
|
| 77209 | 220 |
sealed case class SHA1_Entry(digest: SHA1.Digest, name: String) {
|
221 |
override def toString: String = SHA1.shasum(digest, name).toString |
|
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
222 |
} |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
223 |
|
| 77209 | 224 |
def read_components_sha1(lines: List[String] = Nil): List[SHA1_Entry] = |
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
225 |
(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:
69426
diff
changeset
|
226 |
Word.explode(line) match {
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
227 |
case Nil => None |
| 77209 | 228 |
case List(sha1, name) => Some(SHA1_Entry(SHA1.fake_digest(sha1), name)) |
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
229 |
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:
69426
diff
changeset
|
230 |
}) |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
231 |
|
| 77209 | 232 |
def write_components_sha1(entries: List[SHA1_Entry]): Unit = |
| 75349 | 233 |
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:
69426
diff
changeset
|
234 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
235 |
|
| 73172 | 236 |
/** manage user components **/ |
237 |
||
| 75349 | 238 |
val components_path: Path = Path.explode("$ISABELLE_HOME_USER/etc/components")
|
| 73172 | 239 |
|
240 |
def read_components(): List[String] = |
|
241 |
if (components_path.is_file) Library.trim_split_lines(File.read(components_path)) |
|
242 |
else Nil |
|
243 |
||
| 75393 | 244 |
def write_components(lines: List[String]): Unit = {
|
| 73172 | 245 |
Isabelle_System.make_directory(components_path.dir) |
| 77216 | 246 |
File.write(components_path, terminate_lines(lines)) |
| 73172 | 247 |
} |
248 |
||
| 75393 | 249 |
def update_components(add: Boolean, path0: Path, progress: Progress = new Progress): Unit = {
|
| 73172 | 250 |
val path = path0.expand.absolute |
|
76551
c7996b073524
clarified check: allow to remove bad directories;
wenzelm
parents:
76550
diff
changeset
|
251 |
if (add) Directory(path).check |
| 73172 | 252 |
|
253 |
val lines1 = read_components() |
|
254 |
val lines2 = |
|
255 |
lines1.filter(line => |
|
256 |
line.isEmpty || line.startsWith("#") || !File.eq(Path.explode(line), path))
|
|
257 |
val lines3 = if (add) lines2 ::: List(path.implode) else lines2 |
|
258 |
||
259 |
if (lines1 != lines3) write_components(lines3) |
|
260 |
||
261 |
val prefix = if (lines1 == lines3) "Unchanged" else if (add) "Added" else "Removed" |
|
262 |
progress.echo(prefix + " component " + path) |
|
263 |
} |
|
264 |
||
265 |
||
266 |
/* main entry point */ |
|
267 |
||
| 75393 | 268 |
def main(args: Array[String]): Unit = {
|
| 73172 | 269 |
Command_Line.tool {
|
270 |
for (arg <- args) {
|
|
271 |
val add = |
|
272 |
if (arg.startsWith("+")) true
|
|
273 |
else if (arg.startsWith("-")) false
|
|
274 |
else error("Bad argument: " + quote(arg))
|
|
275 |
val path = Path.explode(arg.substring(1)) |
|
276 |
update_components(add, path, progress = new Console_Progress) |
|
277 |
} |
|
278 |
} |
|
279 |
} |
|
280 |
||
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
281 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
282 |
/** build and publish components **/ |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
283 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
284 |
def build_components( |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
285 |
options: Options, |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
286 |
components: List[Path], |
|
71726
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
wenzelm
parents:
71601
diff
changeset
|
287 |
progress: Progress = new Progress, |
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
288 |
publish: Boolean = false, |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
289 |
force: Boolean = false, |
| 75393 | 290 |
update_components_sha1: Boolean = false |
291 |
): Unit = {
|
|
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
292 |
val archives: List[Path] = |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
293 |
for (path <- components) yield {
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
294 |
path.file_name match {
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
295 |
case Archive(_) => path |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
296 |
case name => |
| 76550 | 297 |
Directory(path).check |
298 |
val component_path = path.expand |
|
299 |
val archive_dir = component_path.dir |
|
300 |
val archive_name = Archive(name) |
|
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
301 |
|
| 76550 | 302 |
val archive = archive_dir + Path.explode(archive_name) |
303 |
if (archive.is_file && !force) {
|
|
304 |
error("Component archive already exists: " + archive)
|
|
305 |
} |
|
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
306 |
|
| 76550 | 307 |
progress.echo("Packaging " + archive_name)
|
308 |
Isabelle_System.gnutar("-czf " + File.bash_path(archive) + " " + Bash.string(name),
|
|
309 |
dir = archive_dir).check |
|
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
310 |
|
| 76550 | 311 |
archive |
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
312 |
} |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
313 |
} |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
314 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
315 |
if ((publish && archives.nonEmpty) || update_components_sha1) {
|
|
76169
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
316 |
val server = options.string("isabelle_components_server")
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
317 |
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:
69426
diff
changeset
|
318 |
|
|
76169
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
319 |
using(SSH.open_session(options, server)) { ssh =>
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
320 |
val components_dir = Path.explode(options.string("isabelle_components_dir"))
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
321 |
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:
69426
diff
changeset
|
322 |
|
|
76169
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
323 |
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:
76122
diff
changeset
|
324 |
error("Bad remote directory: " + dir)
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
325 |
} |
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
326 |
|
|
76169
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
327 |
if (publish) {
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
328 |
for (archive <- archives) {
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
329 |
val archive_name = archive.file_name |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
330 |
val name = Archive.get_name(archive_name) |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
331 |
val remote_component = components_dir + archive.base |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
332 |
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:
69426
diff
changeset
|
333 |
|
|
76169
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
334 |
// component archive |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
335 |
if (ssh.is_file(remote_component) && !force) {
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
336 |
error("Remote component archive already exists: " + remote_component)
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
337 |
} |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
338 |
progress.echo("Uploading " + archive_name)
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
339 |
ssh.write_file(remote_component, archive) |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
340 |
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
341 |
// contrib directory |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
342 |
val is_standard_component = |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
343 |
Isabelle_System.with_tmp_dir("component") { tmp_dir =>
|
|
76540
83de6e9ae983
clarified signature: prefer Scala functions instead of shell scripts;
wenzelm
parents:
76519
diff
changeset
|
344 |
Isabelle_System.extract(archive, tmp_dir) |
| 76550 | 345 |
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:
69426
diff
changeset
|
346 |
} |
|
76169
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
347 |
if (is_standard_component) {
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
348 |
if (ssh.is_dir(remote_contrib)) {
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
349 |
if (force) ssh.rm_tree(remote_contrib) |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
350 |
else error("Remote component directory already exists: " + remote_contrib)
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
351 |
} |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
352 |
progress.echo("Unpacking remote " + archive_name)
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
353 |
ssh.execute("tar -C " + ssh.bash_path(contrib_dir) + " -xzf " +
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
354 |
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:
69426
diff
changeset
|
355 |
} |
|
76169
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
356 |
else {
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
357 |
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:
69426
diff
changeset
|
358 |
} |
| 75394 | 359 |
} |
|
76169
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
360 |
} |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
361 |
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
362 |
// remote SHA1 digests |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
363 |
if (update_components_sha1) {
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
364 |
val lines = |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
365 |
for {
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
366 |
entry <- ssh.read_dir(components_dir) |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
367 |
if ssh.is_file(components_dir + Path.basic(entry)) && |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
368 |
entry.endsWith(Archive.suffix) |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
369 |
} |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
370 |
yield {
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
371 |
progress.echo("Digesting remote " + entry)
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
372 |
ssh.execute("cd " + ssh.bash_path(components_dir) +
|
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
373 |
"; sha1sum " + Bash.string(entry)).check.out |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
374 |
} |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
375 |
write_components_sha1(read_components_sha1(lines)) |
|
a3c694039fd6
discontinued pointless SSH.Target: OpenSSH client can handle user@host directly;
wenzelm
parents:
76122
diff
changeset
|
376 |
} |
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
377 |
} |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
378 |
} |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
379 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
380 |
// local SHA1 digests |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
381 |
{
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
382 |
val new_entries = |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
383 |
for (archive <- archives) |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
384 |
yield {
|
| 75349 | 385 |
val name = archive.file_name |
386 |
progress.echo("Digesting local " + name)
|
|
| 77209 | 387 |
SHA1_Entry(SHA1.digest(archive), name) |
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
388 |
} |
| 75349 | 389 |
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:
69426
diff
changeset
|
390 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
391 |
write_components_sha1( |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
392 |
new_entries ::: |
| 75349 | 393 |
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:
69426
diff
changeset
|
394 |
} |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
395 |
} |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
396 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
397 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
398 |
/* Isabelle tool wrapper */ |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
399 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
400 |
private val relevant_options = |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
401 |
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:
69426
diff
changeset
|
402 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
403 |
val isabelle_tool = |
| 72763 | 404 |
Isabelle_Tool("build_components", "build and publish Isabelle components",
|
| 75394 | 405 |
Scala_Project.here, |
406 |
{ args =>
|
|
407 |
var publish = false |
|
408 |
var update_components_sha1 = false |
|
409 |
var force = false |
|
410 |
var options = Options.init() |
|
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
411 |
|
| 75394 | 412 |
def show_options: String = |
|
75844
7d27944d7141
clarified signature: avoid public representation;
wenzelm
parents:
75394
diff
changeset
|
413 |
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:
69426
diff
changeset
|
414 |
|
| 75394 | 415 |
val getopts = Getopts("""
|
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
416 |
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:
69426
diff
changeset
|
417 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
418 |
Options are: |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
419 |
-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:
69426
diff
changeset
|
420 |
-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:
69426
diff
changeset
|
421 |
-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:
69426
diff
changeset
|
422 |
-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:
69426
diff
changeset
|
423 |
|
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
424 |
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:
69426
diff
changeset
|
425 |
depending on system options: |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
426 |
|
| 73736 | 427 |
""" + Library.indent_lines(2, show_options) + "\n", |
| 75394 | 428 |
"P" -> (_ => publish = true), |
429 |
"f" -> (_ => force = true), |
|
430 |
"o:" -> (arg => options = options + arg), |
|
431 |
"u" -> (_ => update_components_sha1 = true)) |
|
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
432 |
|
| 75394 | 433 |
val more_args = getopts(args) |
434 |
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:
69426
diff
changeset
|
435 |
|
| 75394 | 436 |
val progress = new Console_Progress |
|
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
437 |
|
| 75394 | 438 |
build_components(options, more_args.map(Path.explode), progress = progress, |
439 |
publish = publish, force = force, update_components_sha1 = update_components_sha1) |
|
440 |
}) |
|
| 69395 | 441 |
} |