| author | wenzelm | 
| Tue, 06 Sep 2022 21:06:20 +0200 | |
| changeset 76074 | 2456721602b2 | 
| parent 75796 | e5c3353df22e | 
| child 76156 | e73025785dc7 | 
| permissions | -rw-r--r-- | 
| 64202 | 1 | /* Title: Pure/Admin/build_release.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Build full Isabelle distribution from repository. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 75393 | 10 | object Build_Release {
 | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 11 | /** release context **/ | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 12 | |
| 73629 | 13 | private def execute(dir: Path, script: String): Unit = | 
| 14 | Isabelle_System.bash(script, cwd = dir.file).check | |
| 15 | ||
| 16 | private def execute_tar(dir: Path, args: String, strip: Int = 0): Unit = | |
| 17 | Isabelle_System.gnutar(args, dir = dir, strip = strip).check | |
| 18 | ||
| 74857 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 19 | private def bash_java_opens(args: String*): String = | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 20 |     Bash.strings(args.toList.flatMap(arg => List("--add-opens", arg + "=ALL-UNNAMED")))
 | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 21 | |
| 75393 | 22 |   object Release_Context {
 | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 23 | def apply( | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 24 | target_dir: Path, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 25 | release_name: String = "", | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 26 | components_base: Path = Components.default_components_base, | 
| 75393 | 27 | progress: Progress = new Progress | 
| 28 |     ): Release_Context = {
 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 29 | val date = Date.now() | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 30 |       val dist_name = proper_string(release_name) getOrElse ("Isabelle_" + Date.Format.date(date))
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 31 |       val dist_dir = (target_dir + Path.explode("dist-" + dist_name)).absolute
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 32 | new Release_Context(release_name, dist_name, dist_dir, components_base, progress) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 33 | } | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 34 | } | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 35 | |
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 36 | class Release_Context private[Build_Release]( | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 37 | val release_name: String, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 38 | val dist_name: String, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 39 | val dist_dir: Path, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 40 | val components_base: Path, | 
| 75393 | 41 | val progress: Progress | 
| 42 |   ) {
 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 43 | override def toString: String = dist_name | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 44 | |
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 45 | val isabelle: Path = Path.explode(dist_name) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 46 | val isabelle_dir: Path = dist_dir + isabelle | 
| 73629 | 47 | val isabelle_archive: Path = dist_dir + isabelle.tar.gz | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 48 | val isabelle_library_archive: Path = dist_dir + Path.explode(dist_name + "_library.tar.gz") | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 49 | |
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 50 | def other_isabelle(dir: Path): Other_Isabelle = | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 51 | Other_Isabelle(dir + isabelle, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 52 | isabelle_identifier = dist_name + "-build", | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 53 | progress = progress) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 54 | |
| 75393 | 55 |     def make_announce(id: String): Unit = {
 | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 56 |       if (release_name.isEmpty) {
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 57 |         File.write(isabelle_dir + Path.explode("ANNOUNCE"),
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 58 | """ | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 59 | IMPORTANT NOTE | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 60 | ============== | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 61 | |
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 62 | This is a snapshot of Isabelle/""" + id + """ from the repository. | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 63 | """) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 64 | } | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 65 | } | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 66 | |
| 75393 | 67 |     def make_contrib(): Unit = {
 | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 68 | Isabelle_System.make_directory(Components.contrib(isabelle_dir)) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 69 | File.write(Components.contrib(isabelle_dir, name = "README"), | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 70 | """This directory contains add-on components that contribute to the main | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 71 | Isabelle distribution. Separate licensing conditions apply, see each | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 72 | directory individually. | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 73 | """) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 74 | } | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 75 | |
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 76 | def bundle_info(platform: Platform.Family.Value): Bundle_Info = | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 77 |       platform match {
 | 
| 73637 | 78 | case Platform.Family.linux_arm => | 
| 79 | Bundle_Info(platform, "Linux (ARM)", dist_name + "_linux_arm.tar.gz") | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 80 | case Platform.Family.linux => Bundle_Info(platform, "Linux", dist_name + "_linux.tar.gz") | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 81 | case Platform.Family.macos => Bundle_Info(platform, "macOS", dist_name + "_macos.tar.gz") | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 82 | case Platform.Family.windows => Bundle_Info(platform, "Windows", dist_name + ".exe") | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 83 | } | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 84 | } | 
| 69174 | 85 | |
| 86 | sealed case class Bundle_Info( | |
| 69410 | 87 | platform: Platform.Family.Value, | 
| 69174 | 88 | platform_description: String, | 
| 75393 | 89 | name: String | 
| 90 |   ) {
 | |
| 69432 
d072f3287ffa
discontinued somewhat point dmg: plain .tar.gz is smaller and more convenient to install;
 wenzelm parents: 
69425diff
changeset | 91 | def path: Path = Path.explode(name) | 
| 69174 | 92 | } | 
| 93 | ||
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 94 | |
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 95 | |
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 96 | /** release archive **/ | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 97 | |
| 73629 | 98 |   val ISABELLE: Path = Path.basic("Isabelle")
 | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 99 |   val ISABELLE_ID: Path = Path.explode("etc/ISABELLE_ID")
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 100 |   val ISABELLE_TAGS: Path = Path.explode("etc/ISABELLE_TAGS")
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 101 |   val ISABELLE_IDENTIFIER: Path = Path.explode("etc/ISABELLE_IDENTIFIER")
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 102 | |
| 75393 | 103 |   object Release_Archive {
 | 
| 104 |     def make(bytes: Bytes, rename: String = ""): Release_Archive = {
 | |
| 75491 | 105 |       Isabelle_System.with_tmp_dir("build_release")(dir =>
 | 
| 75394 | 106 |         Isabelle_System.with_tmp_file("archive", ext = "tar.gz") { archive_path =>
 | 
| 73629 | 107 | val isabelle_dir = Isabelle_System.make_directory(dir + ISABELLE) | 
| 108 | ||
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 109 | Bytes.write(archive_path, bytes) | 
| 73629 | 110 | execute_tar(isabelle_dir, "-xzf " + File.bash_path(archive_path), strip = 1) | 
| 111 | ||
| 112 | val id = File.read(isabelle_dir + ISABELLE_ID) | |
| 113 | val tags = File.read(isabelle_dir + ISABELLE_TAGS) | |
| 114 | val identifier = File.read(isabelle_dir + ISABELLE_IDENTIFIER) | |
| 69174 | 115 | |
| 73629 | 116 | val (bytes1, identifier1) = | 
| 117 | if (rename.isEmpty || rename == identifier) (bytes, identifier) | |
| 118 |             else {
 | |
| 119 | File.write(isabelle_dir + ISABELLE_IDENTIFIER, rename) | |
| 120 | Isabelle_System.move_file(isabelle_dir, dir + Path.basic(rename)) | |
| 121 | execute_tar(dir, "-czf " + File.bash_path(archive_path) + " " + Bash.string(rename)) | |
| 122 | (Bytes.read(archive_path), rename) | |
| 123 | } | |
| 124 | new Release_Archive(bytes1, id, tags, identifier1) | |
| 75394 | 125 | } | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 126 | ) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 127 | } | 
| 69174 | 128 | |
| 73629 | 129 | def read(path: Path, rename: String = ""): Release_Archive = | 
| 130 | make(Bytes.read(path), rename = rename) | |
| 131 | ||
| 75393 | 132 | def get( | 
| 133 | url: String, | |
| 134 | rename: String = "", | |
| 135 | progress: Progress = new Progress | |
| 136 |     ): Release_Archive = {
 | |
| 73629 | 137 | val bytes = | 
| 138 | if (Path.is_wellformed(url)) Bytes.read(Path.explode(url)) | |
| 139 | else Isabelle_System.download(url, progress = progress).bytes | |
| 140 | make(bytes, rename = rename) | |
| 141 | } | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 142 | } | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 143 | |
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 144 | case class Release_Archive private[Build_Release]( | 
| 75393 | 145 |     bytes: Bytes, id: String, tags: String, identifier: String) {
 | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 146 | override def toString: String = identifier | 
| 69174 | 147 | } | 
| 148 | ||
| 149 | ||
| 150 | ||
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 151 | /** generated content **/ | 
| 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 152 | |
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 153 | /* bundled components */ | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 154 | |
| 75393 | 155 |   class Bundled(platform: Option[Platform.Family.Value] = None) {
 | 
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 156 | def detect(s: String): Boolean = | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 157 |       s.startsWith("#bundled") && !s.startsWith("#bundled ")
 | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 158 | |
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 159 | def apply(name: String): String = | 
| 69410 | 160 |       "#bundled" + (platform match { case None => "" case Some(plat) => "-" + plat }) + ":" + name
 | 
| 69391 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 161 | |
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 162 |     private val Pattern1 = ("""^#bundled:(.*)$""").r
 | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 163 |     private val Pattern2 = ("""^#bundled-(.*):(.*)$""").r
 | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 164 | |
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 165 | def unapply(s: String): Option[String] = | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 166 |       s match {
 | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 167 | case Pattern1(name) => Some(name) | 
| 69410 | 168 | case Pattern2(Platform.Family(plat), name) if platform == Some(plat) => Some(name) | 
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 169 | case _ => None | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 170 | } | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 171 | } | 
| 69391 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 172 | |
| 75393 | 173 |   def record_bundled_components(dir: Path): Unit = {
 | 
| 69391 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 174 | val catalogs = | 
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 175 |       List("main", "bundled").map((_, new Bundled())) :::
 | 
| 73637 | 176 | Platform.Family.list.flatMap(platform => | 
| 69410 | 177 | List(platform.toString, "bundled-" + platform.toString). | 
| 178 | map((_, new Bundled(platform = Some(platform))))) | |
| 69391 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 179 | |
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 180 | File.append(Components.components(dir), | 
| 69391 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 181 |       terminate_lines("#bundled components" ::
 | 
| 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 182 |         (for {
 | 
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 183 | (catalog, bundled) <- catalogs.iterator | 
| 71601 | 184 | path = Components.admin(dir) + Path.basic(catalog) | 
| 69391 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 185 | if path.is_file | 
| 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 186 | line <- split_lines(File.read(path)) | 
| 73987 
fc363a3b690a
build.props for isabelle.jar, including isabelle.jedit;
 wenzelm parents: 
73894diff
changeset | 187 |           if line.nonEmpty && !line.startsWith("#")
 | 
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 188 | } yield bundled(line)).toList)) | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 189 | } | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 190 | |
| 75393 | 191 |   def get_bundled_components(dir: Path, platform: Platform.Family.Value): (List[String], String) = {
 | 
| 69410 | 192 | val Bundled = new Bundled(platform = Some(platform)) | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 193 | val components = | 
| 73987 
fc363a3b690a
build.props for isabelle.jar, including isabelle.jedit;
 wenzelm parents: 
73894diff
changeset | 194 |       for { Bundled(name) <- Components.read_components(dir) } yield name
 | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 195 | val jdk_component = | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 196 |       components.find(_.startsWith("jdk")) getOrElse error("Missing jdk component")
 | 
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 197 | (components, jdk_component) | 
| 69391 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 198 | } | 
| 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 199 | |
| 73340 | 200 | def activate_components( | 
| 75393 | 201 |     dir: Path, platform: Platform.Family.Value, more_names: List[String]): Unit = {
 | 
| 69413 | 202 | def contrib_name(name: String): String = | 
| 203 | Components.contrib(name = name).implode | |
| 204 | ||
| 69410 | 205 | val Bundled = new Bundled(platform = Some(platform)) | 
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 206 | Components.write_components(dir, | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 207 | Components.read_components(dir).flatMap(line => | 
| 69391 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 208 |         line match {
 | 
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 209 | case Bundled(name) => | 
| 69413 | 210 | if (Components.check_dir(Components.contrib(dir, name))) Some(contrib_name(name)) | 
| 69395 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 211 | else None | 
| 
d1c4a1dee9e7
more explicit support for Isabelle system components;
 wenzelm parents: 
69392diff
changeset | 212 | case _ => if (Bundled.detect(line)) None else Some(line) | 
| 71601 | 213 | }) ::: more_names.map(contrib_name)) | 
| 69391 
a3c776b9d3dd
manage components similar to makedist_bundle (still inactive);
 wenzelm parents: 
69390diff
changeset | 214 | } | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 215 | |
| 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 216 | |
| 70099 
9b9c1192f972
support for platform-specific builds on remote server;
 wenzelm parents: 
