| author | Fabian Huch <huch@in.tum.de> | 
| Mon, 10 Jun 2024 17:08:47 +0200 | |
| changeset 80340 | 992bd899a027 | 
| parent 79724 | 54d0f6edfe3a | 
| permissions | -rw-r--r-- | 
| 67904 | 1 | (*:maxLineLen=78:*) | 
| 2 | ||
| 3 | theory Server | |
| 4 | imports Base | |
| 5 | begin | |
| 6 | ||
| 7 | chapter \<open>The Isabelle server\<close> | |
| 8 | ||
| 9 | text \<open> | |
| 10 | An Isabelle session requires at least two processes, which are both rather | |
| 11 | heavy: Isabelle/Scala for the system infrastructure and Isabelle/ML for the | |
| 12 | logic session (e.g.\ HOL). In principle, these processes can be invoked | |
| 13 |   directly on the command-line, e.g.\ via @{tool java}, @{tool scala}, @{tool
 | |
| 14 |   process}, @{tool console}, but this approach is inadequate for reactive
 | |
| 15 | applications that require quick responses from the prover. | |
| 16 | ||
| 17 | In contrast, the Isabelle server exposes Isabelle/Scala as a | |
| 18 | ``terminate-stay-resident'' application that manages multiple logic | |
| 75161 | 19 | \<^emph>\<open>sessions\<close> and concurrent tasks to use \<^emph>\<open>theories\<close>. This is analogous to | 
| 20 | \<^ML>\<open>Thy_Info.use_theories\<close> in Isabelle/ML, with proper support for | |
| 21 | concurrent invocations. | |
| 67904 | 22 | |
| 23 | The client/server arrangement via TCP sockets also opens possibilities for | |
| 24 | remote Isabelle services that are accessed by local applications, e.g.\ via | |
| 25 | an SSH tunnel. | |
| 26 | \<close> | |
| 27 | ||
| 28 | ||
| 29 | section \<open>Command-line tools\<close> | |
| 30 | ||
| 31 | subsection \<open>Server \label{sec:tool-server}\<close>
 | |
| 32 | ||
| 33 | text \<open> | |
| 34 |   The @{tool_def server} tool manages resident server processes:
 | |
| 35 |   @{verbatim [display]
 | |
| 36 | \<open>Usage: isabelle server [OPTIONS] | |
| 37 | ||
| 38 | Options are: | |
| 39 | -L FILE logging on FILE | |
| 40 | -c console interaction with specified server | |
| 41 | -l list servers (alternative operation) | |
| 42 | -n NAME explicit server name (default: isabelle) | |
| 43 | -p PORT explicit server port | |
| 44 | -s assume existing server, no implicit startup | |
| 45 | -x exit specified server (alternative operation) | |
| 46 | ||
| 47 | Manage resident Isabelle servers.\<close>} | |
| 48 | ||
| 49 | The main operation of \<^verbatim>\<open>isabelle server\<close> is to ensure that a named server is | |
| 50 | running, either by finding an already running process (according to the | |
| 69593 | 51 | central database file \<^path>\<open>$ISABELLE_HOME_USER/servers.db\<close>) or by | 
| 67904 | 52 | becoming itself a new server that accepts connections on a particular TCP | 
| 53 | socket. The server name and its address are printed as initial output line. | |
| 54 | If another server instance is already running, the current | |
| 55 | \<^verbatim>\<open>isabelle server\<close> process will terminate; otherwise, it keeps running as a | |
| 56 | new server process until an explicit \<^verbatim>\<open>shutdown\<close> command is received. | |
| 57 | Further details of the server socket protocol are explained in | |
| 58 |   \secref{sec:server-protocol}.
 | |
| 59 | ||
| 60 | Other server management operations are invoked via options \<^verbatim>\<open>-l\<close> and \<^verbatim>\<open>-x\<close> | |
| 61 | (see below). | |
| 62 | ||
| 63 | \<^medskip> | |
| 64 | Option \<^verbatim>\<open>-n\<close> specifies an alternative server name: at most one process for | |
| 65 | each name may run, but each server instance supports multiple connections | |
| 66 | and logic sessions. | |
| 67 | ||
| 68 | \<^medskip> | |
| 69 | Option \<^verbatim>\<open>-p\<close> specifies an explicit TCP port for the server socket (which is | |
| 70 | always on \<^verbatim>\<open>localhost\<close>): the default is to let the operating system assign a | |
| 71 | free port number. | |
| 72 | ||
| 73 | \<^medskip> | |
| 74 | Option \<^verbatim>\<open>-s\<close> strictly assumes that the specified server process is already | |
| 75 | running, skipping the optional server startup phase. | |
| 76 | ||
| 77 | \<^medskip> | |
| 78 | Option \<^verbatim>\<open>-c\<close> connects the console in/out channels after the initial check | |
| 67918 | 79 |   for a suitable server process. Also note that the @{tool client} tool
 | 
| 80 |   (\secref{sec:tool-client}) provides a command-line editor to interact with
 | |
| 81 | the server. | |
| 67904 | 82 | |
| 83 | \<^medskip> | |
| 84 | Option \<^verbatim>\<open>-L\<close> specifies a log file for exceptional output of internal server | |
| 85 | and session operations. | |
| 67918 | 86 | |
| 87 | \<^medskip> | |
| 88 | Operation \<^verbatim>\<open>-l\<close> lists all active server processes with their connection | |
| 89 | details. | |
| 90 | ||
| 91 | \<^medskip> | |
| 92 | Operation \<^verbatim>\<open>-x\<close> exits the specified server process by sending it a | |
| 93 | \<^verbatim>\<open>shutdown\<close> command. | |
| 67904 | 94 | \<close> | 
| 95 | ||
| 96 | ||
| 97 | subsection \<open>Client \label{sec:tool-client}\<close>
 | |
| 98 | ||
| 99 | text \<open> | |
| 67918 | 100 |   The @{tool_def client} tool provides console interaction for Isabelle
 | 
| 101 | servers: | |
| 67904 | 102 |   @{verbatim [display]
 | 
| 103 | \<open>Usage: isabelle client [OPTIONS] | |
| 104 | ||
| 105 | Options are: | |
| 106 | -n NAME explicit server name | |
| 107 | -p PORT explicit server port | |
| 108 | ||
| 109 | Console interaction for Isabelle server (with line-editor).\<close>} | |
| 110 | ||
| 67918 | 111 | This is a wrapper to \<^verbatim>\<open>isabelle server -s -c\<close> for interactive | 
| 112 |   experimentation, which uses @{setting ISABELLE_LINE_EDITOR} if available.
 | |
| 67904 | 113 | The server name is sufficient for identification, as the client can | 
| 114 | determine the connection details from the local database of active servers. | |
| 115 | ||
| 116 | \<^medskip> | |
| 117 |   Option \<^verbatim>\<open>-n\<close> specifies an explicit server name as in @{tool server}.
 | |
| 118 | ||
| 119 | \<^medskip> | |
| 120 |   Option \<^verbatim>\<open>-p\<close> specifies an explicit server port as in @{tool server}.
 | |
| 121 | \<close> | |
| 122 | ||
| 123 | ||
| 124 | subsection \<open>Examples\<close> | |
| 125 | ||
| 126 | text \<open> | |
| 127 | Ensure that a particular server instance is running in the background: | |
| 128 |   @{verbatim [display] \<open>isabelle server -n test &\<close>}
 | |
| 129 | ||
| 67918 | 130 | The first line of output presents the connection details:\<^footnote>\<open>This information | 
| 131 | may be used in other TCP clients, without access to Isabelle/Scala and the | |
| 132 | underlying database of running servers.\<close> | |
| 67904 | 133 |   @{verbatim [display] \<open>server "test" = 127.0.0.1:4711 (password "XYZ")\<close>}
 | 
| 134 | ||
| 135 | \<^medskip> | |
| 136 | List available server processes: | |
| 137 |   @{verbatim [display] \<open>isabelle server -l\<close>}
 | |
| 138 | ||
| 139 | \<^medskip> | |
| 140 | Connect the command-line client to the above test server: | |
| 141 |   @{verbatim [display] \<open>isabelle client -n test\<close>}
 | |
| 142 | ||
| 143 | Interaction now works on a line-by-line basis, with commands like \<^verbatim>\<open>help\<close> or | |
| 144 | \<^verbatim>\<open>echo\<close>. For example, some JSON values may be echoed like this: | |
| 145 | ||
| 146 |   @{verbatim [display]
 | |
| 147 | \<open>echo 42 | |
| 148 | echo [1, 2, 3] | |
| 149 | echo {"a": "text", "b": true, "c": 42}\<close>}
 | |
| 150 | ||
| 151 | Closing the connection (via CTRL-D) leaves the server running: it is | |
| 152 | possible to reconnect again, and have multiple connections at the same time. | |
| 153 | ||
| 154 | \<^medskip> | |
| 67918 | 155 | Exit the named server on the command-line: | 
| 67904 | 156 |   @{verbatim [display] \<open>isabelle server -n test -x\<close>}
 | 
| 157 | \<close> | |
| 158 | ||
| 159 | ||
| 160 | section \<open>Protocol messages \label{sec:server-protocol}\<close>
 | |
| 161 | ||
| 162 | text \<open> | |
| 163 | The Isabelle server listens on a regular TCP socket, using a line-oriented | |
| 67918 | 164 | protocol of structured messages. Input \<^emph>\<open>commands\<close> and output \<^emph>\<open>results\<close> | 
| 67904 | 165 | (via \<^verbatim>\<open>OK\<close> or \<^verbatim>\<open>ERROR\<close>) are strictly alternating on the toplevel, but | 
| 166 | commands may also return a \<^emph>\<open>task\<close> identifier to indicate an ongoing | |
| 167 | asynchronous process that is joined later (via \<^verbatim>\<open>FINISHED\<close> or \<^verbatim>\<open>FAILED\<close>). | |
| 168 | Asynchronous \<^verbatim>\<open>NOTE\<close> messages may occur at any time: they are independent of | |
| 67918 | 169 | the main command-result protocol. | 
| 67904 | 170 | |
| 171 | For example, the synchronous \<^verbatim>\<open>echo\<close> command immediately returns its | |
| 172 | argument as \<^verbatim>\<open>OK\<close> result. In contrast, the asynchronous \<^verbatim>\<open>session_build\<close> | |
| 173 |   command returns \<^verbatim>\<open>OK {"task":\<close>\<open>id\<close>\<^verbatim>\<open>}\<close> and continues in the background. It
 | |
