src/Doc/System/Environment.thy
changeset 68003 9b89d831dc80
parent 67904 465f43a9f780
child 68219 c0341c0080e2
--- a/src/Doc/System/Environment.thy	Thu Apr 19 12:02:59 2018 +0200
+++ b/src/Doc/System/Environment.thy	Thu Apr 19 12:34:52 2018 +0200
@@ -118,38 +118,37 @@
   \<^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
   platform-dependent tools usually need to refer to the more specific
-  identification according to @{setting ISABELLE_PLATFORM} etc.
+  identification according to @{setting ISABELLE_PLATFORM64}, @{setting
+  ISABELLE_PLATFORM32}, @{setting ISABELLE_WINDOWS_PLATFORM64}, @{setting
+  ISABELLE_WINDOWS_PLATFORM32}.
 
-  \<^descr>[@{setting_def ISABELLE_PLATFORM32}\<open>\<^sup>*\<close>, @{setting_def
-  ISABELLE_PLATFORM64}\<open>\<^sup>*\<close>, @{setting_def ISABELLE_PLATFORM}\<open>\<^sup>*\<close>] indicate the
-  standard Posix platform: \<^verbatim>\<open>x86\<close> for 32 bit and \<^verbatim>\<open>x86_64\<close> for 64 bit,
-  together with a symbolic name for the operating system (\<^verbatim>\<open>linux\<close>, \<^verbatim>\<open>darwin\<close>,
-  \<^verbatim>\<open>cygwin\<close>). Some platforms support both 32 bit and 64 bit executables, but
-  this depends on various side-conditions.
+  \<^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.
 
-  In GNU bash scripts, it is possible to use the following expressions
-  (including the quotes) to specify a preference of 64 bit over 32 bit:
+  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 32 bit variant; this is
-  how @{setting ISABELLE_PLATFORM} is defined:
+  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_WINDOWS_PLATFORM32}\<open>\<^sup>*\<close>, @{setting_def
-  ISABELLE_WINDOWS_PLATFORM64}\<open>\<^sup>*\<close>,] @{setting_def
-  ISABELLE_WINDOWS_PLATFORM}\<open>\<^sup>*\<close> indicate the native Windows platform. These
-  settings are analogous (but independent) of those for the standard Posix
-  subsystem: @{setting ISABELLE_PLATFORM32}, @{setting ISABELLE_PLATFORM64},
-  @{setting ISABELLE_PLATFORM}.
+  \<^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}.
 
   In GNU bash scripts, a preference for native Windows platform variants may
-  be specified like this:
+  be specified like this (first 64 bit, second 32 bit):
 
-  @{verbatim [display] \<open>"${ISABELLE_WINDOWS_PLATFORM:-$ISABELLE_PLATFORM}"\<close>}
-
-  @{verbatim [display] \<open>"${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:-${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}}"\<close>}
+  @{verbatim [display] \<open>"${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:-
+  ${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM32}}}"\<close>}
 
   \<^descr>[@{setting ISABELLE_TOOL}\<open>\<^sup>*\<close>] is automatically set to the full path name
   of the @{executable isabelle} executable.