70098diff
changeset | 217 | /** build release **/ | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 218 | |
| 73634 | 219 | /* build heaps */ | 
| 70099 
9b9c1192f972
support for platform-specific builds on remote server;
 wenzelm parents: 
70098diff
changeset | 220 | |
| 73634 | 221 | private def build_heaps( | 
| 70101 | 222 | options: Options, | 
| 223 | platform: Platform.Family.Value, | |
| 224 | build_sessions: List[String], | |
| 75796 | 225 | local_dir: Path, | 
| 226 | progress: Progress = new Progress, | |
| 75393 | 227 |   ): Unit = {
 | 
| 72341 | 228 | val server_option = "build_host_" + platform.toString | 
| 75796 | 229 | val server = options.string(server_option) | 
| 230 |     progress.echo("Building heaps " + commas_quote(build_sessions) +
 | |
| 231 |       " (" + server_option + " = " + quote(server) + ") ...")
 | |
| 232 | ||
| 73634 | 233 | val ssh = | 
| 75796 | 234 |        server match {
 | 
| 73634 | 235 | case "" => | 
| 236 | if (Platform.family == platform) SSH.Local | |
| 237 |           else error("Undefined option " + server_option + ": cannot build heaps")
 | |
| 238 | case SSH.Target(user, host) => | |
| 239 | SSH.open_session(options, host = host, user = user) | |
| 75796 | 240 |         case _ => error("Malformed option " + server_option + ": " + quote(server))
 | 
| 73634 | 241 | } | 
| 242 |     try {
 | |
| 75394 | 243 |       Isabelle_System.with_tmp_file("tmp", ext = "tar") { local_tmp_tar =>
 | 
| 73634 | 244 | execute_tar(local_dir, "-cf " + File.bash_path(local_tmp_tar) + " .") | 
| 75394 | 245 |         ssh.with_tmp_dir { remote_dir =>
 | 
| 73634 | 246 |           val remote_tmp_tar = remote_dir + Path.basic("tmp.tar")
 | 
| 247 | ssh.write_file(remote_tmp_tar, local_tmp_tar) | |
| 248 | val remote_commands = | |
| 249 | List( | |
| 250 | "cd " + File.bash_path(remote_dir), | |
| 251 | "tar -xf tmp.tar", | |
| 252 | "bin/isabelle build -o system_heaps -b -- " + Bash.strings(build_sessions), | |
| 253 | "tar -cf tmp.tar heaps") | |
| 254 |           ssh.execute(remote_commands.mkString(" && "), settings = false).check
 | |
| 255 | ssh.read_file(remote_tmp_tar, local_tmp_tar) | |
| 75394 | 256 | } | 
| 73634 | 257 | execute_tar(local_dir, "-xf " + File.bash_path(local_tmp_tar)) | 
| 75394 | 258 | } | 
| 70099 
9b9c1192f972
support for platform-specific builds on remote server;
 wenzelm parents: 
70098diff
changeset | 259 | } | 
| 73634 | 260 |     finally { ssh.close() }
 | 
| 70099 
9b9c1192f972
support for platform-specific builds on remote server;
 wenzelm parents: 
70098diff
changeset | 261 | } | 
| 
9b9c1192f972
support for platform-specific builds on remote server;
 wenzelm parents: 
