| author | wenzelm | 
| Wed, 27 Dec 2023 13:17:55 +0100 | |
| changeset 79368 | a2d8ecb13d3f | 
| parent 79059 | ae682b2aab03 | 
| child 79496 | 4d82b743c5e1 | 
| 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: 
41512 
diff
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>  | 
|
| 
77567
 
b975f5aaf6b8
renamed "isabelle build_docker" to "isabelle docker_build" (unrelated to "isabelle build");
 
wenzelm 
parents: 
76184 
diff
changeset
 | 
34  | 
  The @{tool_def docker_build} tool builds docker images from a standard
 | 
| 71579 | 35  | 
Isabelle application archive for Linux:  | 
36  | 
||
37  | 
  @{verbatim [display]
 | 
|
| 
77567
 
b975f5aaf6b8
renamed "isabelle build_docker" to "isabelle docker_build" (unrelated to "isabelle build");
 
wenzelm 
parents: 
76184 
diff
changeset
 | 
38  | 
\<open>Usage: isabelle docker_build [OPTIONS] APP_ARCHIVE  | 
| 71579 | 39  | 
|
40  | 
Options are:  | 
|
| 76184 | 41  | 
-B NAME base image (default "ubuntu:22.04")  | 
| 72525 | 42  | 
-E set Isabelle/bin/isabelle as entrypoint  | 
| 71579 | 43  | 
    -P NAME      additional Ubuntu package collection ("X11", "latex")
 | 
| 
76178
 
1f95e9424341
more robust: snap version of docker cannot access /tmp;
 
wenzelm 
parents: 
76136 
diff
changeset
 | 
44  | 
-W DIR working directory that is accessible to docker,  | 
| 
 
1f95e9424341
more robust: snap version of docker cannot access /tmp;
 
wenzelm 
parents: 
76136 
diff
changeset
 | 
45  | 
potentially via snap (default: ".")  | 
| 71579 | 46  | 
-l NAME default logic (default ISABELLE_LOGIC="HOL")  | 
47  | 
-n no docker build  | 
|
48  | 
-o FILE output generated Dockerfile  | 
|
49  | 
-p NAME additional Ubuntu package  | 
|
50  | 
-t TAG docker build tag  | 
|
51  | 
-v verbose  | 
|
52  | 
||
53  | 
Build Isabelle docker image with default logic image, using a standard  | 
|
54  | 
Isabelle application archive for Linux (local file or remote URL).\<close>}  | 
|
55  | 
||
56  | 
Option \<^verbatim>\<open>-E\<close> sets \<^verbatim>\<open>bin/isabelle\<close> of the contained Isabelle distribution as  | 
|
57  | 
the standard entry point of the Docker image. Thus \<^verbatim>\<open>docker run\<close> will  | 
|
58  | 
  imitate the \<^verbatim>\<open>isabelle\<close> command-line tool (\secref{sec:isabelle-tool}) of a
 | 
|
59  | 
regular local installation, but it lacks proper GUI support: \<^verbatim>\<open>isabelle jedit\<close>  | 
|
60  | 
will not work without further provisions. Note that the default entrypoint  | 
|
61  | 
may be changed later via \<^verbatim>\<open>docker run --entrypoint="..."\<close>.  | 
|
62  | 
||
63  | 
Option \<^verbatim>\<open>-t\<close> specifies the Docker image tag: this a symbolic name within the  | 
|
64  | 
local Docker name space, but also relevant for Docker  | 
|
65  | 
Hub\<^footnote>\<open>\<^url>\<open>https://hub.docker.com\<close>\<close>.  | 
|
66  | 
||
67  | 
Option \<^verbatim>\<open>-l\<close> specifies the default logic image of the Isabelle distribution  | 
|
68  | 
contained in the Docker environment: it will be produced by \<^verbatim>\<open>isabelle build -b\<close>  | 
|
69  | 
  as usual (\secref{sec:tool-build}) and stored within the image.
 | 
|
70  | 
||
71  | 
\<^medskip>  | 
|
72  | 
Option \<^verbatim>\<open>-B\<close> specifies the Docker image taken as starting point for the  | 
|
| 76184 | 73  | 
Isabelle installation: it needs to be a suitable version of Ubuntu Linux,  | 
74  | 
see also \<^url>\<open>https://hub.docker.com/_/ubuntu\<close>. The default for Isabelle2022  | 
|
75  | 
is \<^verbatim>\<open>ubuntu:22.04\<close>, but other versions often work as well, after some  | 
|
76  | 
experimentation with packages.  | 
|
| 71579 | 77  | 
|
78  | 
Option \<^verbatim>\<open>-p\<close> includes additional Ubuntu packages, using the terminology  | 
|
79  | 
of \<^verbatim>\<open>apt-get install\<close> within the underlying Linux distribution.  | 
|
80  | 
||
81  | 
Option \<^verbatim>\<open>-P\<close> refers to high-level package collections: \<^verbatim>\<open>X11\<close> or \<^verbatim>\<open>latex\<close> as  | 
|
| 
77567
 
