src/Doc/System/Sessions.thy
changeset 66745 e7ac579b883c
parent 66737 2edc0c42c883
child 66748 3efac90a11a7
--- a/src/Doc/System/Sessions.thy	Sun Oct 01 20:50:26 2017 +0200
+++ b/src/Doc/System/Sessions.thy	Mon Oct 02 11:43:17 2017 +0200
@@ -284,6 +284,7 @@
     -D DIR       include session directory and select its sessions
     -N           cyclic shuffling of NUMA CPU nodes (performance tuning)
     -R           operate on requirements of selected sessions
+    -S           soft build: only observe changes of sources, not heap images
     -X NAME      exclude sessions from group NAME and all descendants
     -a           select all sessions
     -b           build heap images
@@ -350,6 +351,11 @@
   in the given directories.
 
   \<^medskip>
+  Option \<^verbatim>\<open>-S\<close> indicates a ``soft build'': the selection is restricted to
+  those sessions that have changed sources (according to actually imported
+  theories). The status of heap images is ignored.
+
+  \<^medskip>
   The build process depends on additional options
   (\secref{sec:system-options}) that are passed to the prover eventually. The
   settings variable @{setting_ref ISABELLE_BUILD_OPTIONS} allows to provide