70098diff
changeset | 262 | |
| 
9b9c1192f972
support for platform-specific builds on remote server;
 wenzelm parents: 
70098diff
changeset | 263 | |
| 73066 | 264 | /* Isabelle application */ | 
| 73060 | 265 | |
| 75393 | 266 |   def make_isabelle_options(path: Path, options: List[String], line_ending: String = "\n"): Unit = {
 | 
| 73068 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 267 | val title = "# Java runtime options" | 
| 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 268 | File.write(path, (title :: options).map(_ + line_ending).mkString) | 
| 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 269 | } | 
| 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 270 | |
| 73065 | 271 | def make_isabelle_app( | 
| 73193 | 272 | platform: Platform.Family.Value, | 
| 273 | isabelle_target: Path, | |
| 274 | isabelle_name: String, | |
| 73065 | 275 | jdk_component: String, | 
| 73095 | 276 | classpath: List[Path], | 
| 75393 | 277 |     dock_icon: Boolean = false): Unit = {
 | 
| 73064 | 278 | val script = """#!/usr/bin/env bash | 
| 73060 | 279 | # | 
| 280 | # Author: Makarius | |
| 281 | # | |
| 282 | # Main Isabelle application script. | |
| 283 | ||
| 284 | # minimal Isabelle environment | |
| 285 | ||
| 73193 | 286 | ISABELLE_HOME="$(cd "$(dirname "$0")"; cd "$(pwd -P)/../.."; pwd)" | 
| 73060 | 287 | source "$ISABELLE_HOME/lib/scripts/isabelle-platform" | 
| 288 | ||
| 73063 | 289 | #paranoia settings -- avoid intrusion of alien options | 
| 290 | unset "_JAVA_OPTIONS" | |
| 291 | unset "JAVA_TOOL_OPTIONS" | |
| 292 | ||
| 293 | #paranoia settings -- avoid problems of Java/Swing versus XIM/IBus etc. | |
| 294 | unset XMODIFIERS | |
| 295 | ||
| 73062 | 296 | COMPONENT="$ISABELLE_HOME/contrib/""" + jdk_component + """" | 
| 297 | source "$COMPONENT/etc/settings" | |
| 298 | ||
| 73060 | 299 | |
| 73063 | 300 | # main | 
| 73060 | 301 | |
| 73703 | 302 | declare -a JAVA_OPTIONS=($(grep -v '^#' "$ISABELLE_HOME/Isabelle.options")) | 
| 73060 | 303 | |
| 73197 
d967f6643f5e
proper Isabelle environment (amending 31fbde3baa97);
 wenzelm parents: 
73193diff
changeset | 304 | "$ISABELLE_HOME/bin/isabelle" env "$ISABELLE_HOME/lib/scripts/java-gui-setup" | 
| 73153 
96d87b9c2b42
workaround for Big Sur fullscreen mode: better support for JDialog windows (e.g. Find on top of main View);
 wenzelm parents: 
73152diff
changeset | 305 | |
| 73062 | 306 | exec "$ISABELLE_JDK_HOME/bin/java" \ | 
| 73060 | 307 |   "-Disabelle.root=$ISABELLE_HOME" "${JAVA_OPTIONS[@]}" \
 | 
| 73061 | 308 |   -classpath """" + classpath.map(p => "$ISABELLE_HOME/" + p.implode).mkString(":") + """" \
 | 
| 73060 | 309 | "-splash:$ISABELLE_HOME/lib/logo/isabelle.gif" \ | 
| 73095 | 310 | """ + (if (dock_icon) """"-Xdock:icon=$ISABELLE_HOME/lib/logo/isabelle_transparent-128.png" \ | 
| 75291 | 311 | """ else "") + """isabelle.jedit.JEdit_Main "$@" | 
| 73060 | 312 | """ | 
| 73193 | 313 |     val script_path = isabelle_target + Path.explode("lib/scripts/Isabelle_app")
 | 
| 314 | File.write(script_path, script) | |
| 315 | File.set_executable(script_path, true) | |
| 316 | ||
| 317 |     val component_dir = isabelle_target + Path.explode("contrib/Isabelle_app")
 | |
| 73317 | 318 | Isabelle_System.move_file( | 
| 73637 | 319 |       component_dir + Path.explode(Platform.Family.standard(platform)) + Path.explode("Isabelle"),
 | 
| 73193 | 320 | isabelle_target + Path.explode(isabelle_name)) | 
| 321 | Isabelle_System.rm_tree(component_dir) | |
| 73064 | 322 | } | 
| 73060 | 323 | |
| 324 | ||
| 75393 | 325 |   def make_isabelle_plist(path: Path, isabelle_name: String, isabelle_rev: String): Unit = {
 | 
| 73066 | 326 | File.write(path, """<?xml version="1.0" ?> | 
| 327 | <!DOCTYPE plist PUBLIC "-//Apple//DTD PLIST 1.0//EN" "http://www.apple.com/DTDs/PropertyList-1.0.dtd"> | |
| 328 | <plist version="1.0"> | |
| 329 | <dict> | |
| 330 | <key>CFBundleDevelopmentRegion</key> | |
| 331 | <string>English</string> | |
| 73075 
893310d6d76d
recovered bundle icons (not application) from macos_app;
 wenzelm parents: 
73074diff
changeset | 332 | <key>CFBundleIconFile</key> | 
| 
893310d6d76d
recovered bundle icons (not application) from macos_app;
 wenzelm parents: 
73074diff
changeset | 333 | <string>isabelle.icns</string> | 
| 73066 | 334 | <key>CFBundleIdentifier</key> | 
| 73152 
5a954fd5f078
clarified app identification, potentially relevant for macOS "defaults";
 wenzelm parents: 
73112diff
changeset | 335 | <string>de.tum.in.isabelle</string> | 
| 73066 | 336 | <key>CFBundleDisplayName</key> | 
| 337 | <string>""" + isabelle_name + """</string> | |
| 338 | <key>CFBundleInfoDictionaryVersion</key> | |
| 339 | <string>6.0</string> | |
| 340 | <key>CFBundleName</key> | |
| 341 | <string>""" + isabelle_name + """</string> | |
| 342 | <key>CFBundlePackageType</key> | |
| 343 | <string>APPL</string> | |
| 344 | <key>CFBundleShortVersionString</key> | |
| 73152 
5a954fd5f078
clarified app identification, potentially relevant for macOS "defaults";
 wenzelm parents: 
73112diff
changeset | 345 | <string>""" + isabelle_name + """</string> | 
| 73066 | 346 | <key>CFBundleSignature</key> | 
| 347 | <string>????</string> | |
| 348 | <key>CFBundleVersion</key> | |
| 73152 
5a954fd5f078
clarified app identification, potentially relevant for macOS "defaults";
 wenzelm parents: 
73112diff
changeset | 349 | <string>""" + isabelle_rev + """</string> | 
| 73066 | 350 | <key>NSHumanReadableCopyright</key> | 
| 351 | <string></string> | |
| 352 | <key>LSMinimumSystemVersion</key> | |
| 73085 | 353 | <string>10.11</string> | 
| 73066 | 354 | <key>LSApplicationCategoryType</key> | 
| 355 | <string>public.app-category.developer-tools</string> | |
| 356 | <key>NSHighResolutionCapable</key> | |
| 357 | <string>true</string> | |
| 358 | <key>NSSupportsAutomaticGraphicsSwitching</key> | |
| 359 | <string>true</string> | |
| 360 | <key>CFBundleDocumentTypes</key> | |
| 361 | <array> | |
| 362 | <dict> | |
| 363 | <key>CFBundleTypeExtensions</key> | |
| 364 | <array> | |
| 365 | <string>thy</string> | |
| 366 | </array> | |
| 73077 | 367 | <key>CFBundleTypeIconFile</key> | 
| 368 | <string>theory.icns</string> | |
| 73066 | 369 | <key>CFBundleTypeName</key> | 
| 370 | <string>Isabelle theory file</string> | |
| 371 | <key>CFBundleTypeRole</key> | |
| 372 | <string>Editor</string> | |
| 373 | <key>LSTypeIsPackage</key> | |
| 374 | <false/> | |
| 375 | </dict> | |
| 376 | </array> | |
| 377 | </dict> | |
| 378 | </plist> | |
| 379 | """) | |
| 380 | } | |
| 381 | ||
| 382 | ||
| 73064 | 383 | /* main */ | 
| 70099 
9b9c1192f972
support for platform-specific builds on remote server;
 wenzelm parents: 
