| author | wenzelm | 
| Tue, 25 Aug 2020 14:54:41 +0200 | |
| changeset 72205 | bc71db05abe3 | 
| parent 71579 | 9b49538845cc | 
| child 72309 | 564012e31db1 | 
| permissions | -rw-r--r-- | 
| 61656 | 1 | (*:maxLineLen=78:*) | 
| 61575 | 2 | |
| 28224 | 3 | theory Misc | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
41512diff
changeset | 4 | imports Base | 
| 28224 | 5 | begin | 
| 6 | ||
| 58618 | 7 | chapter \<open>Miscellaneous tools \label{ch:tools}\<close>
 | 
| 28224 | 8 | |
| 58618 | 9 | text \<open> | 
| 61575 | 10 | Subsequently we describe various Isabelle related utilities, given in | 
| 11 | alphabetical order. | |
| 58618 | 12 | \<close> | 
| 28224 | 13 | |
| 71579 | 14 | section \<open>Building Isabelle docker images\<close> | 
| 15 | ||
| 16 | text \<open> | |
| 17 | Docker\<^footnote>\<open>\<^url>\<open>https://docs.docker.com\<close>\<close> provides a self-contained environment | |
| 18 | for complex applications called \<^emph>\<open>container\<close>, although it does not fully | |
| 19 | contain the program in a strict sense of the word. This includes basic | |
| 20 | operating system services (usually based on Linux), shared libraries and | |
| 21 | other required packages. Thus Docker is a light-weight alternative to | |
| 22 | regular virtual machines, or a heavy-weight alternative to conventional | |
| 23 | self-contained applications. | |
| 24 | ||
| 25 | Although Isabelle can be easily run on a variety of OS environments without | |
| 26 | extra containment, Docker images may occasionally be useful when a | |
| 27 | standardized Linux environment is required, even on | |
| 28 | Windows\<^footnote>\<open>\<^url>\<open>https://docs.docker.com/docker-for-windows\<close>\<close> and | |
| 29 | macOS\<^footnote>\<open>\<^url>\<open>https://docs.docker.com/docker-for-mac\<close>\<close>. Further uses are in | |
| 30 | common cloud computing environments, where applications need to be submitted | |
| 31 | as Docker images in the first place. | |
| 32 | ||
| 33 | \<^medskip> | |
| 34 |   The @{tool_def build_docker} tool builds docker images from a standard
 | |
| 35 | Isabelle application archive for Linux: | |
| 36 | ||
| 37 |   @{verbatim [display]
 | |
| 38 | \<open>Usage: isabelle build_docker [OPTIONS] APP_ARCHIVE | |
| 39 | ||
| 40 | Options are: | |
| 41 | -B NAME base image (default "ubuntu") | |
| 42 | -E set bin/isabelle as entrypoint | |
| 43 |     -P NAME      additional Ubuntu package collection ("X11", "latex")
 | |
| 44 | -l NAME default logic (default ISABELLE_LOGIC="HOL") | |
| 45 | -n no docker build | |
| 46 | -o FILE output generated Dockerfile | |
| 47 | -p NAME additional Ubuntu package | |
| 48 | -t TAG docker build tag | |
| 49 | -v verbose | |
| 50 | ||
| 51 | Build Isabelle docker image with default logic image, using a standard | |
| 52 | Isabelle application archive for Linux (local file or remote URL).\<close>} | |
| 53 | ||
| 54 | Option \<^verbatim>\<open>-E\<close> sets \<^verbatim>\<open>bin/isabelle\<close> of the contained Isabelle distribution as | |
| 55 | the standard entry point of the Docker image. Thus \<^verbatim>\<open>docker run\<close> will | |
| 56 |   imitate the \<^verbatim>\<open>isabelle\<close> command-line tool (\secref{sec:isabelle-tool}) of a
 | |
