diff -r 1722384ffd4a -r 9e668ae81f97 src/Doc/System/Server.thy --- a/src/Doc/System/Server.thy Thu Mar 22 15:11:14 2018 +0100 +++ b/src/Doc/System/Server.thy Thu Mar 22 16:20:53 2018 +0100 @@ -632,7 +632,7 @@ \quad~~\preferences?: string,\ & \<^bold>\default:\ server preferences \\ \quad~~\options?: [string],\ \\ \quad~~\dirs?: [string],\ \\ - \quad~~\all_known?: bool,\ \\ + \quad~~\include_sessions: [string],\ \\ \quad~~\system_mode?: bool,\ \\ \quad~~\verbose?: bool}\ \\[2ex] \end{tabular} @@ -694,13 +694,12 @@ sessions; see also option \<^verbatim>\-d\ in @{tool build}. \<^medskip> - The \all_known\ field set to \<^verbatim>\true\ includes all sessions from reachable - ROOT entries into the name space of theories. This is relevant for proper - session-qualified names, instead of referring to theory file names. The - option enables the \<^verbatim>\use_theories\ command - (\secref{sec:command-use-theories}) to refer to arbitrary theories from - other sessions, but considerable time is required to explore all accessible - session directories and theory dependencies. + The \include_sessions\ field specifies sessions whose theories should be + included in the overall name space of session-qualified theory names. This + corresponds to a \<^bold>\sessions\ specification in ROOT files + (\secref{sec:session-root}). It enables the \<^verbatim>\use_theories\ command + (\secref{sec:command-use-theories}) to refer to sources from other sessions + in a robust manner, instead of relying on directory locations. \<^medskip> The \system_mode\ field set to \<^verbatim>\true\ stores resulting session images and @@ -989,24 +988,20 @@ @{verbatim [display] \use_theories {"session_id": ..., "theories": ["~~/src/HOL/Isar_Examples/Drinker"]}\} \<^medskip> - Process some example theories that import other theories via - session-qualified theory names: - - @{verbatim [display] \session_start {"session": "HOL", "all_known": true} -use_theories {"session_id": ..., "theories": ["~~/src/HOL/Unix/Unix"]} -session_stop {"session_id": ...}\} - - The option \all_known\ has been used here to the full name space of - session-qualified theory names accessible in this session. This is also the - default in Isabelle/jEdit, although it requires significant startup time. - - \<^medskip> Process some example theories in the context of their (single) parent session: @{verbatim [display] \session_start {"session": "HOL-Library"} use_theories {"session_id": ..., "theories": ["~~/src/HOL/Unix/Unix"]} session_stop {"session_id": ...}\} + + \<^medskip> + Process some example theories that import other theories via + session-qualified theory names: + + @{verbatim [display] \session_start {"session": "HOL", "include_sessions": ["HOL-Unix"]} +use_theories {"session_id": ..., "theories": ["HOL-Unix.Unix"]} +session_stop {"session_id": ...}\} \ end