| 174 |   will eventually produce \<^verbatim>\<open>FINISHED {"task":\<close>\<open>id\<close>\<^verbatim>\<open>,\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close> or
 | |
| 175 |   \<^verbatim>\<open>FAILED {"task":\<close>\<open>id\<close>\<^verbatim>\<open>,\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close> with the final result. Intermediately, it
 | |
| 176 |   may emit asynchronous messages of the form \<^verbatim>\<open>NOTE {"task":\<close>\<open>id\<close>\<^verbatim>\<open>,\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close>
 | |
| 177 | to inform about its progress. Due to the explicit task identifier, the | |
| 178 | client can show these messages in the proper context, e.g.\ a GUI window for | |
| 67918 | 179 | this particular session build job. | 
| 67904 | 180 | |
| 181 | \medskip Subsequently, the protocol message formats are described in further | |
| 182 | detail. | |
| 183 | \<close> | |
| 184 | ||
| 185 | ||
| 69464 | 186 | subsection \<open>Byte messages \label{sec:byte-messages}\<close>
 | 
| 67904 | 187 | |
| 188 | text \<open> | |
| 189 | The client-server connection is a raw byte-channel for bidirectional | |
| 190 | communication, but the Isabelle server always works with messages of a | |
| 191 | particular length. Messages are written as a single chunk that is flushed | |
| 192 | immediately. | |
| 193 | ||
| 67918 | 194 | Message boundaries are determined as follows: | 
| 67904 | 195 | |
| 196 | \<^item> A \<^emph>\<open>short message\<close> consists of a single line: it is a sequence of | |
| 197 | arbitrary bytes excluding CR (13) and LF (10), and terminated by CR-LF or | |
| 198 | just LF. | |
| 199 | ||
| 75161 | 200 | \<^item> A \<^emph>\<open>long message\<close> starts with a single line consisting of decimal | 
| 67918 | 201 | digits: these are interpreted as length of the subsequent block of | 
| 202 | arbitrary bytes. A final line-terminator (as above) may be included here, | |
| 67904 | 203 | but is not required. | 
| 204 | ||
| 205 | Messages in JSON format (see below) always fit on a single line, due to | |
| 206 | escaping of newline characters within string literals. This is convenient | |
| 207 | for interactive experimentation, but it can impact performance for very long | |
| 208 | messages. If the message byte-length is given on the preceding line, the | |
| 67918 | 209 | server can read the message more efficiently as a single block. | 
| 67904 | 210 | \<close> | 
| 211 | ||
| 212 | ||
| 213 | subsection \<open>Text messages\<close> | |
| 214 | ||
| 215 | text \<open> | |
| 216 | Messages are read and written as byte streams (with byte lengths), but the | |
| 217 | content is always interpreted as plain text in terms of the UTF-8 | |
| 218 | encoding.\<^footnote>\<open>See also the ``UTF-8 Everywhere Manifesto'' | |
| 68224 | 219 | \<^url>\<open>https://utf8everywhere.org\<close>.\<close> | 
| 67904 | 220 | |
| 221 | Note that line-endings and other formatting characters are invariant wrt. | |
| 222 | UTF-8 representation of text: thus implementations are free to determine the | |
| 223 | overall message structure before or after applying the text encoding. | |
| 224 | \<close> | |
| 225 | ||
| 226 | ||
| 67917 | 227 | subsection \<open>Input and output messages \label{sec:input-output-messages}\<close>
 | 
| 67904 | 228 | |
| 229 | text \<open> | |
| 75161 | 230 | The uniform format for server input and output messages is \<open>name argument\<close>, | 
| 231 | such that: | |
| 67904 | 232 | |
| 233 | \<^item> \<open>name\<close> is the longest prefix consisting of ASCII letters, digits, | |
| 67918 | 234 | ``\<^verbatim>\<open>_\<close>'', ``\<^verbatim>\<open>.\<close>'', | 
| 67904 | 235 | |
| 236 | \<^item> the separator between \<open>name\<close> and \<open>argument\<close> is the longest possible | |
| 237 | sequence of ASCII blanks (it could be empty, e.g.\ when the argument | |
| 238 | starts with a quote or bracket), | |
| 239 | ||
| 67918 | 240 | \<^item> \<open>argument\<close> is the rest of the message without line terminator. | 
| 67904 | 241 | |
| 242 | \<^medskip> | |
| 243 | Input messages are sent from the client to the server. Here the \<open>name\<close> | |
| 244 | specifies a \<^emph>\<open>server command\<close>: the list of known commands may be | |
| 245 | retrieved via the \<^verbatim>\<open>help\<close> command. | |
| 246 | ||
| 247 | \<^medskip> | |
| 248 | Output messages are sent from the server to the client. Here the \<open>name\<close> | |
| 249 | specifies the \<^emph>\<open>server reply\<close>, which always has a specific meaning as | |
| 250 | follows: | |
| 251 | ||
| 252 | \<^item> synchronous results: \<^verbatim>\<open>OK\<close> or \<^verbatim>\<open>ERROR\<close> | |
| 253 | \<^item> asynchronous results: \<^verbatim>\<open>FINISHED\<close> or \<^verbatim>\<open>FAILED\<close> | |
| 254 | \<^item> intermediate notifications: \<^verbatim>\<open>NOTE\<close> | |
| 255 | ||
| 256 | \<^medskip> | |
| 257 | The \<open>argument\<close> format is uniform for both input and output messages: | |
| 258 | ||
| 259 | \<^item> empty argument (Scala type \<^verbatim>\<open>Unit\<close>) | |
| 260 | \<^item> XML element in YXML notation (Scala type \<^verbatim>\<open>XML.Elem\<close>) | |
| 261 | \<^item> JSON value (Scala type \<^verbatim>\<open>JSON.T\<close>) | |
| 262 | ||
| 263 | JSON values may consist of objects (records), arrays (lists), strings, | |
| 75161 | 264 | numbers, bools, or null.\<^footnote>\<open>See also the official specification | 
| 67904 | 265 | \<^url>\<open>https://www.json.org\<close> and unofficial explorations ``Parsing JSON is a | 
| 266 | Minefield'' \<^url>\<open>http://seriot.ch/parsing_json.php\<close>.\<close> Since JSON requires | |
| 267 | explicit quotes and backslash-escapes to represent arbitrary text, the YXML | |
| 268 |   notation for XML trees (\secref{sec:yxml-vs-xml}) works better
 | |
| 269 | for large messages with a lot of PIDE markup. | |
| 270 | ||
| 271 | Nonetheless, the most common commands use JSON by default: big chunks of | |
| 272 | text (theory sources etc.) are taken from the underlying file-system and | |
| 273 | results are pre-formatted for plain-text output, without PIDE markup | |
| 274 | information. This is a concession to simplicity: the server imitates the | |
| 275 | appearance of command-line tools on top of the Isabelle/PIDE infrastructure. | |
| 276 | \<close> | |
| 277 | ||
| 278 | ||
| 279 | subsection \<open>Initial password exchange\<close> | |
| 280 | ||
| 281 | text \<open> | |
| 282 | Whenever a new client opens the server socket, the initial message needs to | |
| 69464 | 283 | be its unique password as a single line, without length indication (i.e.\ a | 
| 284 |   ``short message'' in the sense of \secref{sec:byte-messages}).
 | |
| 285 | ||
| 286 | The server replies either with \<^verbatim>\<open>OK\<close> (and some information about the | |
| 287 | Isabelle version) or by silent disconnection of what is considered an | |
| 288 |   illegal connection attempt. Note that @{tool client} already presents the
 | |
| 289 | correct password internally. | |
| 67904 | 290 | |
| 291 | Server passwords are created as Universally Unique Identifier (UUID) in | |
| 292 | Isabelle/Scala and stored in a per-user database, with restricted | |
| 293 | file-system access only for the current user. The Isabelle/Scala server | |
| 294 | implementation is careful to expose the password only on private output | |
| 295 | channels, and not on a process command-line (which is accessible to other | |
| 296 | users, e.g.\ via the \<^verbatim>\<open>ps\<close> command). | |
| 297 | \<close> | |
| 298 | ||
| 299 | ||
| 300 | subsection \<open>Synchronous commands\<close> | |
| 301 | ||
| 302 | text \<open> | |
| 303 | A \<^emph>\<open>synchronous command\<close> corresponds to regular function application in | |
| 304 | Isabelle/Scala, with single argument and result (regular or error). Both the | |
| 305 | argument and the result may consist of type \<^verbatim>\<open>Unit\<close>, \<^verbatim>\<open>XML.Elem\<close>, \<^verbatim>\<open>JSON.T\<close>. | |
| 306 | An error result typically consists of a JSON object with error message and | |
| 67918 | 307 | potentially further result fields (this resembles exceptions in Scala). | 
| 67904 | 308 | |
| 309 | These are the protocol exchanges for both cases of command execution: | |
| 310 |   \begin{center}
 | |
| 311 |   \begin{tabular}{rl}
 | |
| 312 | \<^bold>\<open>input:\<close> & \<open>command argument\<close> \\ | |
| 313 | (a) regular \<^bold>\<open>output:\<close> & \<^verbatim>\<open>OK\<close> \<open>result\<close> \\ | |
| 314 | (b) error \<^bold>\<open>output:\<close> & \<^verbatim>\<open>ERROR\<close> \<open>result\<close> \\ | |
| 315 |   \end{tabular}
 | |
| 316 |   \end{center}
 | |
| 317 | \<close> | |
| 318 | ||
| 319 | ||
| 320 | subsection \<open>Asynchronous commands\<close> | |
| 321 | ||
| 322 | text \<open> | |
| 323 | An \<^emph>\<open>asynchronous command\<close> corresponds to an ongoing process that finishes | |
| 67918 | 324 | or fails eventually, while emitting arbitrary notifications in between. | 
| 325 | Formally, it starts as synchronous command with immediate result \<^verbatim>\<open>OK\<close> | |
| 326 | giving the \<^verbatim>\<open>task\<close> identifier, or an immediate \<^verbatim>\<open>ERROR\<close> that indicates bad | |
| 327 | command syntax. For a running task, the termination is indicated later by | |
| 328 | \<^verbatim>\<open>FINISHED\<close> or \<^verbatim>\<open>FAILED\<close>, together with its ultimate result value. | |
| 67904 | 329 | |
| 330 | These are the protocol exchanges for various cases of command task | |
| 331 | execution: | |
| 332 | ||
| 333 |   \begin{center}
 | |
| 334 |   \begin{tabular}{rl}
 | |
| 335 | \<^bold>\<open>input:\<close> & \<open>command argument\<close> \\ | |
| 336 |   immediate \<^bold>\<open>output:\<close> & \<^verbatim>\<open>OK {"task":\<close>\<open>id\<close>\<^verbatim>\<open>}\<close> \\
 | |
| 337 |   intermediate \<^bold>\<open>output:\<close> & \<^verbatim>\<open>NOTE {"task":\<close>\<open>id\<close>\<^verbatim>\<open>,\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close> \\
 | |
| 338 |   (a) regular \<^bold>\<open>output:\<close> & \<^verbatim>\<open>FINISHED {"task":\<close>\<open>id\<close>\<^verbatim>\<open>,\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close> \\
 | |
| 339 |   (b) error \<^bold>\<open>output:\<close> & \<^verbatim>\<open>FAILED {"task":\<close>\<open>id\<close>\<^verbatim>\<open>,\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close> \\[3ex]
 | |
| 340 | \<^bold>\<open>input:\<close> & \<open>command argument\<close> \\ | |
| 341 | immediate \<^bold>\<open>output:\<close> & \<^verbatim>\<open>ERROR\<close>~\<open>\<dots>\<close> \\ | |
| 342 |   \end{tabular}
 | |
| 343 |   \end{center}
 | |
| 344 | ||
| 345 | All asynchronous messages are decorated with the task identifier that was | |
| 67918 | 346 | revealed in the immediate (synchronous) result. Thus the client can | 
| 347 | invoke further asynchronous commands and still dispatch the resulting stream of | |
| 67904 | 348 | asynchronous messages properly. | 
| 349 | ||
| 67920 | 350 |   The synchronous command \<^verbatim>\<open>cancel {"task":\<close>~\<open>id\<close>\<^verbatim>\<open>}\<close> tells the specified task
 | 
| 351 | to terminate prematurely: usually causing a \<^verbatim>\<open>FAILED\<close> result, but this is | |
| 352 | not guaranteed: the cancel event may come too late or the running process | |
| 353 | may just ignore it. | |
| 67917 | 354 | \<close> | 
| 355 | ||
| 356 | ||
| 67918 | 357 | section \<open>Types for JSON values \label{sec:json-types}\<close>
 | 
| 67917 | 358 | |
| 359 | text \<open> | |
| 360 | In order to specify concrete JSON types for command arguments and result | |
| 361 | messages, the following type definition language shall be used: | |
| 362 | ||
| 363 | \<^rail>\<open> | |
| 364 |     @{syntax type_def}: @'type' @{syntax name} '=' @{syntax type}
 | |
| 365 | ; | |
| 366 |     @{syntax type}: @{syntax name} | @{syntax value} | 'any' | 'null' |
 | |
| 367 |       'bool' | 'int' | 'long' | 'double' | 'string' | '[' @{syntax type} ']' |
 | |
| 368 |       '{' (@{syntax field_type} ',' *) '}' |
 | |
| 369 |       @{syntax type} '\<oplus>' @{syntax type} |
 | |
| 67918 | 370 |       @{syntax type} '|' @{syntax type} |
 | 
| 371 |       '(' @{syntax type} ')'
 | |
| 67917 | 372 | ; | 
| 373 |     @{syntax field_type}: @{syntax name} ('?'?) ':' @{syntax type}
 | |
| 374 | \<close> | |
| 375 | ||
| 67918 | 376 | This is a simplified variation of TypeScript | 
| 377 | interfaces.\<^footnote>\<open>\<^url>\<open>https://www.typescriptlang.org/docs/handbook/interfaces.html\<close>\<close> | |
| 378 | The meaning of these types is specified wrt. the Isabelle/Scala | |
| 379 | implementation as follows. | |
| 67917 | 380 | |
| 381 | \<^item> A \<open>name\<close> refers to a type defined elsewhere. The environment of type | |
| 382 | definitions is given informally: put into proper foundational order, it | |
| 67918 | 383 | needs to specify a strongly normalizing system of syntactic abbreviations; | 
| 384 | type definitions may not be recursive. | |
| 67917 | 385 | |
| 386 | \<^item> A \<open>value\<close> in JSON notation represents the singleton type of the given | |
| 67918 | 387 | item. For example, the string \<^verbatim>\<open>"error"\<close> can be used as type for a slot that | 
| 388 | is guaranteed to contain that constant. | |
| 67917 | 389 | |
| 390 | \<^item> Type \<open>any\<close> is the super type of all other types: it is an untyped slot in | |
| 391 | the specification and corresponds to \<^verbatim>\<open>Any\<close> or \<^verbatim>\<open>JSON.T\<close> in Isabelle/Scala. | |
| 392 | ||
| 393 | \<^item> Type \<open>null\<close> is the type of the improper value \<open>null\<close>; it corresponds to | |
| 394 | type \<^verbatim>\<open>Null\<close> in Scala and is normally not used in Isabelle/Scala.\<^footnote>\<open>See also | |
| 395 | ``Null References: The Billion Dollar Mistake'' by Tony Hoare | |
| 396 | \<^url>\<open>https://www.infoq.com/presentations/Null-References-The-Billion-Dollar-Mistake-Tony-Hoare\<close>.\<close> | |
| 397 | ||
| 67918 | 398 | \<^item> Type \<open>bool\<close> is the type of the truth values \<^verbatim>\<open>true\<close> and \<^verbatim>\<open>false\<close>; it | 
| 67917 | 399 | corresponds to \<^verbatim>\<open>Boolean\<close> in Scala. | 
| 400 | ||
| 401 | \<^item> Types \<open>int\<close>, \<open>long\<close>, \<open>double\<close> are specific versions of the generic | |
| 402 | \<open>number\<close> type, corresponding to \<^verbatim>\<open>Int\<close>, \<^verbatim>\<open>Long\<close>, \<^verbatim>\<open>Double\<close> in Scala, but | |
| 403 | \<^verbatim>\<open>Long\<close> is limited to 53 bit precision.\<^footnote>\<open>Implementations of JSON typically | |
| 404 | standardize \<open>number\<close> to \<^verbatim>\<open>Double\<close>, which can absorb \<^verbatim>\<open>Int\<close> faithfully, but | |
| 405 | not all of \<^verbatim>\<open>Long\<close>.\<close> | |
| 406 | ||
| 407 | \<^item> Type \<open>string\<close> represents Unicode text; it corresponds to type \<^verbatim>\<open>String\<close> in | |
| 408 | Scala. | |
| 409 | ||
| 410 | \<^item> Type \<open>[t]\<close> is the array (or list) type over \<open>t\<close>; it corresponds to | |
| 411 | \<^verbatim>\<open>List[t]\<close> in Scala. The list type is co-variant as usual (i.e.\ monotonic | |
| 67918 | 412 | wrt. the subtype relation). | 
| 67917 | 413 | |
| 414 | \<^item> Object types describe the possible content of JSON records, with field | |
| 415 | names and types. A question mark after a field name means that it is | |
| 67918 | 416 | optional. In Scala this could refer to an explicit type \<^verbatim>\<open>Option[t]\<close>, e.g.\ | 
| 417 |   \<open>{a: int, b?: string}\<close> corresponding to a Scala case class with arguments
 | |
| 418 | \<^verbatim>\<open>a: Int\<close>, \<^verbatim>\<open>b: Option[String]\<close>. | |
| 67917 | 419 | |
| 420 | Alternatively, optional fields can have a default value. If nothing else is | |
| 67918 | 421 | specified, a standard ``empty value'' is used for each type, i.e.\ \<^verbatim>\<open>0\<close> for | 
| 422 | the number types, \<^verbatim>\<open>false\<close> for \<open>bool\<close>, or the empty string, array, object | |
| 423 | etc. | |
| 67917 | 424 | |
| 425 | Object types are \<^emph>\<open>permissive\<close> in the sense that only the specified field | |
| 426 | names need to conform to the given types, but unspecified fields may be | |
| 427 | present as well. | |
| 428 | ||
| 429 | \<^item> The type expression \<open>t\<^sub>1 \<oplus> t\<^sub>2\<close> only works for two object types with | |
| 430 |   disjoint field names: it is the concatenation of the respective @{syntax
 | |
| 431 |   field_type} specifications taken together. For example: \<open>{task: string} \<oplus>
 | |
| 432 |   {ok: bool}\<close> is the equivalent to \<open>{task: string, ok: bool}\<close>.
 | |
| 433 | ||
| 434 | \<^item> The type expression \<open>t\<^sub>1 | t\<^sub>2\<close> is the disjoint union of two types, either | |
| 435 | one of the two cases may occur. | |
| 436 | ||
| 67918 | 437 | \<^item> Parentheses \<open>(t)\<close> merely group type expressions syntactically. | 
| 438 | ||
| 67917 | 439 | |
| 440 | These types correspond to JSON values in an obvious manner, which is not | |
| 441 | further described here. For example, the JSON array \<^verbatim>\<open>[1, 2, 3]\<close> conforms to | |
| 442 | types \<open>[int]\<close>, \<open>[long]\<close>, \<open>[double]\<close>, \<open>[any]\<close>, \<open>any\<close>. | |
| 443 | ||
| 444 | Note that JSON objects require field names to be quoted, but the type | |
| 67918 | 445 |   language omits quotes for clarity. Thus the object \<^verbatim>\<open>{"a": 42, "b": "xyz"}\<close>
 | 
| 446 |   conforms to the type \<open>{a: int, b: string}\<close>, for example.
 | |
| 67917 | 447 | |
| 448 | \<^medskip> | |
| 67918 | 449 | The absence of an argument or result is represented by the Scala type | 
| 450 | \<^verbatim>\<open>Unit\<close>: it is written as empty text in the message \<open>argument\<close> | |
| 67917 | 451 |   (\secref{sec:input-output-messages}). This is not part of the JSON language.
 | 
| 452 | ||
| 453 | Server replies have name tags like \<^verbatim>\<open>OK\<close>, \<^verbatim>\<open>ERROR\<close>: these are used literally | |
| 454 | together with type specifications to indicate the particular name with the | |
| 455 | type of its argument, e.g.\ \<^verbatim>\<open>OK\<close>~\<open>[string]\<close> for a regular result that is a | |
| 456 | list (JSON array) of strings. | |
| 457 | ||
| 458 | \<^bigskip> | |
| 67918 | 459 | Here are some common type definitions, for use in particular specifications | 
| 460 | of command arguments and results. | |
| 67917 | 461 | |
| 462 |   \<^item> \<^bold>\<open>type\<close>~\<open>position = {line?: int, offset?: int, end_offset?: int, file?:
 | |
| 67918 | 463 | string, id?: long}\<close> describes a source position within Isabelle text. Only | 
| 464 | the \<open>line\<close> and \<open>file\<close> fields make immediate sense to external programs. | |
| 67917 | 465 | Detailed \<open>offset\<close> and \<open>end_offset\<close> positions are counted according to | 
| 76987 | 466 | Isabelle symbols, see \<^ML_type>\<open>Symbol.symbol\<close> in Isabelle/ML \<^cite>\<open>"isabelle-implementation"\<close>. The position \<open>id\<close> belongs to the representation | 
| 67918 | 467 | of command transactions in the Isabelle/PIDE protocol: it normally does not | 
| 468 | occur in externalized positions. | |
| 67917 | 469 | |
| 67931 
f7917c15b566
field "kind" is always present, with default "writeln";
 wenzelm parents: 
67927diff
changeset | 470 |   \<^item> \<^bold>\<open>type\<close>~\<open>message = {kind: string, message: string, pos?: position}\<close> where
 | 
| 67917 | 471 | the \<open>kind\<close> provides some hint about the role and importance of the message. | 
| 67923 
3e072441c96a
clarified exported messages, e.g. suppress "information", "tracing";
 wenzelm parents: 
67922diff
changeset | 472 | The main message kinds are \<^verbatim>\<open>writeln\<close> (for regular output), \<^verbatim>\<open>warning\<close>, | 
| 67931 
f7917c15b566
field "kind" is always present, with default "writeln";
 wenzelm parents: 
67927diff
changeset | 473 | \<^verbatim>\<open>error\<close>. | 
| 67917 | 474 | |
| 475 |   \<^item> \<^bold>\<open>type\<close>~\<open>error_message = {kind:\<close>~\<^verbatim>\<open>"error"\<close>\<open>, message: string}\<close> refers to
 | |
| 476 | error messages in particular. These occur routinely with \<^verbatim>\<open>ERROR\<close> or | |
| 477 | \<^verbatim>\<open>FAILED\<close> replies, but also as initial command syntax errors (which are | |
| 478 | omitted in the command specifications below). | |
| 479 | ||
| 480 |   \<^item> \<^bold>\<open>type\<close>~\<open>theory_progress = {kind:\<close>~\<^verbatim>\<open>"writeln"\<close>\<open>, message: string, theory:
 | |
| 68957 | 481 | string, session: string, percentage?: int}\<close> reports formal progress in | 
| 482 | loading theories (e.g.\ when building a session image). Apart from a regular | |
| 483 | output message, it also reveals the formal theory name (e.g.\ \<^verbatim>\<open>"HOL.Nat"\<close>) | |
| 484 | and session name (e.g.\ \<^verbatim>\<open>"HOL"\<close>). Note that some rare theory names lack a | |
| 485 | proper session prefix, e.g.\ theory \<^verbatim>\<open>"Main"\<close> in session \<^verbatim>\<open>"HOL"\<close>. The | |
| 486 | optional percentage has the same meaning as in \<^bold>\<open>type\<close>~\<open>node_status\<close> below. | |
| 67917 | 487 | |
| 488 |   \<^item> \<^bold>\<open>type\<close>~\<open>timing = {elapsed: double, cpu: double, gc: double}\<close> refers to
 | |
| 489 | common Isabelle timing information in seconds, usually with a precision of | |
| 490 | three digits after the point (whole milliseconds). | |
| 491 | ||
| 492 | \<^item> \<^bold>\<open>type\<close>~\<open>uuid = string\<close> refers to a Universally Unique Identifier (UUID) | |
| 493 | as plain text.\<^footnote>\<open>See \<^url>\<open>https://www.ietf.org/rfc/rfc4122.txt\<close> and | |
| 79724 | 494 | \<^url>\<open>https://docs.oracle.com/en/java/javase/21/docs/api/java.base/java/util/UUID.html\<close>.\<close> Such | 
| 67917 | 495 | identifiers are created as private random numbers of the server and only | 
| 496 | revealed to the client that creates a certain resource (e.g.\ task or | |
| 497 | session). A client may disclose this information for use in a different | |
| 67918 | 498 | client connection: this allows to share sessions between multiple | 
| 67917 | 499 | connections. | 
| 500 | ||
| 501 | Client commands need to provide syntactically wellformed UUIDs: this is | |
| 502 | trivial to achieve by using only identifiers that have been produced by the | |
| 503 | server beforehand. | |
| 504 | ||
| 505 |   \<^item> \<^bold>\<open>type\<close>~\<open>task = {task: uuid}\<close> identifies a newly created asynchronous task
 | |
| 506 | and thus allows the client to control it by the \<^verbatim>\<open>cancel\<close> command. The same | |
| 67918 | 507 | task identification is included in all messages produced by this task. | 
| 508 | ||
| 67921 | 509 |   \<^item> \<^bold>\<open>type\<close> \<open>session_id = {session_id: uuid}\<close> identifies a newly created PIDE
 | 
| 510 | session managed by the server. Sessions are independent of client | |
| 511 | connections and may be shared by different clients, as long as the internal | |
| 512 | session identifier is known. | |
| 513 | ||
| 67943 | 514 |   \<^item> \<^bold>\<open>type\<close> \<open>node = {node_name: string, theory_name: string}\<close> represents the
 | 
| 515 | internal node name of a theory. The \<open>node_name\<close> is derived from the | |
| 71925 
bf085daea304
clarified sessions: "Notable Examples in Isabelle/HOL";
 wenzelm parents: 
71521diff
changeset | 516 | canonical theory file-name (e.g.\ \<^verbatim>\<open>"~~/src/HOL/Examples/Seq.thy"\<close> after | 
| 67943 | 517 | normalization within the file-system). The \<open>theory_name\<close> is the | 
| 71925 
bf085daea304
clarified sessions: "Notable Examples in Isabelle/HOL";
 wenzelm parents: 
71521diff
changeset | 518 | session-qualified theory name (e.g.\ \<^verbatim>\<open>HOL-Examples.Seq\<close>). | 
| 67943 | 519 | |
| 67918 | 520 |   \<^item> \<^bold>\<open>type\<close> \<open>node_status = {ok: bool, total: int, unprocessed: int, running:
 | 
| 68885 | 521 | int, warned: int, failed: int, finished: int, canceled: bool, consolidated: | 
| 68898 | 522 | bool, percentage: int}\<close> represents a formal theory node status of the PIDE | 
| 523 | document model as follows. | |
| 68882 | 524 | |
| 525 | \<^item> Fields \<open>total\<close>, \<open>unprocessed\<close>, \<open>running\<close>, \<open>warned\<close>, \<open>failed\<close>, \<open>finished\<close> | |
| 526 | account for individual commands within a theory node; \<open>ok\<close> is an | |
| 527 | abstraction for \<open>failed = 0\<close>. | |
| 528 | ||
| 529 | \<^item> The \<open>canceled\<close> flag tells if some command in the theory has been | |
| 530 | spontaneously canceled (by an Interrupt exception that could also | |
| 531 | indicate resource problems). | |
| 532 | ||
| 533 | \<^item> The \<open>consolidated\<close> flag indicates whether the outermost theory command | |
| 534 | structure has finished (or failed) and the final \<^theory_text>\<open>end\<close> command has been | |
| 535 | checked. | |
| 68898 | 536 | |
| 537 | \<^item> The \<open>percentage\<close> field tells how far the node has been processed. It | |
| 538 | ranges between 0 and 99 in normal operation, and reaches 100 when the node | |
| 539 | has been formally consolidated as described above. | |
| 67918 | 540 | \<close> | 
| 541 | ||
| 542 | ||
| 68898 | 543 | |
| 67918 | 544 | section \<open>Server commands and results\<close> | 
| 545 | ||
| 546 | text \<open> | |
| 547 | Here follows an overview of particular Isabelle server commands with their | |
| 548 | results, which are usually represented as JSON values with types according | |
| 549 |   to \secref{sec:json-types}. The general format of input and output messages
 | |
| 550 |   is described in \secref{sec:input-output-messages}. The relevant
 | |
| 551 | Isabelle/Scala source files are: | |
| 552 | ||
| 553 | \<^medskip> | |
| 554 |   \begin{tabular}{l}
 | |
| 555 | \<^file>\<open>$ISABELLE_HOME/src/Pure/Tools/server_commands.scala\<close> \\ | |
| 556 | \<^file>\<open>$ISABELLE_HOME/src/Pure/Tools/server.scala\<close> \\ | |
| 557 | \<^file>\<open>$ISABELLE_HOME/src/Pure/General/json.scala\<close> \\ | |
| 558 |   \end{tabular}
 | |
| 67917 | 559 | \<close> | 
| 560 | ||
| 561 | ||
| 562 | subsection \<open>Command \<^verbatim>\<open>help\<close>\<close> | |
| 563 | ||
| 564 | text \<open> | |
| 565 |   \begin{tabular}{ll}
 | |
| 566 | regular result: & \<^verbatim>\<open>OK\<close> \<open>[string]\<close> \\ | |
| 567 |   \end{tabular}
 | |
| 568 | \<^medskip> | |
| 569 | ||
| 67918 | 570 | The \<^verbatim>\<open>help\<close> command has no argument and returns the list of server command | 
| 571 | names. This is occasionally useful for interactive experimentation (see also | |
| 572 |   @{tool client} in \secref{sec:tool-client}).
 | |
| 67917 | 573 | \<close> | 
| 574 | ||
| 575 | ||
| 576 | subsection \<open>Command \<^verbatim>\<open>echo\<close>\<close> | |
| 577 | ||
| 578 | text \<open> | |
| 579 |   \begin{tabular}{ll}
 | |
| 580 | argument: & \<open>any\<close> \\ | |
| 581 | regular result: & \<^verbatim>\<open>OK\<close> \<open>any\<close> \\ | |
| 582 |   \end{tabular}
 | |
| 583 | \<^medskip> | |
| 584 | ||
| 585 | The \<^verbatim>\<open>echo\<close> command is the identity function: it returns its argument as | |
| 586 | regular result. This is occasionally useful for testing and interactive | |
| 587 |   experimentation (see also @{tool client} in \secref{sec:tool-client}).
 | |
| 588 | ||
| 67918 | 589 | The Scala type of \<^verbatim>\<open>echo\<close> is actually more general than given above: | 
| 590 | \<^verbatim>\<open>Unit\<close>, \<^verbatim>\<open>XML.Elem\<close>, \<^verbatim>\<open>JSON.T\<close> work uniformly. Note that \<^verbatim>\<open>XML.Elem\<close> might | |
| 67917 | 591 | be difficult to type on the console in its YXML syntax | 
| 592 |   (\secref{sec:yxml-vs-xml}).
 | |
| 593 | \<close> | |
| 594 | ||
| 595 | ||
| 596 | subsection \<open>Command \<^verbatim>\<open>shutdown\<close>\<close> | |
| 597 | ||
| 598 | text \<open> | |
| 599 |   \begin{tabular}{ll}
 | |
| 600 | regular result: & \<^verbatim>\<open>OK\<close> \\ | |
| 601 |   \end{tabular}
 | |
| 602 | \<^medskip> | |
| 603 | ||
| 604 | The \<^verbatim>\<open>shutdown\<close> command has no argument and result value. It forces a | |
| 605 | shutdown of the connected server process, stopping all open sessions and | |
| 606 | closing the server socket. This may disrupt pending commands on other | |
| 607 | connections! | |
| 608 | ||
| 609 | \<^medskip> | |
| 67918 | 610 | The command-line invocation \<^verbatim>\<open>isabelle server -x\<close> opens a server connection | 
| 611 |   and issues a \<^verbatim>\<open>shutdown\<close> command (see also \secref{sec:tool-server}).
 | |
| 67917 | 612 | \<close> | 
| 613 | ||
| 614 | ||
| 615 | subsection \<open>Command \<^verbatim>\<open>cancel\<close>\<close> | |
| 616 | ||
| 617 | text \<open> | |
| 618 |   \begin{tabular}{ll}
 | |
| 67920 | 619 | argument: & \<open>task\<close> \\ | 
| 67917 | 620 | regular result: & \<^verbatim>\<open>OK\<close> \\ | 
| 621 |   \end{tabular}
 | |
| 622 | \<^medskip> | |
| 623 | ||
| 67921 | 624 |   The command \<^verbatim>\<open>cancel {"task":\<close>~\<open>id\<close>\<^verbatim>\<open>}\<close> attempts to cancel the specified
 | 
| 625 | task. | |
| 67917 | 626 | |
| 627 | Cancellation is merely a hint that the client prefers an ongoing process to | |
| 628 | be stopped. The command always succeeds formally, but it may get ignored by | |
| 67918 | 629 | a task that is still running; it might also refer to a non-existing or | 
| 630 | no-longer existing task (without producing an error). | |
| 67917 | 631 | |
| 632 | Successful cancellation typically leads to an asynchronous failure of type | |
| 67918 | 633 |   \<^verbatim>\<open>FAILED {\<close>\<open>task: uuid, message:\<close>~\<^verbatim>\<open>"Interrupt"}\<close>. A different message is
 | 
| 634 | also possible, depending how the task handles the event. | |
| 67917 | 635 | \<close> | 
| 636 | ||
| 637 | ||
| 638 | subsection \<open>Command \<^verbatim>\<open>session_build\<close> \label{sec:command-session-build}\<close>
 | |
| 639 | ||
| 640 | text \<open> | |
| 641 |   \begin{tabular}{lll}
 | |
| 642 | argument: & \<open>session_build_args\<close> \\ | |
| 643 | immediate result: & \<^verbatim>\<open>OK\<close> \<open>task\<close> \\ | |
| 644 | notifications: & \<^verbatim>\<open>NOTE\<close> \<open>task \<oplus> (theory_progress | message)\<close> \\ | |
| 645 | regular result: & \<^verbatim>\<open>FINISHED\<close> \<open>task \<oplus> session_build_results\<close> \\ | |
| 646 | error result: & \<^verbatim>\<open>FAILED\<close> \<open>task \<oplus> error_message \<oplus> session_build_results\<close> \\[2ex] | |
| 647 |   \end{tabular}
 | |
| 648 | ||
| 649 |   \begin{tabular}{lll}
 | |
| 650 | \<^bold>\<open>type\<close> \<open>session_build_args =\<close> \\ | |
| 651 |   \quad\<open>{session: string,\<close> \\
 | |
| 652 | \quad~~\<open>preferences?: string,\<close> & \<^bold>\<open>default:\<close> server preferences \\ | |
| 653 | \quad~~\<open>options?: [string],\<close> \\ | |
| 654 | \quad~~\<open>dirs?: [string],\<close> \\ | |
| 67922 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 655 | \quad~~\<open>include_sessions: [string],\<close> \\ | 
| 67917 | 656 | \quad~~\<open>verbose?: bool}\<close> \\[2ex] | 
| 67919 
dd90faed43b2
clarified signature: do not expose somewhat accidental internal options;
 wenzelm parents: 
67918diff
changeset | 657 |   \end{tabular}
 | 
| 67917 | 658 | |
| 67919 
dd90faed43b2
clarified signature: do not expose somewhat accidental internal options;
 wenzelm parents: 
67918diff
changeset | 659 |   \begin{tabular}{ll}
 | 
| 67917 | 660 | \<^bold>\<open>type\<close> \<open>session_build_result =\<close> \\ | 
| 661 |   \quad\<open>{session: string,\<close> \\
 | |
| 662 | \quad~~\<open>ok: bool,\<close> \\ | |
| 663 | \quad~~\<open>return_code: int,\<close> \\ | |
| 664 | \quad~~\<open>timeout: bool,\<close> \\ | |
| 67918 | 665 | \quad~~\<open>timing: timing}\<close> \\[2ex] | 
| 666 | ||
| 667 | \<^bold>\<open>type\<close> \<open>session_build_results =\<close> \\ | |
| 668 |   \quad\<open>{ok: bool,\<close> \\
 | |
| 669 | \quad~~\<open>return_code: int,\<close> \\ | |
| 670 | \quad~~\<open>sessions: [session_build_result]}\<close> \\ | |
| 67917 | 671 |   \end{tabular}
 | 
| 672 | \<close> | |
| 673 | ||
| 674 | text \<open> | |
| 67918 | 675 | The \<^verbatim>\<open>session_build\<close> command prepares a session image for interactive use of | 
| 676 |   theories. This is a limited version of command-line tool @{tool build}
 | |
| 677 |   (\secref{sec:tool-build}), with specific options to request a formal context
 | |
| 678 | for an interactive PIDE session. | |
| 67917 | 679 | |
| 680 | The build process is asynchronous, with notifications that inform about the | |
| 67918 | 681 | progress of loaded theories. Some further informative messages are output as | 
| 682 | well. | |
| 67917 | 683 | |
| 684 | Coordination of independent build processes is at the discretion of the | |
| 67918 | 685 |   client (or end-user), just as for @{tool build} and @{tool jedit}. There is
 | 
| 686 | no built-in coordination of conflicting builds with overlapping hierarchies | |
| 687 | of session images. In the worst case, a session image produced by one task | |
| 688 | may get overwritten by another task! | |
| 67917 | 689 | \<close> | 
| 690 | ||
| 691 | ||
| 692 | subsubsection \<open>Arguments\<close> | |
| 693 | ||
| 694 | text \<open> | |
| 67919 
dd90faed43b2
clarified signature: do not expose somewhat accidental internal options;
 wenzelm parents: 
67918diff
changeset | 695 | The \<open>session\<close> field specifies the target session name. The build process | 
| 
dd90faed43b2
clarified signature: do not expose somewhat accidental internal options;
 wenzelm parents: 
67918diff
changeset | 696 | will produce all required ancestor images according to the overall session | 
| 
dd90faed43b2
clarified signature: do not expose somewhat accidental internal options;
 wenzelm parents: 
67918diff
changeset | 697 | graph. | 
| 67917 | 698 | |
| 699 | \<^medskip> | |
| 700 | The environment of Isabelle system options is determined from \<open>preferences\<close> | |
| 701 | that are augmented by \<open>options\<close>, which is a list individual updates of the | |
| 702 | form the \<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close> or \<open>name\<close> (the latter abbreviates | |
| 703 |   \<open>name\<close>\<^verbatim>\<open>=true\<close>); see also command-line option \<^verbatim>\<open>-o\<close> for @{tool build}. The
 | |
| 704 | preferences are loaded from the file | |
| 705 | \<^path>\<open>$ISABELLE_HOME_USER/etc/preferences\<close> by default, but the client may | |
| 706 | provide alternative contents for it (as text, not a file-name). This could | |
| 707 | be relevant in situations where client and server run in different | |
| 708 | operating-system contexts. | |
| 709 | ||
| 710 | \<^medskip> | |
| 67918 | 711 | The \<open>dirs\<close> field specifies additional directories for session ROOT and ROOTS | 
| 712 |   files (\secref{sec:session-root}). This augments the name space of available
 | |
| 713 |   sessions; see also option \<^verbatim>\<open>-d\<close> in @{tool build}.
 | |
| 67917 | 714 | |
| 715 | \<^medskip> | |
| 67922 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 716 | The \<open>include_sessions\<close> field specifies sessions whose theories should be | 
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 717 | included in the overall name space of session-qualified theory names. This | 
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 718 | corresponds to a \<^bold>\<open>sessions\<close> specification in ROOT files | 
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 719 |   (\secref{sec:session-root}). It enables the \<^verbatim>\<open>use_theories\<close> command
 | 
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 720 |   (\secref{sec:command-use-theories}) to refer to sources from other sessions
 | 
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 721 | in a robust manner, instead of relying on directory locations. | 
| 67917 | 722 | \<close> | 
| 723 | ||
| 724 | ||
| 725 | subsubsection \<open>Intermediate output\<close> | |
| 726 | ||
| 727 | text \<open> | |
| 728 | The asynchronous notifications of command \<^verbatim>\<open>session_build\<close> mainly serve as | |
| 729 | progress indicator: the output resembles that of the session build window of | |
| 76987 | 730 | Isabelle/jEdit after startup \<^cite>\<open>"isabelle-jedit"\<close>. | 
| 67917 | 731 | |
| 732 | For the client it is usually sufficient to print the messages in plain text, | |
| 67918 | 733 | but note that \<open>theory_progress\<close> also reveals formal \<open>theory\<close> and | 
| 734 | \<open>session\<close> names directly. | |
| 67917 | 735 | \<close> | 
| 736 | ||
| 737 | ||
| 738 | subsubsection \<open>Results\<close> | |
| 739 | ||
| 740 | text \<open> | |
| 67918 | 741 | The overall \<open>session_build_results\<close> contain both a summary and an entry | 
| 742 | \<open>session_build_result\<close> for each session in the build hierarchy. The result | |
| 743 | is always provided, independently of overall success (\<^verbatim>\<open>FINISHED\<close> task) or | |
| 744 | failure (\<^verbatim>\<open>FAILED\<close> task). | |
| 67917 | 745 | |
| 746 | The \<open>ok\<close> field tells abstractly, whether all required session builds came | |
| 67918 | 747 | out as \<open>ok\<close>, i.e.\ with zero \<open>return_code\<close>. A non-zero \<open>return_code\<close> | 
| 748 | indicates an error according to usual POSIX conventions for process exit. | |
| 749 | ||
| 750 | The individual \<open>session_build_result\<close> entries provide extra fields: | |
| 67917 | 751 | |
| 67918 | 752 | \<^item> \<open>timeout\<close> tells if the build process was aborted after running too long, | 
| 753 | ||
| 754 | \<^item> \<open>timing\<close> gives the overall process timing in the usual Isabelle format | |
| 755 | with elapsed, CPU, GC time. | |
| 67917 | 756 | \<close> | 
| 757 | ||
| 758 | ||
| 759 | subsubsection \<open>Examples\<close> | |
| 760 | ||
| 761 | text \<open> | |
| 762 | Build of a session image from the Isabelle distribution: | |
| 72515 
c7038c397ae3
moved most material from session HOL-Word to Word_Lib in the AFP
 haftmann parents: 
71925diff
changeset | 763 |   @{verbatim [display] \<open>session_build {"session": "HOL-Algebra"}\<close>}
 | 
| 67917 | 764 | |
| 765 | Build a session image from the Archive of Formal Proofs: | |
| 766 |   @{verbatim [display] \<open>session_build {"session": "Coinductive", "dirs": ["$AFP_BASE/thys"]}\<close>}
 | |
| 767 | \<close> | |
| 768 | ||
| 769 | ||
| 770 | subsection \<open>Command \<^verbatim>\<open>session_start\<close> \label{sec:command-session-start}\<close>
 | |
| 771 | ||
| 772 | text \<open> | |
| 773 |   \begin{tabular}{lll}
 | |
| 774 |   argument: & \<open>session_build_args \<oplus> {print_mode?: [string]}\<close> \\
 | |
| 775 | immediate result: & \<^verbatim>\<open>OK\<close> \<open>task\<close> \\ | |
| 776 | notifications: & \<^verbatim>\<open>NOTE\<close> \<open>task \<oplus> (theory_progress | message)\<close> \\ | |
| 67925 | 777 |   regular result: & \<^verbatim>\<open>FINISHED\<close> \<open>task \<oplus> session_id \<oplus> {tmp_dir: string}\<close> \\
 | 
| 67917 | 778 | error result: & \<^verbatim>\<open>FAILED\<close> \<open>task \<oplus> error_message\<close> \\[2ex] | 
| 779 |   \end{tabular}
 | |
| 780 | ||
| 781 | \<^medskip> | |
| 782 | The \<^verbatim>\<open>session_start\<close> command starts a new Isabelle/PIDE session with | |
| 783 | underlying Isabelle/ML process, based on a session image that it produces on | |
| 67918 | 784 | demand using \<^verbatim>\<open>session_build\<close>. Thus it accepts all \<open>session_build_args\<close> and | 
| 785 | produces similar notifications, but the detailed \<open>session_build_results\<close> are | |
| 786 | omitted. | |
| 67917 | 787 | |
| 788 | The session build and startup process is asynchronous: when the task is | |
| 67918 | 789 | finished, the session remains active for commands, until a \<^verbatim>\<open>session_stop\<close> | 
| 790 | or \<^verbatim>\<open>shutdown\<close> command is sent to the server. | |
| 67917 | 791 | |
| 792 | Sessions are independent of client connections: it is possible to start a | |
| 793 | session and later apply \<^verbatim>\<open>use_theories\<close> on different connections, as long as | |
| 67918 | 794 | the internal session identifier is known: shared theory imports will be used | 
| 795 | only once (and persist until purged explicitly). | |
| 67917 | 796 | \<close> | 
| 797 | ||
| 798 | ||
| 799 | subsubsection \<open>Arguments\<close> | |
| 800 | ||
| 801 | text \<open> | |
| 67918 | 802 | Most arguments are shared with \<^verbatim>\<open>session_build\<close> | 
| 67917 | 803 |   (\secref{sec:command-session-build}).
 | 
| 804 | ||
| 805 | \<^medskip> | |
| 806 | The \<open>print_mode\<close> field adds identifiers of print modes to be made active for | |
| 807 | this session. For example, \<^verbatim>\<open>"print_mode": ["ASCII"]\<close> prefers ASCII | |
| 808 | replacement syntax over mathematical Isabelle symbols. See also option \<^verbatim>\<open>-m\<close> | |
| 809 |   in @{tool process} (\secref{sec:tool-process}).
 | |
| 810 | \<close> | |
| 811 | ||
| 812 | ||
| 813 | subsubsection \<open>Results\<close> | |
| 814 | ||
| 815 | text \<open> | |
| 67921 | 816 | The \<open>session_id\<close> provides the internal identification of the session object | 
| 75161 | 817 | within the server process. It can remain active as long as the server is | 
| 67921 | 818 | running, independently of the current client connection. | 
| 67925 | 819 | |
| 820 | \<^medskip> | |
| 67946 | 821 | The \<open>tmp_dir\<close> field refers to a temporary directory that is specifically | 
| 67925 | 822 | created for this session and deleted after it has been stopped. This may | 
| 823 | serve as auxiliary file-space for the \<^verbatim>\<open>use_theories\<close> command, but | |
| 824 | concurrent use requires some care in naming temporary files, e.g.\ by | |
| 825 | using sub-directories with globally unique names. | |
| 67946 | 826 | |
| 827 | As \<open>tmp_dir\<close> is the default \<open>master_dir\<close> for commands \<^verbatim>\<open>use_theories\<close> and | |
| 828 | \<^verbatim>\<open>purge_theories\<close>, theory files copied there may be used without further | |
| 829 | path specification. | |
| 67917 | 830 | \<close> | 
| 831 | ||
| 832 | ||
| 68145 | 833 | subsubsection \<open>Examples\<close> | 
| 67917 | 834 | |
| 835 | text \<open> | |
| 836 | Start a default Isabelle/HOL session: | |
| 837 |   @{verbatim [display] \<open>session_start {"session": "HOL"}\<close>}
 | |
| 67918 | 838 | |
| 839 | Start a session from the Archive of Formal Proofs: | |
| 840 |   @{verbatim [display] \<open>session_start {"session": "Coinductive", "dirs": ["$AFP_BASE/thys"]}\<close>}
 | |
| 79067 
212c94edae2b
more reactive headless server, in contrast to 15656ad28691 (when "isabelle dump" was important to export AFP content);
 wenzelm parents: 
77519diff
changeset | 841 | |
| 
212c94edae2b
more reactive headless server, in contrast to 15656ad28691 (when "isabelle dump" was important to export AFP content);
 wenzelm parents: 
77519diff
changeset | 842 | Start a session with fine-tuning of options: | 
| 
212c94edae2b
more reactive headless server, in contrast to 15656ad28691 (when "isabelle dump" was important to export AFP content);
 wenzelm parents: 
77519diff
changeset | 843 |   @{verbatim [display] \<open>session_start {"session": "HOL",
 | 
| 
212c94edae2b
more reactive headless server, in contrast to 15656ad28691 (when "isabelle dump" was important to export AFP content);
 wenzelm parents: 
77519diff
changeset | 844 | "options": ["headless_consolidate_delay=0.5", "headless_prune_delay=5"]}\<close>} | 
| 67917 | 845 | \<close> | 
| 846 | ||
| 847 | ||
| 848 | subsection \<open>Command \<^verbatim>\<open>session_stop\<close>\<close> | |
| 849 | ||
| 850 | text \<open> | |
| 851 |   \begin{tabular}{ll}
 | |
| 67921 | 852 | argument: & \<open>session_id\<close> \\ | 
| 67917 | 853 | immediate result: & \<^verbatim>\<open>OK\<close> \<open>task\<close> \\ | 
| 854 | regular result: & \<^verbatim>\<open>FINISHED\<close> \<open>task \<oplus> session_stop_result\<close> \\ | |
| 855 | error result: & \<^verbatim>\<open>FAILED\<close> \<open>task \<oplus> error_message \<oplus> session_stop_result\<close> \\[2ex] | |
| 856 |   \end{tabular}
 | |
| 857 | ||
| 858 |   \begin{tabular}{l}
 | |
| 859 |   \<^bold>\<open>type\<close> \<open>session_stop_result = {ok: bool, return_code: int}\<close>
 | |
| 860 |   \end{tabular}
 | |
| 861 | ||
| 862 | \<^medskip> | |
| 67918 | 863 | The \<^verbatim>\<open>session_stop\<close> command forces a shutdown of the identified PIDE | 
| 864 | session. This asynchronous tasks usually finishes quickly. Failure only | |
| 865 | happens in unusual situations, according to the return code of the | |
| 67917 | 866 | underlying Isabelle/ML process. | 
| 867 | \<close> | |
| 868 | ||
| 869 | ||
| 870 | subsubsection \<open>Arguments\<close> | |
| 871 | ||
| 872 | text \<open> | |
| 67921 | 873 | The \<open>session_id\<close> provides the UUID originally created by the server for this | 
| 874 | session. | |
| 67917 | 875 | \<close> | 
| 876 | ||
| 877 | ||
| 878 | subsubsection \<open>Results\<close> | |
| 879 | ||
| 880 | text \<open> | |
| 67918 | 881 | The \<open>ok\<close> field tells abstractly, whether the Isabelle/ML process has | 
| 882 | terminated properly. | |
| 883 | ||
| 884 | The \<open>return_code\<close> field expresses this information according to usual POSIX | |
| 885 | conventions for process exit. | |
| 67917 | 886 | \<close> | 
| 887 | ||
| 888 | ||
| 889 | subsection \<open>Command \<^verbatim>\<open>use_theories\<close> \label{sec:command-use-theories}\<close>
 | |
| 890 | ||
| 891 | text \<open> | |
| 892 |   \begin{tabular}{ll}
 | |
| 893 | argument: & \<open>use_theories_arguments\<close> \\ | |
| 67941 | 894 | immediate result: & \<^verbatim>\<open>OK\<close> \<open>task\<close> \\ | 
| 895 | regular result: & \<^verbatim>\<open>FINISHED\<close> \<open>use_theories_results\<close> \\ | |
| 67917 | 896 |   \end{tabular}
 | 
| 897 | ||
| 898 |   \begin{tabular}{ll}
 | |
| 899 | \<^bold>\<open>type\<close> \<open>use_theories_arguments =\<close> \\ | |
| 900 |   \quad\<open>{session_id: uuid,\<close> \\
 | |
| 67940 
b4e80f062fbf
clarified signature -- eliminated somewhat pointless positions;
 wenzelm parents: 
67938diff
changeset | 901 | \quad~~\<open>theories: [string],\<close> \\ | 
| 67946 | 902 | \quad~~\<open>master_dir?: string,\<close> & \<^bold>\<open>default:\<close> session \<open>tmp_dir\<close> \\ | 
| 68152 
619de043389f
guard result exports via export_pattern -- avoid bombing client via huge blobs;
 wenzelm parents: 
68145diff
changeset | 903 | \quad~~\<open>pretty_margin?: double,\<close> & \<^bold>\<open>default:\<close> \<^verbatim>\<open>76\<close> \\ | 
| 
619de043389f
guard result exports via export_pattern -- avoid bombing client via huge blobs;
 wenzelm parents: 
68145diff
changeset | 904 | \quad~~\<open>unicode_symbols?: bool,\<close> \\ | 
| 68694 
03e104be99af
added check_delay / check_limit for more robust treatment of structurally broken theory sources (or genuine non-termination);
 wenzelm parents: 
68523diff
changeset | 905 | \quad~~\<open>export_pattern?: string,\<close> \\ | 
| 
03e104be99af
added check_delay / check_limit for more robust treatment of structurally broken theory sources (or genuine non-termination);
 wenzelm parents: 
68523diff
changeset | 906 | \quad~~\<open>check_delay?: double,\<close> & \<^bold>\<open>default:\<close> \<^verbatim>\<open>0.5\<close> \\ | 
| 68770 
add44e2b8cb0
optional notification of nodes_status (via progress);
 wenzelm parents: 
68695diff
changeset | 907 | \quad~~\<open>check_limit?: int,\<close> \\ | 
| 68947 | 908 | \quad~~\<open>watchdog_timeout?: double,\<close> & \<^bold>\<open>default:\<close> \<^verbatim>\<open>600.0\<close> \\ | 
| 68770 
add44e2b8cb0
optional notification of nodes_status (via progress);
 wenzelm parents: 
68695diff
changeset | 909 | \quad~~\<open>nodes_status_delay?: double}\<close> & \<^bold>\<open>default:\<close> \<^verbatim>\<open>-1.0\<close> \\ | 
| 67946 | 910 |   \end{tabular}
 | 
| 67917 | 911 | |
| 67946 | 912 |   \begin{tabular}{ll}
 | 
| 68106 | 913 | \<^bold>\<open>type\<close> \<open>export =\<close> \\ | 
| 914 |   \quad~~\<open>{name: string, base64: bool, body: string}\<close> \\
 | |
| 915 | \<^bold>\<open>type\<close> \<open>node_results =\<close> \\ | |
| 916 |   \quad~~\<open>{status: node_status, messages: [message], exports: [export]}\<close> \\
 | |
| 68770 
add44e2b8cb0
optional notification of nodes_status (via progress);
 wenzelm parents: 
68695diff
changeset | 917 | \<^bold>\<open>type\<close> \<open>nodes_status =\<close> \\ | 
| 
add44e2b8cb0
optional notification of nodes_status (via progress);
 wenzelm parents: 
68695diff
changeset | 918 |   \quad~~\<open>[node \<oplus> {status: node_status}]\<close> \\
 | 
| 67917 | 919 | \<^bold>\<open>type\<close> \<open>use_theories_results =\<close> \\ | 
| 920 |   \quad\<open>{ok: bool,\<close> \\
 | |
| 921 | \quad~~\<open>errors: [message],\<close> \\ | |
| 68106 | 922 | \quad~~\<open>nodes: [node \<oplus> node_results]}\<close> \\ | 
| 67917 | 923 |   \end{tabular}
 | 
| 924 | ||
| 925 | \<^medskip> | |
| 67918 | 926 | The \<^verbatim>\<open>use_theories\<close> command updates the identified session by adding the | 
| 927 | current version of theory files to it, while dependencies are resolved | |
| 68883 
3653b3ad729e
clarified Thy_Resources.Session.use_theories: "terminated" node status is sufficient;
 wenzelm parents: 
68882diff
changeset | 928 | implicitly. The command succeeds eventually, when all theories have status | 
| 
3653b3ad729e
clarified Thy_Resources.Session.use_theories: "terminated" node status is sufficient;
 wenzelm parents: 
68882diff
changeset | 929 | \<^emph>\<open>terminated\<close> or \<^emph>\<open>consolidated\<close> in the sense of \<open>node_status\<close> | 
| 
3653b3ad729e
clarified Thy_Resources.Session.use_theories: "terminated" node status is sufficient;
 wenzelm parents: 
68882diff
changeset | 930 |   (\secref{sec:json-types}).
 | 
| 67917 | 931 | |
| 67941 | 932 | Already used theories persist in the session until purged explicitly | 
| 933 |   (\secref{sec:command-purge-theories}). This also means that repeated
 | |
| 934 | invocations of \<^verbatim>\<open>use_theories\<close> are idempotent: it could make sense to do | |
| 935 | that with different values for \<open>pretty_margin\<close> or \<open>unicode_symbols\<close> to get | |
| 936 | different formatting for \<open>errors\<close> or \<open>messages\<close>. | |
| 68106 | 937 | |
| 68152 
619de043389f
guard result exports via export_pattern -- avoid bombing client via huge blobs;
 wenzelm parents: 
68145diff
changeset | 938 | \<^medskip> A non-empty \<open>export_pattern\<close> means that theory \<open>exports\<close> are retrieved | 
| 
619de043389f
guard result exports via export_pattern -- avoid bombing client via huge blobs;
 wenzelm parents: 
68145diff
changeset | 939 |   (see \secref{sec:tool-export}). An \<open>export\<close> \<open>name\<close> roughly follows
 | 
| 
619de043389f
guard result exports via export_pattern -- avoid bombing client via huge blobs;
 wenzelm parents: 
68145diff
changeset | 940 | file-system standards: ``\<^verbatim>\<open>/\<close>'' separated list of base names (excluding | 
| 
619de043389f
guard result exports via export_pattern -- avoid bombing client via huge blobs;
 wenzelm parents: 
68145diff
changeset | 941 | special names like ``\<^verbatim>\<open>.\<close>'' or ``\<^verbatim>\<open>..\<close>''). The \<open>base64\<close> field specifies the | 
| 
619de043389f
guard result exports via export_pattern -- avoid bombing client via huge blobs;
 wenzelm parents: 
68145diff
changeset | 942 | format of the \<open>body\<close> string: it is true for a byte vector that cannot be | 
| 
619de043389f
guard result exports via export_pattern -- avoid bombing client via huge blobs;
 wenzelm parents: 
68145diff
changeset | 943 | represented as plain text in UTF-8 encoding, which means the string needs to | 
| 
619de043389f
guard result exports via export_pattern -- avoid bombing client via huge blobs;
 wenzelm parents: 
68145diff
changeset | 944 | be decoded as in \<^verbatim>\<open>java.util.Base64.getDecoder.decode(String)\<close>. | 
| 68694 
03e104be99af
added check_delay / check_limit for more robust treatment of structurally broken theory sources (or genuine non-termination);
 wenzelm parents: 
68523diff
changeset | 945 | |
| 68883 
3653b3ad729e
clarified Thy_Resources.Session.use_theories: "terminated" node status is sufficient;
 wenzelm parents: 
68882diff
changeset | 946 | \<^medskip> The status of PIDE processing is checked every \<open>check_delay\<close> seconds, and | 
| 
3653b3ad729e
clarified Thy_Resources.Session.use_theories: "terminated" node status is sufficient;
 wenzelm parents: 
68882diff
changeset | 947 | bounded by \<open>check_limit\<close> attempts (default: 0, i.e.\ unbounded). A | 
| 68908 | 948 | \<open>check_limit > 0\<close> effectively specifies a global timeout of \<open>check_delay \<times> | 
| 68883 
3653b3ad729e
clarified Thy_Resources.Session.use_theories: "terminated" node status is sufficient;
 wenzelm parents: 
68882diff
changeset | 949 | check_limit\<close> seconds. | 
| 68770 
add44e2b8cb0
optional notification of nodes_status (via progress);
 wenzelm parents: 
68695diff
changeset | 950 | |
| 68908 | 951 | \<^medskip> If \<open>watchdog_timeout\<close> is greater than 0, it specifies the timespan (in | 
| 952 | seconds) after the last command status change of Isabelle/PIDE, before | |
| 953 | finishing with a potentially non-terminating or deadlocked execution. | |
| 954 | ||
| 68770 
add44e2b8cb0
optional notification of nodes_status (via progress);
 wenzelm parents: 
68695diff
changeset | 955 | \<^medskip> A non-negative \<open>nodes_status_delay\<close> enables continuous notifications of | 
| 
add44e2b8cb0
optional notification of nodes_status (via progress);
 wenzelm parents: 
68695diff
changeset | 956 | kind \<open>nodes_status\<close>, with a field of name and type \<open>nodes_status\<close>. The time | 
| 
add44e2b8cb0
optional notification of nodes_status (via progress);
 wenzelm parents: 
68695diff
changeset | 957 | interval is specified in seconds; by default it is negative and thus | 
| 
add44e2b8cb0
optional notification of nodes_status (via progress);
 wenzelm parents: 
68695diff
changeset | 958 | disabled. | 
| 67917 | 959 | \<close> | 
| 960 | ||
| 961 | ||
| 962 | subsubsection \<open>Arguments\<close> | |
| 963 | ||
| 964 | text \<open> | |
| 965 | The \<open>session_id\<close> is the identifier provided by the server, when the session | |
| 966 | was created (possibly on a different client connection). | |
| 967 | ||
| 968 | \<^medskip> | |
| 67918 | 969 | The \<open>theories\<close> field specifies theory names as in theory \<^theory_text>\<open>imports\<close> or in | 
| 67940 
b4e80f062fbf
clarified signature -- eliminated somewhat pointless positions;
 wenzelm parents: 
67938diff
changeset | 970 | ROOT \<^bold>\<open>theories\<close>. | 
| 67917 | 971 | |
| 67941 | 972 | \<^medskip> | 
| 973 | The \<open>master_dir\<close> field specifies the master directory of imported theories: | |
| 974 | it acts like the ``current working directory'' for locating theory files. | |
| 67946 | 975 | This is irrelevant for \<open>theories\<close> with an absolute path name (e.g.\ | 
| 71925 
bf085daea304
clarified sessions: "Notable Examples in Isabelle/HOL";
 wenzelm parents: 
71521diff
changeset | 976 | \<^verbatim>\<open>"~~/src/HOL/Examples/Seq.thy"\<close>) or session-qualified theory name (e.g.\ | 
| 
bf085daea304
clarified sessions: "Notable Examples in Isabelle/HOL";
 wenzelm parents: 
71521diff
changeset | 977 | \<^verbatim>\<open>"HOL-Examples.Seq"\<close>). | 
| 67917 | 978 | |
| 979 | \<^medskip> | |
| 980 | The \<open>pretty_margin\<close> field specifies the line width for pretty-printing. The | |
| 67918 | 981 | default is suitable for classic console output. Formatting happens at the | 
| 982 | end of \<^verbatim>\<open>use_theories\<close>, when all prover messages are exported to the client. | |
| 67917 | 983 | |
| 984 | \<^medskip> | |
| 67918 | 985 | The \<open>unicode_symbols\<close> field set to \<^verbatim>\<open>true\<close> renders message output for direct | 
| 986 | output on a Unicode capable channel, ideally with the Isabelle fonts as in | |
| 987 | Isabelle/jEdit. The default is to keep the symbolic representation of | |
| 988 | Isabelle text, e.g.\ \<^verbatim>\<open>\<forall>\<close> instead of its rendering as \<open>\<forall>\<close>. This means the | |
| 989 | client needs to perform its own rendering before presenting it to the | |
| 990 | end-user. | |
| 67917 | 991 | \<close> | 
| 992 | ||
| 67918 | 993 | |
| 67917 | 994 | subsubsection \<open>Results\<close> | 
| 995 | ||
| 996 | text \<open> | |
| 997 | The \<open>ok\<close> field indicates overall success of processing the specified | |
| 998 | theories with all their dependencies. | |
| 999 | ||
| 67918 | 1000 | When \<open>ok\<close> is \<^verbatim>\<open>false\<close>, the \<open>errors\<close> field lists all errors cumulatively | 
| 1001 | (including imported theories). The messages contain position information for | |
| 1002 | the original theory nodes. | |
| 67917 | 1003 | |
| 1004 | \<^medskip> | |
| 67918 | 1005 | The \<open>nodes\<close> field provides detailed information about each imported theory | 
| 1006 | node. The individual fields are as follows: | |
| 67917 | 1007 | |
| 67943 | 1008 | \<^item> \<open>node_name\<close>: the canonical name for the theory node, based on its | 
| 67918 | 1009 | file-system location; | 
| 67917 | 1010 | |
| 67943 | 1011 | \<^item> \<open>theory_name\<close>: the logical theory name; | 
| 67917 | 1012 | |
| 1013 | \<^item> \<open>status\<close>: the overall node status, e.g.\ see the visualization in the | |
| 76987 | 1014 | \<open>Theories\<close> panel of Isabelle/jEdit \<^cite>\<open>"isabelle-jedit"\<close>; | 
| 67917 | 1015 | |
| 67918 | 1016 | \<^item> \<open>messages\<close>: the main bulk of prover messages produced in this theory | 
| 67931 
f7917c15b566
field "kind" is always present, with default "writeln";
 wenzelm parents: 
67927diff
changeset | 1017 | (with kind \<^verbatim>\<open>writeln\<close>, \<^verbatim>\<open>warning\<close>, \<^verbatim>\<open>error\<close>). | 
| 67917 | 1018 | \<close> | 
| 1019 | ||
| 1020 | ||
| 1021 | subsubsection \<open>Examples\<close> | |
| 1022 | ||
| 1023 | text \<open> | |
| 1024 | Process some example theory from the Isabelle distribution, within the | |
| 1025 | context of an already started session for Isabelle/HOL (see also | |
| 1026 |   \secref{sec:command-session-start}):
 | |
| 71925 
bf085daea304
clarified sessions: "Notable Examples in Isabelle/HOL";
 wenzelm parents: 
71521diff
changeset | 1027 |   @{verbatim [display] \<open>use_theories {"session_id": ..., "theories": ["~~/src/HOL/Examples/Seq"]}\<close>}
 | 
| 67917 | 1028 | |
| 1029 | \<^medskip> | |
| 1030 | Process some example theories in the context of their (single) parent | |
| 1031 | session: | |
| 1032 | ||
| 1033 |   @{verbatim [display] \<open>session_start {"session": "HOL-Library"}
 | |
| 1034 | use_theories {"session_id": ..., "theories": ["~~/src/HOL/Unix/Unix"]}
 | |
| 1035 | session_stop {"session_id": ...}\<close>}
 | |
| 67922 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 1036 | |
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 1037 | \<^medskip> | 
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 1038 | Process some example theories that import other theories via | 
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 1039 | session-qualified theory names: | 
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 1040 | |
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 1041 |   @{verbatim [display] \<open>session_start {"session": "HOL", "include_sessions": ["HOL-Unix"]}
 | 
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 1042 | use_theories {"session_id": ..., "theories": ["HOL-Unix.Unix"]}
 | 
| 
9e668ae81f97
clarified signature: prefer selective include_sessions;
 wenzelm parents: 
67921diff
changeset | 1043 | session_stop {"session_id": ...}\<close>}
 | 
| 67904 | 1044 | \<close> | 
| 1045 | ||
| 67941 | 1046 | |
| 1047 | subsection \<open>Command \<^verbatim>\<open>purge_theories\<close> \label{sec:command-purge-theories}\<close>
 | |
| 1048 | ||
| 1049 | text \<open> | |
| 1050 |   \begin{tabular}{ll}
 | |
| 1051 | argument: & \<open>purge_theories_arguments\<close> \\ | |
| 1052 | regular result: & \<^verbatim>\<open>OK\<close> \<open>purge_theories_result\<close> \\ | |
| 1053 |   \end{tabular}
 | |
| 1054 | ||
| 67946 | 1055 |   \begin{tabular}{lll}
 | 
| 67941 | 1056 | \<^bold>\<open>type\<close> \<open>purge_theories_arguments =\<close> \\ | 
| 1057 |   \quad\<open>{session_id: uuid,\<close> \\
 | |
| 1058 | \quad~~\<open>theories: [string],\<close> \\ | |
| 67946 | 1059 | \quad~~\<open>master_dir?: string,\<close> & \<^bold>\<open>default:\<close> session \<open>tmp_dir\<close> \\ | 
| 67941 | 1060 | \quad~~\<open>all?: bool}\<close> \\[2ex] | 
| 1061 |   \end{tabular}
 | |
| 1062 | ||
| 1063 |   \begin{tabular}{ll}
 | |
| 1064 |   \<^bold>\<open>type\<close> \<open>purge_theories_result = {purged: [string]}\<close> \\
 | |
| 1065 |   \end{tabular}
 | |
| 1066 | ||
| 1067 | \<^medskip> | |
| 1068 | The \<^verbatim>\<open>purge_theories\<close> command updates the identified session by removing | |
| 1069 | theories that are no longer required: theories that are used in pending | |
| 1070 | \<^verbatim>\<open>use_theories\<close> tasks or imported by other theories are retained. | |
| 1071 | \<close> | |
| 1072 | ||
| 1073 | ||
| 1074 | subsubsection \<open>Arguments\<close> | |
| 1075 | ||
| 1076 | text \<open> | |
| 1077 | The \<open>session_id\<close> is the identifier provided by the server, when the session | |
| 1078 | was created (possibly on a different client connection). | |
| 1079 | ||
| 1080 | \<^medskip> | |
| 1081 | The \<open>theories\<close> field specifies theory names to be purged: imported | |
| 1082 | dependencies are \<^emph>\<open>not\<close> completed. Instead it is possible to provide the | |
| 1083 | already completed import graph returned by \<^verbatim>\<open>use_theories\<close> as \<open>nodes\<close> / | |
| 1084 | \<open>node_name\<close>. | |
| 1085 | ||
| 1086 | \<^medskip> | |
| 67943 | 1087 | The \<open>master_dir\<close> field specifies the master directory as in \<^verbatim>\<open>use_theories\<close>. | 
| 67946 | 1088 | This is irrelevant, when passing fully-qualified theory node names (e.g.\ | 
| 67943 | 1089 | \<open>node_name\<close> from \<open>nodes\<close> in \<open>use_theories_results\<close>). | 
| 67941 | 1090 | |
| 1091 | \<^medskip> | |
| 1092 | The \<open>all\<close> field set to \<^verbatim>\<open>true\<close> attempts to purge all presently loaded | |
| 1093 | theories. | |
| 1094 | \<close> | |
| 1095 | ||
| 1096 | ||
| 1097 | subsubsection \<open>Results\<close> | |
| 1098 | ||
| 1099 | text \<open> | |
| 67943 | 1100 | The \<open>purged\<close> field gives the theory nodes that were actually removed. | 
| 1101 | ||
| 1102 | \<^medskip> | |
| 1103 | The \<open>retained\<close> field gives the remaining theory nodes, i.e.\ the complement | |
| 1104 | of \<open>purged\<close>. | |
| 67941 | 1105 | \<close> | 
| 1106 | ||
| 67904 | 1107 | end |