src/Doc/System/Environment.thy
changeset 68219 c0341c0080e2
parent 68003 9b89d831dc80
child 68275 b5d0318757f0
equal deleted inserted replaced
68218:92050155593e 68219:c0341c0080e2
    71   variables. Such advanced features should be added only with great care,
    71   variables. Such advanced features should be added only with great care,
    72   though. In particular, external environment references should be kept at a
    72   though. In particular, external environment references should be kept at a
    73   minimum.
    73   minimum.
    74 
    74 
    75   \<^medskip>
    75   \<^medskip>
    76   A few variables are somewhat special:
    76   A few variables are somewhat special, e.g. @{setting_def ISABELLE_TOOL} is
    77 
    77   set automatically to the absolute path name of the @{executable isabelle}
    78     \<^item> @{setting_def ISABELLE_TOOL} is set automatically to the absolute path
    78   executables.
    79     name of the @{executable isabelle} executables.
       
    80 
       
    81     \<^item> @{setting_ref ISABELLE_OUTPUT} will have the identifiers of the Isabelle
       
    82     distribution (cf.\ @{setting ISABELLE_IDENTIFIER}) and the ML system (cf.\
       
    83     @{setting ML_IDENTIFIER}) appended automatically to its value.
       
    84 
    79 
    85   \<^medskip>
    80   \<^medskip>
    86   Note that the settings environment may be inspected with the @{tool getenv}
    81   Note that the settings environment may be inspected with the @{tool getenv}
    87   tool. This might help to figure out the effect of complex settings scripts.
    82   tool. This might help to figure out the effect of complex settings scripts.
    88 \<close>
    83 \<close>
   177 
   172 
   178   \<^descr>[@{setting_def ISABELLE_JAVA_PLATFORM}] identifies the hardware and
   173   \<^descr>[@{setting_def ISABELLE_JAVA_PLATFORM}] identifies the hardware and
   179   operating system platform for the Java installation of Isabelle. That is
   174   operating system platform for the Java installation of Isabelle. That is
   180   always the (native) 64 bit variant: \<^verbatim>\<open>x86_64-linux\<close>, \<^verbatim>\<open>x86_64-darwin\<close>,
   175   always the (native) 64 bit variant: \<^verbatim>\<open>x86_64-linux\<close>, \<^verbatim>\<open>x86_64-darwin\<close>,
   181   \<^verbatim>\<open>x86_64-windows\<close>.
   176   \<^verbatim>\<open>x86_64-windows\<close>.
   182 
       
   183   \<^descr>[@{setting_def ISABELLE_PATH}] is a list of directories (separated by
       
   184   colons) where Isabelle logic images may reside. When looking up heaps files,
       
   185   the value of @{setting ML_IDENTIFIER} is appended to each component
       
   186   internally.
       
   187 
       
   188   \<^descr>[@{setting_def ISABELLE_OUTPUT}\<open>\<^sup>*\<close>] is a directory where output heap files
       
   189   should be stored by default. The ML system and Isabelle version identifier
       
   190   is appended here, too.
       
   191 
   177 
   192   \<^descr>[@{setting_def ISABELLE_BROWSER_INFO}] is the directory where theory
   178   \<^descr>[@{setting_def ISABELLE_BROWSER_INFO}] is the directory where theory
   193   browser information is stored as HTML and PDF (see also \secref{sec:info}).
   179   browser information is stored as HTML and PDF (see also \secref{sec:info}).
   194   The default value is @{path "$ISABELLE_HOME_USER/browser_info"}.
   180   The default value is @{path "$ISABELLE_HOME_USER/browser_info"}.
   195 
   181