| author | desharna | 
| Fri, 29 Sep 2023 15:36:12 +0200 | |
| changeset 78733 | 70e1c0167ae2 | 
| parent 78610 | fd1fec53665b | 
| child 79536 | d39976404aa7 | 
| permissions | -rw-r--r-- | 
| 31825 | 1  | 
/* Title: Pure/System/platform.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
| 69112 | 4  | 
System platform identification.  | 
| 31825 | 5  | 
*/  | 
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
| 31828 | 9  | 
|
| 75393 | 10  | 
object Platform {
 | 
| 69112 | 11  | 
/* platform family */  | 
| 31828 | 12  | 
|
| 73906 | 13  | 
val is_windows: Boolean = isabelle.setup.Environment.is_windows()  | 
| 71383 | 14  | 
  val is_linux: Boolean = System.getProperty("os.name", "") == "Linux"
 | 
15  | 
  val is_macos: Boolean = System.getProperty("os.name", "") == "Mac OS X"
 | 
|
| 
73602
 
37243ad3ecb6
fast approximation of test for process group (NB: initial process might already be terminated, while background processes are still running);
 
wenzelm 
parents: 
73193 
diff
changeset
 | 
16  | 
val is_unix: Boolean = is_linux || is_macos  | 
| 31825 | 17  | 
|
| 73637 | 18  | 
  def is_arm: Boolean = cpu_arch.startsWith("arm")
 | 
19  | 
||
| 78610 | 20  | 
def family: Family =  | 
| 73637 | 21  | 
if (is_linux && is_arm) Family.linux_arm  | 
22  | 
else if (is_linux) Family.linux  | 
|
| 69410 | 23  | 
else if (is_macos) Family.macos  | 
24  | 
else if (is_windows) Family.windows  | 
|
25  | 
    else error("Failed to determine current platform family")
 | 
|
26  | 
||
| 78610 | 27  | 
  object Family {
 | 
28  | 
val list0: List[Family] = List(Family.linux, Family.windows, Family.macos)  | 
|
29  | 
val list: List[Family] = List(Family.linux, Family.linux_arm, Family.windows, Family.macos)  | 
|
| 69410 | 30  | 
|
| 78610 | 31  | 
def unapply(name: String): Option[Family] =  | 
32  | 
      try { Some(Family.valueOf(name)) }
 | 
|
33  | 
      catch { case _: IllegalArgumentException => None }
 | 
|
| 69410 | 34  | 
|
| 78610 | 35  | 
def parse(name: String): Family =  | 
| 69410 | 36  | 
      unapply(name) getOrElse error("Bad platform family: " + quote(name))
 | 
| 73637 | 37  | 
|
| 78610 | 38  | 
val standard: Family => String =  | 
39  | 
      {
 | 
|
40  | 
case Family.linux_arm => "arm64-linux"  | 
|
41  | 
case Family.linux => "x86_64-linux"  | 
|
42  | 
case Family.macos => "x86_64-darwin"  | 
|
43  | 
case Family.windows => "x86_64-cygwin"  | 
|
44  | 
}  | 
|
| 75083 | 45  | 
|
| 78610 | 46  | 
val native: Family => String =  | 
47  | 
      {
 | 
|
48  | 
case Family.macos => "arm64-darwin"  | 
|
49  | 
case Family.windows => "x86_64-windows"  | 
|
50  | 
case platform => standard(platform)  | 
|
51  | 
}  | 
|
| 69410 | 52  | 
}  | 
53  | 
||
| 78610 | 54  | 
  enum Family { case linux_arm, linux, macos, windows }
 | 
55  | 
||
| 31828 | 56  | 
|
| 69112 | 57  | 
/* platform identifiers */  | 
| 31828 | 58  | 
|
| 69112 | 59  | 
private val X86_64 = """amd64|x86_64""".r  | 
| 72344 | 60  | 
private val Arm64 = """arm64|aarch64""".r  | 
| 31825 | 61  | 
|
| 69726 | 62  | 
def cpu_arch: String =  | 
63  | 
    System.getProperty("os.arch", "") match {
 | 
|
64  | 
case X86_64() => "x86_64"  | 
|
| 72344 | 65  | 
case Arm64() => "arm64"  | 
| 69726 | 66  | 
      case _ => error("Failed to determine CPU architecture")
 | 
67  | 
}  | 
|
68  | 
||
69  | 
def os_name: String =  | 
|
70  | 
    family match {
 | 
|
| 
73639
 
e1432539df35
proper jvm_platform, notably for org.sqlite.lib.path;
 
wenzelm 
parents: 
73638 
diff
changeset
 | 
71  | 
case Family.linux_arm => "linux"  | 
| 69726 | 72  | 
case Family.macos => "darwin"  | 
73  | 
case _ => family.toString  | 
|
74  | 
}  | 
|
75  | 
||
76  | 
lazy val jvm_platform: String = cpu_arch + "-" + os_name  | 
|
| 31828 | 77  | 
|
78  | 
||
| 61001 | 79  | 
/* JVM version */  | 
80  | 
||
| 69112 | 81  | 
private val Version = """1\.(\d+)\.0_(\d+)""".r  | 
| 71383 | 82  | 
lazy val jvm_version: String =  | 
| 61001 | 83  | 
    System.getProperty("java.version") match {
 | 
84  | 
case Version(a, b) => a + "u" + b  | 
|
85  | 
case a => a  | 
|
86  | 
}  | 
|
87  | 
||
88  | 
||
| 41381 | 89  | 
/* JVM name */  | 
90  | 
||
| 53582 | 91  | 
  val jvm_name: String = System.getProperty("java.vm.name", "")
 | 
| 31825 | 92  | 
}  |