| 57 | regular local installation, but it lacks proper GUI support: \<^verbatim>\<open>isabelle jedit\<close> | |
| 58 | will not work without further provisions. Note that the default entrypoint | |
| 59 | may be changed later via \<^verbatim>\<open>docker run --entrypoint="..."\<close>. | |
| 60 | ||
| 61 | Option \<^verbatim>\<open>-t\<close> specifies the Docker image tag: this a symbolic name within the | |
| 62 | local Docker name space, but also relevant for Docker | |
| 63 | Hub\<^footnote>\<open>\<^url>\<open>https://hub.docker.com\<close>\<close>. | |
| 64 | ||
| 65 | Option \<^verbatim>\<open>-l\<close> specifies the default logic image of the Isabelle distribution | |
| 66 | contained in the Docker environment: it will be produced by \<^verbatim>\<open>isabelle build -b\<close> | |
| 67 |   as usual (\secref{sec:tool-build}) and stored within the image.
 | |
| 68 | ||
| 69 | \<^medskip> | |
| 70 | Option \<^verbatim>\<open>-B\<close> specifies the Docker image taken as starting point for the | |
| 71 | Isabelle installation: it needs to be a suitable version of Ubuntu Linux. | |
| 72 | The default \<^verbatim>\<open>ubuntu\<close> refers to the latest LTS version provided by Canonical | |
| 73 | as the official Ubuntu vendor\<^footnote>\<open>\<^url>\<open>https://hub.docker.com/_/ubuntu\<close>\<close>. For | |
| 74 | Isabelle2020 this should be Ubuntu 18.04 LTS. | |
| 75 | ||
| 76 | Option \<^verbatim>\<open>-p\<close> includes additional Ubuntu packages, using the terminology | |
| 77 | of \<^verbatim>\<open>apt-get install\<close> within the underlying Linux distribution. | |
| 78 | ||
| 79 | Option \<^verbatim>\<open>-P\<close> refers to high-level package collections: \<^verbatim>\<open>X11\<close> or \<^verbatim>\<open>latex\<close> as | |
| 80 | provided by \<^verbatim>\<open>isabelle build_docker\<close> (assuming Ubuntu 18.04 LTS). This | |
| 81 | imposes extra weight on the resulting Docker images. Note that \<^verbatim>\<open>X11\<close> will | |
| 82 | only provide remote X11 support according to the modest GUI quality | |
| 83 | standards of the late 1990-ies. | |
| 84 | ||
| 85 | \<^medskip> | |
| 86 | Option \<^verbatim>\<open>-n\<close> suppresses the actual \<^verbatim>\<open>docker build\<close> process. Option \<^verbatim>\<open>-o\<close> | |
| 87 | outputs the generated \<^verbatim>\<open>Dockerfile\<close>. Both options together produce a | |
| 88 | Dockerfile only, which might be useful for informative purposes or other | |
| 89 | tools. | |
| 90 | ||
| 91 | Option \<^verbatim>\<open>-v\<close> disables quiet-mode of the underlying \<^verbatim>\<open>docker build\<close> process. | |
| 92 | \<close> | |
| 93 | ||
| 94 | ||
| 95 | subsubsection \<open>Examples\<close> | |
| 96 | ||
| 97 | text \<open> | |
| 98 | Produce a Dockerfile (without image) from a remote Isabelle distribution: | |
| 99 |   @{verbatim [display]
 | |
| 100 | \<open> isabelle build_docker -E -n -o Dockerfile | |
| 101 | https://isabelle.in.tum.de/website-Isabelle2020/dist/Isabelle2020_linux.tar.gz\<close>} | |
| 102 | ||
| 103 | Build a standard Isabelle Docker image from a local Isabelle distribution, | |
| 104 | with \<^verbatim>\<open>bin/isabelle\<close> as executable entry point: | |
| 105 | ||
| 106 |   @{verbatim [display]
 | |
| 107 | \<open> isabelle build_docker -E -t test/isabelle:Isabelle2020 Isabelle2020_linux.tar.gz\<close>} | |
| 108 | ||
| 109 | Invoke the raw Isabelle/ML process within that image: | |
| 110 |   @{verbatim [display]
 | |
| 111 | \<open> docker run test/isabelle:Isabelle2020 process -e "Session.welcome ()"\<close>} | |
| 112 | ||
| 113 | Invoke a Linux command-line tool within the contained Isabelle system | |
| 114 | environment: | |
| 115 |   @{verbatim [display]
 | |
| 116 | \<open> docker run test/isabelle:Isabelle2020 env uname -a\<close>} | |
| 117 | The latter should always report a Linux operating system, even when running | |
| 118 | on Windows or macOS. | |
| 119 | \<close> | |
| 120 | ||
| 28224 | 121 | |
| 58618 | 122 | section \<open>Resolving Isabelle components \label{sec:tool-components}\<close>
 | 
| 48844 | 123 | |
| 58618 | 124 | text \<open> | 
| 48844 | 125 |   The @{tool_def components} tool resolves Isabelle components:
 | 
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 126 |   @{verbatim [display]
 | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 127 | \<open>Usage: isabelle components [OPTIONS] [COMPONENTS ...] | 
| 48844 | 128 | |
| 129 | Options are: | |
| 50653 
5c85f8b80b95
simplified quick start via "isabelle components -I";
 wenzelm parents: 
50132diff
changeset | 130 | -I init user settings | 
| 48844 | 131 | -R URL component repository | 
| 132 | (default $ISABELLE_COMPONENT_REPOSITORY) | |
| 53435 | 133 | -a resolve all missing components | 
| 48844 | 134 | -l list status | 
| 135 | ||
| 136 | Resolve Isabelle components via download and installation. | |
| 137 | COMPONENTS are identified via base name. | |
| 138 | ||
| 68224 | 139 | ISABELLE_COMPONENT_REPOSITORY="https://isabelle.in.tum.de/components"\<close>} | 
| 48844 | 140 | |
| 61575 | 141 |   Components are initialized as described in \secref{sec:components} in a
 | 
| 142 | permissive manner, which can mark components as ``missing''. This state is | |
| 143 |   amended by letting @{tool "components"} download and unpack components that
 | |
| 63680 | 144 | are published on the default component repository | 
| 68224 | 145 | \<^url>\<open>https://isabelle.in.tum.de/components\<close> in particular. | 
| 48844 | 146 | |
| 61575 | 147 | Option \<^verbatim>\<open>-R\<close> specifies an alternative component repository. Note that | 
| 148 | \<^verbatim>\<open>file:///\<close> URLs can be used for local directories. | |
| 48844 | 149 | |
| 61575 | 150 | Option \<^verbatim>\<open>-a\<close> selects all missing components to be resolved. Explicit | 
| 151 | components may be named as command line-arguments as well. Note that | |
| 152 | components are uniquely identified by their base name, while the | |
| 153 | installation takes place in the location that was specified in the attempt | |
| 154 | to initialize the component before. | |
| 48844 | 155 | |
| 61575 | 156 | Option \<^verbatim>\<open>-l\<close> lists the current state of available and missing components | 
| 157 | with their location (full name) within the file-system. | |
| 50653 
5c85f8b80b95
simplified quick start via "isabelle components -I";
 wenzelm parents: 
50132diff
changeset | 158 | |
| 61575 | 159 | Option \<^verbatim>\<open>-I\<close> initializes the user settings file to subscribe to the standard | 
| 160 | components specified in the Isabelle repository clone --- this does not make | |
| 161 | any sense for regular Isabelle releases. If the file already exists, it | |
| 162 | needs to be edited manually according to the printed explanation. | |
| 58618 | 163 | \<close> | 
| 48844 | 164 | |
| 165 | ||
| 58618 | 166 | section \<open>Displaying documents \label{sec:tool-display}\<close>
 | 
| 28224 | 167 | |
| 61575 | 168 | text \<open> | 
| 169 |   The @{tool_def display} tool displays documents in DVI or PDF format:
 | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 170 |   @{verbatim [display]
 | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 171 | \<open>Usage: isabelle display DOCUMENT | 
| 28224 | 172 | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 173 | Display DOCUMENT (in DVI or PDF format).\<close>} | 
| 28224 | 174 | |
| 61406 | 175 | \<^medskip> | 
| 61575 | 176 |   The settings @{setting DVI_VIEWER} and @{setting PDF_VIEWER} determine the
 | 
| 177 | programs for viewing the corresponding file formats. Normally this opens the | |
| 178 | document via the desktop environment, potentially in an asynchronous manner | |
| 179 | with re-use of previews views. | |
| 58618 | 180 | \<close> | 
| 28224 | 181 | |
| 182 | ||
| 58618 | 183 | section \<open>Viewing documentation \label{sec:tool-doc}\<close>
 | 
| 28224 | 184 | |
| 58618 | 185 | text \<open> | 
| 52444 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 wenzelm parents: 
52052diff
changeset | 186 |   The @{tool_def doc} tool displays Isabelle documentation:
 | 
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 187 |   @{verbatim [display]
 | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 188 | \<open>Usage: isabelle doc [DOC ...] | 
| 28224 | 189 | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 190 | View Isabelle documentation.\<close>} | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 191 | |
| 61575 | 192 | If called without arguments, it lists all available documents. Each line | 
| 193 | starts with an identifier, followed by a short description. Any of these | |
| 194 | identifiers may be specified as arguments, in order to display the | |
| 195 |   corresponding document (see also \secref{sec:tool-display}).
 | |
| 28224 | 196 | |
| 61406 | 197 | \<^medskip> | 
| 61575 | 198 |   The @{setting ISABELLE_DOCS} setting specifies the list of directories
 | 
| 199 | (separated by colons) to be scanned for documentations. | |
| 58618 | 200 | \<close> | 
| 28224 | 201 | |
| 202 | ||
| 58618 | 203 | section \<open>Shell commands within the settings environment \label{sec:tool-env}\<close>
 | 
| 47828 | 204 | |
| 61575 | 205 | text \<open> | 
| 206 |   The @{tool_def env} tool is a direct wrapper for the standard
 | |
| 207 | \<^verbatim>\<open>/usr/bin/env\<close> command on POSIX systems, running within the Isabelle | |
| 208 |   settings environment (\secref{sec:settings}).
 | |
| 47828 | 209 | |
| 61575 | 210 | The command-line arguments are that of the underlying version of \<^verbatim>\<open>env\<close>. For | 
| 211 | example, the following invokes an instance of the GNU Bash shell within the | |
| 212 | Isabelle environment: | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 213 |   @{verbatim [display] \<open>isabelle env bash\<close>}
 | 
| 58618 | 214 | \<close> | 
| 47828 | 215 | |
| 216 | ||
| 58618 | 217 | section \<open>Inspecting the settings environment \label{sec:tool-getenv}\<close>
 | 
| 28224 | 218 | |
| 58618 | 219 | text \<open>The Isabelle settings environment --- as provided by the | 
| 28224 | 220 | site-default and user-specific settings files --- can be inspected | 
| 48602 | 221 |   with the @{tool_def getenv} tool:
 | 
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 222 |   @{verbatim [display]
 | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 223 | \<open>Usage: isabelle getenv [OPTIONS] [VARNAMES ...] | 
| 28224 | 224 | |
| 225 | Options are: | |
| 226 | -a display complete environment | |
| 227 | -b print values only (doesn't work for -a) | |
| 31497 | 228 | -d FILE dump complete environment to FILE | 
| 229 | (null terminated entries) | |
| 28224 | 230 | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 231 | Get value of VARNAMES from the Isabelle settings.\<close>} | 
| 28224 | 232 | |
| 61575 | 233 | With the \<^verbatim>\<open>-a\<close> option, one may inspect the full process environment that | 
| 234 | Isabelle related programs are run in. This usually contains much more | |
| 235 | variables than are actually Isabelle settings. Normally, output is a list of | |
| 236 | lines of the form \<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close>. The \<^verbatim>\<open>-b\<close> option causes only the values | |
| 237 | to be printed. | |
| 31497 | 238 | |
| 61575 | 239 | Option \<^verbatim>\<open>-d\<close> produces a dump of the complete environment to the specified | 
| 240 | file. Entries are terminated by the ASCII null character, i.e.\ the C string | |
| 241 | terminator. | |
| 58618 | 242 | \<close> | 
| 28224 | 243 | |
| 244 | ||
| 58618 | 245 | subsubsection \<open>Examples\<close> | 
| 28224 | 246 | |
| 61575 | 247 | text \<open> | 
| 248 |   Get the location of @{setting ISABELLE_HOME_USER} where user-specific
 | |
| 249 | information is stored: | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 250 |   @{verbatim [display] \<open>isabelle getenv ISABELLE_HOME_USER\<close>}
 | 
| 28224 | 251 | |
| 61406 | 252 | \<^medskip> | 
| 61575 | 253 | Get the value only of the same settings variable, which is particularly | 
| 254 | useful in shell scripts: | |
| 68219 | 255 |   @{verbatim [display] \<open>isabelle getenv -b ISABELLE_HOME_USER\<close>}
 | 
| 58618 | 256 | \<close> | 
| 28224 | 257 | |
| 258 | ||
| 71322 | 259 | section \<open>Mercurial repository setup \label{sec:hg-setup}\<close>
 | 
| 260 | ||
| 261 | text \<open> | |
| 262 |   The @{tool_def hg_setup} tool simplifies the setup of Mercurial
 | |
| 71325 | 263 |   repositories, with hosting via Phabricator (\chref{ch:phabricator}) or SSH
 | 
| 264 | file server access. | |
| 71322 | 265 | |
| 266 |   @{verbatim [display]
 | |
| 267 | \<open>Usage: isabelle hg_setup [OPTIONS] REMOTE LOCAL_DIR | |
| 268 | ||
| 269 | Options are: | |
| 270 | -n NAME remote repository name (default: base name of LOCAL_DIR) | |
| 271 | -p PATH Mercurial path name (default: "default") | |
| 272 | -r assume that remote repository already exists | |
| 273 | ||
| 274 | Setup a remote vs. local Mercurial repository: REMOTE either refers to a | |
| 275 | Phabricator server "user@host" or SSH file server "ssh://user@host/path".\<close>} | |
| 276 | ||
| 277 | The \<^verbatim>\<open>REMOTE\<close> repository specification \<^emph>\<open>excludes\<close> the actual repository | |
| 278 | name: that is given by the base name of \<^verbatim>\<open>LOCAL_DIR\<close>, or via option \<^verbatim>\<open>-n\<close>. | |
| 279 | ||
| 280 | By default, both sides of the repository are created on demand by default. | |
| 281 | In contrast, option \<^verbatim>\<open>-r\<close> assumes that the remote repository already exists: | |
| 282 | it avoids accidental creation of a persistent repository with unintended | |
| 283 | name. | |
| 284 | ||
| 285 | The local \<^verbatim>\<open>.hg/hgrc\<close> file is changed to refer to the remote repository, | |
| 286 | usually via the symbolic path name ``\<^verbatim>\<open>default\<close>''; option \<^verbatim>\<open>-p\<close> allows to | |
| 287 | provided a different name. | |
| 288 | \<close> | |
| 289 | ||
| 71324 | 290 | subsubsection \<open>Examples\<close> | 
| 291 | ||
| 292 | text \<open> | |
| 293 | Setup the current directory as a repository with Phabricator server hosting: | |
| 294 |   @{verbatim [display] \<open>  isabelle hg_setup vcs@vcs.example.org .\<close>}
 | |
| 295 | ||
| 296 | \<^medskip> | |
| 297 | Setup the current directory as a repository with plain SSH server hosting: | |
| 298 |   @{verbatim [display] \<open>  isabelle hg_setup ssh://files.example.org/data/repositories .\<close>}
 | |
| 299 | ||
| 300 | \<^medskip> | |
| 301 | Both variants require SSH access to the target server, via public key | |
| 302 | without password. | |
| 303 | \<close> | |
| 304 | ||
| 71322 | 305 | |
| 58618 | 306 | section \<open>Installing standalone Isabelle executables \label{sec:tool-install}\<close>
 | 
| 28224 | 307 | |
| 61575 | 308 | text \<open> | 
| 309 |   By default, the main Isabelle binaries (@{executable "isabelle"} etc.) are
 | |
| 310 | just run from their location within the distribution directory, probably | |
| 311 |   indirectly by the shell through its @{setting PATH}. Other schemes of
 | |
| 312 |   installation are supported by the @{tool_def install} tool:
 | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 313 |   @{verbatim [display]
 | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 314 | \<open>Usage: isabelle install [OPTIONS] BINDIR | 
| 28224 | 315 | |
| 316 | Options are: | |
| 50132 | 317 | -d DISTDIR refer to DISTDIR as Isabelle distribution | 
| 28224 | 318 | (default ISABELLE_HOME) | 
| 319 | ||
| 50132 | 320 | Install Isabelle executables with absolute references to the | 
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 321 | distribution directory.\<close>} | 
| 28224 | 322 | |
| 61575 | 323 | The \<^verbatim>\<open>-d\<close> option overrides the current Isabelle distribution directory as | 
| 324 |   determined by @{setting ISABELLE_HOME}.
 | |
| 28224 | 325 | |
| 61575 | 326 | The \<open>BINDIR\<close> argument tells where executable wrapper scripts for | 
| 62588 
cd266473b81b
isabelle_process is superseded by "isabelle process" tool;
 wenzelm parents: 
62562diff
changeset | 327 |   @{executable "isabelle"} and @{executable isabelle_scala_script} should be
 | 
| 61575 | 328 |   placed, which is typically a directory in the shell's @{setting PATH}, such
 | 
| 329 | as \<^verbatim>\<open>$HOME/bin\<close>. | |
| 48815 | 330 | |
| 61406 | 331 | \<^medskip> | 
| 61575 | 332 | It is also possible to make symbolic links of the main Isabelle executables | 
| 333 | manually, but making separate copies outside the Isabelle distribution | |
| 334 | directory will not work! | |
| 335 | \<close> | |
| 28224 | 336 | |
| 337 | ||
| 58618 | 338 | section \<open>Creating instances of the Isabelle logo\<close> | 
| 28224 | 339 | |
| 61575 | 340 | text \<open> | 
| 341 |   The @{tool_def logo} tool creates instances of the generic Isabelle logo as
 | |
| 342 |   EPS and PDF, for inclusion in {\LaTeX} documents.
 | |
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 343 |   @{verbatim [display]
 | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 344 | \<open>Usage: isabelle logo [OPTIONS] XYZ | 
| 28224 | 345 | |
| 49072 
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
 wenzelm parents: 
48985diff
changeset | 346 | Create instance XYZ of the Isabelle logo (as EPS and PDF). | 
| 28224 | 347 | |
| 348 | Options are: | |
| 49072 
747835eb2782
"isabelle logo" produces EPS and PDF format simultaneously;
 wenzelm parents: 
48985diff
changeset | 349 | -n NAME alternative output base name (default "isabelle_xyx") | 
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 350 | -q quiet mode\<close>} | 
| 48936 | 351 | |
| 61575 | 352 | Option \<^verbatim>\<open>-n\<close> specifies an alternative (base) name for the generated files. | 
| 353 | The default is \<^verbatim>\<open>isabelle_\<close>\<open>xyz\<close> in lower-case. | |
| 48936 | 354 | |
| 61503 | 355 | Option \<^verbatim>\<open>-q\<close> omits printing of the result file name. | 
| 48936 | 356 | |
| 61406 | 357 | \<^medskip> | 
| 61575 | 358 | Implementors of Isabelle tools and applications are encouraged to make | 
| 359 | derived Isabelle logos for their own projects using this template. | |
| 360 | \<close> | |
| 28224 | 361 | |
| 362 | ||
| 58618 | 363 | section \<open>Output the version identifier of the Isabelle distribution\<close> | 
| 28224 | 364 | |
| 58618 | 365 | text \<open> | 
| 48602 | 366 |   The @{tool_def version} tool displays Isabelle version information:
 | 
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 367 |   @{verbatim [display]
 | 
| 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 368 | \<open>Usage: isabelle version [OPTIONS] | 
| 41511 | 369 | |
| 370 | Options are: | |
| 371 | -i short identification (derived from Mercurial id) | |
| 372 | ||
| 61407 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61406diff
changeset | 373 | Display Isabelle version information.\<close>} | 
| 41511 | 374 | |
| 61406 | 375 | \<^medskip> | 
| 61575 | 376 | The default is to output the full version string of the Isabelle | 
| 71578 | 377 | distribution, e.g.\ ``\<^verbatim>\<open>Isabelle2020: April 2020\<close>. | 
| 41511 | 378 | |
| 61575 | 379 | The \<^verbatim>\<open>-i\<close> option produces a short identification derived from the Mercurial | 
| 380 |   id of the @{setting ISABELLE_HOME} directory.
 | |
| 58618 | 381 | \<close> | 
| 28224 | 382 | |
| 67399 | 383 | end |