src/Doc/System/Environment.thy
changeset 73671 7404f2e1d092
parent 73150 c9a836122739
child 73741 941915a3b811
--- a/src/Doc/System/Environment.thy	Tue May 11 14:04:36 2021 +0200
+++ b/src/Doc/System/Environment.thy	Tue May 11 16:30:24 2021 +0200
@@ -110,39 +110,28 @@
   defaults may be overridden by a private \<^verbatim>\<open>$ISABELLE_HOME_USER/etc/settings\<close>.
 
   \<^descr>[@{setting_def ISABELLE_PLATFORM_FAMILY}\<open>\<^sup>*\<close>] is automatically set to the
-  general platform family: \<^verbatim>\<open>linux\<close>, \<^verbatim>\<open>macos\<close>, \<^verbatim>\<open>windows\<close>. Note that
+  general platform family (\<^verbatim>\<open>linux\<close>, \<^verbatim>\<open>macos\<close>, \<^verbatim>\<open>windows\<close>). Note that
   platform-dependent tools usually need to refer to the more specific
   identification according to @{setting ISABELLE_PLATFORM64}, @{setting
-  ISABELLE_PLATFORM32}, @{setting ISABELLE_WINDOWS_PLATFORM64}, @{setting
-  ISABELLE_WINDOWS_PLATFORM32}.
-
-  \<^descr>[@{setting_def ISABELLE_PLATFORM64}\<open>\<^sup>*\<close>, @{setting_def
-  ISABELLE_PLATFORM32}\<open>\<^sup>*\<close>] indicate the standard Posix platform: \<^verbatim>\<open>x86_64\<close>
-  for 64 bit and \<^verbatim>\<open>x86\<close> for 32 bit, together with a symbolic name for the
-  operating system (\<^verbatim>\<open>linux\<close>, \<^verbatim>\<open>darwin\<close>, \<^verbatim>\<open>cygwin\<close>). All platforms support 64
-  bit executables, some platforms also support 32 bit executables.
+  ISABELLE_WINDOWS_PLATFORM64}, @{setting ISABELLE_APPLE_PLATFORM64}.
 
-  In GNU bash scripts, it is possible to use the following expressions (with
-  quotes) to specify a preference of 64 bit over 32 bit:
-
-  @{verbatim [display] \<open>"${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}"\<close>}
-
-  In contrast, the subsequent expression prefers the old 32 bit variant (which
-  is only relevant for unusual applications):
-
-  @{verbatim [display] \<open>"${ISABELLE_PLATFORM32:-$ISABELLE_PLATFORM64}"\<close>}
+  \<^descr>[@{setting_def ISABELLE_PLATFORM64}\<open>\<^sup>*\<close>] indicates the standard Posix
+  platform (\<^verbatim>\<open>x86_64\<close>, \<^verbatim>\<open>arm64\<close>), together with a symbolic name for the
+  operating system (\<^verbatim>\<open>linux\<close>, \<^verbatim>\<open>darwin\<close>, \<^verbatim>\<open>cygwin\<close>).
 
   \<^descr>[@{setting_def ISABELLE_WINDOWS_PLATFORM64}\<open>\<^sup>*\<close>, @{setting_def
-  ISABELLE_WINDOWS_PLATFORM32}\<open>\<^sup>*\<close>] indicate the native Windows platform.
-  These settings are analogous (but independent) of those for the standard
-  Posix subsystem: @{setting ISABELLE_PLATFORM64}, @{setting
-  ISABELLE_PLATFORM32}.
+  ISABELLE_WINDOWS_PLATFORM32}\<open>\<^sup>*\<close>] indicate the native Windows platform: both
+  64\,bit and 32\,bit executables are supported here.
 
   In GNU bash scripts, a preference for native Windows platform variants may
   be specified like this (first 64 bit, second 32 bit):
 
   @{verbatim [display] \<open>"${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:-
-  ${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}}"\<close>}
+  $ISABELLE_PLATFORM64}}"\<close>}
+
+  \<^descr>[@{setting_def ISABELLE_APPLE_PLATFORM64}\<open>\<^sup>*\<close>] indicates the native Apple
+  Silicon platform (\<^verbatim>\<open>arm64-darwin\<close> if available), instead of Intel emulation
+  via Rosetta (\<^verbatim>\<open>ISABELLE_PLATFORM64=x86_64-darwin\<close>).
 
   \<^descr>[@{setting ISABELLE_TOOL}\<open>\<^sup>*\<close>] is automatically set to the full path name
   of the @{executable isabelle} executable.