src/Doc/System/Misc.thy
changeset 57439 0e41f26a0250
parent 57414 fe1be2844fda
child 57443 577f029fde39
equal deleted inserted replaced
57438:663037c5d848 57439:0e41f26a0250
   228   Isabelle releases.  If the file already exists, it needs to be
   228   Isabelle releases.  If the file already exists, it needs to be
   229   edited manually according to the printed explanation.
   229   edited manually according to the printed explanation.
   230 *}
   230 *}
   231 
   231 
   232 
   232 
       
   233 section {* Raw ML console *}
       
   234 
       
   235 text {*
       
   236   The @{tool_def console} tool runs the Isabelle process with raw ML console:
       
   237 \begin{ttbox}
       
   238 Usage: isabelle console [OPTIONS]
       
   239 
       
   240   Options are:
       
   241     -d DIR       include session directory
       
   242     -l NAME      logic session name (default ISABELLE_LOGIC)
       
   243     -m MODE      add print mode for output
       
   244     -n           no build of session image on startup
       
   245     -o OPTION    override Isabelle system OPTION (via NAME=VAL or NAME)
       
   246     -s           system build mode for session image
       
   247 
       
   248   Run Isabelle process with raw ML console and line editor
       
   249   (default ISABELLE_LINE_EDITOR).
       
   250 \end{ttbox}
       
   251 
       
   252   The @{verbatim "-l"} option specifies the logic session name. By default,
       
   253   its heap image is checked and built on demand, but the option @{verbatim
       
   254   "-n"} skips that.
       
   255 
       
   256   Options @{verbatim "-d"}, @{verbatim "-o"}, @{verbatim "-s"} are passed
       
   257   directly to @{tool build} (\secref{sec:tool-build}).
       
   258 
       
   259   Options @{verbatim "-m"}, @{verbatim "-o"} are passed directly to the
       
   260   underlying Isabelle process (\secref{sec:isabelle-process}).
       
   261 
       
   262   The Isabelle process is run through the line editor that is specified via
       
   263   the settings variable @{setting ISABELLE_LINE_EDITOR} (e.g.\
       
   264   @{executable_def rlwrap} for GNU readline); the fall-back is to use plain
       
   265   standard input/output.
       
   266 
       
   267   Interaction works via the raw ML toplevel loop: this is neither
       
   268   Isabelle/Isar nor Isabelle/ML within the usual formal context. Some useful
       
   269   ML commands at this stage are @{ML cd}, @{ML pwd}, @{ML use}, @{ML use_thy},
       
   270   @{ML use_thys}.
       
   271 *}
       
   272 
       
   273 
   233 section {* Displaying documents \label{sec:tool-display} *}
   274 section {* Displaying documents \label{sec:tool-display} *}
   234 
   275 
   235 text {* The @{tool_def display} tool displays documents in DVI or PDF
   276 text {* The @{tool_def display} tool displays documents in DVI or PDF
   236   format:
   277   format:
   237 \begin{ttbox}
   278 \begin{ttbox}
   418   \medskip Implementors of Isabelle tools and applications are
   459   \medskip Implementors of Isabelle tools and applications are
   419   encouraged to make derived Isabelle logos for their own projects
   460   encouraged to make derived Isabelle logos for their own projects
   420   using this template.  *}
   461   using this template.  *}
   421 
   462 
   422 
   463 
   423 section {* Plain TTY interaction \label{sec:tool-tty} *}
       
   424 
       
   425 text {*
       
   426   The @{tool_def tty} tool runs the Isabelle process interactively
       
   427   within a plain terminal session:
       
   428 \begin{ttbox}
       
   429 Usage: isabelle tty [OPTIONS]
       
   430 
       
   431   Options are:
       
   432     -l NAME      logic image name (default ISABELLE_LOGIC)
       
   433     -m MODE      add print mode for output
       
   434     -o OPTION    override Isabelle system OPTION (via NAME=VAL or NAME)
       
   435     -p NAME      line editor program name (default ISABELLE_LINE_EDITOR)
       
   436 
       
   437   Run Isabelle process with plain tty interaction and line editor.
       
   438 \end{ttbox}
       
   439 
       
   440   The @{verbatim "-l"} option specifies the logic image.  The
       
   441   @{verbatim "-m"} option specifies additional print modes.  The
       
   442   @{verbatim "-p"} option specifies an alternative line editor (such
       
   443   as the @{executable_def rlwrap} wrapper for GNU readline); the
       
   444   fall-back is to use raw standard input.
       
   445 
       
   446   \medskip Option @{verbatim "-o"} allows to override Isabelle system
       
   447   options for this process, see also \secref{sec:system-options}.
       
   448 
       
   449   Regular interaction works via the standard Isabelle/Isar toplevel
       
   450   loop.  The Isar command @{command exit} drops out into the
       
   451   bare-bones ML system, which is occasionally useful for debugging of
       
   452   the Isar infrastructure itself.  Invoking @{ML Isar.loop}~@{verbatim
       
   453   "();"} in ML will return to the Isar toplevel.  *}
       
   454 
       
   455 
       
   456 section {* Output the version identifier of the Isabelle distribution *}
   464 section {* Output the version identifier of the Isabelle distribution *}
   457 
   465 
   458 text {*
   466 text {*
   459   The @{tool_def version} tool displays Isabelle version information:
   467   The @{tool_def version} tool displays Isabelle version information:
   460 \begin{ttbox}
   468 \begin{ttbox}