b975f5aaf6b8
renamed "isabelle build_docker" to "isabelle docker_build" (unrelated to "isabelle build");
 
wenzelm 
parents: 
76184 
diff
changeset
 | 
82  | 
provided by \<^verbatim>\<open>isabelle docker_build\<close> (assuming Ubuntu 20.04 LTS). This  | 
| 71579 | 83  | 
imposes extra weight on the resulting Docker images. Note that \<^verbatim>\<open>X11\<close> will  | 
84  | 
only provide remote X11 support according to the modest GUI quality  | 
|
85  | 
standards of the late 1990-ies.  | 
|
86  | 
||
87  | 
\<^medskip>  | 
|
88  | 
Option \<^verbatim>\<open>-n\<close> suppresses the actual \<^verbatim>\<open>docker build\<close> process. Option \<^verbatim>\<open>-o\<close>  | 
|
89  | 
outputs the generated \<^verbatim>\<open>Dockerfile\<close>. Both options together produce a  | 
|
90  | 
Dockerfile only, which might be useful for informative purposes or other  | 
|
91  | 
tools.  | 
|
92  | 
||
93  | 
Option \<^verbatim>\<open>-v\<close> disables quiet-mode of the underlying \<^verbatim>\<open>docker build\<close> process.  | 
|
| 
76178
 
1f95e9424341
more robust: snap version of docker cannot access /tmp;
 
wenzelm 
parents: 
76136 
diff
changeset
 | 
94  | 
|
| 
 
1f95e9424341
more robust: snap version of docker cannot access /tmp;
 
wenzelm 
parents: 
76136 
diff
changeset
 | 
95  | 
\<^medskip>  | 
| 
 
1f95e9424341
more robust: snap version of docker cannot access /tmp;
 
wenzelm 
parents: 
76136 
diff
changeset
 | 
96  | 
Option \<^verbatim>\<open>-W\<close> specifies an alternative work directory: it needs to be  | 
| 
 
1f95e9424341
more robust: snap version of docker cannot access /tmp;
 
wenzelm 
parents: 
76136 
diff
changeset
 | 
97  | 
accessible to docker, even if this is run via Snap (e.g.\ on Ubuntu 22.04).  | 
| 
 
1f95e9424341
more robust: snap version of docker cannot access /tmp;
 
wenzelm 
parents: 
76136 
diff
changeset
 | 
98  | 
The default ``\<^verbatim>\<open>.\<close>'' usually works, if this is owned by the user: the tool  | 
| 
 
1f95e9424341
more robust: snap version of docker cannot access /tmp;
 
wenzelm 
parents: 
76136 
diff
changeset
 | 
