diff --git a/src/Pure/Admin/build_history.scala b/src/Pure/Admin/build_history.scala --- a/src/Pure/Admin/build_history.scala +++ b/src/Pure/Admin/build_history.scala @@ -1,375 +1,376 @@ /* Title: Pure/Admin/build_history.scala Author: Makarius Build other history versions. */ package isabelle import java.io.{File => JFile} import java.time.format.DateTimeFormatter import java.util.Locale object Build_History { /* log files */ val BUILD_HISTORY = "build_history" val META_INFO_MARKER = "\fmeta_info = " /* augment settings */ def augment_settings( other_isabelle: Other_Isabelle, threads: Int, arch_64: Boolean = false, heap: Int = default_heap, max_heap: Option[Int] = None, more_settings: List[String]): String = { val (ml_platform, ml_settings) = { val windows_32 = "x86-windows" val windows_64 = "x86_64-windows" val platform_32 = other_isabelle("getenv -b ISABELLE_PLATFORM32").check.out val platform_64 = other_isabelle("getenv -b ISABELLE_PLATFORM64").check.out val platform_family = other_isabelle("getenv -b ISABELLE_PLATFORM_FAMILY").check.out val polyml_home = try { Path.explode(other_isabelle("getenv -b ML_HOME").check.out).dir } catch { case ERROR(msg) => error("Bad ML_HOME: " + msg) } def ml_home(platform: String): Path = polyml_home + Path.explode(platform) def err(platform: String): Nothing = error("Platform " + platform + " unavailable on this machine") def check_dir(platform: String): Boolean = platform != "" && ml_home(platform).is_dir val ml_platform = if (Platform.is_windows && arch_64) { if (check_dir(windows_64)) windows_64 else err(windows_64) } else if (Platform.is_windows && !arch_64) { if (check_dir(windows_32)) windows_32 else platform_32 // x86-cygwin } else { val (platform, platform_name) = if (arch_64) (platform_64, "x86_64-" + platform_family) else (platform_32, "x86-" + platform_family) if (check_dir(platform)) platform else err(platform_name) } val ml_options = "--minheap " + heap + (if (max_heap.isDefined) " --maxheap " + max_heap.get else "") + " --gcthreads " + threads + (if (ml_platform.endsWith("-windows")) " --codepage utf8" else "") (ml_platform, List( "ML_HOME=" + File.bash_path(ml_home(ml_platform)), "ML_PLATFORM=" + quote(ml_platform), "ML_OPTIONS=" + quote(ml_options))) } val thread_settings = List( "ISABELLE_JAVA_SYSTEM_OPTIONS=\"$ISABELLE_JAVA_SYSTEM_OPTIONS -Disabelle.threads=" + threads + "\"", "ISABELLE_BUILD_OPTIONS=\"threads=" + threads + "\"") val settings = List(ml_settings, thread_settings) ::: (if (more_settings.isEmpty) Nil else List(more_settings)) File.append(other_isabelle.etc_settings, "\n" + cat_lines(settings.map(terminate_lines(_)))) ml_platform } /** build_history **/ private val default_rev = "tip" private val default_threads = 1 private val default_heap = 1000 private val default_isabelle_identifier = "build_history" def build_history( hg: Mercurial.Repository, progress: Progress = Ignore_Progress, rev: String = default_rev, isabelle_identifier: String = default_isabelle_identifier, components_base: String = "", fresh: Boolean = false, nonfree: Boolean = false, multicore_base: Boolean = false, threads_list: List[Int] = List(default_threads), arch_64: Boolean = false, heap: Int = default_heap, max_heap: Option[Int] = None, more_settings: List[String] = Nil, verbose: Boolean = false, build_args: List[String] = Nil): List[(Process_Result, Path)] = { /* sanity checks */ if (File.eq(Path.explode("~~"), hg.root)) error("Repository coincides with ISABELLE_HOME=" + Path.explode("~~").expand) for (threads <- threads_list if threads < 1) error("Bad threads value < 1: " + threads) if (heap < 100) error("Bad heap value < 100: " + heap) if (max_heap.isDefined && max_heap.get < heap) error("Bad max_heap value < heap: " + max_heap.get) System.getenv("ISABELLE_SETTINGS_PRESENT") match { case null | "" => case _ => error("Cannot run build_history within existing Isabelle settings environment") } /* init repository */ hg.update(rev = rev, clean = true) progress.echo_if(verbose, hg.log(rev, options = "-l1")) val isabelle_version = hg.id(rev) val other_isabelle = new Other_Isabelle(progress, hg.root, isabelle_identifier) /* main */ val build_history_date = Date.now() val build_host = Isabelle_System.hostname() var first_build = true for (threads <- threads_list) yield { /* init settings */ other_isabelle.init_settings(components_base, nonfree) other_isabelle.resolve_components(verbose) val ml_platform = augment_settings(other_isabelle, threads, arch_64, heap, max_heap, more_settings) val isabelle_output = Path.explode(other_isabelle("getenv -b ISABELLE_OUTPUT").check.out) val isabelle_output_log = isabelle_output + Path.explode("log") val isabelle_base_log = isabelle_output + Path.explode("../base_log") if (first_build) { other_isabelle.resolve_components(verbose) other_isabelle.bash( "env PATH=\"" + File.bash_path(Path.explode("~~/lib/dummy_stty").expand) + ":$PATH\" " + "bin/isabelle jedit -b" + (if (fresh) " -f" else ""), redirect = true, echo = verbose).check Isabelle_System.rm_tree(isabelle_base_log) } Isabelle_System.rm_tree(isabelle_output) Isabelle_System.mkdirs(isabelle_output) /* build */ if (multicore_base && !first_build && isabelle_base_log.is_dir) Isabelle_System.copy_dir(isabelle_base_log, isabelle_output_log) val build_start = Date.now() val res = other_isabelle("build -v " + File.bash_args(build_args), redirect = true, echo = verbose) val build_end = Date.now() /* output log */ val log_path = other_isabelle.isabelle_home_user + Build_Log.log_subdir(build_history_date) + Build_Log.log_filename( BUILD_HISTORY, build_history_date, ml_platform, "M" + threads).ext("gz") val build_info = Build_Log.Log_File(log_path.base.implode, res.out_lines).parse_build_info() val meta_info = List(Build_Log.Field.build_engine -> BUILD_HISTORY, Build_Log.Field.build_host -> build_host, Build_Log.Field.build_start -> Build_Log.print_date(build_start), Build_Log.Field.build_end -> Build_Log.print_date(build_end), Build_Log.Field.isabelle_version -> isabelle_version) val ml_statistics = build_info.finished_sessions.flatMap(session_name => { val session_log = isabelle_output_log + Path.explode(session_name).ext("gz") if (session_log.is_file) { Build_Log.Log_File(session_log).parse_session_info(ml_statistics = true). ml_statistics.map(props => (Build_Log.SESSION_NAME -> session_name) :: props) } else Nil }) val heap_sizes = build_info.finished_sessions.flatMap(session_name => { val heap = isabelle_output + Path.explode(session_name) if (heap.is_file) Some("Heap " + session_name + " (" + Value.Long(heap.file.length) + " bytes)") else None }) Isabelle_System.mkdirs(log_path.dir) File.write_gzip(log_path, terminate_lines( Build_Log.Log_File.print_props(META_INFO_MARKER, meta_info) :: res.out_lines ::: ml_statistics.map(Build_Log.Log_File.print_props(Build_Log.ML_STATISTICS_MARKER, _)) ::: heap_sizes)) /* next build */ if (multicore_base && first_build && isabelle_output_log.is_dir) Isabelle_System.copy_dir(isabelle_output_log, isabelle_base_log) first_build = false (res, log_path) } } /* command line entry point */ def main(args: Array[String]) { Command_Line.tool0 { var multicore_base = false var components_base = "" var heap: Option[Int] = None var max_heap: Option[Int] = None var threads_list = List(default_threads) var isabelle_identifier = default_isabelle_identifier var more_settings: List[String] = Nil var fresh = false var arch_64 = false var nonfree = false var rev = default_rev var verbose = false val getopts = Getopts(""" Usage: isabelle build_history [OPTIONS] REPOSITORY [ARGS ...] Options are: -B first multicore build serves as base for scheduling information -C DIR base directory for Isabelle components (default: $ISABELLE_HOME_USER/../contrib) -H SIZE minimal ML heap in MB (default: """ + default_heap + """ for x86, """ + default_heap * 2 + """ for x86_64) -M THREADS multicore configurations (comma-separated list, default: """ + default_threads + """) -N NAME alternative ISABELLE_IDENTIFIER (default: """ + default_isabelle_identifier + """) -U SIZE maximal ML heap in MB (default: unbounded) -e TEXT additional text for generated etc/settings -f fresh build of Isabelle/Scala components (recommended) -m ARCH processor architecture (32=x86, 64=x86_64, default: x86) -n include nonfree components -r REV update to revision (default: """ + default_rev + """) -v verbose Build Isabelle sessions from the history of another REPOSITORY clone, passing ARGS directly to its isabelle build tool. """, "B" -> (_ => multicore_base = true), "C:" -> (arg => components_base = arg), "H:" -> (arg => heap = Some(Value.Int.parse(arg))), "M:" -> (arg => threads_list = space_explode(',', arg).map(Value.Int.parse(_))), "N:" -> (arg => isabelle_identifier = arg), "U:" -> (arg => max_heap = Some(Value.Int.parse(arg))), "e:" -> (arg => more_settings = more_settings ::: List(arg)), "f" -> (_ => fresh = true), "m:" -> { case "32" | "x86" => arch_64 = false case "64" | "x86_64" => arch_64 = true case bad => error("Bad processor architecture: " + quote(bad)) }, "n" -> (_ => nonfree = true), "r:" -> (arg => rev = arg), "v" -> (_ => verbose = true)) val more_args = getopts(args) val (root, build_args) = more_args match { case root :: build_args => (root, build_args) case _ => getopts.usage() } val hg = Mercurial.repository(Path.explode(root)) val progress = new Console_Progress(stderr = true) val results = build_history(hg, progress = progress, rev = rev, isabelle_identifier = isabelle_identifier, components_base = components_base, fresh = fresh, nonfree = nonfree, multicore_base = multicore_base, threads_list = threads_list, arch_64 = arch_64, heap = heap.getOrElse(if (arch_64) default_heap * 2 else default_heap), max_heap = max_heap, more_settings = more_settings, verbose = verbose, build_args = build_args) for ((_, log_path) <- results) Output.writeln(log_path.implode, stdout = true) val rc = (0 /: results) { case (rc, (res, _)) => rc max res.rc } if (rc != 0) sys.exit(rc) } } /** remote build_history -- via command-line **/ def remote_build_history( session: SSH.Session, isabelle_repos_self: Path, isabelle_repos_other: Path, isabelle_repos_source: String = "http://isabelle.in.tum.de/repos/isabelle", self_update: Boolean = false, progress: Progress = Ignore_Progress, options: String = "", args: String = ""): List[(String, Bytes)] = { using(session.sftp())(sftp => { - val isabelle_admin = sftp.path(isabelle_repos_self + Path.explode("Admin")) + val isabelle_admin = sftp.remote_path(isabelle_repos_self + Path.explode("Admin")) /* prepare repository clones */ val isabelle_hg = Mercurial.setup_repository( isabelle_repos_source, isabelle_repos_self, ssh = Some(session)) if (self_update) { isabelle_hg.pull() isabelle_hg.update(clean = true) session.execute(File.bash_string(isabelle_admin + "/build") + " jars_fresh").check } Mercurial.setup_repository( - sftp.path(isabelle_repos_self), isabelle_repos_other, ssh = Some(session)) + sftp.remote_path(isabelle_repos_self), isabelle_repos_other, ssh = Some(session)) /* Admin/build_history */ val result = session.execute( File.bash_string(isabelle_admin + "/build_history") + " " + options + " " + - File.bash_string(sftp.path(isabelle_repos_other)) + " " + args, - progress_stderr = progress.echo(_)) + File.bash_string(sftp.remote_path(isabelle_repos_other)) + " " + args, + progress_stderr = progress.echo(_)).check - result.check.out_lines.map(log => (Path.explode(log).base.implode, sftp.read_bytes(log))) + for (line <- result.out_lines; log = Path.explode(line)) + yield (log.base.implode, sftp.read_bytes(log)) }) } } diff --git a/src/Pure/Admin/remote_dmg.scala b/src/Pure/Admin/remote_dmg.scala --- a/src/Pure/Admin/remote_dmg.scala +++ b/src/Pure/Admin/remote_dmg.scala @@ -1,65 +1,65 @@ /* Title: Pure/Admin/remote_dmg.scala Author: Makarius Build dmg on remote Mac OS X system. */ package isabelle object Remote_DMG { def remote_dmg(session: SSH.Session, tar_gz_file: Path, dmg_file: Path, volume_name: String = "") { session.with_tmp_dir(remote_dir => using(session.sftp())(sftp => { - val cd = "cd " + File.bash_string(remote_dir) + "; " + val cd = "cd " + File.bash_string(sftp.remote_path(remote_dir)) + "; " - sftp.write_file(remote_dir + "/dmg.tar.gz", tar_gz_file) + sftp.write_file(remote_dir + Path.explode("dmg.tar.gz"), tar_gz_file) session.execute(cd + "mkdir root && tar -C root -xzf dmg.tar.gz").check session.execute( cd + "hdiutil create -srcfolder root" + (if (volume_name == "") "" else " -volname " + File.bash_string(volume_name)) + " dmg.dmg").check - sftp.read_file(remote_dir + "/dmg.dmg", dmg_file) + sftp.read_file(remote_dir + Path.explode("dmg.dmg"), dmg_file) })) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("remote_dmg", "build dmg on remote Mac OS X system", args => { Command_Line.tool0 { var port = SSH.default_port var volume_name = "" val getopts = Getopts(""" Usage: isabelle remote_dmg [OPTIONS] USER@HOST TAR_GZ_FILE DMG_FILE Options are: -p PORT alternative SSH port (default: """ + SSH.default_port + """) -V NAME specify volume name Turn the contents of a tar.gz file into a dmg file -- produced on a remote Mac OS X system. """, "p:" -> (arg => port = Value.Int.parse(arg)), "V:" -> (arg => volume_name = arg)) val more_args = getopts(args) val (user, host, tar_gz_file, dmg_file) = more_args match { case List(SSH.Target(user, host), tar_gz_file, dmg_file) => (user, host, Path.explode(tar_gz_file), Path.explode(dmg_file)) case _ => getopts.usage() } val ssh = SSH.init(Options.init) using(ssh.open_session(user = user, host = host, port = port))( remote_dmg(_, tar_gz_file, dmg_file, volume_name)) } }, admin = true) } diff --git a/src/Pure/General/mercurial.scala b/src/Pure/General/mercurial.scala --- a/src/Pure/General/mercurial.scala +++ b/src/Pure/General/mercurial.scala @@ -1,103 +1,103 @@ /* Title: Pure/General/mercurial.scala Author: Makarius Support for Mercurial repositories. */ package isabelle import java.io.{File => JFile} object Mercurial { /* command-line syntax */ def optional(s: String, prefix: String = ""): String = if (s == "") "" else " " + prefix + " " + File.bash_string(s) def opt_flag(flag: String, b: Boolean): String = if (b) " " + flag else "" def opt_rev(s: String): String = optional(s, "--rev") def opt_template(s: String): String = optional(s, "--template") /* repository access */ def repository(root: Path, ssh: Option[SSH.Session] = None): Repository = { val hg = new Repository(root, ssh) hg.command("root").check hg } def clone_repository( source: String, root: Path, options: String = "", ssh: Option[SSH.Session] = None): Repository = { val hg = new Repository(root, ssh) hg.command("clone", File.bash_string(source) + " " + File.bash_path(hg.root), options).check hg } def setup_repository(source: String, root: Path, ssh: Option[SSH.Session] = None): Repository = ssh match { case None => if (root.is_dir) repository(root) else clone_repository(source, root) case Some(session) => using(session.sftp())(sftp => - if (sftp.is_dir(sftp.path(root))) repository(root, ssh = ssh) + if (sftp.is_dir(root)) repository(root, ssh = ssh) else clone_repository(source, root, ssh = ssh)) } class Repository private[Mercurial](root_path: Path, ssh: Option[SSH.Session]) { hg => val root = ssh match { case None => root_path.expand case Some(session) => using(session.sftp())(sftp => root_path.expand_env(sftp.settings)) } override def toString: String = ssh match { case None => root.implode case Some(session) => session.toString + ":" + root.implode } def command(name: String, args: String = "", options: String = ""): Process_Result = { val cmdline = "\"${HG:-hg}\"" + (if (name == "clone") "" else " --repository " + File.bash_path(root)) + " --noninteractive " + name + " " + options + " " + args ssh match { case None => Isabelle_System.bash(cmdline) case Some(session) => session.execute(cmdline) } } def heads(template: String = "{node|short}\n", options: String = ""): List[String] = hg.command("heads", opt_template(template), options).check.out_lines def identify(rev: String = "tip", options: String = ""): String = hg.command("id", opt_rev(rev), options).check.out_lines.headOption getOrElse "" def id(rev: String = "tip"): String = identify(rev, options = "-i") def manifest(rev: String = "", options: String = ""): List[String] = hg.command("manifest", opt_rev(rev), options).check.out_lines def log(rev: String = "", template: String = "", options: String = ""): String = hg.command("log", opt_rev(rev) + opt_template(template), options).check.out def pull(remote: String = "", rev: String = "", options: String = ""): Unit = hg.command("pull", opt_rev(rev) + optional(remote), options).check def update( rev: String = "", clean: Boolean = false, check: Boolean = false, options: String = "") { hg.command("update", opt_rev(rev) + opt_flag("--clean", clean) + opt_flag("--check", check), options).check } } } diff --git a/src/Pure/General/ssh.scala b/src/Pure/General/ssh.scala --- a/src/Pure/General/ssh.scala +++ b/src/Pure/General/ssh.scala @@ -1,348 +1,346 @@ /* Title: Pure/General/ssh.scala Author: Makarius SSH client based on JSch (see also http://www.jcraft.com/jsch/examples). */ package isabelle import java.io.{InputStream, OutputStream, ByteArrayOutputStream} import scala.collection.{mutable, JavaConversions} import com.jcraft.jsch.{JSch, Logger => JSch_Logger, Session => JSch_Session, SftpException, OpenSSHConfig, UserInfo, Channel => JSch_Channel, ChannelExec, ChannelSftp, SftpATTRS} object SSH { /* target machine: user@host syntax */ object Target { val Pattern = "^([^@]+)@(.+)$".r def parse(s: String): (String, String) = s match { case Pattern(user, host) => (user, host) case _ => ("", s) } def unapplySeq(s: String): Option[List[String]] = parse(s) match { case (_, "") => None case (user, host) => Some(List(user, host)) } } val default_port = 22 /* init */ def init(options: Options): SSH = { val config_dir = Path.explode(options.string("ssh_config_dir")) if (!config_dir.is_dir) error("Bad ssh config directory: " + config_dir) val jsch = new JSch val config_file = Path.explode(options.string("ssh_config_file")) if (config_file.is_file) jsch.setConfigRepository(OpenSSHConfig.parseFile(File.platform_path(config_file))) val known_hosts = config_dir + Path.explode("known_hosts") if (!known_hosts.is_file) known_hosts.file.createNewFile jsch.setKnownHosts(File.platform_path(known_hosts)) val identity_files = Library.space_explode(':', options.string("ssh_identity_files")).map(Path.explode(_)) for (identity_file <- identity_files if identity_file.is_file) jsch.addIdentity(File.platform_path(identity_file)) new SSH(options, jsch) } def connect_timeout(options: Options): Int = options.seconds("ssh_connect_timeout").ms.toInt /* logging */ def logging(verbose: Boolean = true, debug: Boolean = false) { JSch.setLogger(if (verbose) new Logger(debug) else null) } private class Logger(debug: Boolean) extends JSch_Logger { def isEnabled(level: Int): Boolean = level != JSch_Logger.DEBUG || debug def log(level: Int, msg: String) { level match { case JSch_Logger.ERROR | JSch_Logger.FATAL => Output.error_message(msg) case JSch_Logger.WARN => Output.warning(msg) case _ => Output.writeln(msg) } } } /* user info */ object No_User_Info extends UserInfo { def getPassphrase: String = null def getPassword: String = null def promptPassword(msg: String): Boolean = false def promptPassphrase(msg: String): Boolean = false def promptYesNo(msg: String): Boolean = false def showMessage(msg: String): Unit = Output.writeln(msg) } /* channel */ class Channel[C <: JSch_Channel] private[SSH]( val session: Session, val kind: String, val channel: C) { override def toString: String = kind + " " + session.toString def close() { channel.disconnect } } /* Sftp channel */ type Attrs = SftpATTRS - sealed case class Dir_Entry(name: String, attrs: Attrs) + sealed case class Dir_Entry(name: Path, attrs: Attrs) { def is_file: Boolean = attrs.isReg def is_dir: Boolean = attrs.isDir } class Sftp private[SSH](session: Session, kind: String, channel: ChannelSftp) extends Channel[ChannelSftp](session, kind, channel) { channel.connect(connect_timeout(session.options)) val settings: Map[String, String] = { val home = channel.getHome Map("HOME" -> home, "USER_HOME" -> home) } - def path(p: Path): String = p.expand_env(settings).implode + def expand_path(path: Path): Path = path.expand_env(settings) + def remote_path(path: Path): String = expand_path(path).implode - def chmod(permissions: Int, remote_path: String) { channel.chmod(permissions, remote_path) } - def mv(remote_path1: String, remote_path2: String): Unit = - channel.rename(remote_path1, remote_path2) - def rm(remote_path: String) { channel.rm(remote_path) } - def mkdir(remote_path: String) { channel.mkdir(remote_path) } - def rmdir(remote_path: String) { channel.rmdir(remote_path) } + def chmod(permissions: Int, path: Path): Unit = channel.chmod(permissions, remote_path(path)) + def mv(path1: Path, path2: Path): Unit = channel.rename(remote_path(path1), remote_path(path2)) + def rm(path: Path): Unit = channel.rm(remote_path(path)) + def mkdir(path: Path): Unit = channel.mkdir(remote_path(path)) + def rmdir(path: Path): Unit = channel.rmdir(remote_path(path)) - def stat(remote_path: String): Option[Dir_Entry] = - try { Some(Dir_Entry(remote_path, channel.stat(remote_path))) } + def stat(path: Path): Option[Dir_Entry] = + try { Some(Dir_Entry(expand_path(path), channel.stat(remote_path(path)))) } catch { case _: SftpException => None } - def is_file(remote_path: String): Boolean = stat(remote_path).map(_.is_file) getOrElse false - def is_dir(remote_path: String): Boolean = stat(remote_path).map(_.is_dir) getOrElse false + def is_file(path: Path): Boolean = stat(path).map(_.is_file) getOrElse false + def is_dir(path: Path): Boolean = stat(path).map(_.is_dir) getOrElse false - def read_dir(remote_path: String): List[Dir_Entry] = + def read_dir(path: Path): List[Dir_Entry] = { - val dir = channel.ls(remote_path) + val dir = channel.ls(remote_path(path)) (for { i <- (0 until dir.size).iterator a = dir.get(i).asInstanceOf[AnyRef] name = Untyped.get[String](a, "filename") attrs = Untyped.get[Attrs](a, "attrs") if name != "." && name != ".." - } yield Dir_Entry(name, attrs)).toList + } yield Dir_Entry(Path.basic(name), attrs)).toList } - def find_files(remote_path: String, pred: Dir_Entry => Boolean = _ => true): List[Dir_Entry] = + def find_files(root: Path, pred: Dir_Entry => Boolean = _ => true): List[Dir_Entry] = { - def find(dir: String): List[Dir_Entry] = + def find(dir: Path): List[Dir_Entry] = read_dir(dir).flatMap(entry => { - val file = dir + "/" + entry.name + val file = dir + entry.name if (entry.is_dir) find(file) else if (pred(entry)) List(entry.copy(name = file)) else Nil }) - find(remote_path) + find(root) } - def open_input(remote_path: String): InputStream = channel.get(remote_path) - def open_output(remote_path: String): OutputStream = channel.put(remote_path) + def open_input(path: Path): InputStream = channel.get(remote_path(path)) + def open_output(path: Path): OutputStream = channel.put(remote_path(path)) - def read_file(remote_path: String, local_path: Path): Unit = - channel.get(remote_path, File.platform_path(local_path)) - def read_bytes(remote_path: String): Bytes = - using(open_input(remote_path))(Bytes.read_stream(_)) - def read(remote_path: String): String = - using(open_input(remote_path))(File.read_stream(_)) + def read_file(path: Path, local_path: Path): Unit = + channel.get(remote_path(path), File.platform_path(local_path)) + def read_bytes(path: Path): Bytes = using(open_input(path))(Bytes.read_stream(_)) + def read(path: Path): String = using(open_input(path))(File.read_stream(_)) - def write_file(remote_path: String, local_path: Path): Unit = - channel.put(File.platform_path(local_path), remote_path) - def write_bytes(remote_path: String, bytes: Bytes): Unit = - using(open_output(remote_path))(bytes.write_stream(_)) - def write(remote_path: String, text: String): Unit = - using(open_output(remote_path))(stream => Bytes(text).write_stream(stream)) + def write_file(path: Path, local_path: Path): Unit = + channel.put(File.platform_path(local_path), remote_path(path)) + def write_bytes(path: Path, bytes: Bytes): Unit = + using(open_output(path))(bytes.write_stream(_)) + def write(path: Path, text: String): Unit = + using(open_output(path))(stream => Bytes(text).write_stream(stream)) } /* exec channel */ private val exec_wait_delay = Time.seconds(0.3) class Exec private[SSH](session: Session, kind: String, channel: ChannelExec) extends Channel[ChannelExec](session, kind, channel) { def kill(signal: String) { channel.sendSignal(signal) } val exit_status: Future[Int] = Future.thread("ssh_wait") { while (!channel.isClosed) Thread.sleep(exec_wait_delay.ms) channel.getExitStatus } val stdin: OutputStream = channel.getOutputStream val stdout: InputStream = channel.getInputStream val stderr: InputStream = channel.getErrStream // connect after preparing streams channel.connect(connect_timeout(session.options)) def result( progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), strict: Boolean = true): Process_Result = { stdin.close def read_lines(stream: InputStream, progress: String => Unit): List[String] = { val result = new mutable.ListBuffer[String] val line_buffer = new ByteArrayOutputStream(100) def line_flush() { val line = line_buffer.toString(UTF8.charset_name) progress(line) result += line line_buffer.reset } var c = 0 var finished = false while (!finished) { while ({ c = stream.read; c != -1 && c != 10 }) line_buffer.write(c) if (c == 10) line_flush() else if (channel.isClosed) { if (line_buffer.size > 0) line_flush() finished = true } else Thread.sleep(exec_wait_delay.ms) } result.toList } val out_lines = Future.thread("ssh_stdout") { read_lines(stdout, progress_stdout) } val err_lines = Future.thread("ssh_stderr") { read_lines(stderr, progress_stderr) } def terminate() { close out_lines.join err_lines.join exit_status.join } val rc = try { exit_status.join } catch { case Exn.Interrupt() => terminate(); Exn.Interrupt.return_code } close if (strict && rc == Exn.Interrupt.return_code) throw Exn.Interrupt() Process_Result(rc, out_lines.join, err_lines.join) } } /* session */ class Session private[SSH](val options: Options, val session: JSch_Session) { def update_options(new_options: Options): Session = new Session(new_options, session) override def toString: String = (if (session.getUserName == null) "" else session.getUserName + "@") + (if (session.getHost == null) "" else session.getHost) + (if (session.getPort == default_port) "" else ":" + session.getPort) + (if (session.isConnected) "" else " (disconnected)") def close() { session.disconnect } def sftp(): Sftp = { val kind = "sftp" val channel = session.openChannel(kind).asInstanceOf[ChannelSftp] new Sftp(this, kind, channel) } def exec(command: String): Exec = { val kind = "exec" val channel = session.openChannel(kind).asInstanceOf[ChannelExec] channel.setCommand("export USER_HOME=\"$HOME\"\n" + command) new Exec(this, kind, channel) } def execute(command: String, progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), strict: Boolean = true): Process_Result = exec(command).result(progress_stdout, progress_stderr, strict) /* tmp dirs */ def rm_tree(remote_dir: String): Unit = execute("rm -r -f " + File.bash_string(remote_dir)).check def tmp_dir(): String = execute("mktemp -d -t tmp.XXXXXXXXXX").check.out - def with_tmp_dir[A](body: String => A): A = + def with_tmp_dir[A](body: Path => A): A = { val remote_dir = tmp_dir() - try { body(remote_dir) } finally { rm_tree(remote_dir) } + try { body(Path.explode(remote_dir)) } finally { rm_tree(remote_dir) } } } } class SSH private(val options: Options, val jsch: JSch) { def update_options(new_options: Options): SSH = new SSH(new_options, jsch) def open_session(host: String, user: String = "", port: Int = SSH.default_port): SSH.Session = { val session = jsch.getSession(if (user == "") null else user, host, port) session.setUserInfo(SSH.No_User_Info) session.setConfig("MaxAuthTries", "3") if (options.bool("ssh_compression")) { session.setConfig("compression.s2c", "zlib@openssh.com,zlib,none") session.setConfig("compression.c2s", "zlib@openssh.com,zlib,none") session.setConfig("compression_level", "9") } session.connect(SSH.connect_timeout(options)) new SSH.Session(options, session) } }