70098diff
changeset | 384 | |
| 73629 | 385 | def use_release_archive( | 
| 386 | context: Release_Context, | |
| 387 | archive: Release_Archive, | |
| 75393 | 388 | id: String = "" | 
| 389 |   ): Unit = {
 | |
| 73629 | 390 |     if (id.nonEmpty && id != archive.id) {
 | 
| 391 |       error("Mismatch of release identification " + id + " vs. archive " + archive.id)
 | |
| 392 | } | |
| 393 | ||
| 394 |     if (!context.isabelle_archive.is_file || Bytes.read(context.isabelle_archive) != archive.bytes) {
 | |
| 395 | Bytes.write(context.isabelle_archive, archive.bytes) | |
| 396 | } | |
| 397 | } | |
| 398 | ||
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 399 | def build_release_archive( | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 400 | context: Release_Context, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 401 | version: String, | 
| 75393 | 402 | parallel_jobs: Int = 1 | 
| 403 |   ): Unit = {
 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 404 | val progress = context.progress | 
| 69174 | 405 | |
| 75506 | 406 | val hg = Mercurial.self_repository() | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 407 | val id = | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 408 |       try { hg.id(version) }
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 409 |       catch { case ERROR(msg) => cat_error("Bad repository version: " + version, msg) }
 | 
| 64202 | 410 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 411 |     if (context.isabelle_archive.is_file) {
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 412 |       progress.echo_warning("Found existing release archive: " + context.isabelle_archive)
 | 
| 73629 | 413 | use_release_archive(context, Release_Archive.read(context.isabelle_archive), id = id) | 
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 414 | } | 
| 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 415 |     else {
 | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 416 |       progress.echo_warning("Preparing release " + context.dist_name + " ...")
 | 
| 64221 | 417 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 418 | Isabelle_System.new_directory(context.dist_dir) | 
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 419 | |
| 75478 | 420 | hg.archive(context.isabelle_dir.expand.implode, rev = id) | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 421 | |
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 422 |       for (name <- List(".hg_archival.txt", ".hgtags", ".hgignore", "README_REPOSITORY")) {
 | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 423 | (context.isabelle_dir + Path.explode(name)).file.delete | 
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 424 | } | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 425 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 426 | File.write(context.isabelle_dir + ISABELLE_ID, id) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 427 | File.write(context.isabelle_dir + ISABELLE_TAGS, hg.tags(rev = id)) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 428 | File.write(context.isabelle_dir + ISABELLE_IDENTIFIER, context.dist_name) | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 429 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 430 | context.make_announce(id) | 
| 73520 
4cba4e250c28
clarified ISABELLE_ID: distribution vs. hg archive vs. hg repos;
 wenzelm parents: 
73519diff
changeset | 431 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 432 | context.make_contrib() | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 433 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 434 | execute(context.isabelle_dir, """find . -print | xargs chmod -f u+rw""") | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 435 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 436 | record_bundled_components(context.isabelle_dir) | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 437 | |
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 438 | |
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 439 | /* build tools and documentation */ | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 440 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 441 | val other_isabelle = context.other_isabelle(context.dist_dir) | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 442 | |
| 69388 | 443 | other_isabelle.init_settings( | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 444 | other_isabelle.init_components( | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 445 |           components_base = context.components_base, catalogs = List("main")))
 | 
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 446 | other_isabelle.resolve_components(echo = true) | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 447 | |
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 448 |       try {
 | 
| 74011 | 449 | other_isabelle.bash( | 
| 450 |           "export CLASSPATH=" + Bash.string(other_isabelle.getenv("ISABELLE_CLASSPATH")) + "\n" +
 | |
| 451 | "bin/isabelle jedit -b", echo = true).check | |
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 452 | } | 
| 71936 | 453 |       catch { case ERROR(msg) => cat_error("Failed to build tools:", msg) }
 | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 454 | |
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 455 |       try {
 | 
| 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 456 | other_isabelle.bash( | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 457 | "bin/isabelle build_doc -a -o system_heaps -j " + parallel_jobs, echo = true).check | 
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 458 | } | 
| 71936 | 459 |       catch { case ERROR(msg) => cat_error("Failed to build documentation:", msg) }
 | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 460 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 461 | other_isabelle.make_news() | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 462 | |
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 463 |       for (name <- List("Admin", "browser_info", "heaps")) {
 | 
| 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 464 | Isabelle_System.rm_tree(other_isabelle.isabelle_home + Path.explode(name)) | 
| 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 465 | } | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 466 | |
| 69171 
710845a85944
more robust release.read_ident: eliminated odd state files ISABELLE_IDENT, ISABELLE_DIST;
 wenzelm parents: 
69170diff
changeset | 467 | other_isabelle.cleanup() | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 468 | |
| 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 469 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 470 |       progress.echo_warning("Creating release archive " + context.isabelle_archive + " ...")
 | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 471 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 472 | execute(context.dist_dir, """chmod -R a+r . && chmod -R u+w . && chmod -R g=o .""") | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 473 | execute(context.dist_dir, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 474 |         """find . -type f "(" -name "*.thy" -o -name "*.ML" -o -name "*.scala" ")" -print | xargs chmod -f u-w""")
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 475 | execute_tar(context.dist_dir, "-czf " + | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 476 | File.bash_path(context.isabelle_archive) + " " + Bash.string(context.dist_name)) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 477 | } | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 478 | } | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 479 | |
| 73642 | 480 | def default_platform_families: List[Platform.Family.Value] = Platform.Family.list0 | 
| 481 | ||
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 482 | def build_release( | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 483 | options: Options, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 484 | context: Release_Context, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 485 | afp_rev: String = "", | 
| 73642 | 486 | platform_families: List[Platform.Family.Value] = default_platform_families, | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 487 | more_components: List[Path] = Nil, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 488 | website: Option[Path] = None, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 489 | build_sessions: List[String] = Nil, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 490 | build_library: Boolean = false, | 
| 75393 | 491 | parallel_jobs: Int = 1 | 
| 492 |   ): Unit = {
 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 493 | val progress = context.progress | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 494 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 495 | |
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 496 | /* release directory */ | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 497 | |
| 73629 | 498 | val archive = Release_Archive.read(context.isabelle_archive) | 
| 499 | ||
| 500 |     for (path <- List(context.isabelle, ISABELLE)) {
 | |
| 501 | Isabelle_System.rm_tree(context.dist_dir + path) | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 502 | } | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 503 | |
| 75394 | 504 |     Isabelle_System.with_tmp_file("archive", ext = "tar.gz") { archive_path =>
 | 
| 73629 | 505 | Bytes.write(archive_path, archive.bytes) | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 506 | val extract = | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 507 |         List("README", "NEWS", "ANNOUNCE", "COPYRIGHT", "CONTRIBUTORS", "doc").
 | 
| 73631 | 508 | map(name => context.dist_name + "/" + name) | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 509 | execute_tar(context.dist_dir, | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 510 | "-xzf " + File.bash_path(archive_path) + " " + Bash.strings(extract)) | 
| 75394 | 511 | } | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 512 | |
| 73629 | 513 | Isabelle_System.symlink(Path.explode(context.dist_name), context.dist_dir + ISABELLE) | 
| 64202 | 514 | |
| 515 | ||
| 516 | /* make application bundles */ | |
| 517 | ||
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 518 | val bundle_infos = platform_families.map(context.bundle_info) | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 519 | |
| 66730 | 520 |     for (bundle_info <- bundle_infos) {
 | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 521 | val isabelle_name = context.dist_name | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 522 | val platform = bundle_info.platform | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 523 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 524 |       progress.echo("\nApplication bundle for " + platform)
 | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 525 | |
| 75394 | 526 |       Isabelle_System.with_tmp_dir("build_release") { tmp_dir =>
 | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 527 | // release archive | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 528 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 529 | execute_tar(tmp_dir, "-xzf " + File.bash_path(context.isabelle_archive)) | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 530 | val other_isabelle = context.other_isabelle(tmp_dir) | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 531 | val isabelle_target = other_isabelle.isabelle_home | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 532 | |
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 533 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 534 | // bundled components | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 535 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 536 |         progress.echo("Bundled components:")
 | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 537 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 538 | val contrib_dir = Components.contrib(isabelle_target) | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 539 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 540 | val (bundled_components, jdk_component) = | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 541 | get_bundled_components(isabelle_target, platform) | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 542 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 543 | Components.resolve(context.components_base, bundled_components, | 
| 70102 | 544 | target_dir = Some(contrib_dir), | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 545 |           copy_dir = Some(context.dist_dir + Path.explode("contrib")),
 | 
| 70102 | 546 | progress = progress) | 
| 69413 | 547 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 548 | val more_components_names = | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 549 | more_components.map(Components.unpack(contrib_dir, _, progress = progress)) | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 550 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 551 | activate_components(isabelle_target, platform, more_components_names) | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 552 | |
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 553 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 554 | // Java parameters | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 555 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 556 | val java_options: List[String] = | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 557 |           (for {
 | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 558 | variable <- | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 559 | List( | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 560 | "ISABELLE_JAVA_SYSTEM_OPTIONS", | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 561 | "JEDIT_JAVA_SYSTEM_OPTIONS", | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 562 | "JEDIT_JAVA_OPTIONS") | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 563 | opt <- Word.explode(other_isabelle.getenv(variable)) | 
| 73068 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 564 | } | 
| 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 565 |           yield {
 | 
| 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 566 | val s = "-Dapple.awt.application.name=" | 
| 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 567 | if (opt.startsWith(s)) s + isabelle_name else opt | 
| 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 568 |           }) ::: List("-Disabelle.jedit_server=" + isabelle_name)
 | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 569 | |
| 75393 | 570 |         val classpath: List[Path] = {
 | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 571 | val base = isabelle_target.absolute | 
| 73988 | 572 |           val classpath1 = Path.split(other_isabelle.getenv("ISABELLE_CLASSPATH"))
 | 
| 573 |           val classpath2 = Path.split(other_isabelle.getenv("ISABELLE_SETUP_CLASSPATH"))
 | |
| 75394 | 574 |           (classpath1 ::: classpath2).map { path =>
 | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 575 | val abs_path = path.absolute | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 576 |             File.relative_path(base, abs_path) match {
 | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 577 | case Some(rel_path) => rel_path | 
| 73988 | 578 |               case None => error("Bad classpath element: " + abs_path)
 | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 579 | } | 
| 75394 | 580 | } | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 581 | } | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 582 | |
| 71542 
e76692ec6e5a
more usable defaults for high resolution on Linux, where the desktop environment usually lacks automatic scaling;
 wenzelm parents: 
71459diff
changeset | 583 |         val jedit_options = Path.explode("src/Tools/jEdit/etc/options")
 | 
| 73987 
fc363a3b690a
build.props for isabelle.jar, including isabelle.jedit;
 wenzelm parents: 
73894diff
changeset | 584 | val jedit_props = | 
| 
fc363a3b690a
build.props for isabelle.jar, including isabelle.jedit;
 wenzelm parents: 
73894diff
changeset | 585 |           Path.explode(other_isabelle.getenv("JEDIT_HOME") + "/properties/jEdit.props")
 | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 586 | |
| 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 587 | |
| 70101 | 588 | // build heaps | 
| 589 | ||
| 590 |         if (build_sessions.nonEmpty) {
 | |
| 75796 | 591 | build_heaps(options, platform, build_sessions, isabelle_target, progress = progress) | 
| 70101 | 592 | } | 
| 593 | ||
| 594 | ||
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 595 | // application bundling | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 596 | |
| 73990 
778ab9983f40
proper cross-platform build: jdk component is required for ISABELLE_SETUP_CLASSPATH in other_isabelle;
 wenzelm parents: 
73988diff
changeset | 597 | Components.purge(contrib_dir, platform) | 
| 
778ab9983f40
proper cross-platform build: jdk component is required for ISABELLE_SETUP_CLASSPATH in other_isabelle;
 wenzelm parents: 
73988diff
changeset | 598 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 599 |         platform match {
 | 
| 73637 | 600 | case Platform.Family.linux_arm | Platform.Family.linux => | 
| 75202 | 601 |             File.change(isabelle_target + jedit_options) {
 | 
| 602 |               _.replaceAll("jedit_reset_font_size : int =.*", "jedit_reset_font_size : int = 24")
 | |
| 603 | } | |
| 71542 
e76692ec6e5a
more usable defaults for high resolution on Linux, where the desktop environment usually lacks automatic scaling;
 wenzelm parents: 
71459diff
changeset | 604 | |
| 75202 | 605 |             File.change(isabelle_target + jedit_props) {
 | 
| 72378 | 606 |               _.replaceAll("console.fontsize=.*", "console.fontsize=18")
 | 
| 607 |                .replaceAll("helpviewer.fontsize=.*", "helpviewer.fontsize=18")
 | |
| 608 |                .replaceAll("metal.primary.fontsize=.*", "metal.primary.fontsize=18")
 | |
| 609 |                .replaceAll("metal.secondary.fontsize=.*", "metal.secondary.fontsize=18")
 | |
| 610 |                .replaceAll("view.fontsize=.*", "view.fontsize=24")
 | |
| 75202 | 611 |                .replaceAll("view.gutter.fontsize=.*", "view.gutter.fontsize=16")
 | 
| 612 | } | |
| 71542 
e76692ec6e5a
more usable defaults for high resolution on Linux, where the desktop environment usually lacks automatic scaling;
 wenzelm parents: 
71459diff
changeset | 613 | |
| 73068 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 614 | make_isabelle_options( | 
| 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 615 |               isabelle_target + Path.explode("Isabelle.options"), java_options)
 | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 616 | |
| 73193 | 617 | make_isabelle_app(platform, isabelle_target, isabelle_name, jdk_component, classpath) | 
| 69417 | 618 | |
| 73637 | 619 |             progress.echo("Packaging " + bundle_info.name + " ...")
 | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 620 | execute_tar(tmp_dir, | 
| 73637 | 621 | "-czf " + File.bash_path(context.dist_dir + bundle_info.path) + " " + | 
| 70242 | 622 | Bash.string(isabelle_name)) | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 623 | |
| 69432 
d072f3287ffa
discontinued somewhat point dmg: plain .tar.gz is smaller and more convenient to install;
 wenzelm parents: 
69425diff
changeset | 624 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 625 | case Platform.Family.macos => | 
| 75202 | 626 |             File.change(isabelle_target + jedit_props) {
 | 
| 73112 | 627 |               _.replaceAll("delete-line.shortcut=.*", "delete-line.shortcut=C+d")
 | 
| 75202 | 628 |                .replaceAll("delete.shortcut2=.*", "delete.shortcut2=A+d")
 | 
| 629 | } | |
| 69417 | 630 | |
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 631 | |
| 73193 | 632 | // macOS application bundle | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 633 | |
| 73193 | 634 |             val app_contents = isabelle_target + Path.explode("Contents")
 | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 635 | |
| 73077 | 636 |             for (icon <- List("lib/logo/isabelle.icns", "lib/logo/theory.icns")) {
 | 
| 73317 | 637 | Isabelle_System.copy_file(isabelle_target + Path.explode(icon), | 
| 73193 | 638 |                 Isabelle_System.make_directory(app_contents + Path.explode("Resources")))
 | 
| 73077 | 639 | } | 
| 73075 
893310d6d76d
recovered bundle icons (not application) from macos_app;
 wenzelm parents: 
73074diff
changeset | 640 | |
| 73083 | 641 | make_isabelle_plist( | 
| 73629 | 642 |               app_contents + Path.explode("Info.plist"), isabelle_name, archive.id)
 | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 643 | |
| 73193 | 644 | make_isabelle_app(platform, isabelle_target, isabelle_name, jdk_component, | 
| 645 | classpath, dock_icon = true) | |
| 646 | ||
| 647 |             val isabelle_options = Path.explode("Isabelle.options")
 | |
| 648 | make_isabelle_options( | |
| 649 | isabelle_target + isabelle_options, | |
| 650 |               java_options ::: List("-Disabelle.app=true"))
 | |
| 651 | ||
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 652 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 653 | // application archive | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 654 | |
| 73637 | 655 |             progress.echo("Packaging " + bundle_info.name + " ...")
 | 
| 73193 | 656 | |
| 657 | val isabelle_app = Path.explode(isabelle_name + ".app") | |
| 73317 | 658 | Isabelle_System.move_file(tmp_dir + Path.explode(isabelle_name), | 
| 659 | tmp_dir + isabelle_app) | |
| 73193 | 660 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 661 | execute_tar(tmp_dir, | 
| 73637 | 662 | "-czf " + File.bash_path(context.dist_dir + bundle_info.path) + " " + | 
| 70242 | 663 | File.bash_path(isabelle_app)) | 
| 69432 
d072f3287ffa
discontinued somewhat point dmg: plain .tar.gz is smaller and more convenient to install;
 wenzelm parents: 
69425diff
changeset | 664 | |
| 69417 | 665 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 666 | case Platform.Family.windows => | 
| 75202 | 667 |             File.change(isabelle_target + jedit_props) {
 | 
| 668 |               _.replaceAll("foldPainter=.*", "foldPainter=Square")
 | |
| 669 | } | |
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 670 | |
| 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 671 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 672 | // application launcher | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 673 | |
| 73317 | 674 |             Isabelle_System.move_file(isabelle_target + Path.explode("contrib/windows_app"), tmp_dir)
 | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 675 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 676 |             val app_template = Path.explode("~~/Admin/Windows/launch4j")
 | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 677 | |
| 73068 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 678 | make_isabelle_options( | 
| 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 679 | isabelle_target + Path.explode(isabelle_name + ".l4j.ini"), | 
| 
a95f5ae5a12a
discontinued macOS JavaAppLauncher: re-use plain shell script;
 wenzelm parents: 
73067diff
changeset | 680 | java_options, line_ending = "\r\n") | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 681 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 682 |             val isabelle_xml = Path.explode("isabelle.xml")
 | 
| 73637 | 683 | val isabelle_exe = bundle_info.path | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 684 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 685 | File.write(tmp_dir + isabelle_xml, | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 686 | File.read(app_template + isabelle_xml) | 
| 72387 | 687 |                 .replace("{ISABELLE_NAME}", isabelle_name)
 | 
| 688 |                 .replace("{OUTFILE}", File.platform_path(isabelle_target + isabelle_exe))
 | |
| 689 |                 .replace("{ICON}",
 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 690 |                   File.platform_path(app_template + Path.explode("isabelle_transparent.ico")))
 | 
| 72387 | 691 |                 .replace("{SPLASH}",
 | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 692 |                   File.platform_path(app_template + Path.explode("isabelle.bmp")))
 | 
| 72387 | 693 |                 .replace("{CLASSPATH}",
 | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 694 | cat_lines(classpath.map(cp => | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 695 |                     "    <cp>%EXEDIR%\\" + File.platform_path(cp).replace('/', '\\') + "</cp>")))
 | 
| 72387 | 696 |                 .replace("\\jdk\\", "\\" + jdk_component + "\\"))
 | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 697 | |
| 74857 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 698 | val java_opts = | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 699 | bash_java_opens( | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 700 | "java.base/java.io", | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 701 | "java.base/java.lang", | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 702 | "java.base/java.lang.reflect", | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 703 | "java.base/java.text", | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 704 | "java.base/java.util", | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 705 | "java.desktop/java.awt.font") | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 706 | val launch4j_jar = | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 707 |               Path.explode("windows_app/launch4j-" + Platform.family + "/launch4j.jar")
 | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 708 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 709 | execute(tmp_dir, | 
| 74857 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 710 | cat_lines(List( | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 711 | "export LAUNCH4J=" + File.bash_platform_path(launch4j_jar), | 
| 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 712 | "isabelle java " + java_opts + " -jar \"$LAUNCH4J\" isabelle.xml"))) | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 713 | |
| 73317 | 714 |             Isabelle_System.copy_file(app_template + Path.explode("manifest.xml"),
 | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 715 |               isabelle_target + isabelle_exe.ext("manifest"))
 | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 716 | |
| 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 717 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 718 | // Cygwin setup | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 719 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 720 |             val cygwin_template = Path.explode("~~/Admin/Windows/Cygwin")
 | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 721 | |
| 73317 | 722 |             Isabelle_System.copy_file(cygwin_template + Path.explode("Cygwin-Terminal.bat"),
 | 
| 723 | isabelle_target) | |
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 724 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 725 | val cygwin_mirror = | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 726 |               File.read(isabelle_target + Path.explode("contrib/cygwin/isabelle/cygwin_mirror"))
 | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 727 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 728 |             val cygwin_bat = Path.explode("Cygwin-Setup.bat")
 | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 729 | File.write(isabelle_target + cygwin_bat, | 
| 72387 | 730 |               File.read(cygwin_template + cygwin_bat).replace("{MIRROR}", cygwin_mirror))
 | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 731 | File.set_executable(isabelle_target + cygwin_bat, true) | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 732 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 733 |             for (name <- List("isabelle/postinstall", "isabelle/rebaseall")) {
 | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 734 | val path = Path.explode(name) | 
| 73317 | 735 | Isabelle_System.copy_file(cygwin_template + path, | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 736 |                 isabelle_target + Path.explode("contrib/cygwin") + path)
 | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 737 | } | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 738 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 739 | execute(isabelle_target, | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 740 | """find . -type f -not -name "*.exe" -not -name "*.dll" """ + | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 741 | (if (Platform.is_macos) "-perm +100" else "-executable") + | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 742 | " -print0 > contrib/cygwin/isabelle/executables") | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 743 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 744 | execute(isabelle_target, | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 745 |               """find . -type l -exec echo "{}" ";" -exec readlink "{}" ";" """ +
 | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 746 | """> contrib/cygwin/isabelle/symlinks""") | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 747 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 748 |             execute(isabelle_target, """find . -type l -exec rm "{}" ";" """)
 | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 749 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 750 |             File.write(isabelle_target + Path.explode("contrib/cygwin/isabelle/uninitialized"), "")
 | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 751 | |
| 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 752 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 753 | // executable archive (self-extracting 7z) | 
| 69424 
840f0cadeba8
clarified application bundling: discontinued redundant archives;
 wenzelm parents: 
69417diff
changeset | 754 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 755 | val archive_name = isabelle_name + ".7z" | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 756 | val exe_archive = tmp_dir + Path.explode(archive_name) | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 757 | exe_archive.file.delete | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 758 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 759 |             progress.echo("Packaging " + archive_name + " ...")
 | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 760 | execute(tmp_dir, | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 761 | "7z -y -bd a " + File.bash_path(exe_archive) + " " + Bash.string(isabelle_name)) | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 762 |             if (!exe_archive.is_file) error("Failed to create archive: " + exe_archive)
 | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 763 | |
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 764 |             val sfx_exe = tmp_dir + Path.explode("windows_app/7zsd_All_x64.sfx")
 | 
| 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 765 | val sfx_txt = | 
| 72387 | 766 |               File.read(Path.explode("~~/Admin/Windows/Installer/sfx.txt"))
 | 
| 767 |                 .replace("{ISABELLE_NAME}", isabelle_name)
 | |
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 768 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 769 | Bytes.write(context.dist_dir + isabelle_exe, | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 770 | Bytes.read(sfx_exe) + Bytes(sfx_txt) + Bytes.read(exe_archive)) | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 771 | File.set_executable(context.dist_dir + isabelle_exe, true) | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 772 | } | 
| 75394 | 773 | } | 
| 70098 
956d2430cb29
more robust: always (re)build platform application bundles;
 wenzelm parents: 
