src/Pure/System/isabelle_platform.scala
changeset 82612 2757f73abda7
parent 82611 cf64152e9f51
child 82717 8d42bf3b821d
--- a/src/Pure/System/isabelle_platform.scala	Tue May 06 17:03:56 2025 +0200
+++ b/src/Pure/System/isabelle_platform.scala	Wed May 07 10:45:09 2025 +0200
@@ -28,6 +28,45 @@
       result.out_lines.map(line =>
         Properties.Eq.unapply(line) getOrElse error("Bad output: " + quote(result.out))))
   }
+
+
+  /* system context for progress/process */
+
+  object Context {
+    def apply(
+      isabelle_platform: Isabelle_Platform = local,
+      mingw: MinGW = MinGW.none,
+      progress: Progress = new Progress
+    ): Context = {
+      val context_platform = isabelle_platform
+      val context_mingw = mingw
+      val context_progress = progress
+      new Context {
+        override def isabelle_platform: Isabelle_Platform = context_platform
+        override def mingw: MinGW = context_mingw
+        override def progress: Progress = context_progress
+      }
+    }
+  }
+
+  trait Context {
+    def isabelle_platform: Isabelle_Platform
+    def mingw: MinGW
+    def progress: Progress
+
+    def standard_path(path: Path): String =
+      mingw.standard_path(File.standard_path(path))
+
+    def execute(cwd: Path, script_lines: String*): Process_Result = {
+      val script = cat_lines("set -e" :: script_lines.toList)
+      val script1 =
+        if (isabelle_platform.is_arm && isabelle_platform.is_macos) {
+          "arch -arch arm64 bash -c " + Bash.string(script)
+        }
+        else mingw.bash_script(script)
+      progress.bash(script1, cwd = cwd, echo = progress.verbose).check
+    }
+  }
 }
 
 class Isabelle_Platform private(val settings: List[(String, String)]) {