src/Doc/System/Environment.thy
changeset 68523 ccacc84e0251
parent 68514 b20980997cd2
child 68541 12b4b3bc585d
--- a/src/Doc/System/Environment.thy	Wed Jun 27 11:16:43 2018 +0200
+++ b/src/Doc/System/Environment.thy	Wed Jun 27 20:31:22 2018 +0200
@@ -175,9 +175,17 @@
   always the (native) 64 bit variant: \<^verbatim>\<open>x86_64-linux\<close>, \<^verbatim>\<open>x86_64-darwin\<close>,
   \<^verbatim>\<open>x86_64-windows\<close>.
 
-  \<^descr>[@{setting_def ISABELLE_BROWSER_INFO}] is the directory where theory
-  browser information is stored as HTML and PDF (see also \secref{sec:info}).
-  The default value is @{path "$ISABELLE_HOME_USER/browser_info"}.
+  \<^descr>[@{setting_def ISABELLE_BROWSER_INFO}] is the directory where HTML and PDF
+  browser information is stored (see also \secref{sec:info}); its default is
+  @{path "$ISABELLE_HOME_USER/browser_info"}. For ``system build mode'' (see
+  \secref{sec:tool-build}), @{setting_def ISABELLE_BROWSER_INFO_SYSTEM} is
+  used instead; its default is @{path "$ISABELLE_HOME/browser_info"}.
+
+  \<^descr>[@{setting_def ISABELLE_HEAPS}] is the directory where session heap images,
+  log files, and build databases are stored; its default is @{path
+  "$ISABELLE_HOME_USER/heaps"}. For ``system build mode'' (see
+  \secref{sec:tool-build}), @{setting_def ISABELLE_HEAPS_SYSTEM} is used
+  instead; its default is @{path "$ISABELLE_HOME/heaps"}.
 
   \<^descr>[@{setting_def ISABELLE_LOGIC}] specifies the default logic to load if none
   is given explicitely by the user. The default value is \<^verbatim>\<open>HOL\<close>.