author | wenzelm |
Tue, 21 Apr 2020 22:19:59 +0200 | |
changeset 71777 | 3875815f5967 |
parent 71726 | a5fda30edae2 |
child 72353 | 1e5516c55b46 |
permissions | -rw-r--r-- |
69395 | 1 |
/* Title: Pure/Admin/components.scala |
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 |
|
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:
69426
diff
changeset
|
20 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
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:
69426
diff
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:
69426
diff
changeset
|
23 |
else name + suffix |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
24 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
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:
69426
diff
changeset
|
26 |
{ |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
27 |
for { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
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:
69426
diff
changeset
|
29 |
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
|
30 |
} yield name |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
31 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
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 |
||
71601 | 41 |
val default_components_base: Path = Path.explode("$ISABELLE_COMPONENTS_BASE") |
69434 | 42 |
|
69395 | 43 |
def admin(dir: Path): Path = dir + Path.explode("Admin/components") |
44 |
||
45 |
def contrib(dir: Path = Path.current, name: String = ""): Path = |
|
46 |
dir + Path.explode("contrib") + Path.explode(name) |
|
47 |
||
71726
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
wenzelm
parents:
71601
diff
changeset
|
48 |
def unpack(dir: Path, archive: Path, progress: Progress = new Progress): String = |
69413 | 49 |
{ |
50 |
val name = Archive.get_name(archive.file_name) |
|
51 |
progress.echo("Unpacking " + name) |
|
69425 | 52 |
Isabelle_System.gnutar("-xzf " + File.bash_path(archive), dir = dir).check |
69413 | 53 |
name |
54 |
} |
|
55 |
||
69401
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
56 |
def resolve(base_dir: Path, names: List[String], |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
57 |
target_dir: Option[Path] = None, |
70102 | 58 |
copy_dir: Option[Path] = None, |
71726
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
wenzelm
parents:
71601
diff
changeset
|
59 |
progress: Progress = new Progress) |
69398
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
wenzelm
parents:
69395
diff
changeset
|
60 |
{ |
69401
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
61 |
Isabelle_System.mkdirs(base_dir) |
69398
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
wenzelm
parents:
69395
diff
changeset
|
62 |
for (name <- names) { |
69413 | 63 |
val archive_name = Archive(name) |
69401
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
64 |
val archive = base_dir + Path.explode(archive_name) |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
65 |
if (!archive.is_file) { |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
66 |
val remote = Isabelle_System.getenv("ISABELLE_COMPONENT_REPOSITORY") + "/" + archive_name |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
67 |
progress.echo("Getting " + remote) |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
68 |
Bytes.write(archive, Url.read_bytes(Url(remote))) |
69398
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
wenzelm
parents:
69395
diff
changeset
|
69 |
} |
70102 | 70 |
for (dir <- copy_dir) { |
71 |
Isabelle_System.mkdirs(dir) |
|
72 |
File.copy(archive, dir) |
|
73 |
} |
|
69413 | 74 |
unpack(target_dir getOrElse base_dir, archive, progress = progress) |
69398
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
wenzelm
parents:
69395
diff
changeset
|
75 |
} |
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
wenzelm
parents:
69395
diff
changeset
|
76 |
} |
0698ded5caf1
Components.download similar to "isabelle components", but without unpacking;
wenzelm
parents:
69395
diff
changeset
|
77 |
|
69410 | 78 |
def purge(dir: Path, platform: Platform.Family.Value) |
69401
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
79 |
{ |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
80 |
def purge_platforms(platforms: String*): Set[String] = |
69703 | 81 |
platforms.flatMap(name => List("x86-" + name, "x86_64_32-" + name, "x86_64-" + name)).toSet + |
82 |
"ppc-darwin" |
|
69401
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
83 |
val purge_set = |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
84 |
platform match { |
69410 | 85 |
case Platform.Family.linux => purge_platforms("darwin", "cygwin", "windows") |
86 |
case Platform.Family.macos => purge_platforms("linux", "cygwin", "windows") |
|
87 |
case Platform.Family.windows => purge_platforms("linux", "darwin") |
|
69401
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
88 |
} |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
89 |
|
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
90 |
File.find_files(dir.file, |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
91 |
(file: JFile) => file.isDirectory && purge_set(file.getName), |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
92 |
include_dirs = true).foreach(Isabelle_System.rm_tree) |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
93 |
} |
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
wenzelm
parents:
69398
diff
changeset
|
94 |
|
69395 | 95 |
|
96 |
/* component directory content */ |
|
97 |
||
69429
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
98 |
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:
69426
diff
changeset
|
99 |
def components(dir: Path = Path.current): Path = dir + Path.explode("etc/components") |
69395 | 100 |
|
101 |
def check_dir(dir: Path): Boolean = |
|
102 |
settings(dir).is_file || components(dir).is_file |
|
103 |
||
104 |
def read_components(dir: Path): List[String] = |
|
105 |
split_lines(File.read(components(dir))).filter(_.nonEmpty) |
|
106 |
||
107 |
def write_components(dir: Path, lines: List[String]): Unit = |
|
108 |
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:
69426
diff
changeset
|
109 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
110 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
111 |
/* component repository content */ |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
112 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
113 |
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
|
114 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
115 |
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:
69426
diff
changeset
|
116 |
{ |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
117 |
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:
69426
diff
changeset
|
118 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
119 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
120 |
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:
69426
diff
changeset
|
121 |
(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
|
122 |
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
|
123 |
case Nil => None |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
124 |
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:
69426
diff
changeset
|
125 |
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
|
126 |
}) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
127 |
|
71601 | 128 |
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:
69426
diff
changeset
|
129 |
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:
69426
diff
changeset
|
130 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
131 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
132 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
133 |
/** 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
|
134 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
135 |
def build_components( |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
136 |
options: Options, |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
137 |
components: List[Path], |
71726
a5fda30edae2
clarified signature: more uniform treatment of stopped/interrupted state;
wenzelm
parents:
71601
diff
changeset
|
138 |
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
|
139 |
publish: Boolean = false, |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
140 |
force: Boolean = false, |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
141 |
update_components_sha1: Boolean = false) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
142 |
{ |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
143 |
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
|
144 |
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
|
145 |
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
|
146 |
case Archive(_) => path |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
147 |
case name => |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
148 |
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:
69426
diff
changeset
|
149 |
else if (!check_dir(path)) { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
150 |
error("Malformed component directory: " + path + |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
151 |
"\n (requires " + settings() + " or " + Components.components() + ")") |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
152 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
153 |
else { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
154 |
val component_path = path.expand |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
155 |
val archive_dir = component_path.dir |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
156 |
val archive_name = Archive(name) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
157 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
158 |
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:
69426
diff
changeset
|
159 |
if (archive.is_file && !force) { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
160 |
error("Component archive already exists: " + archive) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
161 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
162 |
|
69439 | 163 |
progress.echo("Packaging " + archive_name) |
69430 | 164 |
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:
69426
diff
changeset
|
165 |
dir = archive_dir).check |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
166 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
167 |
archive |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
168 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
169 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
170 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
171 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
172 |
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:
69426
diff
changeset
|
173 |
options.string("isabelle_components_server") match { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
174 |
case SSH.Target(user, host) => |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
175 |
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:
69426
diff
changeset
|
176 |
{ |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
177 |
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:
69426
diff
changeset
|
178 |
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:
69426
diff
changeset
|
179 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
180 |
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:
69426
diff
changeset
|
181 |
error("Bad remote directory: " + dir) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
182 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
183 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
184 |
if (publish) { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
185 |
for (archive <- archives) { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
186 |
val archive_name = archive.file_name |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
187 |
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:
69426
diff
changeset
|
188 |
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:
69426
diff
changeset
|
189 |
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:
69426
diff
changeset
|
190 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
191 |
// component archive |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
192 |
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:
69426
diff
changeset
|
193 |
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:
69426
diff
changeset
|
194 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
195 |
progress.echo("Uploading " + archive_name) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
196 |
ssh.write_file(remote_component, archive) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
197 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
198 |
// contrib directory |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
199 |
val is_standard_component = |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
200 |
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:
69426
diff
changeset
|
201 |
{ |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
202 |
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:
69426
diff
changeset
|
203 |
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:
69426
diff
changeset
|
204 |
}) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
205 |
if (is_standard_component) { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
206 |
if (ssh.is_dir(remote_contrib)) { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
207 |
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:
69426
diff
changeset
|
208 |
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:
69426
diff
changeset
|
209 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
210 |
progress.echo("Unpacking remote " + archive_name) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
211 |
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:
69426
diff
changeset
|
212 |
ssh.bash_path(remote_component)).check |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
213 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
214 |
else { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
215 |
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:
69426
diff
changeset
|
216 |
} |
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 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
219 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
220 |
// remote SHA1 digests |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
221 |
if (update_components_sha1) { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
222 |
val lines = |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
223 |
for { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
224 |
entry <- ssh.read_dir(components_dir) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
225 |
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:
69426
diff
changeset
|
226 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
227 |
yield { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
228 |
progress.echo("Digesting remote " + entry.name) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
229 |
Library.trim_line( |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
230 |
ssh.execute("cd " + ssh.bash_path(components_dir) + |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
231 |
"; sha1sum " + Bash.string(entry.name)).check.out) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
232 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
233 |
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:
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 |
}) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
236 |
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:
69426
diff
changeset
|
237 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
238 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
239 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
240 |
// local SHA1 digests |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
241 |
{ |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
242 |
val new_entries = |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
243 |
for (archive <- archives) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
244 |
yield { |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
245 |
val file_name = archive.file_name |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
246 |
progress.echo("Digesting local " + file_name) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
247 |
val sha1 = SHA1.digest(archive).rep |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
248 |
SHA1_Digest(sha1, file_name) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
249 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
250 |
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:
69426
diff
changeset
|
251 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
252 |
write_components_sha1( |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
253 |
new_entries ::: |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
254 |
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:
69426
diff
changeset
|
255 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
256 |
} |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
257 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
258 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
259 |
/* Isabelle tool wrapper */ |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
260 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
261 |
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
|
262 |
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
|
263 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
264 |
val isabelle_tool = |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
265 |
Isabelle_Tool("build_components", "build and publish Isabelle components", args => |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
266 |
{ |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
267 |
var publish = false |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
268 |
var update_components_sha1 = false |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
269 |
var force = false |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
270 |
var options = Options.init() |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
271 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
272 |
def show_options: String = |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
273 |
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:
69426
diff
changeset
|
274 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
275 |
val getopts = Getopts(""" |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
276 |
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
|
277 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
278 |
Options are: |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
279 |
-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
|
280 |
-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
|
281 |
-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
|
282 |
-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
|
283 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
284 |
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
|
285 |
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
|
286 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
287 |
""" + Library.prefix_lines(" ", show_options) + "\n", |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
288 |
"P" -> (_ => publish = true), |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
289 |
"f" -> (_ => force = true), |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
290 |
"o:" -> (arg => options = options + arg), |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
291 |
"u" -> (_ => update_components_sha1 = true)) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
292 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
293 |
val more_args = getopts(args) |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
294 |
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:
69426
diff
changeset
|
295 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
296 |
val progress = new Console_Progress |
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
297 |
|
dc5fbcb07c7b
replaced "isabelle components_checksum" shell script by "isabelle build_components" in Scala, with more functionality;
wenzelm
parents:
69426
diff
changeset
|
298 |
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:
69426
diff
changeset
|
299 |
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:
69426
diff
changeset
|
300 |
}) |
69395 | 301 |
} |