70046diff
changeset | 774 |       progress.echo("DONE")
 | 
| 64202 | 775 | } | 
| 776 | ||
| 777 | ||
| 778 | /* minimal website */ | |
| 779 | ||
| 64361 | 780 |     for (dir <- website) {
 | 
| 781 | val website_platform_bundles = | |
| 782 |         for {
 | |
| 66730 | 783 | bundle_info <- bundle_infos | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 784 | if (context.dist_dir + bundle_info.path).is_file | 
| 69432 
d072f3287ffa
discontinued somewhat point dmg: plain .tar.gz is smaller and more convenient to install;
 wenzelm parents: 
69425diff
changeset | 785 | } yield (bundle_info.name, bundle_info) | 
| 64206 | 786 | |
| 71275 | 787 | val isabelle_link = | 
| 73629 | 788 | HTML.link(Isabelle_System.isabelle_repository.changeset(archive.id), | 
| 789 |           HTML.text("Isabelle/" + archive.id))
 | |
| 64405 
81bac77929d9
just one task to identify Isabelle + AFP repository snapshots and build release;
 wenzelm parents: 
64371diff
changeset | 790 | val afp_link = | 
| 73610 | 791 | HTML.link(Isabelle_System.afp_repository.changeset(afp_rev), | 
| 792 |           HTML.text("AFP/" + afp_rev))
 | |
| 64405 
81bac77929d9
just one task to identify Isabelle + AFP repository snapshots and build release;
 wenzelm parents: 
64371diff
changeset | 793 | |
| 65838 | 794 | HTML.write_document(dir, "index.html", | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 795 | List(HTML.title(context.dist_name)), | 
| 65838 | 796 | List( | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 797 | HTML.section(context.dist_name), | 
| 73626 | 798 |           HTML.subsection("Downloads"),
 | 
| 65838 | 799 | HTML.itemize( | 
| 73626 | 800 |             List(HTML.link(context.dist_name + ".tar.gz", HTML.text("Source archive"))) ::
 | 
| 66730 | 801 |             website_platform_bundles.map({ case (bundle, bundle_info) =>
 | 
| 73626 | 802 | List(HTML.link(bundle, HTML.text(bundle_info.platform_description + " bundle"))) })), | 
| 71275 | 803 |           HTML.subsection("Repositories"),
 | 
| 804 | HTML.itemize( | |
| 805 | List(List(isabelle_link)) ::: (if (afp_rev == "") Nil else List(List(afp_link)))))) | |
| 64202 | 806 | |
| 73626 | 807 | Isabelle_System.copy_file(context.isabelle_archive, dir) | 
| 808 | ||
| 809 |       for ((bundle, _) <- website_platform_bundles) {
 | |
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 810 | Isabelle_System.copy_file(context.dist_dir + Path.explode(bundle), dir) | 
| 73626 | 811 | } | 
| 64361 | 812 | } | 
| 64202 | 813 | |
| 814 | ||
| 815 | /* HTML library */ | |
| 816 | ||
| 64203 | 817 |     if (build_library) {
 | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 818 |       if (context.isabelle_library_archive.is_file) {
 | 
| 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 819 |         progress.echo_warning("Library archive already exists: " + context.isabelle_library_archive)
 | 
| 69168 
68816d1c73a7
eliminated "isabelle makedist" -- prefer Scala over bash/perl scripting;
 wenzelm parents: 
69167diff
changeset | 820 | } | 
| 64203 | 821 |       else {
 | 
| 75394 | 822 |         Isabelle_System.with_tmp_dir("build_release") { tmp_dir =>
 | 
| 823 | val bundle = | |
| 824 | context.dist_dir + Path.explode(context.dist_name + "_" + Platform.family + ".tar.gz") | |
| 825 | execute_tar(tmp_dir, "-xzf " + File.bash_path(bundle)) | |
| 64316 | 826 | |
| 75394 | 827 | val other_isabelle = context.other_isabelle(tmp_dir) | 
| 64316 | 828 | |
| 75394 | 829 | Isabelle_System.make_directory(other_isabelle.etc) | 
| 830 | File.write(other_isabelle.etc_settings, "ML_OPTIONS=\"--minheap 1000 --maxheap 4000\"\n") | |
| 74856 | 831 | |
| 75394 | 832 |           other_isabelle.bash("bin/isabelle build -f -j " + parallel_jobs +
 | 
| 833 | " -o browser_info -o document=pdf -o document_variants=document:outline=/proof,/ML" + | |
| 834 | " -o system_heaps -c -a -d '~~/src/Benchmarks'", echo = true).check | |
| 835 | other_isabelle.isabelle_home_user.file.delete | |
| 64316 | 836 | |
| 75394 | 837 | execute(tmp_dir, "chmod -R a+r " + Bash.string(context.dist_name)) | 
| 838 | execute(tmp_dir, "chmod -R g=o " + Bash.string(context.dist_name)) | |
| 839 | execute_tar(tmp_dir, "-czf " + File.bash_path(context.isabelle_library_archive) + | |
| 840 | " " + Bash.string(context.dist_name + "/browser_info")) | |
| 841 | } | |
| 64203 | 842 | } | 
| 843 | } | |
| 64202 | 844 | } | 
| 845 | ||
| 846 | ||
| 847 | ||
| 848 | /** command line entry point **/ | |
| 849 | ||
| 75393 | 850 |   def main(args: Array[String]): Unit = {
 | 
| 71632 | 851 |     Command_Line.tool {
 | 
| 64405 
81bac77929d9
just one task to identify Isabelle + AFP repository snapshots and build release;
 wenzelm parents: 
64371diff
changeset | 852 | var afp_rev = "" | 
| 69434 | 853 | var components_base: Path = Components.default_components_base | 
| 73607 
fc13738e1933
clarified command-line, following other build_XYZ tools;
 wenzelm parents: 
73582diff
changeset | 854 | var target_dir = Path.current | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 855 | var release_name = "" | 
| 73629 | 856 | var source_archive = "" | 
| 64211 | 857 | var website: Option[Path] = None | 
| 70101 | 858 | var build_sessions: List[String] = Nil | 
| 69413 | 859 | var more_components: List[Path] = Nil | 
| 64202 | 860 | var parallel_jobs = 1 | 
| 861 | var build_library = false | |
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 862 | var options = Options.init() | 
| 73642 | 863 | var platform_families = default_platform_families | 
| 64202 | 864 | var rev = "" | 
| 865 | ||
| 866 |       val getopts = Getopts("""
 | |
| 73894 | 867 | Usage: Admin/build_release [OPTIONS] | 
| 64202 | 868 | |
| 869 | Options are: | |
| 64405 
81bac77929d9
just one task to identify Isabelle + AFP repository snapshots and build release;
 wenzelm parents: 
64371diff
changeset | 870 | -A REV corresponding AFP changeset id | 
| 69434 | 871 | -C DIR base directory for Isabelle components (default: """ + | 
| 872 | Components.default_components_base + """) | |
| 73607 
fc13738e1933
clarified command-line, following other build_XYZ tools;
 wenzelm parents: 
73582diff
changeset | 873 | -D DIR target directory (default ".") | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 874 | -R RELEASE explicit release name | 
| 73629 | 875 | -S ARCHIVE use existing source archive (file or URL) | 
| 64211 | 876 | -W WEBSITE produce minimal website in given directory | 
| 70101 | 877 | -b SESSIONS build platform-specific session images (separated by commas) | 
| 69413 | 878 | -c ARCHIVE clean bundling with additional component .tar.gz archive | 
| 64202 | 879 | -j INT maximum number of parallel jobs (default 1) | 
| 880 | -l build library | |
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 881 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | 
| 73642 | 882 |     -p NAMES     platform families (default: """ + default_platform_families.mkString(",") + """)
 | 
| 73629 | 883 | -r REV Mercurial changeset id (default: ARCHIVE or RELEASE or tip) | 
| 64202 | 884 | |
| 885 | Build Isabelle release in base directory, using the local repository clone. | |
| 886 | """, | |
| 64405 
81bac77929d9
just one task to identify Isabelle + AFP repository snapshots and build release;
 wenzelm parents: 
64371diff
changeset | 887 | "A:" -> (arg => afp_rev = arg), | 
| 69434 | 888 | "C:" -> (arg => components_base = Path.explode(arg)), | 
| 73607 
fc13738e1933
clarified command-line, following other build_XYZ tools;
 wenzelm parents: 
73582diff
changeset | 889 | "D:" -> (arg => target_dir = Path.explode(arg)), | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 890 | "R:" -> (arg => release_name = arg), | 
| 73629 | 891 | "S:" -> (arg => source_archive = arg), | 
| 64211 | 892 | "W:" -> (arg => website = Some(Path.explode(arg))), | 
| 70101 | 893 |         "b:" -> (arg => build_sessions = space_explode(',', arg)),
 | 
| 69413 | 894 | "c:" -> (arg => | 
| 895 |           {
 | |
| 896 | val path = Path.explode(arg) | |
| 897 | Components.Archive.get_name(path.file_name) | |
| 898 | more_components = more_components ::: List(path) | |
| 899 | }), | |
| 64202 | 900 | "j:" -> (arg => parallel_jobs = Value.Int.parse(arg)), | 
| 64316 | 901 | "l" -> (_ => build_library = true), | 
| 69401 
7a1b7b737c02
eliminated old makedist_bundle and remote_dmg: build_release does everything in Scala;
 wenzelm parents: 
69400diff
changeset | 902 | "o:" -> (arg => options = options + arg), | 
| 69410 | 903 |         "p:" -> (arg => platform_families = space_explode(',', arg).map(Platform.Family.parse)),
 | 
| 64204 | 904 | "r:" -> (arg => rev = arg)) | 
| 64202 | 905 | |
| 906 | val more_args = getopts(args) | |
| 73607 
fc13738e1933
clarified command-line, following other build_XYZ tools;
 wenzelm parents: 
73582diff
changeset | 907 | if (more_args.nonEmpty) getopts.usage() | 
| 64202 | 908 | |
| 73629 | 909 |       if (platform_families.contains(Platform.Family.windows) && !Isabelle_System.bash("7z i").ok)
 | 
| 910 |         error("Building for windows requires 7z")
 | |
| 911 | ||
| 73634 | 912 | val progress = new Console_Progress() | 
| 73629 | 913 | def make_context(name: String): Release_Context = | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 914 | Release_Context(target_dir, | 
| 73629 | 915 | release_name = name, | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 916 | components_base = components_base, | 
| 73634 | 917 | progress = progress) | 
| 64202 | 918 | |
| 73629 | 919 | val context = | 
| 920 |         if (source_archive.isEmpty) {
 | |
| 921 | val context = make_context(release_name) | |
| 922 | val version = proper_string(rev) orElse proper_string(release_name) getOrElse "tip" | |
| 923 | build_release_archive(context, version, parallel_jobs = parallel_jobs) | |
| 924 | context | |
| 925 | } | |
| 926 |         else {
 | |
| 73634 | 927 | val archive = | 
| 928 | Release_Archive.get(source_archive, rename = release_name, progress = progress) | |
| 73632 | 929 | val context = make_context(archive.identifier) | 
| 73634 | 930 | Isabelle_System.make_directory(context.dist_dir) | 
| 73629 | 931 | use_release_archive(context, archive, id = rev) | 
| 932 | context | |
| 933 | } | |
| 69415 | 934 | |
| 73630 | 935 | build_release(options, context, afp_rev = afp_rev, platform_families = platform_families, | 
| 74857 
25e9e7088561
address problems with launch4j and jdk-17 (see also 41d009462d3c);
 wenzelm parents: 
74856diff
changeset | 936 | more_components = more_components, build_sessions = build_sessions, | 
| 73625 
f8f065e20837
misc tuning and clarification: more explicit types Release_Context, Release_Archive;
 wenzelm parents: 
73610diff
changeset | 937 | build_library = build_library, parallel_jobs = parallel_jobs, website = website) | 
| 64202 | 938 | } | 
| 939 | } | |
| 940 | } |