99  | 
will create a fresh directory within it, and remove it afterwards.  | 
| 71579 | 100  | 
\<close>  | 
101  | 
||
102  | 
||
103  | 
subsubsection \<open>Examples\<close>  | 
|
104  | 
||
105  | 
text \<open>  | 
|
106  | 
Produce a Dockerfile (without image) from a remote Isabelle distribution:  | 
|
107  | 
  @{verbatim [display]
 | 
|
| 
77567
 
b975f5aaf6b8
renamed "isabelle build_docker" to "isabelle docker_build" (unrelated to "isabelle build");
 
wenzelm 
parents: 
76184 
diff
changeset
 | 
108  | 
\<open> isabelle docker_build -E -n -o Dockerfile  | 
| 76105 | 109  | 
https://isabelle.in.tum.de/website-Isabelle2022/dist/Isabelle2022_linux.tar.gz\<close>}  | 
| 71579 | 110  | 
|
111  | 
Build a standard Isabelle Docker image from a local Isabelle distribution,  | 
|
112  | 
with \<^verbatim>\<open>bin/isabelle\<close> as executable entry point:  | 
|
113  | 
||
114  | 
  @{verbatim [display]
 | 
|
| 
77567
 
b975f5aaf6b8
renamed "isabelle build_docker" to "isabelle docker_build" (unrelated to "isabelle build");
 
wenzelm 
parents: 
76184 
diff
changeset
 | 
115  | 
\<open> isabelle docker_build -E -t test/isabelle:Isabelle2022 Isabelle2022_linux.tar.gz\<close>}  | 
| 71579 | 116  | 
|
117  | 
Invoke the raw Isabelle/ML process within that image:  | 
|
118  | 
  @{verbatim [display]
 | 
|
| 76105 | 119  | 
\<open> docker run test/isabelle:Isabelle2022 process -e "Session.welcome ()"\<close>}  | 
| 71579 | 120  | 
|
121  | 
Invoke a Linux command-line tool within the contained Isabelle system  | 
|
122  | 
environment:  | 
|
123  | 
  @{verbatim [display]
 | 
|
| 76105 | 124  | 
\<open> docker run test/isabelle:Isabelle2022 env uname -a\<close>}  | 
| 71579 | 125  | 
The latter should always report a Linux operating system, even when running  | 
126  | 
on Windows or macOS.  | 
|
127  | 
\<close>  | 
|
128  | 
||
| 28224 | 129  | 
|
| 73172 | 130  | 
section \<open>Managing Isabelle components \label{sec:tool-components}\<close>
 | 
| 48844 | 131  | 
|
| 58618 | 132  | 
text \<open>  | 
| 73172 | 133  | 
  The @{tool_def components} tool manages Isabelle components:
 | 
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
134  | 
  @{verbatim [display]
 | 
| 
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
135  | 
\<open>Usage: isabelle components [OPTIONS] [COMPONENTS ...]  | 
| 48844 | 136  | 
|
137  | 
Options are:  | 
|
| 
50653
 
5c85f8b80b95
simplified quick start via "isabelle components -I";
 
wenzelm 
parents: 
50132 
diff
changeset
 | 
138  | 
-I init user settings  | 
| 73172 | 139  | 
-R URL component repository (default $ISABELLE_COMPONENT_REPOSITORY)  | 
| 53435 | 140  | 
-a resolve all missing components  | 
| 48844 | 141  | 
-l list status  | 
| 73172 | 142  | 
-u DIR update $ISABELLE_HOME_USER/components: add directory  | 
143  | 
-x DIR update $ISABELLE_HOME_USER/components: remove directory  | 
|
| 48844 | 144  | 
|
| 73172 | 145  | 
Resolve Isabelle components via download and installation: given COMPONENTS  | 
146  | 
are identified via base name. Further operations manage etc/settings and  | 
|
147  | 
etc/components in $ISABELLE_HOME_USER.  | 
|
| 48844 | 148  | 
|
| 73172 | 149  | 
ISABELLE_COMPONENT_REPOSITORY="..."  | 
150  | 
ISABELLE_HOME_USER="..."  | 
|
151  | 
\<close>}  | 
|
| 48844 | 152  | 
|
| 61575 | 153  | 
  Components are initialized as described in \secref{sec:components} in a
 | 
154  | 
permissive manner, which can mark components as ``missing''. This state is  | 
|
155  | 
  amended by letting @{tool "components"} download and unpack components that
 | 
|
| 63680 | 156  | 
are published on the default component repository  | 
| 68224 | 157  | 
\<^url>\<open>https://isabelle.in.tum.de/components\<close> in particular.  | 
| 48844 | 158  | 
|
| 61575 | 159  | 
Option \<^verbatim>\<open>-R\<close> specifies an alternative component repository. Note that  | 
160  | 
\<^verbatim>\<open>file:///\<close> URLs can be used for local directories.  | 
|
| 48844 | 161  | 
|
| 61575 | 162  | 
Option \<^verbatim>\<open>-a\<close> selects all missing components to be resolved. Explicit  | 
163  | 
components may be named as command line-arguments as well. Note that  | 
|
164  | 
components are uniquely identified by their base name, while the  | 
|
165  | 
installation takes place in the location that was specified in the attempt  | 
|
166  | 
to initialize the component before.  | 
|
| 48844 | 167  | 
|
| 73172 | 168  | 
\<^medskip>  | 
| 61575 | 169  | 
Option \<^verbatim>\<open>-l\<close> lists the current state of available and missing components  | 
170  | 
with their location (full name) within the file-system.  | 
|
| 
50653
 
5c85f8b80b95
simplified quick start via "isabelle components -I";
 
wenzelm 
parents: 
50132 
diff
changeset
 | 
171  | 
|
| 73172 | 172  | 
\<^medskip>  | 
| 61575 | 173  | 
Option \<^verbatim>\<open>-I\<close> initializes the user settings file to subscribe to the standard  | 
174  | 
components specified in the Isabelle repository clone --- this does not make  | 
|
| 73484 | 175  | 
any sense for regular Isabelle releases. An existing file that does not  | 
176  | 
contain a suitable line ``\<^verbatim>\<open>init_components\<close>\<open>\<dots>\<close>\<^verbatim>\<open>components/main\<close>'' needs  | 
|
177  | 
to be edited according to the printed explanation.  | 
|
| 73172 | 178  | 
|
179  | 
\<^medskip>  | 
|
180  | 
Options \<^verbatim>\<open>-u\<close> and \<^verbatim>\<open>-x\<close> operate on user components listed in  | 
|
| 
75161
 
95612f330c93
misc tuning based on comments by Heiko Eißfeldt;
 
wenzelm 
parents: 
75116 
diff
changeset
 | 
181  | 
\<^path>\<open>$ISABELLE_HOME_USER/etc/components\<close>: this avoids manual editing of  | 
| 73172 | 182  | 
Isabelle configuration files.  | 
| 58618 | 183  | 
\<close>  | 
| 48844 | 184  | 
|
185  | 
||
| 58618 | 186  | 
section \<open>Viewing documentation \label{sec:tool-doc}\<close>
 | 
| 28224 | 187  | 
|
| 58618 | 188  | 
text \<open>  | 
| 
52444
 
2cfe6656d6d6
slightly improved "isabelle doc" based on Isabelle/Scala;
 
wenzelm 
parents: 
52052 
diff
changeset
 | 
189  | 
  The @{tool_def doc} tool displays Isabelle documentation:
 | 
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
190  | 
  @{verbatim [display]
 | 
| 
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
191  | 
\<open>Usage: isabelle doc [DOC ...]  | 
| 28224 | 192  | 
|
| 75116 | 193  | 
View Isabelle PDF documentation.\<close>}  | 
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
194  | 
|
| 61575 | 195  | 
If called without arguments, it lists all available documents. Each line  | 
196  | 
starts with an identifier, followed by a short description. Any of these  | 
|
197  | 
identifiers may be specified as arguments, in order to display the  | 
|
| 
72309
 
564012e31db1
discontinued obsolete DVI document format and related settings/tools;
 
wenzelm 
parents: 
71579 
diff
changeset
 | 
198  | 
corresponding document.  | 
| 28224 | 199  | 
|
| 61406 | 200  | 
\<^medskip>  | 
| 61575 | 201  | 
  The @{setting ISABELLE_DOCS} setting specifies the list of directories
 | 
202  | 
(separated by colons) to be scanned for documentations.  | 
|
| 58618 | 203  | 
\<close>  | 
| 28224 | 204  | 
|
205  | 
||
| 58618 | 206  | 
section \<open>Shell commands within the settings environment \label{sec:tool-env}\<close>
 | 
| 47828 | 207  | 
|
| 61575 | 208  | 
text \<open>  | 
209  | 
  The @{tool_def env} tool is a direct wrapper for the standard
 | 
|
210  | 
\<^verbatim>\<open>/usr/bin/env\<close> command on POSIX systems, running within the Isabelle  | 
|
211  | 
  settings environment (\secref{sec:settings}).
 | 
|
| 47828 | 212  | 
|
| 61575 | 213  | 
The command-line arguments are that of the underlying version of \<^verbatim>\<open>env\<close>. For  | 
214  | 
example, the following invokes an instance of the GNU Bash shell within the  | 
|
215  | 
Isabelle environment:  | 
|
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
216  | 
  @{verbatim [display] \<open>isabelle env bash\<close>}
 | 
| 58618 | 217  | 
\<close>  | 
| 47828 | 218  | 
|
219  | 
||
| 58618 | 220  | 
section \<open>Inspecting the settings environment \label{sec:tool-getenv}\<close>
 | 
| 28224 | 221  | 
|
| 58618 | 222  | 
text \<open>The Isabelle settings environment --- as provided by the  | 
| 28224 | 223  | 
site-default and user-specific settings files --- can be inspected  | 
| 48602 | 224  | 
  with the @{tool_def getenv} tool:
 | 
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
225  | 
  @{verbatim [display]
 | 
| 
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
226  | 
\<open>Usage: isabelle getenv [OPTIONS] [VARNAMES ...]  | 
| 28224 | 227  | 
|
228  | 
Options are:  | 
|
229  | 
-a display complete environment  | 
|
230  | 
-b print values only (doesn't work for -a)  | 
|
| 73581 | 231  | 
-d FILE dump complete environment to file (NUL terminated entries)  | 
| 28224 | 232  | 
|
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
233  | 
Get value of VARNAMES from the Isabelle settings.\<close>}  | 
| 28224 | 234  | 
|
| 61575 | 235  | 
With the \<^verbatim>\<open>-a\<close> option, one may inspect the full process environment that  | 
236  | 
Isabelle related programs are run in. This usually contains much more  | 
|
237  | 
variables than are actually Isabelle settings. Normally, output is a list of  | 
|
238  | 
lines of the form \<open>name\<close>\<^verbatim>\<open>=\<close>\<open>value\<close>. The \<^verbatim>\<open>-b\<close> option causes only the values  | 
|
239  | 
to be printed.  | 
|
| 31497 | 240  | 
|
| 61575 | 241  | 
Option \<^verbatim>\<open>-d\<close> produces a dump of the complete environment to the specified  | 
| 73581 | 242  | 
file. Entries are terminated by the ASCII NUL character, i.e.\ the string  | 
243  | 
terminator in C. Thus the Isabelle/Scala operation  | 
|
244  | 
\<^scala_method>\<open>isabelle.Isabelle_System.init\<close> can import the settings  | 
|
245  | 
environment robustly, and provide its own  | 
|
246  | 
\<^scala_method>\<open>isabelle.Isabelle_System.getenv\<close> function.  | 
|
| 58618 | 247  | 
\<close>  | 
| 28224 | 248  | 
|
249  | 
||
| 58618 | 250  | 
subsubsection \<open>Examples\<close>  | 
| 28224 | 251  | 
|
| 61575 | 252  | 
text \<open>  | 
253  | 
  Get the location of @{setting ISABELLE_HOME_USER} where user-specific
 | 
|
254  | 
information is stored:  | 
|
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
255  | 
  @{verbatim [display] \<open>isabelle getenv ISABELLE_HOME_USER\<close>}
 | 
| 28224 | 256  | 
|
| 61406 | 257  | 
\<^medskip>  | 
| 61575 | 258  | 
Get the value only of the same settings variable, which is particularly  | 
259  | 
useful in shell scripts:  | 
|
| 68219 | 260  | 
  @{verbatim [display] \<open>isabelle getenv -b ISABELLE_HOME_USER\<close>}
 | 
| 58618 | 261  | 
\<close>  | 
| 28224 | 262  | 
|
263  | 
||
| 71322 | 264  | 
section \<open>Mercurial repository setup \label{sec:hg-setup}\<close>
 | 
265  | 
||
266  | 
text \<open>  | 
|
267  | 
  The @{tool_def hg_setup} tool simplifies the setup of Mercurial
 | 
|
| 71325 | 268  | 
  repositories, with hosting via Phabricator (\chref{ch:phabricator}) or SSH
 | 
269  | 
file server access.  | 
|
| 71322 | 270  | 
|
271  | 
  @{verbatim [display]
 | 
|
272  | 
\<open>Usage: isabelle hg_setup [OPTIONS] REMOTE LOCAL_DIR  | 
|
273  | 
||
274  | 
Options are:  | 
|
275  | 
-n NAME remote repository name (default: base name of LOCAL_DIR)  | 
|
276  | 
-p PATH Mercurial path name (default: "default")  | 
|
277  | 
-r assume that remote repository already exists  | 
|
278  | 
||
279  | 
Setup a remote vs. local Mercurial repository: REMOTE either refers to a  | 
|
280  | 
Phabricator server "user@host" or SSH file server "ssh://user@host/path".\<close>}  | 
|
281  | 
||
282  | 
The \<^verbatim>\<open>REMOTE\<close> repository specification \<^emph>\<open>excludes\<close> the actual repository  | 
|
283  | 
name: that is given by the base name of \<^verbatim>\<open>LOCAL_DIR\<close>, or via option \<^verbatim>\<open>-n\<close>.  | 
|
284  | 
||
285  | 
By default, both sides of the repository are created on demand by default.  | 
|
286  | 
In contrast, option \<^verbatim>\<open>-r\<close> assumes that the remote repository already exists:  | 
|
287  | 
it avoids accidental creation of a persistent repository with unintended  | 
|
288  | 
name.  | 
|
289  | 
||
290  | 
The local \<^verbatim>\<open>.hg/hgrc\<close> file is changed to refer to the remote repository,  | 
|
291  | 
usually via the symbolic path name ``\<^verbatim>\<open>default\<close>''; option \<^verbatim>\<open>-p\<close> allows to  | 
|
| 
75161
 
95612f330c93
misc tuning based on comments by Heiko Eißfeldt;
 
wenzelm 
parents: 
75116 
diff
changeset
 | 
292  | 
provide a different name.  | 
| 71322 | 293  | 
\<close>  | 
294  | 
||
| 71324 | 295  | 
subsubsection \<open>Examples\<close>  | 
296  | 
||
297  | 
text \<open>  | 
|
298  | 
Setup the current directory as a repository with Phabricator server hosting:  | 
|
299  | 
  @{verbatim [display] \<open>  isabelle hg_setup vcs@vcs.example.org .\<close>}
 | 
|
300  | 
||
301  | 
\<^medskip>  | 
|
302  | 
Setup the current directory as a repository with plain SSH server hosting:  | 
|
303  | 
  @{verbatim [display] \<open>  isabelle hg_setup ssh://files.example.org/data/repositories .\<close>}
 | 
|
304  | 
||
305  | 
\<^medskip>  | 
|
306  | 
Both variants require SSH access to the target server, via public key  | 
|
307  | 
without password.  | 
|
308  | 
\<close>  | 
|
309  | 
||
| 71322 | 310  | 
|
| 
75555
 
197a5b3a1ea2
promote "isabelle sync" to regular user-space tool, with proper documentation;
 
wenzelm 
parents: 
75511 
diff
changeset
 | 
311  | 
section \<open>Mercurial repository synchronization \label{sec:tool-hg-sync}\<close>
 | 
| 75476 | 312  | 
|
313  | 
text \<open>  | 
|
| 75479 | 314  | 
  The @{tool_def hg_sync} tool synchronizes a local Mercurial repository with
 | 
| 75511 | 315  | 
a target directory.  | 
| 75476 | 316  | 
|
317  | 
  @{verbatim [display]
 | 
|
318  | 
\<open>Usage: isabelle hg_sync [OPTIONS] TARGET  | 
|
319  | 
||
320  | 
Options are:  | 
|
| 75511 | 321  | 
-F RULE add rsync filter RULE  | 
322  | 
(e.g. "protect /foo" to avoid deletion)  | 
|
| 75476 | 323  | 
-R ROOT explicit repository root directory  | 
324  | 
(default: implicit from current directory)  | 
|
| 75493 | 325  | 
-T thorough treatment of file content and directory times  | 
| 75476 | 326  | 
-n no changes: dry-run  | 
| 
77783
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
327  | 
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME)  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
328  | 
-p PORT explicit SSH port  | 
| 75479 | 329  | 
-r REV explicit revision (default: state of working directory)  | 
| 
77783
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
330  | 
-s HOST SSH host name for remote target (default: local)  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
331  | 
-u USER explicit SSH user name  | 
| 75476 | 332  | 
-v verbose  | 
333  | 
||
| 75479 | 334  | 
Synchronize Mercurial repository with TARGET directory,  | 
| 
77783
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
335  | 
which can be local or remote (see options -s -p -u).\<close>}  | 
| 75476 | 336  | 
|
| 
77783
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
337  | 
The \<^verbatim>\<open>TARGET\<close> specifies a directory, which can be local or an a remote SSH  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
338  | 
host; the latter requires option \<^verbatim>\<open>-s\<close> for the host name. The content is  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
339  | 
written directly into the target, \<^emph>\<open>without\<close> creating a separate  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
340  | 
sub-directory. The special sub-directory \<^verbatim>\<open>.hg_sync\<close> within the target  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
341  | 
contains meta data from the original Mercurial repository. Repeated  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
342  | 
synchronization is guarded by the presence of a \<^verbatim>\<open>.hg_sync\<close> sub-directory:  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
343  | 
this sanity check prevents accidental changes (or deletion!) of targets that  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
344  | 
  were not created by @{tool hg_sync}.
 | 
| 75476 | 345  | 
|
| 75479 | 346  | 
\<^medskip> Option \<^verbatim>\<open>-r\<close> specifies an explicit revision of the repository; the default  | 
347  | 
is the current state of the working directory (which might be uncommitted).  | 
|
348  | 
||
| 75476 | 349  | 
\<^medskip> Option \<^verbatim>\<open>-v\<close> enables verbose mode. Option \<^verbatim>\<open>-n\<close> enables ``dry-run'' mode:  | 
| 75511 | 350  | 
operations are only simulated; use it with option \<^verbatim>\<open>-v\<close> to actually see  | 
351  | 
results.  | 
|
| 75476 | 352  | 
|
| 75489 | 353  | 
\<^medskip> Option \<^verbatim>\<open>-F\<close> adds a filter rule to the underlying \<^verbatim>\<open>rsync\<close> command;  | 
354  | 
multiple \<^verbatim>\<open>-F\<close> options may be given to accumulate a list of rules.  | 
|
| 75476 | 355  | 
|
356  | 
\<^medskip> Option \<^verbatim>\<open>-R\<close> specifies an explicit repository root directory. The default  | 
|
| 75479 | 357  | 
is to derive it from the current directory, searching upwards until a  | 
358  | 
suitable \<^verbatim>\<open>.hg\<close> directory is found.  | 
|
| 75487 | 359  | 
|
| 75493 | 360  | 
\<^medskip> Option \<^verbatim>\<open>-T\<close> indicates thorough treatment of file content and directory  | 
361  | 
times. The default is to consider files with equal time and size already as  | 
|
362  | 
equal, and to ignore time stamps on directories.  | 
|
| 75499 | 363  | 
|
| 
77783
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
364  | 
\<^medskip> Options \<^verbatim>\<open>-s\<close>, \<^verbatim>\<open>-p\<close>, \<^verbatim>\<open>-u\<close> specify the SSH connection precisely. If no  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
365  | 
SSH host name is given, the local file-system is used. An explicit port and  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
366  | 
user are only required in special situations.  | 
| 
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
367  | 
|
| 75499 | 368  | 
\<^medskip> Option \<^verbatim>\<open>-p\<close> specifies an explicit port for the SSH connection underlying  | 
| 
76131
 
8b695e59db3f
clarified default: do not override port from ssh_config, which could be different from 22;
 
wenzelm 
parents: 
76105 
diff
changeset
 | 
369  | 
\<^verbatim>\<open>rsync\<close>; the default is taken from the user's \<^path>\<open>ssh_config\<close> file.  | 
| 75476 | 370  | 
\<close>  | 
371  | 
||
372  | 
subsubsection \<open>Examples\<close>  | 
|
373  | 
||
374  | 
text \<open>  | 
|
| 
75490
 
5e37ea93759d
clarified documentation: $ISABELLE_HOME is not a repository for regular releases;
 
wenzelm 
parents: 
75489 
diff
changeset
 | 
375  | 
Synchronize the current repository onto a remote host, with accurate  | 
| 75493 | 376  | 
treatment of all content:  | 
| 
77783
 
fb61887c069a
clarified underlying SSH session of "isabelle hg_sync" and "isabelle sync";
 
wenzelm 
parents: 
77567 
diff
changeset
 | 
377  | 
  @{verbatim [display] \<open>  isabelle hg_sync -T -s remotename test/repos\<close>}
 | 
| 75476 | 378  | 
\<close>  | 
379  | 
||
380  | 
||
| 58618 | 381  | 
section \<open>Installing standalone Isabelle executables \label{sec:tool-install}\<close>
 | 
| 28224 | 382  | 
|
| 61575 | 383  | 
text \<open>  | 
384  | 
  By default, the main Isabelle binaries (@{executable "isabelle"} etc.) are
 | 
|
385  | 
just run from their location within the distribution directory, probably  | 
|
386  | 
  indirectly by the shell through its @{setting PATH}. Other schemes of
 | 
|
387  | 
  installation are supported by the @{tool_def install} tool:
 | 
|
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
388  | 
  @{verbatim [display]
 | 
| 
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
389  | 
\<open>Usage: isabelle install [OPTIONS] BINDIR  | 
| 28224 | 390  | 
|
391  | 
Options are:  | 
|
| 50132 | 392  | 
-d DISTDIR refer to DISTDIR as Isabelle distribution  | 
| 28224 | 393  | 
(default ISABELLE_HOME)  | 
394  | 
||
| 50132 | 395  | 
Install Isabelle executables with absolute references to the  | 
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
396  | 
distribution directory.\<close>}  | 
| 28224 | 397  | 
|
| 61575 | 398  | 
The \<^verbatim>\<open>-d\<close> option overrides the current Isabelle distribution directory as  | 
399  | 
  determined by @{setting ISABELLE_HOME}.
 | 
|
| 28224 | 400  | 
|
| 61575 | 401  | 
The \<open>BINDIR\<close> argument tells where executable wrapper scripts for  | 
| 
79059
 
ae682b2aab03
removed obsolete/broken isabelle_scala_script wrapper (see also abf9fcfa65cf);
 
wenzelm 
parents: 
78665 
diff
changeset
 | 
402  | 
  @{executable "isabelle"} and @{executable isabelle_java} should be
 | 
| 61575 | 403  | 
  placed, which is typically a directory in the shell's @{setting PATH}, such
 | 
404  | 
as \<^verbatim>\<open>$HOME/bin\<close>.  | 
|
| 48815 | 405  | 
|
| 61406 | 406  | 
\<^medskip>  | 
| 61575 | 407  | 
It is also possible to make symbolic links of the main Isabelle executables  | 
408  | 
manually, but making separate copies outside the Isabelle distribution  | 
|
409  | 
directory will not work!  | 
|
410  | 
\<close>  | 
|
| 28224 | 411  | 
|
412  | 
||
| 58618 | 413  | 
section \<open>Creating instances of the Isabelle logo\<close>  | 
| 28224 | 414  | 
|
| 61575 | 415  | 
text \<open>  | 
| 73399 | 416  | 
  The @{tool_def logo} tool creates variants of the Isabelle logo, for
 | 
417  | 
  inclusion in PDF{\LaTeX} documents.
 | 
|
418  | 
||
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
419  | 
  @{verbatim [display]
 | 
| 73399 | 420  | 
\<open>Usage: isabelle logo [OPTIONS] [NAME]  | 
| 28224 | 421  | 
|
422  | 
Options are:  | 
|
| 73399 | 423  | 
-o FILE alternative output file  | 
424  | 
-q quiet mode  | 
|
425  | 
||
426  | 
Create variant NAME of the Isabelle logo as "isabelle_name.pdf".\<close>}  | 
|
| 48936 | 427  | 
|
| 73399 | 428  | 
Option \<^verbatim>\<open>-o\<close> provides an alternative output file, instead of the default in  | 
429  | 
the current directory: \<^verbatim>\<open>isabelle_\<close>\<open>name\<close>\<^verbatim>\<open>.pdf\<close> with the lower-case version  | 
|
430  | 
of the given name.  | 
|
| 48936 | 431  | 
|
| 73399 | 432  | 
\<^medskip>  | 
| 
72316
 
3cc6aa405858
clarified "isabelle logo", after discontinuation of DVI output (see 564012e31db1);
 
wenzelm 
parents: 
72309 
diff
changeset
 | 
433  | 
Option \<^verbatim>\<open>-q\<close> omits printing of the resulting output file name.  | 
| 48936 | 434  | 
|
| 61406 | 435  | 
\<^medskip>  | 
| 61575 | 436  | 
Implementors of Isabelle tools and applications are encouraged to make  | 
| 73399 | 437  | 
derived Isabelle logos for their own projects using this template. The  | 
438  | 
license is the same as for the regular Isabelle distribution (BSD).  | 
|
| 61575 | 439  | 
\<close>  | 
| 28224 | 440  | 
|
441  | 
||
| 58618 | 442  | 
section \<open>Output the version identifier of the Isabelle distribution\<close>  | 
| 28224 | 443  | 
|
| 58618 | 444  | 
text \<open>  | 
| 48602 | 445  | 
  The @{tool_def version} tool displays Isabelle version information:
 | 
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
446  | 
  @{verbatim [display]
 | 
| 
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
447  | 
\<open>Usage: isabelle version [OPTIONS]  | 
| 41511 | 448  | 
|
449  | 
Options are:  | 
|
450  | 
-i short identification (derived from Mercurial id)  | 
|
| 73481 | 451  | 
-t symbolic tags (derived from Mercurial id)  | 
| 41511 | 452  | 
|
| 
61407
 
7ba7b8103565
@{verbatim [display]} supersedes old alltt/ttbox;
 
wenzelm 
parents: 
61406 
diff
changeset
 | 
453  | 
Display Isabelle version information.\<close>}  | 
| 41511 | 454  | 
|
| 61406 | 455  | 
\<^medskip>  | 
| 61575 | 456  | 
The default is to output the full version string of the Isabelle  | 
| 76105 | 457  | 
distribution, e.g.\ ``\<^verbatim>\<open>Isabelle2022: October 2022\<close>.  | 
| 41511 | 458  | 
|
| 73481 | 459  | 
\<^medskip>  | 
460  | 
Option \<^verbatim>\<open>-i\<close> produces a short identification derived from the Mercurial id  | 
|
461  | 
  of the @{setting ISABELLE_HOME} directory; option \<^verbatim>\<open>-t\<close> prints version tags
 | 
|
462  | 
(if available).  | 
|
463  | 
||
464  | 
These options require either a repository clone or a repository archive  | 
|
465  | 
(e.g. download of  | 
|
| 
73480
 
0e880b793db1
support repository archives (without full .hg directory);
 
wenzelm 
parents: 
73399 
diff
changeset
 | 
466  | 
\<^url>\<open>https://isabelle.sketis.net/repos/isabelle/archive/tip.tar.gz\<close>).  | 
| 58618 | 467  | 
\<close>  | 
| 28224 | 468  | 
|
| 78665 | 469  | 
|
470  | 
section \<open>Managed installations of \<^text>\<open>Haskell\<close> and \<^text>\<open>OCaml\<close>\<close>  | 
|
471  | 
||
472  | 
text \<open>  | 
|
473  | 
Code generated in Isabelle \<^cite>\<open>"Haftmann-codegen"\<close> for \<^text>\<open>SML\<close>  | 
|
474  | 
or \<^text>\<open>Scala\<close> integrates easily using Isabelle/ML or Isabelle/Scala  | 
|
475  | 
respectively.  | 
|
476  | 
||
477  | 
To facilitate integration with further target languages, there are  | 
|
478  | 
tools to provide managed installations of the required ecosystems:  | 
|
479  | 
||
480  | 
  \<^item> Tool @{tool_def ghc_setup} provides a basic \<^text>\<open>Haskell\<close> \<^cite>\<open>"Thompson-Haskell"\<close> environment
 | 
|
481  | 
consisting of the Glasgow Haskell Compiler and the Haskell Tool Stack.  | 
|
482  | 
||
483  | 
  \<^item> Tool @{tool_def ghc_stack} provides an interface to that \<^text>\<open>Haskell\<close>
 | 
|
484  | 
environment; use \<^verbatim>\<open>isabelle ghc_stack --help\<close> for elementary  | 
|
485  | 
instructions.  | 
|
486  | 
||
487  | 
  \<^item> Tool @{tool_def ocaml_setup} provides a basic \<^text>\<open>OCaml\<close> \<^cite>\<open>OCaml\<close> environment
 | 
|
488  | 
consisting of the OCaml compiler and the OCaml Package Manager.  | 
|
489  | 
||
490  | 
  \<^item> Tool @{tool_def ocaml_opam} provides an interface to that \<^text>\<open>OCaml\<close>
 | 
|
491  | 
environment; use \<^verbatim>\<open>isabelle ocaml_opam --help\<close> for elementary  | 
|
492  | 
instructions.  | 
|
493  | 
\<close>  | 
|
494  | 
||
| 67399 | 495  | 
end  |