author | wenzelm |
Mon, 24 Sep 2018 15:34:19 +0200 | |
changeset 69055 | dab89758545c |
parent 68737 | a8bef9ff7dc0 |
child 69343 | 395c4fb15ea2 |
permissions | -rw-r--r-- |
61656 | 1 |
(*:maxLineLen=78:*) |
53981 | 2 |
|
53769 | 3 |
theory JEdit |
4 |
imports Base |
|
5 |
begin |
|
6 |
||
58618 | 7 |
chapter \<open>Introduction\<close> |
53769 | 8 |
|
58618 | 9 |
section \<open>Concepts and terminology\<close> |
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
10 |
|
58618 | 11 |
text \<open> |
61574 | 12 |
Isabelle/jEdit is a Prover IDE that integrates \<^emph>\<open>parallel proof checking\<close> |
13 |
@{cite "Wenzel:2009" and "Wenzel:2013:ITP"} with \<^emph>\<open>asynchronous user |
|
14 |
interaction\<close> @{cite "Wenzel:2010" and "Wenzel:2012:UITP-EPTCS" and |
|
15 |
"Wenzel:2014:ITP-PIDE" and "Wenzel:2014:UITP"}, based on a document-oriented |
|
16 |
approach to \<^emph>\<open>continuous proof processing\<close> @{cite "Wenzel:2011:CICM" and |
|
17 |
"Wenzel:2012"}. Many concepts and system components are fit together in |
|
18 |
order to make this work. The main building blocks are as follows. |
|
53769 | 19 |
|
61574 | 20 |
\<^descr>[Isabelle/ML] is the implementation and extension language of Isabelle, |
21 |
see also @{cite "isabelle-implementation"}. It is integrated into the |
|
22 |
logical context of Isabelle/Isar and allows to manipulate logical entities |
|
23 |
directly. Arbitrary add-on tools may be implemented for object-logics such |
|
24 |
as Isabelle/HOL. |
|
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
25 |
|
61574 | 26 |
\<^descr>[Isabelle/Scala] is the system programming language of Isabelle. It |
62183 | 27 |
extends the pure logical environment of Isabelle/ML towards the outer |
28 |
world of graphical user interfaces, text editors, IDE frameworks, web |
|
66683 | 29 |
services, SSH servers, SQL databases etc. Special infrastructure allows to |
30 |
transfer algebraic datatypes and formatted text easily between ML and |
|
31 |
Scala, using asynchronous protocol commands. |
|
54321 | 32 |
|
62183 | 33 |
\<^descr>[PIDE] is a general framework for Prover IDEs based on Isabelle/Scala. It |
34 |
is built around a concept of parallel and asynchronous document |
|
35 |
processing, which is supported natively by the parallel proof engine that |
|
36 |
is implemented in Isabelle/ML. The traditional prover command loop is |
|
37 |
given up; instead there is direct support for editing of source text, with |
|
38 |
rich formal markup for GUI rendering. |
|
39 |
||
63680 | 40 |
\<^descr>[jEdit] is a sophisticated text editor\<^footnote>\<open>\<^url>\<open>http://www.jedit.org\<close>\<close> |
68224 | 41 |
implemented in Java\<^footnote>\<open>\<^url>\<open>https://www.java.com\<close>\<close>. It is easily extensible by |
63680 | 42 |
plugins written in any language that works on the JVM. In the context of |
68224 | 43 |
Isabelle this is always Scala\<^footnote>\<open>\<^url>\<open>https://www.scala-lang.org\<close>\<close>. |
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
44 |
|
62183 | 45 |
\<^descr>[Isabelle/jEdit] is the main application of the PIDE framework and the |
46 |
default user-interface for Isabelle. It targets both beginners and |
|
47 |
experts. Technically, Isabelle/jEdit consists of the original jEdit code |
|
48 |
base with minimal patches and a special plugin for Isabelle. This is |
|
49 |
integrated as a desktop application for the main operating system |
|
50 |
families: Linux, Windows, Mac OS X. |
|
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
51 |
|
62183 | 52 |
End-users of Isabelle download and run a standalone application that exposes |
53 |
jEdit as a text editor on the surface. Thus there is occasionally a tendency |
|
54 |
to apply the name ``jEdit'' to any of the Isabelle Prover IDE aspects, |
|
55 |
without proper differentiation. When discussing these PIDE building blocks |
|
56 |
in public forums, mailing lists, or even scientific publications, it is |
|
57 |
particularly important to distinguish Isabelle/ML versus Standard ML, |
|
58 |
Isabelle/Scala versus Scala, Isabelle/jEdit versus jEdit. |
|
58618 | 59 |
\<close> |
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
60 |
|
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
61 |
|
58618 | 62 |
section \<open>The Isabelle/jEdit Prover IDE\<close> |
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
63 |
|
58618 | 64 |
text \<open> |
62183 | 65 |
\begin{figure}[!htb] |
54331 | 66 |
\begin{center} |
57312 | 67 |
\includegraphics[scale=0.333]{isabelle-jedit} |
54331 | 68 |
\end{center} |
69 |
\caption{The Isabelle/jEdit Prover IDE} |
|
54357 | 70 |
\label{fig:isabelle-jedit} |
54331 | 71 |
\end{figure} |
53773 | 72 |
|
57337 | 73 |
Isabelle/jEdit (\figref{fig:isabelle-jedit}) consists of some plugins for |
74 |
the jEdit text editor, while preserving its general look-and-feel as far as |
|
75 |
possible. The main plugin is called ``Isabelle'' and has its own menu |
|
68067 | 76 |
\<^emph>\<open>Plugins~/ Isabelle\<close> with access to several actions and add-on panels (see |
77 |
also \secref{sec:dockables}), as well as \<^emph>\<open>Plugins~/ Plugin Options~/ |
|
78 |
Isabelle\<close> (see also \secref{sec:options}). |
|
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
79 |
|
62183 | 80 |
The options allow to specify a logic session name, but the same selector is |
81 |
also accessible in the \<^emph>\<open>Theories\<close> panel (\secref{sec:theories}). After |
|
66683 | 82 |
startup of the Isabelle plugin, the selected logic session image is provided |
62183 | 83 |
automatically by the Isabelle build tool @{cite "isabelle-system"}: if it is |
66683 | 84 |
absent or outdated wrt.\ its sources, the build process updates it within |
85 |
the running text editor. Prover IDE functionality is fully activated after |
|
62183 | 86 |
successful termination of the build process. A failure may require changing |
66683 | 87 |
some options and restart of the Isabelle plugin or application. Changing the |
88 |
logic session, or the underlying ML system platform (32\,bit versus 64\,bit) |
|
89 |
requires a restart of the whole application to take effect. |
|
53769 | 90 |
|
61415 | 91 |
\<^medskip> |
61574 | 92 |
The main job of the Prover IDE is to manage sources and their changes, |
93 |
taking the logical structure as a formal document into account (see also |
|
94 |
\secref{sec:document-model}). The editor and the prover are connected |
|
57337 | 95 |
asynchronously in a lock-free manner. The prover is free to organize the |
96 |
checking of the formal text in parallel on multiple cores, and provides |
|
97 |
feedback via markup, which is rendered in the editor via colors, boxes, |
|
57420 | 98 |
squiggly underlines, hyperlinks, popup windows, icons, clickable output etc. |
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
99 |
|
61574 | 100 |
Using the mouse together with the modifier key \<^verbatim>\<open>CONTROL\<close> (Linux, Windows) |
62183 | 101 |
or \<^verbatim>\<open>COMMAND\<close> (Mac OS X) exposes formal content via tooltips and/or |
102 |
hyperlinks (see also \secref{sec:tooltips-hyperlinks}). Output (in popups |
|
103 |
etc.) may be explored recursively, using the same techniques as in the |
|
104 |
editor source buffer. |
|
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
105 |
|
57337 | 106 |
Thus the Prover IDE gives an impression of direct access to formal content |
107 |
of the prover within the editor, but in reality only certain aspects are |
|
62183 | 108 |
exposed, according to the possibilities of the prover and its add-on tools. |
58618 | 109 |
\<close> |
53769 | 110 |
|
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
111 |
|
58618 | 112 |
subsection \<open>Documentation\<close> |
54321 | 113 |
|
58618 | 114 |
text \<open> |
62183 | 115 |
The \<^emph>\<open>Documentation\<close> panel of Isabelle/jEdit provides access to some example |
116 |
theory files and the standard Isabelle documentation. PDF files are opened |
|
117 |
by regular desktop operations of the underlying platform. The section |
|
118 |
``Original jEdit Documentation'' contains the original \<^emph>\<open>User's Guide\<close> of |
|
119 |
this sophisticated text editor. The same is accessible via the \<^verbatim>\<open>Help\<close> menu |
|
120 |
or \<^verbatim>\<open>F1\<close> keyboard shortcut, using the built-in HTML viewer of Java/Swing. |
|
121 |
The latter also includes \<^emph>\<open>Frequently Asked Questions\<close> and documentation of |
|
122 |
individual plugins. |
|
54321 | 123 |
|
62183 | 124 |
Most of the information about jEdit is relevant for Isabelle/jEdit as well, |
66683 | 125 |
but users need to keep in mind that defaults sometimes differ, and the |
62183 | 126 |
official jEdit documentation does not know about the Isabelle plugin with |
127 |
its support for continuous checking of formal source text: jEdit is a plain |
|
128 |
text editor, but Isabelle/jEdit is a Prover IDE. |
|
58618 | 129 |
\<close> |
54321 | 130 |
|
131 |
||
58618 | 132 |
subsection \<open>Plugins\<close> |
54321 | 133 |
|
58618 | 134 |
text \<open> |
61574 | 135 |
The \<^emph>\<open>Plugin Manager\<close> of jEdit allows to augment editor functionality by JVM |
136 |
modules (jars) that are provided by the central plugin repository, which is |
|
137 |
accessible via various mirror sites. |
|
54321 | 138 |
|
57420 | 139 |
Connecting to the plugin server-infrastructure of the jEdit project allows |
140 |
to update bundled plugins or to add further functionality. This needs to be |
|
141 |
done with the usual care for such an open bazaar of contributions. Arbitrary |
|
142 |
combinations of add-on features are apt to cause problems. It is advisable |
|
143 |
to start with the default configuration of Isabelle/jEdit and develop some |
|
64513 | 144 |
sense how it is meant to work, before loading too many other plugins. |
54321 | 145 |
|
61415 | 146 |
\<^medskip> |
66683 | 147 |
The \<^emph>\<open>Isabelle\<close> plugin is responsible for the main Prover IDE functionality |
148 |
of Isabelle/jEdit: it manages the prover session in the background. A few |
|
66462 | 149 |
additional plugins are bundled with Isabelle/jEdit for convenience or out of |
66683 | 150 |
necessity, notably \<^emph>\<open>Console\<close> with its \<^emph>\<open>Scala\<close> sub-plugin |
66462 | 151 |
(\secref{sec:scala-console}) and \<^emph>\<open>SideKick\<close> with some Isabelle-specific |
152 |
parsers for document tree structure (\secref{sec:sidekick}). The |
|
153 |
\<^emph>\<open>Navigator\<close> plugin is particularly important for hyperlinks within the |
|
154 |
formal document-model (\secref{sec:tooltips-hyperlinks}). Further plugins |
|
155 |
(e.g.\ \<^emph>\<open>ErrorList\<close>, \<^emph>\<open>Code2HTML\<close>) are included to saturate the dependencies |
|
66683 | 156 |
of bundled plugins, but have no particular use in Isabelle/jEdit. |
157 |
\<close> |
|
54321 | 158 |
|
159 |
||
58618 | 160 |
subsection \<open>Options \label{sec:options}\<close> |
54321 | 161 |
|
61574 | 162 |
text \<open> |
163 |
Both jEdit and Isabelle have distinctive management of persistent options. |
|
54321 | 164 |
|
61574 | 165 |
Regular jEdit options are accessible via the dialogs \<^emph>\<open>Utilities~/ Global |
166 |
Options\<close> or \<^emph>\<open>Plugins~/ Plugin Options\<close>, with a second chance to flip the |
|
63669 | 167 |
two within the central options dialog. Changes are stored in @{path |
63749 | 168 |
"$JEDIT_SETTINGS/properties"} and @{path "$JEDIT_SETTINGS/keymaps"}. |
57310 | 169 |
|
170 |
Isabelle system options are managed by Isabelle/Scala and changes are stored |
|
63669 | 171 |
in @{path "$ISABELLE_HOME_USER/etc/preferences"}, independently of |
60270 | 172 |
other jEdit properties. See also @{cite "isabelle-system"}, especially the |
57310 | 173 |
coverage of sessions and command-line tools like @{tool build} or @{tool |
54372 | 174 |
options}. |
54321 | 175 |
|
62183 | 176 |
Those Isabelle options that are declared as \<^verbatim>\<open>public\<close> are configurable in |
61574 | 177 |
Isabelle/jEdit via \<^emph>\<open>Plugin Options~/ Isabelle~/ General\<close>. Moreover, there |
66683 | 178 |
are various options for rendering document content, which are configurable |
179 |
via \<^emph>\<open>Plugin Options~/ Isabelle~/ Rendering\<close>. Thus \<^emph>\<open>Plugin Options~/ |
|
180 |
Isabelle\<close> in jEdit provides a view on a subset of Isabelle system options. |
|
181 |
Note that some of these options affect general parameters that are relevant |
|
182 |
outside Isabelle/jEdit as well, e.g.\ @{system_option threads} or |
|
61574 | 183 |
@{system_option parallel_proofs} for the Isabelle build tool @{cite |
184 |
"isabelle-system"}, but it is possible to use the settings variable |
|
185 |
@{setting ISABELLE_BUILD_OPTIONS} to change defaults for batch builds |
|
62183 | 186 |
without affecting the Prover IDE. |
54329 | 187 |
|
57627
65fc7ae1bf66
added action "isabelle.options" (despite problems with initial window size);
wenzelm
parents:
57603
diff
changeset
|
188 |
The jEdit action @{action_def isabelle.options} opens the options dialog for |
65fc7ae1bf66
added action "isabelle.options" (despite problems with initial window size);
wenzelm
parents:
57603
diff
changeset
|
189 |
the Isabelle plugin; it can be mapped to editor GUI elements as usual. |
65fc7ae1bf66
added action "isabelle.options" (despite problems with initial window size);
wenzelm
parents:
57603
diff
changeset
|
190 |
|
61415 | 191 |
\<^medskip> |
192 |
Options are usually loaded on startup and saved on shutdown of |
|
66683 | 193 |
Isabelle/jEdit. Editing the generated @{path "$JEDIT_SETTINGS/properties"} |
194 |
or @{path "$ISABELLE_HOME_USER/etc/preferences"} manually while the |
|
195 |
application is running may cause surprise due to lost updates! |
|
61574 | 196 |
\<close> |
54321 | 197 |
|
198 |
||
58618 | 199 |
subsection \<open>Keymaps\<close> |
54321 | 200 |
|
61574 | 201 |
text \<open> |
62183 | 202 |
Keyboard shortcuts are managed as a separate concept of \<^emph>\<open>keymap\<close> that is |
61574 | 203 |
configurable via \<^emph>\<open>Global Options~/ Shortcuts\<close>. The \<^verbatim>\<open>imported\<close> keymap is |
204 |
derived from the initial environment of properties that is available at the |
|
62183 | 205 |
first start of the editor; afterwards the keymap file takes precedence and |
206 |
is no longer affected by change of default properties. |
|
54321 | 207 |
|
66683 | 208 |
Users may change their keymap later, but this may lead to conflicts with |
209 |
\<^verbatim>\<open>shortcut\<close> properties in \<^file>\<open>$JEDIT_HOME/src/jEdit.props\<close>. |
|
64513 | 210 |
|
211 |
The action @{action_def "isabelle.keymap-merge"} helps to resolve pending |
|
212 |
Isabelle keymap changes that are in conflict with the current jEdit keymap; |
|
66683 | 213 |
while non-conflicting changes are applied implicitly. This action is |
64513 | 214 |
automatically invoked on Isabelle/jEdit startup. |
58618 | 215 |
\<close> |
54321 | 216 |
|
217 |
||
58618 | 218 |
section \<open>Command-line invocation \label{sec:command-line}\<close> |
57320
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
219 |
|
58618 | 220 |
text \<open> |
62183 | 221 |
Isabelle/jEdit is normally invoked as a single-instance desktop application, |
222 |
based on platform-specific executables for Linux, Windows, Mac OS X. |
|
62014 | 223 |
|
62183 | 224 |
It is also possible to invoke the Prover IDE on the command-line, with some |
225 |
extra options and environment settings. The command-line usage of @{tool_def |
|
226 |
jedit} is as follows: |
|
61408
9020a3ba6c9a
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61199
diff
changeset
|
227 |
@{verbatim [display] |
9020a3ba6c9a
@{verbatim [display]} supersedes old alltt/ttbox;
wenzelm
parents:
61199
diff
changeset
|
228 |
\<open>Usage: isabelle jedit [OPTIONS] [FILES ...] |
57320
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
229 |
|
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
230 |
Options are: |
68370 | 231 |
-A NAME ancestor session for options -R and -S (default: parent) |
62039 | 232 |
-D NAME=X set JVM system property |
61132 | 233 |
-J OPTION add JVM runtime option |
66683 | 234 |
(default $JEDIT_JAVA_SYSTEM_OPTIONS $JEDIT_JAVA_OPTIONS) |
68370 | 235 |
-R NAME build image with requirements from other sessions |
236 |
-S NAME like option -R, with focus on selected session |
|
57320
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
237 |
-b build only |
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
238 |
-d DIR include session directory |
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
239 |
-f fresh build |
68541 | 240 |
-i NAME include session in name-space of theories |
61132 | 241 |
-j OPTION add jEdit runtime option |
66683 | 242 |
(default $JEDIT_OPTIONS) |
61132 | 243 |
-l NAME logic image name |
57320
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
244 |
-m MODE add print mode for output |
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
245 |
-n no build of session image on startup |
63987
ac96fe9224f6
just one option is enough -- "isabelle jedit" java process may be prefixed directly in the shell;
wenzelm
parents:
63871
diff
changeset
|
246 |
-p CMD ML process command prefix (process policy) |
57320
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
247 |
-s system build mode for session image |
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
248 |
|
61512
933463440449
more uniform command-line for "isabelle jedit" and the isabelle.Main app wrapper;
wenzelm
parents:
61506
diff
changeset
|
249 |
Start jEdit with Isabelle plugin setup and open FILES |
933463440449
more uniform command-line for "isabelle jedit" and the isabelle.Main app wrapper;
wenzelm
parents:
61506
diff
changeset
|
250 |
(default "$USER_HOME/Scratch.thy" or ":" for empty buffer).\<close>} |
57320
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
251 |
|
61574 | 252 |
The \<^verbatim>\<open>-l\<close> option specifies the session name of the logic image to be used |
253 |
for proof processing. Additional session root directories may be included |
|
66683 | 254 |
via option \<^verbatim>\<open>-d\<close> to augment the session name space (see also @{cite |
68472 | 255 |
"isabelle-system"}). By default, the specified image is checked and built on |
256 |
demand. The \<^verbatim>\<open>-s\<close> option determines where to store the result session image |
|
257 |
of @{tool build}. The \<^verbatim>\<open>-n\<close> option bypasses the implicit build process for |
|
258 |
the selected session image. |
|
57320
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
259 |
|
68472 | 260 |
The \<^verbatim>\<open>-R\<close> option builds an auxiliary logic image with all theories from |
261 |
other sessions that are not already present in its parent; it also opens the |
|
262 |
session \<^verbatim>\<open>ROOT\<close> entry in the editor to facilitate editing of the main |
|
263 |
session. The \<^verbatim>\<open>-S\<close> option is like \<^verbatim>\<open>-R\<close>, with a focus on the selected |
|
264 |
session and its descendants: the namespace of accessible theories is |
|
265 |
restricted accordingly. This reduces startup time for big projects, notably |
|
266 |
the ``Archive of Formal Proofs''. The \<^verbatim>\<open>-A\<close> option specifies and alternative |
|
267 |
ancestor session for options \<^verbatim>\<open>-R\<close> and \<^verbatim>\<open>-S\<close>: this allows to restructure the |
|
268 |
hierarchy of session images on the spot. |
|
66987
352b23c97ac8
support focus_session, for much faster startup of Isabelle/jEdit;
wenzelm
parents:
66977
diff
changeset
|
269 |
|
68541 | 270 |
The \<^verbatim>\<open>-i\<close> option includes additional sessions into the name-space of |
271 |
theories: multiple occurrences are possible. |
|
272 |
||
61574 | 273 |
The \<^verbatim>\<open>-m\<close> option specifies additional print modes for the prover process. |
274 |
Note that the system option @{system_option_ref jedit_print_mode} allows to |
|
275 |
do the same persistently (e.g.\ via the \<^emph>\<open>Plugin Options\<close> dialog of |
|
276 |
Isabelle/jEdit), without requiring command-line invocation. |
|
57320
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
277 |
|
66683 | 278 |
The \<^verbatim>\<open>-J\<close> and \<^verbatim>\<open>-j\<close> options pass additional low-level options to the JVM or |
279 |
jEdit, respectively. The defaults are provided by the Isabelle settings |
|
280 |
environment @{cite "isabelle-system"}, but note that these only work for the |
|
281 |
command-line tool described here, and not the regular application. |
|
57320
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
282 |
|
62039 | 283 |
The \<^verbatim>\<open>-D\<close> option allows to define JVM system properties; this is passed |
284 |
directly to the underlying \<^verbatim>\<open>java\<close> process. |
|
285 |
||
61574 | 286 |
The \<^verbatim>\<open>-b\<close> and \<^verbatim>\<open>-f\<close> options control the self-build mechanism of |
287 |
Isabelle/jEdit. This is only relevant for building from sources, which also |
|
63680 | 288 |
requires an auxiliary \<^verbatim>\<open>jedit_build\<close> component from |
68224 | 289 |
\<^url>\<open>https://isabelle.in.tum.de/components\<close>. The official Isabelle release |
62014 | 290 |
already includes a pre-built version of Isabelle/jEdit. |
291 |
||
62183 | 292 |
\<^bigskip> |
62014 | 293 |
It is also possible to connect to an already running Isabelle/jEdit process |
294 |
via @{tool_def jedit_client}: |
|
295 |
@{verbatim [display] |
|
296 |
\<open>Usage: isabelle jedit_client [OPTIONS] [FILES ...] |
|
297 |
||
298 |
Options are: |
|
299 |
-c only check presence of server |
|
300 |
-n only report server name |
|
301 |
-s NAME server name (default Isabelle) |
|
302 |
||
303 |
Connect to already running Isabelle/jEdit instance and open FILES\<close>} |
|
304 |
||
305 |
The \<^verbatim>\<open>-c\<close> option merely checks the presence of the server, producing a |
|
62183 | 306 |
process return code accordingly. |
62014 | 307 |
|
308 |
The \<^verbatim>\<open>-n\<close> option reports the server name, and the \<^verbatim>\<open>-s\<close> option provides a |
|
309 |
different server name. The default server name is the official distribution |
|
68394 | 310 |
name (e.g.\ \<^verbatim>\<open>Isabelle2018\<close>). Thus @{tool jedit_client} can connect to the |
62183 | 311 |
Isabelle desktop application without further options. |
62014 | 312 |
|
63987
ac96fe9224f6
just one option is enough -- "isabelle jedit" java process may be prefixed directly in the shell;
wenzelm
parents:
63871
diff
changeset
|
313 |
The \<^verbatim>\<open>-p\<close> option allows to override the implicit default of the system |
ac96fe9224f6
just one option is enough -- "isabelle jedit" java process may be prefixed directly in the shell;
wenzelm
parents:
63871
diff
changeset
|
314 |
option @{system_option_ref ML_process_policy} for ML processes started by |
ac96fe9224f6
just one option is enough -- "isabelle jedit" java process may be prefixed directly in the shell;
wenzelm
parents:
63871
diff
changeset
|
315 |
the Prover IDE, e.g. to control CPU affinity on multiprocessor systems. |
ac96fe9224f6
just one option is enough -- "isabelle jedit" java process may be prefixed directly in the shell;
wenzelm
parents:
63871
diff
changeset
|
316 |
|
62036 | 317 |
The JVM system property \<^verbatim>\<open>isabelle.jedit_server\<close> provides a different server |
62183 | 318 |
name, e.g.\ use \<^verbatim>\<open>isabelle jedit -Disabelle.jedit_server=\<close>\<open>name\<close> and |
62036 | 319 |
\<^verbatim>\<open>isabelle jedit_client -s\<close>~\<open>name\<close> to connect later on. |
62014 | 320 |
\<close> |
57320
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
321 |
|
00f2c8d1aa0b
more on command-line invocation -- moved material from system manual;
wenzelm
parents:
57319
diff
changeset
|
322 |
|
60291 | 323 |
section \<open>GUI rendering\<close> |
324 |
||
325 |
subsection \<open>Look-and-feel \label{sec:look-and-feel}\<close> |
|
54321 | 326 |
|
60291 | 327 |
text \<open> |
66683 | 328 |
jEdit is a Java/AWT/Swing application with the ambition to support |
60291 | 329 |
``native'' look-and-feel on all platforms, within the limits of what Oracle |
330 |
as Java provider and major operating system distributors allow (see also |
|
331 |
\secref{sec:problems}). |
|
54321 | 332 |
|
54329 | 333 |
Isabelle/jEdit enables platform-specific look-and-feel by default as |
57420 | 334 |
follows. |
54321 | 335 |
|
61574 | 336 |
\<^descr>[Linux:] The platform-independent \<^emph>\<open>Metal\<close> is used by default. |
54321 | 337 |
|
66683 | 338 |
The Linux-specific \<^emph>\<open>GTK+\<close> often works as well, but the overall GTK theme |
339 |
and options need to be selected to suite Java/AWT/Swing. Note that Java |
|
340 |
Virtual Machine has no direct influence of GTK rendering. |
|
54321 | 341 |
|
62183 | 342 |
\<^descr>[Windows:] Regular \<^emph>\<open>Windows\<close> is used by default. |
54372 | 343 |
|
61574 | 344 |
\<^descr>[Mac OS X:] Regular \<^emph>\<open>Mac OS X\<close> is used by default. |
54321 | 345 |
|
61574 | 346 |
The bundled \<^emph>\<open>MacOSX\<close> plugin provides various functions that are expected |
347 |
from applications on that particular platform: quit from menu or dock, |
|
348 |
preferences menu, drag-and-drop of text files on the application, |
|
349 |
full-screen mode for main editor windows. It is advisable to have the |
|
350 |
\<^emph>\<open>MacOSX\<close> plugin enabled all the time on that platform. |
|
54321 | 351 |
|
62183 | 352 |
Users may experiment with different Swing look-and-feels, but need to keep |
66683 | 353 |
in mind that this extra variance of GUI functionality often causes problems. |
354 |
The platform-independent \<^emph>\<open>Metal\<close> and \<^emph>\<open>Nimbus\<close> should always work on all |
|
355 |
platforms, although they are technically and stylistically outdated. The |
|
356 |
historic \<^emph>\<open>CDE/Motif\<close> should be ignored. |
|
54372 | 357 |
|
66683 | 358 |
Changing the look-and-feel in \<^emph>\<open>Global Options~/ Appearance\<close> may update the |
359 |
GUI only partially: proper restart of Isabelle/jEdit is usually required. |
|
60291 | 360 |
\<close> |
361 |
||
362 |
||
363 |
subsection \<open>Displays with very high resolution \label{sec:hdpi}\<close> |
|
364 |
||
365 |
text \<open> |
|
62183 | 366 |
In distant past, displays with $1024 \times 768$ or $1280 \times 1024$ |
367 |
pixels were considered ``high resolution'' and bitmap fonts with 12 or 14 |
|
68736 | 368 |
pixels as adequate for text rendering. After 2016, we have routinely seen |
369 |
much higher resolutions, e.g. ``Full HD'' at $1920 \times 1080$ pixels or |
|
370 |
``Ultra HD'' / ``4K'' at $3840 \times 2160$. |
|
60291 | 371 |
|
62183 | 372 |
GUI frameworks are usually lagging behind, with hard-wired icon sizes and |
373 |
tiny fonts. Java and jEdit do provide reasonable support for very high |
|
374 |
resolution, but this requires manual adjustments as described below. |
|
60291 | 375 |
|
61415 | 376 |
\<^medskip> |
62183 | 377 |
The \<^bold>\<open>operating-system\<close> usually provides some configuration for global |
378 |
scaling of text fonts, e.g.\ $120\%$--$250\%$ on Windows. This impacts |
|
379 |
regular GUI elements, when used with native look-and-feel: Linux \<^emph>\<open>GTK+\<close>, |
|
380 |
\<^emph>\<open>Windows\<close>, \<^emph>\<open>Mac OS X\<close>, respectively. Alternatively, it is possible to use |
|
381 |
the platform-independent \<^emph>\<open>Metal\<close> look-and-feel and readjust its main font |
|
382 |
sizes via jEdit options explained below. The Isabelle/jEdit \<^bold>\<open>application\<close> |
|
383 |
provides further options to adjust font sizes in particular GUI elements. |
|
384 |
Here is a summary of all relevant font properties: |
|
60291 | 385 |
|
61574 | 386 |
\<^item> \<^emph>\<open>Global Options / Text Area / Text font\<close>: the main text area font, |
387 |
which is also used as reference point for various derived font sizes, |
|
62185 | 388 |
e.g.\ the \<^emph>\<open>Output\<close> (\secref{sec:output}) and \<^emph>\<open>State\<close> |
62183 | 389 |
(\secref{sec:state-output}) panels. |
60291 | 390 |
|
61574 | 391 |
\<^item> \<^emph>\<open>Global Options / Gutter / Gutter font\<close>: the font for the gutter area |
392 |
left of the main text area, e.g.\ relevant for display of line numbers |
|
393 |
(disabled by default). |
|
60291 | 394 |
|
61574 | 395 |
\<^item> \<^emph>\<open>Global Options / Appearance / Button, menu and label font\<close> as well as |
396 |
\<^emph>\<open>List and text field font\<close>: this specifies the primary and secondary font |
|
62183 | 397 |
for the \<^emph>\<open>Metal\<close> look-and-feel (\secref{sec:look-and-feel}). |
60291 | 398 |
|
61574 | 399 |
\<^item> \<^emph>\<open>Plugin Options / Isabelle / General / Reset Font Size\<close>: the main text |
400 |
area font size for action @{action_ref "isabelle.reset-font-size"}, e.g.\ |
|
62183 | 401 |
relevant for quick scaling like in common web browsers. |
60291 | 402 |
|
61574 | 403 |
\<^item> \<^emph>\<open>Plugin Options / Console / General / Font\<close>: the console window font, |
404 |
e.g.\ relevant for Isabelle/Scala command-line. |
|
60291 | 405 |
|
61574 | 406 |
In \figref{fig:isabelle-jedit-hdpi} the \<^emph>\<open>Metal\<close> look-and-feel is configured |
407 |
with custom fonts at 30 pixels, and the main text area and console at 36 |
|
62183 | 408 |
pixels. This leads to decent rendering quality, despite the old-fashioned |
409 |
appearance of \<^emph>\<open>Metal\<close>. |
|
60291 | 410 |
|
62183 | 411 |
\begin{figure}[!htb] |
60291 | 412 |
\begin{center} |
413 |
\includegraphics[width=\textwidth]{isabelle-jedit-hdpi} |
|
414 |
\end{center} |
|
415 |
\caption{Metal look-and-feel with custom fonts for very high resolution} |
|
416 |
\label{fig:isabelle-jedit-hdpi} |
|
417 |
\end{figure} |
|
418 |
\<close> |
|
54321 | 419 |
|
420 |
||
62183 | 421 |
chapter \<open>Augmented jEdit functionality\<close> |
422 |
||
58618 | 423 |
section \<open>Dockable windows \label{sec:dockables}\<close> |
57316 | 424 |
|
58618 | 425 |
text \<open> |
61574 | 426 |
In jEdit terminology, a \<^emph>\<open>view\<close> is an editor window with one or more \<^emph>\<open>text |
427 |
areas\<close> that show the content of one or more \<^emph>\<open>buffers\<close>. A regular view may |
|
428 |
be surrounded by \<^emph>\<open>dockable windows\<close> that show additional information in |
|
429 |
arbitrary format, not just text; a \<^emph>\<open>plain view\<close> does not allow dockables. |
|
430 |
The \<^emph>\<open>dockable window manager\<close> of jEdit organizes these dockable windows, |
|
431 |
either as \<^emph>\<open>floating\<close> windows, or \<^emph>\<open>docked\<close> panels within one of the four |
|
432 |
margins of the view. There may be any number of floating instances of some |
|
433 |
dockable window, but at most one docked instance; jEdit actions that address |
|
434 |
\<^emph>\<open>the\<close> dockable window of a particular kind refer to the unique docked |
|
435 |
instance. |
|
57316 | 436 |
|
437 |
Dockables are used routinely in jEdit for important functionality like |
|
61574 | 438 |
\<^emph>\<open>HyperSearch Results\<close> or the \<^emph>\<open>File System Browser\<close>. Plugins often provide |
62184 | 439 |
a central dockable to access their main functionality, which may be opened |
440 |
by the user on demand. The Isabelle/jEdit plugin takes this approach to the |
|
61574 | 441 |
extreme: its plugin menu provides the entry-points to many panels that are |
442 |
managed as dockable windows. Some important panels are docked by default, |
|
62184 | 443 |
e.g.\ \<^emph>\<open>Documentation\<close>, \<^emph>\<open>State\<close>, \<^emph>\<open>Theories\<close> \<^emph>\<open>Output\<close>, \<^emph>\<open>Query\<close>. The user |
444 |
can change this arrangement easily and persistently. |
|
57316 | 445 |
|
446 |
Compared to plain jEdit, dockable window management in Isabelle/jEdit is |
|
447 |
slightly augmented according to the the following principles: |
|
448 |
||
61477 | 449 |
\<^item> Floating windows are dependent on the main window as \<^emph>\<open>dialog\<close> in |
57316 | 450 |
the sense of Java/AWT/Swing. Dialog windows always stay on top of the view, |
451 |
which is particularly important in full-screen mode. The desktop environment |
|
452 |
of the underlying platform may impose further policies on such dependent |
|
453 |
dialogs, in contrast to fully independent windows, e.g.\ some window |
|
454 |
management functions may be missing. |
|
455 |
||
61415 | 456 |
\<^item> Keyboard focus of the main view vs.\ a dockable window is carefully |
57316 | 457 |
managed according to the intended semantics, as a panel mainly for output or |
62184 | 458 |
input. For example, activating the \<^emph>\<open>Output\<close> (\secref{sec:output}) or State |
459 |
(\secref{sec:state-output}) panel via the dockable window manager returns |
|
460 |
keyboard focus to the main text area, but for \<^emph>\<open>Query\<close> (\secref{sec:query}) |
|
461 |
or \<^emph>\<open>Sledgehammer\<close> \secref{sec:sledgehammer} the focus is given to the main |
|
462 |
input field of that panel. |
|
57316 | 463 |
|
61415 | 464 |
\<^item> Panels that provide their own text area for output have an additional |
61477 | 465 |
dockable menu item \<^emph>\<open>Detach\<close>. This produces an independent copy of the |
466 |
current output as a floating \<^emph>\<open>Info\<close> window, which displays that content |
|
57316 | 467 |
independently of ongoing changes of the PIDE document-model. Note that |
468 |
Isabelle/jEdit popup windows (\secref{sec:tooltips-hyperlinks}) provide a |
|
61477 | 469 |
similar \<^emph>\<open>Detach\<close> operation as an icon. |
58618 | 470 |
\<close> |
57316 | 471 |
|
472 |
||
58618 | 473 |
section \<open>Isabelle symbols \label{sec:symbols}\<close> |
57319 | 474 |
|
58618 | 475 |
text \<open> |
61477 | 476 |
Isabelle sources consist of \<^emph>\<open>symbols\<close> that extend plain ASCII to allow |
57420 | 477 |
infinitely many mathematical symbols within the formal sources. This works |
478 |
without depending on particular encodings and varying Unicode |
|
61574 | 479 |
standards.\<^footnote>\<open>Raw Unicode characters within formal sources would compromise |
480 |
portability and reliability in the face of changing interpretation of |
|
481 |
special features of Unicode, such as Combining Characters or Bi-directional |
|
62184 | 482 |
Text.\<close> See @{cite "Wenzel:2011:CICM"}. |
57319 | 483 |
|
57420 | 484 |
For the prover back-end, formal text consists of ASCII characters that are |
61574 | 485 |
grouped according to some simple rules, e.g.\ as plain ``\<^verbatim>\<open>a\<close>'' or symbolic |
486 |
``\<^verbatim>\<open>\<alpha>\<close>''. For the editor front-end, a certain subset of symbols is rendered |
|
487 |
physically via Unicode glyphs, in order to show ``\<^verbatim>\<open>\<alpha>\<close>'' as ``\<open>\<alpha>\<close>'', for |
|
488 |
example. This symbol interpretation is specified by the Isabelle system |
|
63680 | 489 |
distribution in \<^file>\<open>$ISABELLE_HOME/etc/symbols\<close> and may be augmented by the |
490 |
user in @{path "$ISABELLE_HOME_USER/etc/symbols"}. |
|
57319 | 491 |
|
58554 | 492 |
The appendix of @{cite "isabelle-isar-ref"} gives an overview of the |
57319 | 493 |
standard interpretation of finitely many symbols from the infinite |
58554 | 494 |
collection. Uninterpreted symbols are displayed literally, e.g.\ |
61503 | 495 |
``\<^verbatim>\<open>\<foobar>\<close>''. Overlap of Unicode characters used in symbol |
58554 | 496 |
interpretation with informal ones (which might appear e.g.\ in comments) |
497 |
needs to be avoided. Raw Unicode characters within prover source files |
|
498 |
should be restricted to informal parts, e.g.\ to write text in non-latin |
|
499 |
alphabets in comments. |
|
61506 | 500 |
\<close> |
57319 | 501 |
|
61506 | 502 |
paragraph \<open>Encoding.\<close> |
67304
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
503 |
|
62184 | 504 |
text \<open>Technically, the Unicode interpretation of Isabelle symbols is an |
505 |
\<^emph>\<open>encoding\<close> called \<^verbatim>\<open>UTF-8-Isabelle\<close> in jEdit (\<^emph>\<open>not\<close> in the underlying |
|
66462 | 506 |
JVM). It is provided by the Isabelle Base plugin and enabled by default for |
67304
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
507 |
all source files in Isabelle/jEdit. |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
508 |
|
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
509 |
Sometimes such defaults are reset accidentally, or malformed UTF-8 sequences |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
510 |
in the text force jEdit to fall back on a different encoding like |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
511 |
\<^verbatim>\<open>ISO-8859-15\<close>. In that case, verbatim ``\<^verbatim>\<open>\<alpha>\<close>'' will be shown in the text |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
512 |
buffer instead of its Unicode rendering ``\<open>\<alpha>\<close>''. The jEdit menu operation |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
513 |
\<^emph>\<open>File~/ Reload with Encoding~/ UTF-8-Isabelle\<close> helps to resolve such |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
514 |
problems (after repairing malformed parts of the text). |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
515 |
|
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
516 |
If the loaded text already contains Unicode sequences that are in conflict |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
517 |
with the Isabelle symbol encoding, the fallback-encoding UTF-8 is used and |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
518 |
Isabelle symbols remain in literal \<^verbatim>\<open>\<symbol>\<close> form. The jEdit menu |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
519 |
operation \<^emph>\<open>Utilities~/ Buffer Options~/ Character encoding\<close> allows to |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
520 |
enforce the UTF-8-Isabelle, but this will also change original Unicode |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
521 |
text into Isabelle symbols when saving the file! |
3cf05d7cf174
more robust treatment of conflicts with existing Unicode text;
wenzelm
parents:
67262
diff
changeset
|
522 |
\<close> |
57319 | 523 |
|
61506 | 524 |
paragraph \<open>Font.\<close> |
525 |
text \<open>Correct rendering via Unicode requires a font that contains glyphs for |
|
62184 | 526 |
the corresponding codepoints. There are also various unusual symbols with |
527 |
particular purpose in Isabelle, e.g.\ control symbols and very long arrows. |
|
528 |
Isabelle/jEdit prefers its own application fonts \<^verbatim>\<open>IsabelleText\<close>, which |
|
529 |
ensures that standard collection of Isabelle symbols is actually shown on |
|
530 |
the screen (or printer) as expected. |
|
57319 | 531 |
|
57420 | 532 |
Note that a Java/AWT/Swing application can load additional fonts only if |
533 |
they are not installed on the operating system already! Some outdated |
|
61574 | 534 |
version of \<^verbatim>\<open>IsabelleText\<close> that happens to be provided by the operating |
535 |
system would prevent Isabelle/jEdit to use its bundled version. This could |
|
536 |
lead to missing glyphs (black rectangles), when the system version of |
|
537 |
\<^verbatim>\<open>IsabelleText\<close> is older than the application version. This problem can be |
|
538 |
avoided by refraining to ``install'' any version of \<^verbatim>\<open>IsabelleText\<close> in the |
|
62184 | 539 |
first place, although it might be tempting to use the same font in other |
540 |
applications. |
|
541 |
||
542 |
HTML pages generated by Isabelle refer to the same \<^verbatim>\<open>IsabelleText\<close> font as a |
|
543 |
server-side resource. Thus a web-browser can use that without requiring a |
|
544 |
locally installed copy. |
|
61506 | 545 |
\<close> |
57319 | 546 |
|
61506 | 547 |
paragraph \<open>Input methods.\<close> |
548 |
text \<open>In principle, Isabelle/jEdit could delegate the problem to produce |
|
549 |
Isabelle symbols in their Unicode rendering to the underlying operating |
|
550 |
system and its \<^emph>\<open>input methods\<close>. Regular jEdit also provides various ways to |
|
551 |
work with \<^emph>\<open>abbreviations\<close> to produce certain non-ASCII characters. Since |
|
552 |
none of these standard input methods work satisfactorily for the |
|
553 |
mathematical characters required for Isabelle, various specific |
|
554 |
Isabelle/jEdit mechanisms are provided. |
|
57319 | 555 |
|
57420 | 556 |
This is a summary for practically relevant input methods for Isabelle |
557 |
symbols. |
|
57319 | 558 |
|
61504 | 559 |
\<^enum> The \<^emph>\<open>Symbols\<close> panel: some GUI buttons allow to insert certain symbols in |
560 |
the text buffer. There are also tooltips to reveal the official Isabelle |
|
561 |
representation with some additional information about \<^emph>\<open>symbol |
|
562 |
abbreviations\<close> (see below). |
|
57319 | 563 |
|
61504 | 564 |
\<^enum> Copy/paste from decoded source files: text that is rendered as Unicode |
565 |
already can be re-used to produce further text. This also works between |
|
566 |
different applications, e.g.\ Isabelle/jEdit and some web browser or mail |
|
62184 | 567 |
client, as long as the same Unicode interpretation of Isabelle symbols is |
568 |
used. |
|
57319 | 569 |
|
61504 | 570 |
\<^enum> Copy/paste from prover output within Isabelle/jEdit. The same principles |
571 |
as for text buffers apply, but note that \<^emph>\<open>copy\<close> in secondary Isabelle/jEdit |
|
572 |
windows works via the keyboard shortcuts \<^verbatim>\<open>C+c\<close> or \<^verbatim>\<open>C+INSERT\<close>, while jEdit |
|
573 |
menu actions always refer to the primary text area! |
|
57319 | 574 |
|
62184 | 575 |
\<^enum> Completion provided by the Isabelle plugin (see \secref{sec:completion}). |
61504 | 576 |
Isabelle symbols have a canonical name and optional abbreviations. This can |
577 |
be used with the text completion mechanism of Isabelle/jEdit, to replace a |
|
578 |
prefix of the actual symbol like \<^verbatim>\<open>\<lambda>\<close>, or its name preceded by backslash |
|
579 |
\<^verbatim>\<open>\lambda\<close>, or its ASCII abbreviation \<^verbatim>\<open>%\<close> by the Unicode rendering. |
|
57319 | 580 |
|
581 |
The following table is an extract of the information provided by the |
|
63680 | 582 |
standard \<^file>\<open>$ISABELLE_HOME/etc/symbols\<close> file: |
57319 | 583 |
|
61415 | 584 |
\<^medskip> |
57319 | 585 |
\begin{tabular}{lll} |
61477 | 586 |
\<^bold>\<open>symbol\<close> & \<^bold>\<open>name with backslash\<close> & \<^bold>\<open>abbreviation\<close> \\\hline |
61503 | 587 |
\<open>\<lambda>\<close> & \<^verbatim>\<open>\lambda\<close> & \<^verbatim>\<open>%\<close> \\ |
588 |
\<open>\<Rightarrow>\<close> & \<^verbatim>\<open>\Rightarrow\<close> & \<^verbatim>\<open>=>\<close> \\ |
|
589 |
\<open>\<Longrightarrow>\<close> & \<^verbatim>\<open>\Longrightarrow\<close> & \<^verbatim>\<open>==>\<close> \\[0.5ex] |
|
590 |
\<open>\<And>\<close> & \<^verbatim>\<open>\And\<close> & \<^verbatim>\<open>!!\<close> \\ |
|
591 |
\<open>\<equiv>\<close> & \<^verbatim>\<open>\equiv\<close> & \<^verbatim>\<open>==\<close> \\[0.5ex] |
|
592 |
\<open>\<forall>\<close> & \<^verbatim>\<open>\forall\<close> & \<^verbatim>\<open>!\<close> \\ |
|
593 |
\<open>\<exists>\<close> & \<^verbatim>\<open>\exists\<close> & \<^verbatim>\<open>?\<close> \\ |
|
594 |
\<open>\<longrightarrow>\<close> & \<^verbatim>\<open>\longrightarrow\<close> & \<^verbatim>\<open>-->\<close> \\ |
|
595 |
\<open>\<and>\<close> & \<^verbatim>\<open>\and\<close> & \<^verbatim>\<open>&\<close> \\ |
|
596 |
\<open>\<or>\<close> & \<^verbatim>\<open>\or\<close> & \<^verbatim>\<open>|\<close> \\ |
|
597 |
\<open>\<not>\<close> & \<^verbatim>\<open>\not\<close> & \<^verbatim>\<open>~\<close> \\ |
|
598 |
\<open>\<noteq>\<close> & \<^verbatim>\<open>\noteq\<close> & \<^verbatim>\<open>~=\<close> \\ |
|
599 |
\<open>\<in>\<close> & \<^verbatim>\<open>\in\<close> & \<^verbatim>\<open>:\<close> \\ |
|
600 |
\<open>\<notin>\<close> & \<^verbatim>\<open>\notin\<close> & \<^verbatim>\<open>~:\<close> \\ |
|
57319 | 601 |
\end{tabular} |
61415 | 602 |
\<^medskip> |
61574 | 603 |
|
57420 | 604 |
Note that the above abbreviations refer to the input method. The logical |
605 |
notation provides ASCII alternatives that often coincide, but sometimes |
|
62184 | 606 |
deviate. This occasionally causes user confusion with old-fashioned Isabelle |
607 |
source that use ASCII replacement notation like \<^verbatim>\<open>!\<close> or \<^verbatim>\<open>ALL\<close> directly in |
|
608 |
the text. |
|
57319 | 609 |
|
610 |
On the other hand, coincidence of symbol abbreviations with ASCII |
|
61574 | 611 |
replacement syntax syntax helps to update old theory sources via explicit |
612 |
completion (see also \<^verbatim>\<open>C+b\<close> explained in \secref{sec:completion}). |
|
61506 | 613 |
\<close> |
57319 | 614 |
|
61506 | 615 |
paragraph \<open>Control symbols.\<close> |
616 |
text \<open>There are some special control symbols to modify the display style of a |
|
617 |
single symbol (without nesting). Control symbols may be applied to a region |
|
618 |
of selected text, either using the \<^emph>\<open>Symbols\<close> panel or keyboard shortcuts or |
|
619 |
jEdit actions. These editor operations produce a separate control symbol for |
|
620 |
each symbol in the text, in order to make the whole text appear in a certain |
|
621 |
style. |
|
57319 | 622 |
|
61415 | 623 |
\<^medskip> |
57319 | 624 |
\begin{tabular}{llll} |
61477 | 625 |
\<^bold>\<open>style\<close> & \<^bold>\<open>symbol\<close> & \<^bold>\<open>shortcut\<close> & \<^bold>\<open>action\<close> \\\hline |
61503 | 626 |
superscript & \<^verbatim>\<open>\<^sup>\<close> & \<^verbatim>\<open>C+e UP\<close> & @{action_ref "isabelle.control-sup"} \\ |
627 |
subscript & \<^verbatim>\<open>\<^sub>\<close> & \<^verbatim>\<open>C+e DOWN\<close> & @{action_ref "isabelle.control-sub"} \\ |
|
628 |
bold face & \<^verbatim>\<open>\<^bold>\<close> & \<^verbatim>\<open>C+e RIGHT\<close> & @{action_ref "isabelle.control-bold"} \\ |
|
629 |
emphasized & \<^verbatim>\<open>\<^emph>\<close> & \<^verbatim>\<open>C+e LEFT\<close> & @{action_ref "isabelle.control-emph"} \\ |
|
630 |
reset & & \<^verbatim>\<open>C+e BACK_SPACE\<close> & @{action_ref "isabelle.control-reset"} \\ |
|
57319 | 631 |
\end{tabular} |
61415 | 632 |
\<^medskip> |
61483 | 633 |
|
634 |
To produce a single control symbol, it is also possible to complete on |
|
61504 | 635 |
\<^verbatim>\<open>\sup\<close>, \<^verbatim>\<open>\sub\<close>, \<^verbatim>\<open>\bold\<close>, \<^verbatim>\<open>\emph\<close> as for regular symbols. |
61483 | 636 |
|
62184 | 637 |
The emphasized style only takes effect in document output (when used with a |
638 |
cartouche), but not in the editor. |
|
58618 | 639 |
\<close> |
57319 | 640 |
|
641 |
||
58618 | 642 |
section \<open>Scala console \label{sec:scala-console}\<close> |
57319 | 643 |
|
58618 | 644 |
text \<open> |
61574 | 645 |
The \<^emph>\<open>Console\<close> plugin manages various shells (command interpreters), e.g.\ |
646 |
\<^emph>\<open>BeanShell\<close>, which is the official jEdit scripting language, and the |
|
647 |
cross-platform \<^emph>\<open>System\<close> shell. Thus the console provides similar |
|
648 |
functionality than the Emacs buffers \<^verbatim>\<open>*scratch*\<close> and \<^verbatim>\<open>*shell*\<close>. |
|
57319 | 649 |
|
61574 | 650 |
Isabelle/jEdit extends the repertoire of the console by \<^emph>\<open>Scala\<close>, which is |
651 |
the regular Scala toplevel loop running inside the same JVM process as |
|
57420 | 652 |
Isabelle/jEdit itself. This means the Scala command interpreter has access |
57603 | 653 |
to the JVM name space and state of the running Prover IDE application. The |
61503 | 654 |
default environment imports the full content of packages \<^verbatim>\<open>isabelle\<close> and |
655 |
\<^verbatim>\<open>isabelle.jedit\<close>. |
|
57603 | 656 |
|
61574 | 657 |
For example, \<^verbatim>\<open>PIDE\<close> refers to the Isabelle/jEdit plugin object, and \<^verbatim>\<open>view\<close> |
658 |
to the current editor view of jEdit. The Scala expression |
|
659 |
\<^verbatim>\<open>PIDE.snapshot(view)\<close> makes a PIDE document snapshot of the current buffer |
|
660 |
within the current editor view. |
|
57319 | 661 |
|
662 |
This helps to explore Isabelle/Scala functionality interactively. Some care |
|
663 |
is required to avoid interference with the internals of the running |
|
62184 | 664 |
application. |
58618 | 665 |
\<close> |
57319 | 666 |
|
667 |
||
58618 | 668 |
section \<open>File-system access\<close> |
57318 | 669 |
|
58618 | 670 |
text \<open> |
57420 | 671 |
File specifications in jEdit follow various formats and conventions |
61477 | 672 |
according to \<^emph>\<open>Virtual File Systems\<close>, which may be also provided by |
61503 | 673 |
additional plugins. This allows to access remote files via the \<^verbatim>\<open>http:\<close> |
61574 | 674 |
protocol prefix, for example. Isabelle/jEdit attempts to work with the |
675 |
file-system model of jEdit as far as possible. In particular, theory sources |
|
676 |
are passed directly from the editor to the prover, without indirection via |
|
677 |
physical files. |
|
57318 | 678 |
|
57420 | 679 |
Despite the flexibility of URLs in jEdit, local files are particularly |
62184 | 680 |
important and are accessible without protocol prefix. The file path notation |
57420 | 681 |
is that of the Java Virtual Machine on the underlying platform. On Windows |
62184 | 682 |
the preferred form uses backslashes, but happens to accept forward slashes |
683 |
like Unix/POSIX as well. Further differences arise due to Windows drive |
|
60257 | 684 |
letters and network shares. |
57318 | 685 |
|
57331 | 686 |
The Java notation for files needs to be distinguished from the one of |
61477 | 687 |
Isabelle, which uses POSIX notation with forward slashes on \<^emph>\<open>all\<close> |
62184 | 688 |
platforms. Isabelle/ML on Windows uses Unix-style path notation, too, and |
689 |
driver letter representation as in Cygwin (e.g.\ \<^verbatim>\<open>/cygdrive/c\<close>). Moreover, |
|
690 |
environment variables from the Isabelle process may be used freely, e.g.\ |
|
63680 | 691 |
\<^file>\<open>$ISABELLE_HOME/etc/symbols\<close> or \<^file>\<open>$POLYML_HOME/README\<close>. There are special |
692 |
shortcuts: \<^dir>\<open>~\<close> for \<^dir>\<open>$USER_HOME\<close> and \<^dir>\<open>~~\<close> for \<^dir>\<open>$ISABELLE_HOME\<close>. |
|
57318 | 693 |
|
61415 | 694 |
\<^medskip> |
695 |
Since jEdit happens to support environment variables within file |
|
57420 | 696 |
specifications as well, it is natural to use similar notation within the |
697 |
editor, e.g.\ in the file-browser. This does not work in full generality, |
|
698 |
though, due to the bias of jEdit towards platform-specific notation and of |
|
699 |
Isabelle towards POSIX. Moreover, the Isabelle settings environment is not |
|
700 |
yet active when starting Isabelle/jEdit via its standard application |
|
60257 | 701 |
wrapper, in contrast to @{tool jedit} run from the command line |
702 |
(\secref{sec:command-line}). |
|
57318 | 703 |
|
63749 | 704 |
Isabelle/jEdit imitates important system settings within the Java process |
705 |
environment, in order to allow easy access to these important places from |
|
706 |
the editor: \<^verbatim>\<open>$ISABELLE_HOME\<close>, \<^verbatim>\<open>$ISABELLE_HOME_USER\<close>, \<^verbatim>\<open>$JEDIT_HOME\<close>, |
|
707 |
\<^verbatim>\<open>$JEDIT_SETTINGS\<close>. The file browser of jEdit also includes \<^emph>\<open>Favorites\<close> for |
|
708 |
these two important locations. |
|
57318 | 709 |
|
61415 | 710 |
\<^medskip> |
61574 | 711 |
Path specifications in prover input or output usually include formal markup |
712 |
that turns it into a hyperlink (see also \secref{sec:tooltips-hyperlinks}). |
|
713 |
This allows to open the corresponding file in the text editor, independently |
|
62184 | 714 |
of the path notation. If the path refers to a directory, the jEdit file |
715 |
browser is opened on it. |
|
57318 | 716 |
|
717 |
Formally checked paths in prover input are subject to completion |
|
61574 | 718 |
(\secref{sec:completion}): partial specifications are resolved via directory |
719 |
content and possible completions are offered in a popup. |
|
58618 | 720 |
\<close> |
57318 | 721 |
|
722 |
||
64515 | 723 |
section \<open>Indentation\<close> |
724 |
||
725 |
text \<open> |
|
726 |
Isabelle/jEdit augments the existing indentation facilities of jEdit to take |
|
727 |
the structure of theory and proof texts into account. There is also special |
|
728 |
support for unstructured proof scripts. |
|
729 |
||
730 |
\<^descr>[Syntactic indentation] follows the outer syntax of Isabelle/Isar. |
|
731 |
||
732 |
Action @{action "indent-lines"} (shortcut \<^verbatim>\<open>C+i\<close>) indents the current line |
|
733 |
according to command keywords and some command substructure: this |
|
734 |
approximation may need further manual tuning. |
|
735 |
||
736 |
Action @{action "isabelle.newline"} (shortcut \<^verbatim>\<open>ENTER\<close>) indents the old |
|
737 |
and the new line according to command keywords only: this leads to precise |
|
738 |
alignment of the main Isar language elements. This depends on option |
|
739 |
@{system_option_def "jedit_indent_newline"} (enabled by default). |
|
740 |
||
66681 | 741 |
Regular input (via keyboard or completion) indents the current line |
66683 | 742 |
whenever an new keyword is emerging at the start of the line. This depends |
743 |
on option @{system_option_def "jedit_indent_input"} (enabled by default). |
|
66681 | 744 |
|
64515 | 745 |
\<^descr>[Semantic indentation] adds additional white space to unstructured proof |
746 |
scripts (\<^theory_text>\<open>apply\<close> etc.) via number of subgoals. This requires information |
|
747 |
of ongoing document processing and may thus lag behind, when the user is |
|
748 |
editing too quickly; see also option @{system_option_def |
|
749 |
"jedit_script_indent"} and @{system_option_def |
|
750 |
"jedit_script_indent_limit"}. |
|
751 |
||
752 |
The above options are accessible in the menu \<^emph>\<open>Plugins / Plugin Options / |
|
66681 | 753 |
Isabelle / General\<close>. A prerequisite for advanced indentation is \<^emph>\<open>Utilities |
754 |
/ Buffer Options / Automatic indentation\<close>: it needs to be set to \<^verbatim>\<open>full\<close> |
|
755 |
(default). |
|
64515 | 756 |
\<close> |
757 |
||
758 |
||
62184 | 759 |
section \<open>SideKick parsers \label{sec:sidekick}\<close> |
760 |
||
761 |
text \<open> |
|
762 |
The \<^emph>\<open>SideKick\<close> plugin provides some general services to display buffer |
|
763 |
structure in a tree view. Isabelle/jEdit provides SideKick parsers for its |
|
64513 | 764 |
main mode for theory files, ML files, as well as some minor modes for the |
765 |
\<^verbatim>\<open>NEWS\<close> file (see \figref{fig:sidekick}), session \<^verbatim>\<open>ROOT\<close> files, system |
|
766 |
\<^verbatim>\<open>options\<close>, and Bib{\TeX} files (\secref{sec:bibtex}). |
|
62184 | 767 |
|
768 |
\begin{figure}[!htb] |
|
769 |
\begin{center} |
|
770 |
\includegraphics[scale=0.333]{sidekick} |
|
771 |
\end{center} |
|
772 |
\caption{The Isabelle NEWS file with SideKick tree view} |
|
773 |
\label{fig:sidekick} |
|
774 |
\end{figure} |
|
775 |
||
64513 | 776 |
The default SideKick parser for theory files is \<^verbatim>\<open>isabelle\<close>: it provides a |
777 |
tree-view on the formal document structure, with section headings at the top |
|
778 |
and formal specification elements at the bottom. The alternative parser |
|
779 |
\<^verbatim>\<open>isabelle-context\<close> shows nesting of context blocks according to \<^theory_text>\<open>begin \<dots> |
|
780 |
end\<close> structure. |
|
781 |
||
782 |
\<^medskip> |
|
783 |
Isabelle/ML files are structured according to semi-formal comments that are |
|
784 |
explained in @{cite "isabelle-implementation"}. This outline is turned into |
|
785 |
a tree-view by default, by using the \<^verbatim>\<open>isabelle-ml\<close> parser. There is also a |
|
786 |
folding mode of the same name, for hierarchic text folds within ML files. |
|
787 |
||
788 |
\<^medskip> |
|
62184 | 789 |
The special SideKick parser \<^verbatim>\<open>isabelle-markup\<close> exposes the uninterpreted |
790 |
markup tree of the PIDE document model of the current buffer. This is |
|
791 |
occasionally useful for informative purposes, but the amount of displayed |
|
792 |
information might cause problems for large buffers. |
|
793 |
\<close> |
|
794 |
||
795 |
||
58618 | 796 |
chapter \<open>Prover IDE functionality \label{sec:document-model}\<close> |
57315 | 797 |
|
58618 | 798 |
section \<open>Document model \label{sec:document-model}\<close> |
57322 | 799 |
|
58618 | 800 |
text \<open> |
57322 | 801 |
The document model is central to the PIDE architecture: the editor and the |
802 |
prover have a common notion of structured source text with markup, which is |
|
803 |
produced by formal processing. The editor is responsible for edits of |
|
804 |
document source, as produced by the user. The prover is responsible for |
|
805 |
reports of document markup, as produced by its processing in the background. |
|
806 |
||
807 |
Isabelle/jEdit handles classic editor events of jEdit, in order to connect |
|
808 |
the physical world of the GUI (with its singleton state) to the mathematical |
|
809 |
world of multiple document versions (with timeless and stateless updates). |
|
58618 | 810 |
\<close> |
57322 | 811 |
|
54322 | 812 |
|
58618 | 813 |
subsection \<open>Editor buffers and document nodes \label{sec:buffer-node}\<close> |
57322 | 814 |
|
58618 | 815 |
text \<open> |
61477 | 816 |
As a regular text editor, jEdit maintains a collection of \<^emph>\<open>buffers\<close> to |
57322 | 817 |
store text files; each buffer may be associated with any number of visible |
61574 | 818 |
\<^emph>\<open>text areas\<close>. Buffers are subject to an \<^emph>\<open>edit mode\<close> that is determined |
819 |
from the file name extension. The following modes are treated specifically |
|
820 |
in Isabelle/jEdit: |
|
57322 | 821 |
|
61415 | 822 |
\<^medskip> |
57322 | 823 |
\begin{tabular}{lll} |
62185 | 824 |
\<^bold>\<open>mode\<close> & \<^bold>\<open>file name\<close> & \<^bold>\<open>content\<close> \\\hline |
825 |
\<^verbatim>\<open>isabelle\<close> & \<^verbatim>\<open>*.thy\<close> & theory source \\ |
|
826 |
\<^verbatim>\<open>isabelle-ml\<close> & \<^verbatim>\<open>*.ML\<close> & Isabelle/ML source \\ |
|
827 |
\<^verbatim>\<open>sml\<close> & \<^verbatim>\<open>*.sml\<close> or \<^verbatim>\<open>*.sig\<close> & Standard ML source \\ |
|
828 |
\<^verbatim>\<open>isabelle-root\<close> & \<^verbatim>\<open>ROOT\<close> & session root \\ |
|
829 |
\<^verbatim>\<open>isabelle-options\<close> & & Isabelle options \\ |
|
830 |
\<^verbatim>\<open>isabelle-news\<close> & & Isabelle NEWS \\ |
|
57322 | 831 |
\end{tabular} |
61415 | 832 |
\<^medskip> |
54321 | 833 |
|
57322 | 834 |
All jEdit buffers are automatically added to the PIDE document-model as |
61574 | 835 |
\<^emph>\<open>document nodes\<close>. The overall document structure is defined by the theory |
836 |
nodes in two dimensions: |
|
57322 | 837 |
|
61574 | 838 |
\<^enum> via \<^bold>\<open>theory imports\<close> that are specified in the \<^emph>\<open>theory header\<close> using |
839 |
concrete syntax of the @{command_ref theory} command @{cite |
|
840 |
"isabelle-isar-ref"}; |
|
57322 | 841 |
|
66683 | 842 |
\<^enum> via \<^bold>\<open>auxiliary files\<close> that are included into a theory by \<^emph>\<open>load |
61574 | 843 |
commands\<close>, notably @{command_ref ML_file} and @{command_ref SML_file} |
844 |
@{cite "isabelle-isar-ref"}. |
|
54322 | 845 |
|
57322 | 846 |
In any case, source files are managed by the PIDE infrastructure: the |
847 |
physical file-system only plays a subordinate role. The relevant version of |
|
60257 | 848 |
source text is passed directly from the editor to the prover, using internal |
57322 | 849 |
communication channels. |
58618 | 850 |
\<close> |
57322 | 851 |
|
852 |
||
58618 | 853 |
subsection \<open>Theories \label{sec:theories}\<close> |
57322 | 854 |
|
58618 | 855 |
text \<open> |
61574 | 856 |
The \<^emph>\<open>Theories\<close> panel (see also \figref{fig:theories}) provides an overview |
857 |
of the status of continuous checking of theory nodes within the document |
|
66683 | 858 |
model. |
57322 | 859 |
|
62183 | 860 |
\begin{figure}[!htb] |
57339 | 861 |
\begin{center} |
862 |
\includegraphics[scale=0.333]{theories} |
|
863 |
\end{center} |
|
62185 | 864 |
\caption{Theories panel with an overview of the document-model, and jEdit |
865 |
text areas as editable views on some of the document nodes} |
|
57339 | 866 |
\label{fig:theories} |
867 |
\end{figure} |
|
868 |
||
64842 | 869 |
Theory imports are resolved automatically by the PIDE document model: all |
870 |
required files are loaded and stored internally, without the need to open |
|
871 |
corresponding jEdit buffers. Opening or closing editor buffers later on has |
|
66683 | 872 |
no direct impact on the formal document content: it only affects visibility. |
64842 | 873 |
|
66683 | 874 |
In contrast, auxiliary files (e.g.\ from @{command ML_file} commands) are |
875 |
\<^emph>\<open>not\<close> resolved within the editor by default, but the prover process takes |
|
876 |
care of that. This may be changed by enabling the system option |
|
877 |
@{system_option jedit_auto_resolve}: it ensures that all files are uniformly |
|
878 |
provided by the editor. |
|
54321 | 879 |
|
61415 | 880 |
\<^medskip> |
61574 | 881 |
The visible \<^emph>\<open>perspective\<close> of Isabelle/jEdit is defined by the collective |
882 |
view on theory buffers via open text areas. The perspective is taken as a |
|
883 |
hint for document processing: the prover ensures that those parts of a |
|
884 |
theory where the user is looking are checked, while other parts that are |
|
885 |
presently not required are ignored. The perspective is changed by opening or |
|
886 |
closing text area windows, or scrolling within a window. |
|
54322 | 887 |
|
61574 | 888 |
The \<^emph>\<open>Theories\<close> panel provides some further options to influence the process |
889 |
of continuous checking: it may be switched off globally to restrict the |
|
890 |
prover to superficial processing of command syntax. It is also possible to |
|
891 |
indicate theory nodes as \<^emph>\<open>required\<close> for continuous checking: this means |
|
892 |
such nodes and all their imports are always processed independently of the |
|
893 |
visibility status (if continuous checking is enabled). Big theory libraries |
|
62185 | 894 |
that are marked as required can have significant impact on performance! |
54322 | 895 |
|
66683 | 896 |
The \<^emph>\<open>Purge\<close> button restricts the document model to theories that are |
897 |
required for open editor buffers: inaccessible theories are removed and will |
|
898 |
be rechecked when opened or imported later. |
|
899 |
||
61415 | 900 |
\<^medskip> |
61574 | 901 |
Formal markup of checked theory content is turned into GUI rendering, based |
62185 | 902 |
on a standard repertoire known from mainstream IDEs for programming |
903 |
languages: colors, icons, highlighting, squiggly underlines, tooltips, |
|
904 |
hyperlinks etc. For outer syntax of Isabelle/Isar there is some traditional |
|
905 |
syntax-highlighting via static keywords and tokenization within the editor; |
|
906 |
this buffer syntax is determined from theory imports. In contrast, the |
|
907 |
painting of inner syntax (term language etc.)\ uses semantic information |
|
908 |
that is reported dynamically from the logical context. Thus the prover can |
|
909 |
provide additional markup to help the user to understand the meaning of |
|
910 |
formal text, and to produce more text with some add-on tools (e.g.\ |
|
911 |
information messages with \<^emph>\<open>sendback\<close> markup by automated provers or |
|
912 |
disprovers in the background). \<close> |
|
57322 | 913 |
|
914 |
||
58618 | 915 |
subsection \<open>Auxiliary files \label{sec:aux-files}\<close> |
57322 | 916 |
|
58618 | 917 |
text \<open> |
57329 | 918 |
Special load commands like @{command_ref ML_file} and @{command_ref |
58554 | 919 |
SML_file} @{cite "isabelle-isar-ref"} refer to auxiliary files within some |
57329 | 920 |
theory. Conceptually, the file argument of the command extends the theory |
921 |
source by the content of the file, but its editor buffer may be loaded~/ |
|
922 |
changed~/ saved separately. The PIDE document model propagates changes of |
|
923 |
auxiliary file content to the corresponding load command in the theory, to |
|
924 |
update and process it accordingly: changes of auxiliary file content are |
|
925 |
treated as changes of the corresponding load command. |
|
57323 | 926 |
|
61415 | 927 |
\<^medskip> |
61574 | 928 |
As a concession to the massive amount of ML files in Isabelle/HOL itself, |
929 |
the content of auxiliary files is only added to the PIDE document-model on |
|
930 |
demand, the first time when opened explicitly in the editor. There are |
|
931 |
further tricks to manage markup of ML files, such that Isabelle/HOL may be |
|
932 |
edited conveniently in the Prover IDE on small machines with only 8\,GB of |
|
933 |
main memory. Using \<^verbatim>\<open>Pure\<close> as logic session image, the exploration may start |
|
63680 | 934 |
at the top \<^file>\<open>$ISABELLE_HOME/src/HOL/Main.thy\<close> or the bottom |
66683 | 935 |
\<^file>\<open>$ISABELLE_HOME/src/HOL/HOL.thy\<close>, for example. It is also possible to |
936 |
explore the Isabelle/Pure bootstrap process (a virtual copy) by opening |
|
937 |
\<^file>\<open>$ISABELLE_HOME/src/Pure/ROOT.ML\<close> like a theory in the Prover IDE. |
|
57323 | 938 |
|
939 |
Initially, before an auxiliary file is opened in the editor, the prover |
|
940 |
reads its content from the physical file-system. After the file is opened |
|
941 |
for the first time in the editor, e.g.\ by following the hyperlink |
|
942 |
(\secref{sec:tooltips-hyperlinks}) for the argument of its @{command |
|
943 |
ML_file} command, the content is taken from the jEdit buffer. |
|
944 |
||
945 |
The change of responsibility from prover to editor counts as an update of |
|
946 |
the document content, so subsequent theory sources need to be re-checked. |
|
57420 | 947 |
When the buffer is closed, the responsibility remains to the editor: the |
948 |
file may be opened again without causing another document update. |
|
57323 | 949 |
|
950 |
A file that is opened in the editor, but its theory with the load command is |
|
951 |
not, is presently inactive in the document model. A file that is loaded via |
|
952 |
multiple load commands is associated to an arbitrary one: this situation is |
|
953 |
morally unsupported and might lead to confusion. |
|
954 |
||
61415 | 955 |
\<^medskip> |
61574 | 956 |
Output that refers to an auxiliary file is combined with that of the |
957 |
corresponding load command, and shown whenever the file or the command are |
|
958 |
active (see also \secref{sec:output}). |
|
57323 | 959 |
|
960 |
Warnings, errors, and other useful markup is attached directly to the |
|
62185 | 961 |
positions in the auxiliary file buffer, in the manner of standard IDEs. By |
63680 | 962 |
using the load command @{command SML_file} as explained in |
963 |
\<^file>\<open>$ISABELLE_HOME/src/Tools/SML/Examples.thy\<close>, Isabelle/jEdit may be used as |
|
57323 | 964 |
fully-featured IDE for Standard ML, independently of theory or proof |
61574 | 965 |
development: the required theory merely serves as some kind of project file |
966 |
for a collection of SML source modules. |
|
58618 | 967 |
\<close> |
54322 | 968 |
|
54352 | 969 |
|
58618 | 970 |
section \<open>Output \label{sec:output}\<close> |
54353 | 971 |
|
58618 | 972 |
text \<open> |
61574 | 973 |
Prover output consists of \<^emph>\<open>markup\<close> and \<^emph>\<open>messages\<close>. Both are directly |
974 |
attached to the corresponding positions in the original source text, and |
|
975 |
visualized in the text area, e.g.\ as text colours for free and bound |
|
976 |
variables, or as squiggly underlines for warnings, errors etc.\ (see also |
|
977 |
\figref{fig:output}). In the latter case, the corresponding messages are |
|
978 |
shown by hovering with the mouse over the highlighted text --- although in |
|
979 |
many situations the user should already get some clue by looking at the |
|
62185 | 980 |
position of the text highlighting, without seeing the message body itself. |
54357 | 981 |
|
62183 | 982 |
\begin{figure}[!htb] |
54357 | 983 |
\begin{center} |
57312 | 984 |
\includegraphics[scale=0.333]{output} |
54357 | 985 |
\end{center} |
62185 | 986 |
\caption{Multiple views on prover output: gutter with icon, text area with |
987 |
popup, text overview column, \<^emph>\<open>Theories\<close> panel, \<^emph>\<open>Output\<close> panel} |
|
54357 | 988 |
\label{fig:output} |
989 |
\end{figure} |
|
54353 | 990 |
|
62185 | 991 |
The ``gutter'' on the left-hand-side of the text area uses icons to |
992 |
provide a summary of the messages within the adjacent text line. Message |
|
61574 | 993 |
priorities are used to prefer errors over warnings, warnings over |
62185 | 994 |
information messages; other output is ignored. |
54353 | 995 |
|
62185 | 996 |
The ``text overview column'' on the right-hand-side of the text area uses |
997 |
similar information to paint small rectangles for the overall status of the |
|
998 |
whole text buffer. The graphics is scaled to fit the logical buffer length |
|
999 |
into the given window height. Mouse clicks on the overview area move the |
|
1000 |
cursor approximately to the corresponding text line in the buffer. |
|
54353 | 1001 |
|
62185 | 1002 |
The \<^emph>\<open>Theories\<close> panel provides another course-grained overview, but without |
1003 |
direct correspondence to text positions. The coloured rectangles represent |
|
1004 |
the amount of messages of a certain kind (warnings, errors, etc.) and the |
|
66683 | 1005 |
execution status of commands. The border of each rectangle indicates the |
1006 |
overall status of processing: a thick border means it is \<^emph>\<open>finished\<close> or |
|
1007 |
\<^emph>\<open>failed\<close> (with color for errors). A double-click on one of the theory |
|
1008 |
entries with their status overview opens the corresponding text buffer, |
|
1009 |
without moving the cursor to a specific point. |
|
54353 | 1010 |
|
61415 | 1011 |
\<^medskip> |
62185 | 1012 |
The \<^emph>\<open>Output\<close> panel displays prover messages that correspond to a given |
1013 |
command, within a separate window. The cursor position in the presently |
|
1014 |
active text area determines the prover command whose cumulative message |
|
1015 |
output is appended and shown in that window (in canonical order according to |
|
1016 |
the internal execution of the command). There are also control elements to |
|
1017 |
modify the update policy of the output wrt.\ continued editor movements: |
|
1018 |
\<^emph>\<open>Auto update\<close> and \<^emph>\<open>Update\<close>. This is particularly useful for multiple |
|
1019 |
instances of the \<^emph>\<open>Output\<close> panel to look at different situations. |
|
1020 |
Alternatively, the panel can be turned into a passive \<^emph>\<open>Info\<close> window via the |
|
1021 |
\<^emph>\<open>Detach\<close> menu item. |
|
54353 | 1022 |
|
62185 | 1023 |
Proof state is handled separately (\secref{sec:state-output}), but it is |
1024 |
also possible to tick the corresponding checkbox to append it to regular |
|
1025 |
output (\figref{fig:output-including-state}). This is a globally persistent |
|
1026 |
option: it affects all open panels and future editor sessions. |
|
54353 | 1027 |
|
62185 | 1028 |
\begin{figure}[!htb] |
1029 |
\begin{center} |
|
1030 |
\includegraphics[scale=0.333]{output-including-state} |
|
1031 |
\end{center} |
|
1032 |
\caption{Proof state display within the regular output panel} |
|
1033 |
\label{fig:output-including-state} |
|
1034 |
\end{figure} |
|
54353 | 1035 |
|
61415 | 1036 |
\<^medskip> |
62185 | 1037 |
Following the IDE principle, regular messages are attached to the original |
1038 |
source in the proper place and may be inspected on demand via popups. This |
|
1039 |
excludes messages that are somehow internal to the machinery of proof |
|
1040 |
checking, notably \<^emph>\<open>proof state\<close> and \<^emph>\<open>tracing\<close>. |
|
1041 |
||
1042 |
In any case, the same display technology is used for small popups and big |
|
1043 |
output windows. The formal text contains markup that may be explored |
|
1044 |
recursively via further popups and hyperlinks (see |
|
61574 | 1045 |
\secref{sec:tooltips-hyperlinks}), or clicked directly to initiate certain |
1046 |
actions (see \secref{sec:auto-tools} and \secref{sec:sledgehammer}). |
|
1047 |
\<close> |
|
54353 | 1048 |
|
1049 |
||
62185 | 1050 |
section \<open>Proof state \label{sec:state-output}\<close> |
62154 | 1051 |
|
1052 |
text \<open> |
|
62185 | 1053 |
The main purpose of the Prover IDE is to help the user editing proof |
1054 |
documents, with ongoing formal checking by the prover in the background. |
|
1055 |
This can be done to some extent in the main text area alone, especially for |
|
1056 |
well-structured Isar proofs. |
|
1057 |
||
1058 |
Nonetheless, internal proof state needs to be inspected in many situations |
|
1059 |
of exploration and ``debugging''. The \<^emph>\<open>State\<close> panel shows exclusively such |
|
1060 |
proof state messages without further distraction, while all other messages |
|
1061 |
are displayed in \<^emph>\<open>Output\<close> (\secref{sec:output}). |
|
1062 |
\Figref{fig:output-and-state} shows a typical GUI layout where both panels |
|
1063 |
are open. |
|
62154 | 1064 |
|
62183 | 1065 |
\begin{figure}[!htb] |
62154 | 1066 |
\begin{center} |
1067 |
\includegraphics[scale=0.333]{output-and-state} |
|
1068 |
\end{center} |
|
1069 |
\caption{Separate proof state display (right) and other output (bottom).} |
|
1070 |
\label{fig:output-and-state} |
|
1071 |
\end{figure} |
|
1072 |
||
62185 | 1073 |
Another typical arrangement has more than one \<^emph>\<open>State\<close> panel open (as |
1074 |
floating windows), with \<^emph>\<open>Auto update\<close> disabled to look at an old situation |
|
1075 |
while the proof text in the vicinity is changed. The \<^emph>\<open>Update\<close> button |
|
1076 |
triggers an explicit one-shot update; this operation is also available via |
|
1077 |
the action @{action "isabelle.update-state"} (keyboard shortcut \<^verbatim>\<open>S+ENTER\<close>). |
|
1078 |
||
1079 |
On small screens, it is occasionally useful to have all messages |
|
1080 |
concatenated in the regular \<^emph>\<open>Output\<close> panel, e.g.\ see |
|
1081 |
\figref{fig:output-including-state}. |
|
1082 |
||
1083 |
\<^medskip> |
|
1084 |
The mechanics of \<^emph>\<open>Output\<close> versus \<^emph>\<open>State\<close> are slightly different: |
|
1085 |
||
1086 |
\<^item> \<^emph>\<open>Output\<close> shows information that is continuously produced and already |
|
1087 |
present when the GUI wants to show it. This is implicitly controlled by |
|
1088 |
the visible perspective on the text. |
|
1089 |
||
1090 |
\<^item> \<^emph>\<open>State\<close> initiates a real-time query on demand, with a full round trip |
|
1091 |
including a fresh print operation on the prover side. This is controlled |
|
1092 |
explicitly when the cursor is moved to the next command (\<^emph>\<open>Auto update\<close>) |
|
1093 |
or the \<^emph>\<open>Update\<close> operation is triggered. |
|
1094 |
||
1095 |
This can make a difference in GUI responsibility and resource usage within |
|
1096 |
the prover process. Applications with very big proof states that are only |
|
1097 |
inspected in isolation work better with the \<^emph>\<open>State\<close> panel. |
|
62154 | 1098 |
\<close> |
1099 |
||
62185 | 1100 |
|
58618 | 1101 |
section \<open>Query \label{sec:query}\<close> |
57311 | 1102 |
|
58618 | 1103 |
text \<open> |
61574 | 1104 |
The \<^emph>\<open>Query\<close> panel provides various GUI forms to request extra information |
62249 | 1105 |
from the prover, as a replacement of old-style diagnostic commands like |
1106 |
@{command find_theorems}. There are input fields and buttons for a |
|
1107 |
particular query command, with output in a dedicated text area. |
|
57311 | 1108 |
|
62249 | 1109 |
The main query modes are presented as separate tabs: \<^emph>\<open>Find Theorems\<close>, |
1110 |
\<^emph>\<open>Find Constants\<close>, \<^emph>\<open>Print Context\<close>, e.g.\ see \figref{fig:query}. As usual |
|
1111 |
in jEdit, multiple \<^emph>\<open>Query\<close> windows may be active at the same time: any |
|
1112 |
number of floating instances, but at most one docked instance (which is used |
|
1113 |
by default). |
|
57313 | 1114 |
|
62183 | 1115 |
\begin{figure}[!htb] |
57313 | 1116 |
\begin{center} |
1117 |
\includegraphics[scale=0.333]{query} |
|
1118 |
\end{center} |
|
62154 | 1119 |
\caption{An instance of the Query panel: find theorems} |
57313 | 1120 |
\label{fig:query} |
1121 |
\end{figure} |
|
57311 | 1122 |
|
61415 | 1123 |
\<^medskip> |
1124 |
The following GUI elements are common to all query modes: |
|
57311 | 1125 |
|
61574 | 1126 |
\<^item> The spinning wheel provides feedback about the status of a pending query |
1127 |
wrt.\ the evaluation of its context and its own operation. |
|
57311 | 1128 |
|
61574 | 1129 |
\<^item> The \<^emph>\<open>Apply\<close> button attaches a fresh query invocation to the current |
1130 |
context of the command where the cursor is pointing in the text. |
|
57311 | 1131 |
|
61574 | 1132 |
\<^item> The \<^emph>\<open>Search\<close> field allows to highlight query output according to some |
1133 |
regular expression, in the notation that is commonly used on the Java |
|
63680 | 1134 |
platform.\<^footnote>\<open>\<^url>\<open>https://docs.oracle.com/javase/8/docs/api/java/util/regex/Pattern.html\<close>\<close> |
61574 | 1135 |
This may serve as an additional visual filter of the result. |
57311 | 1136 |
|
61574 | 1137 |
\<^item> The \<^emph>\<open>Zoom\<close> box controls the font size of the output area. |
57311 | 1138 |
|
1139 |
All query operations are asynchronous: there is no need to wait for the |
|
1140 |
evaluation of the document for the query context, nor for the query |
|
61477 | 1141 |
operation itself. Query output may be detached as independent \<^emph>\<open>Info\<close> |
57311 | 1142 |
window, using a menu operation of the dockable window manager. The printed |
1143 |
result usually provides sufficient clues about the original query, with some |
|
1144 |
hyperlink to its context (via markup of its head line). |
|
58618 | 1145 |
\<close> |
57311 | 1146 |
|
1147 |
||
58618 | 1148 |
subsection \<open>Find theorems\<close> |
57311 | 1149 |
|
58618 | 1150 |
text \<open> |
61574 | 1151 |
The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Find Theorems\<close> mode retrieves facts from the theory |
1152 |
or proof context matching all of given criteria in the \<^emph>\<open>Find\<close> text field. A |
|
68042 | 1153 |
single criterion has the following syntax: |
57311 | 1154 |
|
57313 | 1155 |
@{rail \<open> |
62969 | 1156 |
('-'?) ('name' ':' @{syntax name} | 'intro' | 'elim' | 'dest' | |
57313 | 1157 |
'solves' | 'simp' ':' @{syntax term} | @{syntax term}) |
1158 |
\<close>} |
|
1159 |
||
61574 | 1160 |
See also the Isar command @{command_ref find_theorems} in @{cite |
1161 |
"isabelle-isar-ref"}. |
|
58618 | 1162 |
\<close> |
57311 | 1163 |
|
1164 |
||
58618 | 1165 |
subsection \<open>Find constants\<close> |
57311 | 1166 |
|
58618 | 1167 |
text \<open> |
61574 | 1168 |
The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Find Constants\<close> mode prints all constants whose type |
1169 |
meets all of the given criteria in the \<^emph>\<open>Find\<close> text field. A single |
|
68042 | 1170 |
criterion has the following syntax: |
57313 | 1171 |
|
1172 |
@{rail \<open> |
|
1173 |
('-'?) |
|
62969 | 1174 |
('name' ':' @{syntax name} | 'strict' ':' @{syntax type} | @{syntax type}) |
57313 | 1175 |
\<close>} |
1176 |
||
58554 | 1177 |
See also the Isar command @{command_ref find_consts} in @{cite |
1178 |
"isabelle-isar-ref"}. |
|
58618 | 1179 |
\<close> |
57311 | 1180 |
|
1181 |
||
58618 | 1182 |
subsection \<open>Print context\<close> |
57311 | 1183 |
|
58618 | 1184 |
text \<open> |
61574 | 1185 |
The \<^emph>\<open>Query\<close> panel in \<^emph>\<open>Print Context\<close> mode prints information from the |
1186 |
theory or proof context, or proof state. See also the Isar commands |
|
57329 | 1187 |
@{command_ref print_context}, @{command_ref print_cases}, @{command_ref |
62249 | 1188 |
print_term_bindings}, @{command_ref print_theorems}, described in @{cite |
1189 |
"isabelle-isar-ref"}. |
|
58618 | 1190 |
\<close> |
57311 | 1191 |
|
1192 |
||
58618 | 1193 |
section \<open>Tooltips and hyperlinks \label{sec:tooltips-hyperlinks}\<close> |
54352 | 1194 |
|
58618 | 1195 |
text \<open> |
62249 | 1196 |
Formally processed text (prover input or output) contains rich markup that |
1197 |
can be explored by using the \<^verbatim>\<open>CONTROL\<close> modifier key on Linux and Windows, |
|
1198 |
or \<^verbatim>\<open>COMMAND\<close> on Mac OS X. Hovering with the mouse while the modifier is |
|
1199 |
pressed reveals a \<^emph>\<open>tooltip\<close> (grey box over the text with a yellow popup) |
|
1200 |
and/or a \<^emph>\<open>hyperlink\<close> (black rectangle over the text with change of mouse |
|
1201 |
pointer); see also \figref{fig:tooltip}. |
|
54331 | 1202 |
|
62183 | 1203 |
\begin{figure}[!htb] |
54331 | 1204 |
\begin{center} |
57312 | 1205 |
\includegraphics[scale=0.5]{popup1} |
54331 | 1206 |
\end{center} |
54356 | 1207 |
\caption{Tooltip and hyperlink for some formal entity} |
54350 | 1208 |
\label{fig:tooltip} |
54331 | 1209 |
\end{figure} |
1210 |
||
62249 | 1211 |
Tooltip popups use the same rendering technology as the main text area, and |
61574 | 1212 |
further tooltips and/or hyperlinks may be exposed recursively by the same |
1213 |
mechanism; see \figref{fig:nested-tooltips}. |
|
54323 | 1214 |
|
62183 | 1215 |
\begin{figure}[!htb] |
54331 | 1216 |
\begin{center} |
57312 | 1217 |
\includegraphics[scale=0.5]{popup2} |
54331 | 1218 |
\end{center} |
54356 | 1219 |
\caption{Nested tooltips over formal entities} |
54350 | 1220 |
\label{fig:nested-tooltips} |
54331 | 1221 |
\end{figure} |
54350 | 1222 |
|
61574 | 1223 |
The tooltip popup window provides some controls to \<^emph>\<open>close\<close> or \<^emph>\<open>detach\<close> the |
1224 |
window, turning it into a separate \<^emph>\<open>Info\<close> window managed by jEdit. The |
|
1225 |
\<^verbatim>\<open>ESCAPE\<close> key closes \<^emph>\<open>all\<close> popups, which is particularly relevant when |
|
1226 |
nested tooltips are stacking up. |
|
54352 | 1227 |
|
61415 | 1228 |
\<^medskip> |
61574 | 1229 |
A black rectangle in the text indicates a hyperlink that may be followed by |
1230 |
a mouse click (while the \<^verbatim>\<open>CONTROL\<close> or \<^verbatim>\<open>COMMAND\<close> modifier key is still |
|
1231 |
pressed). Such jumps to other text locations are recorded by the |
|
1232 |
\<^emph>\<open>Navigator\<close> plugin, which is bundled with Isabelle/jEdit and enabled by |
|
62249 | 1233 |
default. There are usually navigation arrows in the main jEdit toolbar. |
54352 | 1234 |
|
62249 | 1235 |
Note that the link target may be a file that is itself not subject to formal |
1236 |
document processing of the editor session and thus prevents further |
|
61574 | 1237 |
exploration: the chain of hyperlinks may end in some source file of the |
62249 | 1238 |
underlying logic image, or within the ML bootstrap sources of Isabelle/Pure. |
61574 | 1239 |
\<close> |
54321 | 1240 |
|
1241 |
||
64514 | 1242 |
section \<open>Formal scopes and semantic selection\<close> |
1243 |
||
1244 |
text \<open> |
|
1245 |
Formal entities are semantically annotated in the source text as explained |
|
1246 |
in \secref{sec:tooltips-hyperlinks}. A \<^emph>\<open>formal scope\<close> consists of the |
|
1247 |
defining position with all its referencing positions. This correspondence is |
|
1248 |
highlighted in the text according to the cursor position, see also |
|
1249 |
\figref{fig:scope1}. Here the referencing positions are rendered with an |
|
1250 |
additional border, in reminiscence to a hyperlink: clicking there moves the |
|
1251 |
cursor to the original defining position. |
|
1252 |
||
1253 |
\begin{figure}[!htb] |
|
1254 |
\begin{center} |
|
1255 |
\includegraphics[scale=0.5]{scope1} |
|
1256 |
\end{center} |
|
1257 |
\caption{Scope of formal entity: defining vs.\ referencing positions} |
|
1258 |
\label{fig:scope1} |
|
1259 |
\end{figure} |
|
1260 |
||
1261 |
The action @{action_def "isabelle.select-entity"} (shortcut \<^verbatim>\<open>CS+ENTER\<close>) |
|
1262 |
supports semantic selection of all occurrences of the formal entity at the |
|
1263 |
caret position. This facilitates systematic renaming, using regular jEdit |
|
1264 |
editing of a multi-selection, see also \figref{fig:scope2}. |
|
1265 |
||
1266 |
\begin{figure}[!htb] |
|
1267 |
\begin{center} |
|
1268 |
\includegraphics[scale=0.5]{scope2} |
|
1269 |
\end{center} |
|
1270 |
\caption{The result of semantic selection and systematic renaming} |
|
1271 |
\label{fig:scope2} |
|
1272 |
\end{figure} |
|
1273 |
\<close> |
|
1274 |
||
1275 |
||
58618 | 1276 |
section \<open>Completion \label{sec:completion}\<close> |
57324 | 1277 |
|
58618 | 1278 |
text \<open> |
61477 | 1279 |
Smart completion of partial input is the IDE functionality \<^emph>\<open>par |
1280 |
excellance\<close>. Isabelle/jEdit combines several sources of information to |
|
57328 | 1281 |
achieve that. Despite its complexity, it should be possible to get some idea |
1282 |
how completion works by experimentation, based on the overview of completion |
|
57335 | 1283 |
varieties in \secref{sec:completion-varieties}. The remaining subsections |
1284 |
explain concepts around completion more systematically. |
|
57325 | 1285 |
|
61415 | 1286 |
\<^medskip> |
61477 | 1287 |
\<^emph>\<open>Explicit completion\<close> is triggered by the action @{action_ref |
61574 | 1288 |
"isabelle.complete"}, which is bound to the keyboard shortcut \<^verbatim>\<open>C+b\<close>, and |
1289 |
thus overrides the jEdit default for @{action_ref "complete-word"}. |
|
57335 | 1290 |
|
61574 | 1291 |
\<^emph>\<open>Implicit completion\<close> hooks into the regular keyboard input stream of the |
1292 |
editor, with some event filtering and optional delays. |
|
54361 | 1293 |
|
61415 | 1294 |
\<^medskip> |
61574 | 1295 |
Completion options may be configured in \<^emph>\<open>Plugin Options~/ Isabelle~/ |
1296 |
General~/ Completion\<close>. These are explained in further detail below, whenever |
|
1297 |
relevant. There is also a summary of options in |
|
57328 | 1298 |
\secref{sec:completion-options}. |
1299 |
||
57335 | 1300 |
The asynchronous nature of PIDE interaction means that information from the |
1301 |
prover is delayed --- at least by a full round-trip of the document update |
|
1302 |
protocol. The default options already take this into account, with a |
|
57324 | 1303 |
sufficiently long completion delay to speculate on the availability of all |
57335 | 1304 |
relevant information from the editor and the prover, before completing text |
1305 |
immediately or producing a popup. Although there is an inherent danger of |
|
1306 |
non-deterministic behaviour due to such real-time parameters, the general |
|
1307 |
completion policy aims at determined results as far as possible. |
|
58618 | 1308 |
\<close> |
57324 | 1309 |
|
1310 |
||
58618 | 1311 |
subsection \<open>Varieties of completion \label{sec:completion-varieties}\<close> |
57324 | 1312 |
|
58618 | 1313 |
subsubsection \<open>Built-in templates\<close> |
57324 | 1314 |
|
58618 | 1315 |
text \<open> |
57327 | 1316 |
Isabelle is ultimately a framework of nested sub-languages of different |
57328 | 1317 |
kinds and purposes. The completion mechanism supports this by the following |
1318 |
built-in templates: |
|
1319 |
||
64513 | 1320 |
\<^descr> \<^verbatim>\<open>`\<close> (single ASCII back-quote) or \<^verbatim>\<open>"\<close> (double ASCII quote) support |
1321 |
\<^emph>\<open>quotations\<close> via text cartouches. There are three selections, which are |
|
1322 |
always presented in the same order and do not depend on any context |
|
1323 |
information. The default choice produces a template ``\<open>\<open>\<box>\<close>\<close>'', where the |
|
1324 |
box indicates the cursor position after insertion; the other choices help |
|
1325 |
to repair the block structure of unbalanced text cartouches. |
|
57324 | 1326 |
|
61574 | 1327 |
\<^descr> \<^verbatim>\<open>@{\<close> is completed to the template ``\<open>@{\<box>}\<close>'', where the box indicates |
1328 |
the cursor position after insertion. Here it is convenient to use the |
|
1329 |
wildcard ``\<^verbatim>\<open>__\<close>'' or a more specific name prefix to let semantic |
|
1330 |
completion of name-space entries propose antiquotation names. |
|
57335 | 1331 |
|
1332 |
With some practice, input of quoted sub-languages and antiquotations of |
|
1333 |
embedded languages should work fluently. Note that national keyboard layouts |
|
64513 | 1334 |
might cause problems with back-quote as dead key, but double quote can be |
1335 |
used instead. |
|
58618 | 1336 |
\<close> |
57335 | 1337 |
|
57327 | 1338 |
|
58618 | 1339 |
subsubsection \<open>Syntax keywords\<close> |
57335 | 1340 |
|
58618 | 1341 |
text \<open> |
57335 | 1342 |
Syntax completion tables are determined statically from the keywords of the |
1343 |
``outer syntax'' of the underlying edit mode: for theory files this is the |
|
60257 | 1344 |
syntax of Isar commands according to the cumulative theory imports. |
57327 | 1345 |
|
57335 | 1346 |
Keywords are usually plain words, which means the completion mechanism only |
1347 |
inserts them directly into the text for explicit completion |
|
1348 |
(\secref{sec:completion-input}), but produces a popup |
|
1349 |
(\secref{sec:completion-popup}) otherwise. |
|
1350 |
||
1351 |
At the point where outer syntax keywords are defined, it is possible to |
|
1352 |
specify an alternative replacement string to be inserted instead of the |
|
1353 |
keyword itself. An empty string means to suppress the keyword altogether, |
|
1354 |
which is occasionally useful to avoid confusion, e.g.\ the rare keyword |
|
61493 | 1355 |
@{command simproc_setup} vs.\ the frequent name-space entry \<open>simp\<close>. |
58618 | 1356 |
\<close> |
57324 | 1357 |
|
1358 |
||
58618 | 1359 |
subsubsection \<open>Isabelle symbols\<close> |
57324 | 1360 |
|
58618 | 1361 |
text \<open> |
57325 | 1362 |
The completion tables for Isabelle symbols (\secref{sec:symbols}) are |
63680 | 1363 |
determined statically from \<^file>\<open>$ISABELLE_HOME/etc/symbols\<close> and @{path |
1364 |
"$ISABELLE_HOME_USER/etc/symbols"} for each symbol specification as follows: |
|
57325 | 1365 |
|
61415 | 1366 |
\<^medskip> |
57325 | 1367 |
\begin{tabular}{ll} |
61477 | 1368 |
\<^bold>\<open>completion entry\<close> & \<^bold>\<open>example\<close> \\\hline |
61503 | 1369 |
literal symbol & \<^verbatim>\<open>\<forall>\<close> \\ |
1370 |
symbol name with backslash & \<^verbatim>\<open>\\<close>\<^verbatim>\<open>forall\<close> \\ |
|
1371 |
symbol abbreviation & \<^verbatim>\<open>ALL\<close> or \<^verbatim>\<open>!\<close> \\ |
|
57325 | 1372 |
\end{tabular} |
61415 | 1373 |
\<^medskip> |
57325 | 1374 |
|
57335 | 1375 |
When inserted into the text, the above examples all produce the same Unicode |
61503 | 1376 |
rendering \<open>\<forall>\<close> of the underlying symbol \<^verbatim>\<open>\<forall>\<close>. |
57325 | 1377 |
|
61574 | 1378 |
A symbol abbreviation that is a plain word, like \<^verbatim>\<open>ALL\<close>, is treated like a |
1379 |
syntax keyword. Non-word abbreviations like \<^verbatim>\<open>-->\<close> are inserted more |
|
1380 |
aggressively, except for single-character abbreviations like \<^verbatim>\<open>!\<close> above. |
|
57324 | 1381 |
|
62250 | 1382 |
Completion via abbreviations like \<^verbatim>\<open>ALL\<close> or \<^verbatim>\<open>-->\<close> depends on the semantic |
1383 |
language context (\secref{sec:completion-context}). In contrast, backslash |
|
68736 | 1384 |
sequences like \<^verbatim>\<open>\forall\<close> \<^verbatim>\<open>\<forall>\<close> are always possible, but require additional |
1385 |
interaction to confirm (via popup). This is important in ambiguous |
|
1386 |
situations, e.g.\ for Isabelle document source, which may contain formal |
|
1387 |
symbols or informal {\LaTeX} macros. Backslash sequences also help when |
|
1388 |
input is broken, and thus escapes its normal semantic context: e.g.\ |
|
1389 |
antiquotations or string literals in ML, which do not allow arbitrary |
|
1390 |
backslash sequences. |
|
62250 | 1391 |
|
68736 | 1392 |
Special symbols like \<^verbatim>\<open>\<comment>\<close> or control symbols like \<^verbatim>\<open>\<^cancel>\<close>, |
1393 |
\<^verbatim>\<open>\<^latex>\<close>, \<^verbatim>\<open>\<^binding>\<close> can have an argument: completing on a name |
|
1394 |
prefix offers a template with an empty cartouche. Thus completion of \<^verbatim>\<open>\co\<close> |
|
1395 |
or \<^verbatim>\<open>\ca\<close> allows to compose formal document comments quickly.\<^footnote>\<open>It is |
|
1396 |
customary to put a space between \<^verbatim>\<open>\<comment>\<close> and its argument, while |
|
1397 |
control symbols do \<^emph>\<open>not\<close> allow extra space here.\<close> |
|
58618 | 1398 |
\<close> |
57324 | 1399 |
|
1400 |
||
64513 | 1401 |
subsubsection \<open>User-defined abbreviations\<close> |
1402 |
||
1403 |
text \<open> |
|
1404 |
The theory header syntax supports abbreviations via the \<^theory_text>\<open>abbrevs\<close> keyword |
|
1405 |
@{cite "isabelle-isar-ref"}. This is a slight generalization of built-in |
|
1406 |
templates and abbreviations for Isabelle symbols, as explained above. |
|
1407 |
Examples may be found in the Isabelle sources, by searching for |
|
1408 |
``\<^verbatim>\<open>abbrevs\<close>'' in \<^verbatim>\<open>*.thy\<close> files. |
|
1409 |
||
1410 |
The \<^emph>\<open>Symbols\<close> panel shows the abbreviations that are available in the |
|
1411 |
current theory buffer (according to its \<^theory_text>\<open>imports\<close>) in the \<^verbatim>\<open>Abbrevs\<close> tab. |
|
1412 |
\<close> |
|
1413 |
||
1414 |
||
58618 | 1415 |
subsubsection \<open>Name-space entries\<close> |
57324 | 1416 |
|
58618 | 1417 |
text \<open> |
57324 | 1418 |
This is genuine semantic completion, using information from the prover, so |
61477 | 1419 |
it requires some delay. A \<^emph>\<open>failed name-space lookup\<close> produces an error |
57335 | 1420 |
message that is annotated with a list of alternative names that are legal. |
1421 |
The list of results is truncated according to the system option |
|
1422 |
@{system_option_ref completion_limit}. The completion mechanism takes this |
|
1423 |
into account when collecting information on the prover side. |
|
57324 | 1424 |
|
61574 | 1425 |
Already recognized names are \<^emph>\<open>not\<close> completed further, but completion may be |
1426 |
extended by appending a suffix of underscores. This provokes a failed |
|
57328 | 1427 |
lookup, and another completion attempt while ignoring the underscores. For |
61574 | 1428 |
example, in a name space where \<^verbatim>\<open>foo\<close> and \<^verbatim>\<open>foobar\<close> are known, the input |
1429 |
\<^verbatim>\<open>foo\<close> remains unchanged, but \<^verbatim>\<open>foo_\<close> may be completed to \<^verbatim>\<open>foo\<close> or |
|
1430 |
\<^verbatim>\<open>foobar\<close>. |
|
57324 | 1431 |
|
61574 | 1432 |
The special identifier ``\<^verbatim>\<open>__\<close>'' serves as a wild-card for arbitrary |
1433 |
completion: it exposes the name-space content to the completion mechanism |
|
1434 |
(truncated according to @{system_option completion_limit}). This is |
|
1435 |
occasionally useful to explore an unknown name-space, e.g.\ in some |
|
57324 | 1436 |
template. |
58618 | 1437 |
\<close> |
57324 | 1438 |
|
1439 |
||
58618 | 1440 |
subsubsection \<open>File-system paths\<close> |
57324 | 1441 |
|
58618 | 1442 |
text \<open> |
62250 | 1443 |
Depending on prover markup about file-system paths in the source text, e.g.\ |
1444 |
for the argument of a load command (\secref{sec:aux-files}), the completion |
|
1445 |
mechanism explores the directory content and offers the result as completion |
|
1446 |
popup. Relative path specifications are understood wrt.\ the \<^emph>\<open>master |
|
1447 |
directory\<close> of the document node (\secref{sec:buffer-node}) of the enclosing |
|
1448 |
editor buffer; this requires a proper theory, not an auxiliary file. |
|
57324 | 1449 |
|
1450 |
A suffix of slashes may be used to continue the exploration of an already |
|
1451 |
recognized directory name. |
|
58618 | 1452 |
\<close> |
57324 | 1453 |
|
1454 |
||
58618 | 1455 |
subsubsection \<open>Spell-checking\<close> |
57328 | 1456 |
|
58618 | 1457 |
text \<open> |
57328 | 1458 |
The spell-checker combines semantic markup from the prover (regions of plain |
1459 |
words) with static dictionaries (word lists) that are known to the editor. |
|
1460 |
||
57333 | 1461 |
Unknown words are underlined in the text, using @{system_option_ref |
57328 | 1462 |
spell_checker_color} (blue by default). This is not an error, but a hint to |
57335 | 1463 |
the user that some action may be taken. The jEdit context menu provides |
1464 |
various actions, as far as applicable: |
|
57328 | 1465 |
|
61415 | 1466 |
\<^medskip> |
57328 | 1467 |
\begin{tabular}{l} |
57329 | 1468 |
@{action_ref "isabelle.complete-word"} \\ |
1469 |
@{action_ref "isabelle.exclude-word"} \\ |
|
1470 |
@{action_ref "isabelle.exclude-word-permanently"} \\ |
|
1471 |
@{action_ref "isabelle.include-word"} \\ |
|
1472 |
@{action_ref "isabelle.include-word-permanently"} \\ |
|
57328 | 1473 |
\end{tabular} |
61415 | 1474 |
\<^medskip> |
57328 | 1475 |
|
57329 | 1476 |
Instead of the specific @{action_ref "isabelle.complete-word"}, it is also |
1477 |
possible to use the generic @{action_ref "isabelle.complete"} with its |
|
61503 | 1478 |
default keyboard shortcut \<^verbatim>\<open>C+b\<close>. |
57328 | 1479 |
|
61415 | 1480 |
\<^medskip> |
61574 | 1481 |
Dictionary lookup uses some educated guesses about lower-case, upper-case, |
1482 |
and capitalized words. This is oriented on common use in English, where this |
|
62250 | 1483 |
aspect is not decisive for proper spelling (in contrast to German, for |
1484 |
example). |
|
58618 | 1485 |
\<close> |
57328 | 1486 |
|
1487 |
||
58618 | 1488 |
subsection \<open>Semantic completion context \label{sec:completion-context}\<close> |
57325 | 1489 |
|
58618 | 1490 |
text \<open> |
57325 | 1491 |
Completion depends on a semantic context that is provided by the prover, |
1492 |
although with some delay, because at least a full PIDE protocol round-trip |
|
1493 |
is required. Until that information becomes available in the PIDE |
|
1494 |
document-model, the default context is given by the outer syntax of the |
|
1495 |
editor mode (see also \secref{sec:buffer-node}). |
|
1496 |
||
61477 | 1497 |
The semantic \<^emph>\<open>language context\<close> provides information about nested |
62250 | 1498 |
sub-languages of Isabelle: keywords are only completed for outer syntax, and |
1499 |
antiquotations for languages that support them. Symbol abbreviations only |
|
1500 |
work for specific sub-languages: e.g.\ ``\<^verbatim>\<open>=>\<close>'' is \<^emph>\<open>not\<close> completed in |
|
1501 |
regular ML source, but is completed within ML strings, comments, |
|
1502 |
antiquotations. Backslash representations of symbols like ``\<^verbatim>\<open>\foobar\<close>'' or |
|
1503 |
``\<^verbatim>\<open>\<foobar>\<close>'' work in any context --- after additional confirmation. |
|
57325 | 1504 |
|
61574 | 1505 |
The prover may produce \<^emph>\<open>no completion\<close> markup in exceptional situations, to |
1506 |
tell that some language keywords should be excluded from further completion |
|
62250 | 1507 |
attempts. For example, ``\<^verbatim>\<open>:\<close>'' within accepted Isar syntax looses its |
1508 |
meaning as abbreviation for symbol ``\<open>\<in>\<close>''. |
|
58618 | 1509 |
\<close> |
57325 | 1510 |
|
1511 |
||
58618 | 1512 |
subsection \<open>Input events \label{sec:completion-input}\<close> |
57324 | 1513 |
|
58618 | 1514 |
text \<open> |
57332 | 1515 |
Completion is triggered by certain events produced by the user, with |
1516 |
optional delay after keyboard input according to @{system_option |
|
1517 |
jedit_completion_delay}. |
|
57325 | 1518 |
|
61574 | 1519 |
\<^descr>[Explicit completion] works via action @{action_ref "isabelle.complete"} |
1520 |
with keyboard shortcut \<^verbatim>\<open>C+b\<close>. This overrides the shortcut for @{action_ref |
|
1521 |
"complete-word"} in jEdit, but it is possible to restore the original jEdit |
|
1522 |
keyboard mapping of @{action "complete-word"} via \<^emph>\<open>Global Options~/ |
|
1523 |
Shortcuts\<close> and invent a different one for @{action "isabelle.complete"}. |
|
57325 | 1524 |
|
61439 | 1525 |
\<^descr>[Explicit spell-checker completion] works via @{action_ref |
57332 | 1526 |
"isabelle.complete-word"}, which is exposed in the jEdit context menu, if |
1527 |
the mouse points to a word that the spell-checker can complete. |
|
1528 |
||
61574 | 1529 |
\<^descr>[Implicit completion] works via regular keyboard input of the editor. It |
1530 |
depends on further side-conditions: |
|
57325 | 1531 |
|
61574 | 1532 |
\<^enum> The system option @{system_option_ref jedit_completion} needs to be |
1533 |
enabled (default). |
|
57325 | 1534 |
|
61574 | 1535 |
\<^enum> Completion of syntax keywords requires at least 3 relevant characters in |
1536 |
the text. |
|
57325 | 1537 |
|
61574 | 1538 |
\<^enum> The system option @{system_option_ref jedit_completion_delay} determines |
1539 |
an additional delay (0.5 by default), before opening a completion popup. |
|
1540 |
The delay gives the prover a chance to provide semantic completion |
|
61458 | 1541 |
information, notably the context (\secref{sec:completion-context}). |
57325 | 1542 |
|
61458 | 1543 |
\<^enum> The system option @{system_option_ref jedit_completion_immediate} |
1544 |
(enabled by default) controls whether replacement text should be inserted |
|
1545 |
immediately without popup, regardless of @{system_option |
|
61574 | 1546 |
jedit_completion_delay}. This aggressive mode of completion is restricted |
62250 | 1547 |
to symbol abbreviations that are not plain words (\secref{sec:symbols}). |
57325 | 1548 |
|
61574 | 1549 |
\<^enum> Completion of symbol abbreviations with only one relevant character in |
1550 |
the text always enforces an explicit popup, regardless of |
|
1551 |
@{system_option_ref jedit_completion_immediate}. |
|
58618 | 1552 |
\<close> |
57324 | 1553 |
|
1554 |
||
58618 | 1555 |
subsection \<open>Completion popup \label{sec:completion-popup}\<close> |
57324 | 1556 |
|
58618 | 1557 |
text \<open> |
61574 | 1558 |
A \<^emph>\<open>completion popup\<close> is a minimally invasive GUI component over the text |
1559 |
area that offers a selection of completion items to be inserted into the |
|
1560 |
text, e.g.\ by mouse clicks. Items are sorted dynamically, according to the |
|
1561 |
frequency of selection, with persistent history. The popup may interpret |
|
1562 |
special keys \<^verbatim>\<open>ENTER\<close>, \<^verbatim>\<open>TAB\<close>, \<^verbatim>\<open>ESCAPE\<close>, \<^verbatim>\<open>UP\<close>, \<^verbatim>\<open>DOWN\<close>, \<^verbatim>\<open>PAGE_UP\<close>, |
|
1563 |
\<^verbatim>\<open>PAGE_DOWN\<close>, but all other key events are passed to the underlying text |
|
1564 |
area. This allows to ignore unwanted completions most of the time and |
|
1565 |
continue typing quickly. Thus the popup serves as a mechanism of |
|
62250 | 1566 |
confirmation of proposed items, while the default is to continue without |
61574 | 1567 |
completion. |
57324 | 1568 |
|
1569 |
The meaning of special keys is as follows: |
|
1570 |
||
61415 | 1571 |
\<^medskip> |
57324 | 1572 |
\begin{tabular}{ll} |
61477 | 1573 |
\<^bold>\<open>key\<close> & \<^bold>\<open>action\<close> \\\hline |
61503 | 1574 |
\<^verbatim>\<open>ENTER\<close> & select completion (if @{system_option jedit_completion_select_enter}) \\ |
1575 |
\<^verbatim>\<open>TAB\<close> & select completion (if @{system_option jedit_completion_select_tab}) \\ |
|
1576 |
\<^verbatim>\<open>ESCAPE\<close> & dismiss popup \\ |
|
1577 |
\<^verbatim>\<open>UP\<close> & move up one item \\ |
|
1578 |
\<^verbatim>\<open>DOWN\<close> & move down one item \\ |
|
1579 |
\<^verbatim>\<open>PAGE_UP\<close> & move up one page of items \\ |
|
1580 |
\<^verbatim>\<open>PAGE_DOWN\<close> & move down one page of items \\ |
|
57324 | 1581 |
\end{tabular} |
61415 | 1582 |
\<^medskip> |
57324 | 1583 |
|
61574 | 1584 |
Movement within the popup is only active for multiple items. Otherwise the |
1585 |
corresponding key event retains its standard meaning within the underlying |
|
1586 |
text area. |
|
58618 | 1587 |
\<close> |
57324 | 1588 |
|
1589 |
||
58618 | 1590 |
subsection \<open>Insertion \label{sec:completion-insert}\<close> |
57324 | 1591 |
|
58618 | 1592 |
text \<open> |
57333 | 1593 |
Completion may first propose replacements to be selected (via a popup), or |
1594 |
replace text immediately in certain situations and depending on certain |
|
1595 |
options like @{system_option jedit_completion_immediate}. In any case, |
|
57420 | 1596 |
insertion works uniformly, by imitating normal jEdit text insertion, |
61477 | 1597 |
depending on the state of the \<^emph>\<open>text selection\<close>. Isabelle/jEdit tries to |
57420 | 1598 |
accommodate the most common forms of advanced selections in jEdit, but not |
1599 |
all combinations make sense. At least the following important cases are |
|
1600 |
well-defined: |
|
57333 | 1601 |
|
61574 | 1602 |
\<^descr>[No selection.] The original is removed and the replacement inserted, |
1603 |
depending on the caret position. |
|
57324 | 1604 |
|
61574 | 1605 |
\<^descr>[Rectangular selection of zero width.] This special case is treated by |
1606 |
jEdit as ``tall caret'' and insertion of completion imitates its normal |
|
1607 |
behaviour: separate copies of the replacement are inserted for each line |
|
1608 |
of the selection. |
|
57333 | 1609 |
|
61574 | 1610 |
\<^descr>[Other rectangular selection or multiple selections.] Here the original |
1611 |
is removed and the replacement is inserted for each line (or segment) of |
|
1612 |
the selection. |
|
57333 | 1613 |
|
61574 | 1614 |
Support for multiple selections is particularly useful for \<^emph>\<open>HyperSearch\<close>: |
1615 |
clicking on one of the items in the \<^emph>\<open>HyperSearch Results\<close> window makes |
|
1616 |
jEdit select all its occurrences in the corresponding line of text. Then |
|
1617 |
explicit completion can be invoked via \<^verbatim>\<open>C+b\<close>, e.g.\ to replace occurrences |
|
1618 |
of \<^verbatim>\<open>-->\<close> by \<open>\<longrightarrow>\<close>. |
|
57333 | 1619 |
|
61415 | 1620 |
\<^medskip> |
61574 | 1621 |
Insertion works by removing and inserting pieces of text from the buffer. |
1622 |
This counts as one atomic operation on the jEdit history. Thus unintended |
|
1623 |
completions may be reverted by the regular @{action undo} action of jEdit. |
|
1624 |
According to normal jEdit policies, the recovered text after @{action undo} |
|
1625 |
is selected: \<^verbatim>\<open>ESCAPE\<close> is required to reset the selection and to continue |
|
1626 |
typing more text. |
|
58618 | 1627 |
\<close> |
57324 | 1628 |
|
1629 |
||
58618 | 1630 |
subsection \<open>Options \label{sec:completion-options}\<close> |
57324 | 1631 |
|
58618 | 1632 |
text \<open> |
57324 | 1633 |
This is a summary of Isabelle/Scala system options that are relevant for |
61574 | 1634 |
completion. They may be configured in \<^emph>\<open>Plugin Options~/ Isabelle~/ General\<close> |
1635 |
as usual. |
|
57332 | 1636 |
|
61415 | 1637 |
\<^item> @{system_option_def completion_limit} specifies the maximum number of |
60257 | 1638 |
items for various semantic completion operations (name-space entries etc.) |
57332 | 1639 |
|
61415 | 1640 |
\<^item> @{system_option_def jedit_completion} guards implicit completion via |
57335 | 1641 |
regular jEdit key events (\secref{sec:completion-input}): it allows to |
1642 |
disable implicit completion altogether. |
|
57324 | 1643 |
|
61574 | 1644 |
\<^item> @{system_option_def jedit_completion_select_enter} and @{system_option_def |
1645 |
jedit_completion_select_tab} enable keys to select a completion item from |
|
1646 |
the popup (\secref{sec:completion-popup}). Note that a regular mouse click |
|
1647 |
on the list of items is always possible. |
|
57833
2c2bae3da1c2
completion popup supports both ENTER and TAB (default);
wenzelm
parents:
57822
diff
changeset
|
1648 |
|
61415 | 1649 |
\<^item> @{system_option_def jedit_completion_context} specifies whether the |
57335 | 1650 |
language context provided by the prover should be used at all. Disabling |
1651 |
that option makes completion less ``semantic''. Note that incomplete or |
|
1652 |
severely broken input may cause some disagreement of the prover and the user |
|
1653 |
about the intended language context. |
|
57332 | 1654 |
|
61415 | 1655 |
\<^item> @{system_option_def jedit_completion_delay} and @{system_option_def |
57333 | 1656 |
jedit_completion_immediate} determine the handling of keyboard events for |
1657 |
implicit completion (\secref{sec:completion-input}). |
|
57332 | 1658 |
|
61574 | 1659 |
A @{system_option jedit_completion_delay}~\<^verbatim>\<open>> 0\<close> postpones the processing of |
1660 |
key events, until after the user has stopped typing for the given time span, |
|
62250 | 1661 |
but @{system_option jedit_completion_immediate}~\<^verbatim>\<open>= true\<close> means that |
61574 | 1662 |
abbreviations of Isabelle symbols are handled nonetheless. |
57332 | 1663 |
|
66158 | 1664 |
\<^item> @{system_option_def completion_path_ignore} specifies ``glob'' |
57335 | 1665 |
patterns to ignore in file-system path completion (separated by colons), |
1666 |
e.g.\ backup files ending with tilde. |
|
57332 | 1667 |
|
61574 | 1668 |
\<^item> @{system_option_def spell_checker} is a global guard for all spell-checker |
1669 |
operations: it allows to disable that mechanism altogether. |
|
57332 | 1670 |
|
61415 | 1671 |
\<^item> @{system_option_def spell_checker_dictionary} determines the current |
57335 | 1672 |
dictionary, taken from the colon-separated list in the settings variable |
57333 | 1673 |
@{setting_def JORTHO_DICTIONARIES}. There are jEdit actions to specify local |
1674 |
updates to a dictionary, by including or excluding words. The result of |
|
63669 | 1675 |
permanent dictionary updates is stored in the directory @{path |
57335 | 1676 |
"$ISABELLE_HOME_USER/dictionaries"}, in a separate file for each dictionary. |
57332 | 1677 |
|
67395
b39d596b77ce
more accurate spell-checking for nested quotations / antiquotations, notably in formal comments;
wenzelm
parents:
67304
diff
changeset
|
1678 |
\<^item> @{system_option_def spell_checker_include} specifies a comma-separated |
61574 | 1679 |
list of markup elements that delimit words in the source that is subject to |
1680 |
spell-checking, including various forms of comments. |
|
67395
b39d596b77ce
more accurate spell-checking for nested quotations / antiquotations, notably in formal comments;
wenzelm
parents:
67304
diff
changeset
|
1681 |
|
b39d596b77ce
more accurate spell-checking for nested quotations / antiquotations, notably in formal comments;
wenzelm
parents:
67304
diff
changeset
|
1682 |
\<^item> @{system_option_def spell_checker_exclude} specifies a comma-separated |
b39d596b77ce
more accurate spell-checking for nested quotations / antiquotations, notably in formal comments;
wenzelm
parents:
67304
diff
changeset
|
1683 |
list of markup elements that disable spell-checking (e.g.\ in nested |
b39d596b77ce
more accurate spell-checking for nested quotations / antiquotations, notably in formal comments;
wenzelm
parents:
67304
diff
changeset
|
1684 |
antiquotations). |
58618 | 1685 |
\<close> |
54361 | 1686 |
|
1687 |
||
58618 | 1688 |
section \<open>Automatically tried tools \label{sec:auto-tools}\<close> |
54353 | 1689 |
|
58618 | 1690 |
text \<open> |
57325 | 1691 |
Continuous document processing works asynchronously in the background. |
1692 |
Visible document source that has been evaluated may get augmented by |
|
62251 | 1693 |
additional results of \<^emph>\<open>asynchronous print functions\<close>. An example for that |
1694 |
is proof state output, if that is enabled in the Output panel |
|
1695 |
(\secref{sec:output}). More heavy-weight print functions may be applied as |
|
1696 |
well, e.g.\ to prove or disprove parts of the formal text by other means. |
|
54354 | 1697 |
|
61574 | 1698 |
Isabelle/HOL provides various automatically tried tools that operate on |
1699 |
outermost goal statements (e.g.\ @{command lemma}, @{command theorem}), |
|
1700 |
independently of the state of the current proof attempt. They work |
|
1701 |
implicitly without any arguments. Results are output as \<^emph>\<open>information |
|
1702 |
messages\<close>, which are indicated in the text area by blue squiggles and a blue |
|
1703 |
information sign in the gutter (see \figref{fig:auto-tools}). The message |
|
1704 |
content may be shown as for other output (see also \secref{sec:output}). |
|
1705 |
Some tools produce output with \<^emph>\<open>sendback\<close> markup, which means that clicking |
|
62251 | 1706 |
on certain parts of the text inserts that into the source in the proper |
1707 |
place. |
|
54356 | 1708 |
|
62183 | 1709 |
\begin{figure}[!htb] |
54356 | 1710 |
\begin{center} |
57312 | 1711 |
\includegraphics[scale=0.333]{auto-tools} |
54356 | 1712 |
\end{center} |
57312 | 1713 |
\caption{Result of automatically tried tools} |
54356 | 1714 |
\label{fig:auto-tools} |
1715 |
\end{figure} |
|
54354 | 1716 |
|
61415 | 1717 |
\<^medskip> |
61574 | 1718 |
The following Isabelle system options control the behavior of automatically |
1719 |
tried tools (see also the jEdit dialog window \<^emph>\<open>Plugin Options~/ Isabelle~/ |
|
1720 |
General~/ Automatically tried tools\<close>): |
|
54354 | 1721 |
|
61574 | 1722 |
\<^item> @{system_option_ref auto_methods} controls automatic use of a combination |
1723 |
of standard proof methods (@{method auto}, @{method simp}, @{method blast}, |
|
1724 |
etc.). This corresponds to the Isar command @{command_ref "try0"} @{cite |
|
1725 |
"isabelle-isar-ref"}. |
|
54354 | 1726 |
|
1727 |
The tool is disabled by default, since unparameterized invocation of |
|
61574 | 1728 |
standard proof methods often consumes substantial CPU resources without |
1729 |
leading to success. |
|
54354 | 1730 |
|
61574 | 1731 |
\<^item> @{system_option_ref auto_nitpick} controls a slightly reduced version of |
1732 |
@{command_ref nitpick}, which tests for counterexamples using first-order |
|
1733 |
relational logic. See also the Nitpick manual @{cite "isabelle-nitpick"}. |
|
54354 | 1734 |
|
61574 | 1735 |
This tool is disabled by default, due to the extra overhead of invoking an |
1736 |
external Java process for each attempt to disprove a subgoal. |
|
54354 | 1737 |
|
61415 | 1738 |
\<^item> @{system_option_ref auto_quickcheck} controls automatic use of |
61574 | 1739 |
@{command_ref quickcheck}, which tests for counterexamples using a series of |
1740 |
assignments for free variables of a subgoal. |
|
54354 | 1741 |
|
61574 | 1742 |
This tool is \<^emph>\<open>enabled\<close> by default. It requires little overhead, but is a |
1743 |
bit weaker than @{command nitpick}. |
|
54354 | 1744 |
|
61574 | 1745 |
\<^item> @{system_option_ref auto_sledgehammer} controls a significantly reduced |
1746 |
version of @{command_ref sledgehammer}, which attempts to prove a subgoal |
|
1747 |
using external automatic provers. See also the Sledgehammer manual @{cite |
|
1748 |
"isabelle-sledgehammer"}. |
|
54354 | 1749 |
|
61574 | 1750 |
This tool is disabled by default, due to the relatively heavy nature of |
1751 |
Sledgehammer. |
|
54354 | 1752 |
|
61415 | 1753 |
\<^item> @{system_option_ref auto_solve_direct} controls automatic use of |
61574 | 1754 |
@{command_ref solve_direct}, which checks whether the current subgoals can |
1755 |
be solved directly by an existing theorem. This also helps to detect |
|
1756 |
duplicate lemmas. |
|
54354 | 1757 |
|
61477 | 1758 |
This tool is \<^emph>\<open>enabled\<close> by default. |
54354 | 1759 |
|
1760 |
||
61574 | 1761 |
Invocation of automatically tried tools is subject to some global policies |
1762 |
of parallel execution, which may be configured as follows: |
|
54354 | 1763 |
|
61574 | 1764 |
\<^item> @{system_option_ref auto_time_limit} (default 2.0) determines the timeout |
1765 |
(in seconds) for each tool execution. |
|
54354 | 1766 |
|
61574 | 1767 |
\<^item> @{system_option_ref auto_time_start} (default 1.0) determines the start |
1768 |
delay (in seconds) for automatically tried tools, after the main command |
|
1769 |
evaluation is finished. |
|
54354 | 1770 |
|
1771 |
||
61574 | 1772 |
Each tool is submitted independently to the pool of parallel execution tasks |
1773 |
in Isabelle/ML, using hardwired priorities according to its relative |
|
1774 |
``heaviness''. The main stages of evaluation and printing of proof states |
|
1775 |
take precedence, but an already running tool is not canceled and may thus |
|
1776 |
reduce reactivity of proof document processing. |
|
54354 | 1777 |
|
61574 | 1778 |
Users should experiment how the available CPU resources (number of cores) |
1779 |
are best invested to get additional feedback from prover in the background, |
|
1780 |
by using a selection of weaker or stronger tools. |
|
58618 | 1781 |
\<close> |
54353 | 1782 |
|
1783 |
||
58618 | 1784 |
section \<open>Sledgehammer \label{sec:sledgehammer}\<close> |
54353 | 1785 |
|
61574 | 1786 |
text \<open> |
1787 |
The \<^emph>\<open>Sledgehammer\<close> panel (\figref{fig:sledgehammer}) provides a view on |
|
1788 |
some independent execution of the Isar command @{command_ref sledgehammer}, |
|
1789 |
with process indicator (spinning wheel) and GUI elements for important |
|
1790 |
Sledgehammer arguments and options. Any number of Sledgehammer panels may be |
|
1791 |
active, according to the standard policies of Dockable Window Management in |
|
1792 |
jEdit. Closing such windows also cancels the corresponding prover tasks. |
|
54356 | 1793 |
|
62183 | 1794 |
\begin{figure}[!htb] |
54356 | 1795 |
\begin{center} |
57312 | 1796 |
\includegraphics[scale=0.333]{sledgehammer} |
54356 | 1797 |
\end{center} |
1798 |
\caption{An instance of the Sledgehammer panel} |
|
1799 |
\label{fig:sledgehammer} |
|
1800 |
\end{figure} |
|
54355 | 1801 |
|
61574 | 1802 |
The \<^emph>\<open>Apply\<close> button attaches a fresh invocation of @{command sledgehammer} |
1803 |
to the command where the cursor is pointing in the text --- this should be |
|
1804 |
some pending proof problem. Further buttons like \<^emph>\<open>Cancel\<close> and \<^emph>\<open>Locate\<close> |
|
1805 |
help to manage the running process. |
|
54355 | 1806 |
|
61574 | 1807 |
Results appear incrementally in the output window of the panel. Proposed |
1808 |
proof snippets are marked-up as \<^emph>\<open>sendback\<close>, which means a single mouse |
|
1809 |
click inserts the text into a suitable place of the original source. Some |
|
1810 |
manual editing may be required nonetheless, say to remove earlier proof |
|
1811 |
attempts. |
|
1812 |
\<close> |
|
54353 | 1813 |
|
1814 |
||
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1815 |
chapter \<open>Isabelle document preparation\<close> |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1816 |
|
61574 | 1817 |
text \<open> |
1818 |
The ultimate purpose of Isabelle is to produce nicely rendered documents |
|
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1819 |
with the Isabelle document preparation system, which is based on {\LaTeX}; |
60270 | 1820 |
see also @{cite "isabelle-system" and "isabelle-isar-ref"}. Isabelle/jEdit |
61574 | 1821 |
provides some additional support for document editing. |
1822 |
\<close> |
|
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1823 |
|
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1824 |
|
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1825 |
section \<open>Document outline\<close> |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1826 |
|
61574 | 1827 |
text \<open> |
1828 |
Theory sources may contain document markup commands, such as @{command_ref |
|
1829 |
chapter}, @{command_ref section}, @{command subsection}. The Isabelle |
|
1830 |
SideKick parser (\secref{sec:sidekick}) represents this document outline as |
|
1831 |
structured tree view, with formal statements and proofs nested inside; see |
|
1832 |
\figref{fig:sidekick-document}. |
|
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1833 |
|
62183 | 1834 |
\begin{figure}[!htb] |
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1835 |
\begin{center} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1836 |
\includegraphics[scale=0.333]{sidekick-document} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1837 |
\end{center} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1838 |
\caption{Isabelle document outline via SideKick tree view} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1839 |
\label{fig:sidekick-document} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1840 |
\end{figure} |
60264 | 1841 |
|
1842 |
It is also possible to use text folding according to this structure, by |
|
61574 | 1843 |
adjusting \<^emph>\<open>Utilities / Buffer Options / Folding mode\<close> of jEdit. The default |
1844 |
mode \<^verbatim>\<open>isabelle\<close> uses the structure of formal definitions, statements, and |
|
1845 |
proofs. The alternative mode \<^verbatim>\<open>sidekick\<close> uses the document structure of the |
|
1846 |
SideKick parser, as explained above. |
|
1847 |
\<close> |
|
60264 | 1848 |
|
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1849 |
|
62154 | 1850 |
section \<open>Markdown structure\<close> |
1851 |
||
1852 |
text \<open> |
|
62251 | 1853 |
Document text is internally structured in paragraphs and nested lists, using |
63680 | 1854 |
notation that is similar to Markdown\<^footnote>\<open>\<^url>\<open>http://commonmark.org\<close>\<close>. There are |
1855 |
special control symbols for items of different kinds of lists, corresponding |
|
1856 |
to \<^verbatim>\<open>itemize\<close>, \<^verbatim>\<open>enumerate\<close>, \<^verbatim>\<open>description\<close> in {\LaTeX}. This is illustrated |
|
1857 |
in for \<^verbatim>\<open>itemize\<close> in \figref{fig:markdown-document}. |
|
62154 | 1858 |
|
62183 | 1859 |
\begin{figure}[!htb] |
62154 | 1860 |
\begin{center} |
1861 |
\includegraphics[scale=0.333]{markdown-document} |
|
1862 |
\end{center} |
|
1863 |
\caption{Markdown structure within document text} |
|
1864 |
\label{fig:markdown-document} |
|
1865 |
\end{figure} |
|
62251 | 1866 |
|
1867 |
Items take colour according to the depth of nested lists. This helps to |
|
1868 |
explore the implicit rules for list structure interactively. There is also |
|
1869 |
markup for individual paragraphs in the text: it may be explored via mouse |
|
1870 |
hovering with \<^verbatim>\<open>CONTROL\<close> / \<^verbatim>\<open>COMMAND\<close> as usual |
|
1871 |
(\secref{sec:tooltips-hyperlinks}). |
|
62154 | 1872 |
\<close> |
1873 |
||
1874 |
||
62184 | 1875 |
section \<open>Citations and Bib{\TeX} entries \label{sec:bibtex}\<close> |
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1876 |
|
61574 | 1877 |
text \<open> |
1878 |
Citations are managed by {\LaTeX} and Bib{\TeX} in \<^verbatim>\<open>.bib\<close> files. The |
|
1879 |
Isabelle session build process and the @{tool latex} tool @{cite |
|
60270 | 1880 |
"isabelle-system"} are smart enough to assemble the result, based on the |
60257 | 1881 |
session directory layout. |
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1882 |
|
61493 | 1883 |
The document antiquotation \<open>@{cite}\<close> is described in @{cite |
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1884 |
"isabelle-isar-ref"}. Within the Prover IDE it provides semantic markup for |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1885 |
tooltips, hyperlinks, and completion for Bib{\TeX} database entries. |
61574 | 1886 |
Isabelle/jEdit does \<^emph>\<open>not\<close> know about the actual Bib{\TeX} environment used |
1887 |
in {\LaTeX} batch-mode, but it can take citations from those \<^verbatim>\<open>.bib\<close> files |
|
1888 |
that happen to be open in the editor; see \figref{fig:cite-completion}. |
|
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1889 |
|
62183 | 1890 |
\begin{figure}[!htb] |
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1891 |
\begin{center} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1892 |
\includegraphics[scale=0.333]{cite-completion} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1893 |
\end{center} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1894 |
\caption{Semantic completion of citations from open Bib{\TeX} files} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1895 |
\label{fig:cite-completion} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1896 |
\end{figure} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1897 |
|
68736 | 1898 |
Isabelle/jEdit also provides IDE support for editing \<^verbatim>\<open>.bib\<close> files |
61574 | 1899 |
themselves. There is syntax highlighting based on entry types (according to |
1900 |
standard Bib{\TeX} styles), a context-menu to compose entries |
|
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1901 |
systematically, and a SideKick tree view of the overall content; see |
68736 | 1902 |
\figref{fig:bibtex-mode}. Semantic checking with errors and warnings is |
1903 |
performed by the original \<^verbatim>\<open>bibtex\<close> tool using style \<^verbatim>\<open>plain\<close>: different |
|
1904 |
Bib{\TeX} styles may produce slightly different results. |
|
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1905 |
|
62183 | 1906 |
\begin{figure}[!htb] |
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1907 |
\begin{center} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1908 |
\includegraphics[scale=0.333]{bibtex-mode} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1909 |
\end{center} |
68737 | 1910 |
\caption{Bib{\TeX} mode with context menu, SideKick tree view, and |
1911 |
semantic output from the \<^verbatim>\<open>bibtex\<close> tool} |
|
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1912 |
\label{fig:bibtex-mode} |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1913 |
\end{figure} |
68736 | 1914 |
|
1915 |
Regular document preview (\secref{sec:document-preview}) of \<^verbatim>\<open>.bib\<close> files |
|
1916 |
approximates the usual {\LaTeX} bibliography output in HTML (using style |
|
1917 |
\<^verbatim>\<open>unsort\<close>). |
|
60255
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1918 |
\<close> |
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1919 |
|
0466bd194d74
more on Isabelle document preparation and bibtex files;
wenzelm
parents:
58618
diff
changeset
|
1920 |
|
68736 | 1921 |
section \<open>Document preview \label{sec:document-preview}\<close> |
66683 | 1922 |
|
1923 |
text \<open> |
|
1924 |
The action @{action_def isabelle.preview} opens an HTML preview of the |
|
67246
4cedf44f2af1
isabelle.preview presents auxiliary text files as well;
wenzelm
parents:
67092
diff
changeset
|
1925 |
current document node in the default web browser. The content is derived |
66683 | 1926 |
from the semantic markup produced by the prover, and thus depends on the |
1927 |
status of formal processing. |
|
67262 | 1928 |
|
1929 |
Action @{action_def isabelle.draft} is similar to @{action |
|
1930 |
isabelle.preview}, but shows a plain-text document draft. |
|
66683 | 1931 |
\<close> |
1932 |
||
1933 |
||
62253 | 1934 |
chapter \<open>ML debugging within the Prover IDE\<close> |
62154 | 1935 |
|
1936 |
text \<open> |
|
68224 | 1937 |
Isabelle/ML is based on Poly/ML\<^footnote>\<open>\<^url>\<open>https://www.polyml.org\<close>\<close> and thus |
62253 | 1938 |
benefits from the source-level debugger of that implementation of Standard |
1939 |
ML. The Prover IDE provides the \<^emph>\<open>Debugger\<close> dockable to connect to running |
|
1940 |
ML threads, inspect the stack frame with local ML bindings, and evaluate ML |
|
1941 |
expressions in a particular run-time context. A typical debugger session is |
|
1942 |
shown in \figref{fig:ml-debugger}. |
|
1943 |
||
1944 |
ML debugging depends on the following pre-requisites. |
|
1945 |
||
1946 |
\<^enum> ML source needs to be compiled with debugging enabled. This may be |
|
1947 |
controlled for particular chunks of ML sources using any of the subsequent |
|
1948 |
facilities. |
|
1949 |
||
1950 |
\<^enum> The system option @{system_option_ref ML_debugger} as implicit state |
|
1951 |
of the Isabelle process. It may be changed in the menu \<^emph>\<open>Plugins / |
|
1952 |
Plugin Options / Isabelle / General\<close>. ML modules need to be reloaded and |
|
1953 |
recompiled to pick up that option as intended. |
|
1954 |
||
1955 |
\<^enum> The configuration option @{attribute_ref ML_debugger}, with an |
|
1956 |
attribute of the same name, to update a global or local context (e.g.\ |
|
1957 |
with the @{command declare} command). |
|
1958 |
||
1959 |
\<^enum> Commands that modify @{attribute ML_debugger} state for individual |
|
1960 |
files: @{command_ref ML_file_debug}, @{command_ref ML_file_no_debug}, |
|
1961 |
@{command_ref SML_file_debug}, @{command_ref SML_file_no_debug}. |
|
1962 |
||
1963 |
The instrumentation of ML code for debugging causes minor run-time |
|
1964 |
overhead. ML modules that implement critical system infrastructure may |
|
1965 |
lead to deadlocks or other undefined behaviour, when put under debugger |
|
1966 |
control! |
|
1967 |
||
1968 |
\<^enum> The \<^emph>\<open>Debugger\<close> panel needs to be active, otherwise the program ignores |
|
1969 |
debugger instrumentation of the compiler and runs unmanaged. It is also |
|
1970 |
possible to start debugging with the panel open, and later undock it, to |
|
1971 |
let the program continue unhindered. |
|
1972 |
||
1973 |
\<^enum> The ML program needs to be stopped at a suitable breakpoint, which may |
|
1974 |
be activated individually or globally as follows. |
|
1975 |
||
1976 |
For ML sources that have been compiled with debugger support, the IDE |
|
1977 |
visualizes possible breakpoints in the text. A breakpoint may be toggled |
|
1978 |
by pointing accurately with the mouse, with a right-click to activate |
|
1979 |
jEdit's context menu and its \<^emph>\<open>Toggle Breakpoint\<close> item. Alternatively, the |
|
1980 |
\<^emph>\<open>Break\<close> checkbox in the \<^emph>\<open>Debugger\<close> panel may be enabled to stop ML |
|
1981 |
threads always at the next possible breakpoint. |
|
1982 |
||
1983 |
Note that the state of individual breakpoints \<^emph>\<open>gets lost\<close> when the |
|
1984 |
coresponding ML source is re-compiled! This may happen unintentionally, |
|
1985 |
e.g.\ when following hyperlinks into ML modules that have not been loaded |
|
1986 |
into the IDE before. |
|
62154 | 1987 |
|
62183 | 1988 |
\begin{figure}[!htb] |
62154 | 1989 |
\begin{center} |
1990 |
\includegraphics[scale=0.333]{ml-debugger} |
|
1991 |
\end{center} |
|
62253 | 1992 |
\caption{ML debugger session} |
62154 | 1993 |
\label{fig:ml-debugger} |
1994 |
\end{figure} |
|
62253 | 1995 |
|
1996 |
The debugger panel (\figref{fig:ml-debugger}) shows a list of all threads |
|
1997 |
that are presently stopped. Each thread shows a stack of all function |
|
1998 |
invocations that lead to the current breakpoint at the top. |
|
1999 |
||
2000 |
It is possible to jump between stack positions freely, by clicking on this |
|
2001 |
list. The current situation is displayed in the big output window, as a |
|
2002 |
local ML environment with names and printed values. |
|
2003 |
||
2004 |
ML expressions may be evaluated in the current context by entering snippets |
|
2005 |
of source into the text fields labeled \<open>Context\<close> and \<open>ML\<close>, and pushing the |
|
2006 |
\<open>Eval\<close> button. By default, the source is interpreted as Isabelle/ML with the |
|
2007 |
usual support for antiquotations (like @{command ML}, @{command ML_file}). |
|
2008 |
Alternatively, strict Standard ML may be enforced via the \<^emph>\<open>SML\<close> checkbox |
|
2009 |
(like @{command SML_file}). |
|
2010 |
||
2011 |
The context for Isabelle/ML is optional, it may evaluate to a value of type |
|
2012 |
@{ML_type theory}, @{ML_type Proof.context}, or @{ML_type Context.generic}. |
|
2013 |
Thus the given ML expression (with its antiquotations) may be subject to the |
|
2014 |
intended dynamic run-time context, instead of the static compile-time |
|
2015 |
context. |
|
2016 |
||
2017 |
\<^medskip> |
|
2018 |
The buttons labeled \<^emph>\<open>Continue\<close>, \<^emph>\<open>Step\<close>, \<^emph>\<open>Step over\<close>, \<^emph>\<open>Step out\<close> |
|
2019 |
recommence execution of the program, with different policies concerning |
|
2020 |
nested function invocations. The debugger always moves the cursor within the |
|
2021 |
ML source to the next breakpoint position, and offers new stack frames as |
|
2022 |
before. |
|
62154 | 2023 |
\<close> |
2024 |
||
2025 |
||
58618 | 2026 |
chapter \<open>Miscellaneous tools\<close> |
54358 | 2027 |
|
58618 | 2028 |
section \<open>Timing\<close> |
54359 | 2029 |
|
61574 | 2030 |
text \<open> |
2031 |
Managed evaluation of commands within PIDE documents includes timing |
|
2032 |
information, which consists of elapsed (wall-clock) time, CPU time, and GC |
|
2033 |
(garbage collection) time. Note that in a multithreaded system it is |
|
2034 |
difficult to measure execution time precisely: elapsed time is closer to the |
|
2035 |
real requirements of runtime resources than CPU or GC time, which are both |
|
2036 |
subject to influences from the parallel environment that are outside the |
|
2037 |
scope of the current command transaction. |
|
54359 | 2038 |
|
61574 | 2039 |
The \<^emph>\<open>Timing\<close> panel provides an overview of cumulative command timings for |
2040 |
each document node. Commands with elapsed time below the given threshold are |
|
2041 |
ignored in the grand total. Nodes are sorted according to their overall |
|
2042 |
timing. For the document node that corresponds to the current buffer, |
|
2043 |
individual command timings are shown as well. A double-click on a theory |
|
2044 |
node or command moves the editor focus to that particular source position. |
|
54359 | 2045 |
|
61574 | 2046 |
It is also possible to reveal individual timing information via some tooltip |
2047 |
for the corresponding command keyword, using the technique of mouse hovering |
|
62251 | 2048 |
with \<^verbatim>\<open>CONTROL\<close>~/ \<^verbatim>\<open>COMMAND\<close> modifier (\secref{sec:tooltips-hyperlinks}). |
2049 |
Actual display of timing depends on the global option @{system_option_ref |
|
2050 |
jedit_timing_threshold}, which can be configured in \<^emph>\<open>Plugin Options~/ |
|
2051 |
Isabelle~/ General\<close>. |
|
54360 | 2052 |
|
61415 | 2053 |
\<^medskip> |
61574 | 2054 |
The \<^emph>\<open>Monitor\<close> panel visualizes various data collections about recent |
2055 |
activity of the Isabelle/ML task farm and the underlying ML runtime system. |
|
2056 |
The display is continuously updated according to @{system_option_ref |
|
57869 | 2057 |
editor_chart_delay}. Note that the painting of the chart takes considerable |
2058 |
runtime itself --- on the Java Virtual Machine that runs Isabelle/Scala, not |
|
61503 | 2059 |
Isabelle/ML. Internally, the Isabelle/Scala module \<^verbatim>\<open>isabelle.ML_Statistics\<close> |
2060 |
provides further access to statistics of Isabelle/ML. |
|
2061 |
\<close> |
|
54359 | 2062 |
|
2063 |
||
58618 | 2064 |
section \<open>Low-level output\<close> |
54358 | 2065 |
|
61574 | 2066 |
text \<open> |
62251 | 2067 |
Prover output is normally shown directly in the main text area or specific |
2068 |
panels like \<^emph>\<open>Output\<close> (\secref{sec:output}) or \<^emph>\<open>State\<close> |
|
2069 |
(\secref{sec:state-output}). Beyond this, it is occasionally useful to |
|
2070 |
inspect low-level output channels via some of the following additional |
|
2071 |
panels: |
|
54358 | 2072 |
|
61574 | 2073 |
\<^item> \<^emph>\<open>Protocol\<close> shows internal messages between the Isabelle/Scala and |
2074 |
Isabelle/ML side of the PIDE document editing protocol. Recording of |
|
2075 |
messages starts with the first activation of the corresponding dockable |
|
2076 |
window; earlier messages are lost. |
|
54358 | 2077 |
|
61574 | 2078 |
Actual display of protocol messages causes considerable slowdown, so it is |
2079 |
important to undock all \<^emph>\<open>Protocol\<close> panels for production work. |
|
54358 | 2080 |
|
61503 | 2081 |
\<^item> \<^emph>\<open>Raw Output\<close> shows chunks of text from the \<^verbatim>\<open>stdout\<close> and \<^verbatim>\<open>stderr\<close> |
61574 | 2082 |
channels of the prover process. Recording of output starts with the first |
2083 |
activation of the corresponding dockable window; earlier output is lost. |
|
54358 | 2084 |
|
61574 | 2085 |
The implicit stateful nature of physical I/O channels makes it difficult to |
2086 |
relate raw output to the actual command from where it was originating. |
|
2087 |
Parallel execution may add to the confusion. Peeking at physical process I/O |
|
2088 |
is only the last resort to diagnose problems with tools that are not PIDE |
|
2089 |
compliant. |
|
54358 | 2090 |
|
57310 | 2091 |
Under normal circumstances, prover output always works via managed message |
2092 |
channels (corresponding to @{ML writeln}, @{ML warning}, @{ML |
|
57420 | 2093 |
Output.error_message} in Isabelle/ML), which are displayed by regular means |
60257 | 2094 |
within the document model (\secref{sec:output}). Unhandled Isabelle/ML |
2095 |
exceptions are printed by the system via @{ML Output.error_message}. |
|
54358 | 2096 |
|
61477 | 2097 |
\<^item> \<^emph>\<open>Syslog\<close> shows system messages that might be relevant to diagnose |
60257 | 2098 |
problems with the startup or shutdown phase of the prover process; this also |
61574 | 2099 |
includes raw output on \<^verbatim>\<open>stderr\<close>. Isabelle/ML also provides an explicit @{ML |
2100 |
Output.system_message} operation, which is occasionally useful for |
|
2101 |
diagnostic purposes within the system infrastructure itself. |
|
54358 | 2102 |
|
61574 | 2103 |
A limited amount of syslog messages are buffered, independently of the |
2104 |
docking state of the \<^emph>\<open>Syslog\<close> panel. This allows to diagnose serious |
|
2105 |
problems with Isabelle/PIDE process management, outside of the actual |
|
2106 |
protocol layer. |
|
54358 | 2107 |
|
61574 | 2108 |
Under normal situations, such low-level system output can be ignored. |
58618 | 2109 |
\<close> |
54358 | 2110 |
|
2111 |
||
58618 | 2112 |
chapter \<open>Known problems and workarounds \label{sec:problems}\<close> |
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
2113 |
|
58618 | 2114 |
text \<open> |
61574 | 2115 |
\<^item> \<^bold>\<open>Problem:\<close> Odd behavior of some diagnostic commands with global |
2116 |
side-effects, like writing a physical file. |
|
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
2117 |
|
61574 | 2118 |
\<^bold>\<open>Workaround:\<close> Copy/paste complete command text from elsewhere, or disable |
2119 |
continuous checking temporarily. |
|
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
2120 |
|
61574 | 2121 |
\<^item> \<^bold>\<open>Problem:\<close> Keyboard shortcuts \<^verbatim>\<open>C+PLUS\<close> and \<^verbatim>\<open>C+MINUS\<close> for adjusting the |
2122 |
editor font size depend on platform details and national keyboards. |
|
2123 |
||
2124 |
\<^bold>\<open>Workaround:\<close> Rebind keys via \<^emph>\<open>Global Options~/ Shortcuts\<close>. |
|
54330 | 2125 |
|
61503 | 2126 |
\<^item> \<^bold>\<open>Problem:\<close> The Mac OS X key sequence \<^verbatim>\<open>COMMAND+COMMA\<close> for application |
61574 | 2127 |
\<^emph>\<open>Preferences\<close> is in conflict with the jEdit default keyboard shortcut for |
2128 |
\<^emph>\<open>Incremental Search Bar\<close> (action @{action_ref "quick-search"}). |
|
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
2129 |
|
61574 | 2130 |
\<^bold>\<open>Workaround:\<close> Rebind key via \<^emph>\<open>Global Options~/ Shortcuts\<close> according to |
2131 |
national keyboard, e.g.\ \<^verbatim>\<open>COMMAND+SLASH\<close> on English ones. |
|
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
2132 |
|
61522
4108f91ca810
workaround for problem with C-1, C-2, C-3 seen on Slovak QWERTY keyboard;
wenzelm
parents:
61512
diff
changeset
|
2133 |
\<^item> \<^bold>\<open>Problem:\<close> On Mac OS X with native Apple look-and-feel, some exotic |
4108f91ca810
workaround for problem with C-1, C-2, C-3 seen on Slovak QWERTY keyboard;
wenzelm
parents:
61512
diff
changeset
|
2134 |
national keyboards may cause a conflict of menu accelerator keys with |
4108f91ca810
workaround for problem with C-1, C-2, C-3 seen on Slovak QWERTY keyboard;
wenzelm
parents:
61512
diff
changeset
|
2135 |
regular jEdit key bindings. This leads to duplicate execution of the |
4108f91ca810
workaround for problem with C-1, C-2, C-3 seen on Slovak QWERTY keyboard;
wenzelm
parents:
61512
diff
changeset
|
2136 |
corresponding jEdit action. |
4108f91ca810
workaround for problem with C-1, C-2, C-3 seen on Slovak QWERTY keyboard;
wenzelm
parents:
61512
diff
changeset
|
2137 |
|
4108f91ca810
workaround for problem with C-1, C-2, C-3 seen on Slovak QWERTY keyboard;
wenzelm
parents:
61512
diff
changeset
|
2138 |
\<^bold>\<open>Workaround:\<close> Disable the native Apple menu bar via Java runtime option |
4108f91ca810
workaround for problem with C-1, C-2, C-3 seen on Slovak QWERTY keyboard;
wenzelm
parents:
61512
diff
changeset
|
2139 |
\<^verbatim>\<open>-Dapple.laf.useScreenMenuBar=false\<close>. |
4108f91ca810
workaround for problem with C-1, C-2, C-3 seen on Slovak QWERTY keyboard;
wenzelm
parents:
61512
diff
changeset
|
2140 |
|
61574 | 2141 |
\<^item> \<^bold>\<open>Problem:\<close> Mac OS X system fonts sometimes lead to character drop-outs in |
2142 |
the main text area. |
|
54349 | 2143 |
|
62265 | 2144 |
\<^bold>\<open>Workaround:\<close> Use the default \<^verbatim>\<open>IsabelleText\<close> font. |
2145 |
||
61574 | 2146 |
\<^item> \<^bold>\<open>Problem:\<close> Some Linux/X11 input methods such as IBus tend to disrupt key |
2147 |
event handling of Java/AWT/Swing. |
|
54329 | 2148 |
|
61574 | 2149 |
\<^bold>\<open>Workaround:\<close> Do not use X11 input methods. Note that environment variable |
2150 |
\<^verbatim>\<open>XMODIFIERS\<close> is reset by default within Isabelle settings. |
|
2151 |
||
2152 |
\<^item> \<^bold>\<open>Problem:\<close> Some Linux/X11 window managers that are not ``re-parenting'' |
|
2153 |
cause problems with additional windows opened by Java. This affects either |
|
2154 |
historic or neo-minimalistic window managers like \<^verbatim>\<open>awesome\<close> or \<^verbatim>\<open>xmonad\<close>. |
|
54329 | 2155 |
|
61477 | 2156 |
\<^bold>\<open>Workaround:\<close> Use a regular re-parenting X11 window manager. |
54329 | 2157 |
|
61574 | 2158 |
\<^item> \<^bold>\<open>Problem:\<close> Various forks of Linux/X11 window managers and desktop |
2159 |
environments (like Gnome) disrupt the handling of menu popups and mouse |
|
2160 |
positions of Java/AWT/Swing. |
|
54329 | 2161 |
|
62183 | 2162 |
\<^bold>\<open>Workaround:\<close> Use suitable version of Linux desktops. |
60291 | 2163 |
|
61477 | 2164 |
\<^item> \<^bold>\<open>Problem:\<close> Full-screen mode via jEdit action @{action_ref |
61574 | 2165 |
"toggle-full-screen"} (default keyboard shortcut \<^verbatim>\<open>F11\<close>) works on Windows, |
2166 |
but not on Mac OS X or various Linux/X11 window managers. |
|
54349 | 2167 |
|
61574 | 2168 |
\<^bold>\<open>Workaround:\<close> Use native full-screen control of the window manager (notably |
2169 |
on Mac OS X). |
|
64512 | 2170 |
|
2171 |
\<^item> \<^bold>\<open>Problem:\<close> Heap space of the JVM may fill up and render the Prover IDE |
|
2172 |
unresponsive, e.g.\ when editing big Isabelle sessions with many theories. |
|
2173 |
||
67092 | 2174 |
\<^bold>\<open>Workaround:\<close> Increase JVM heap parameters by editing platform-specific |
2175 |
files (for ``properties'' or ``options'') that are associated with the main |
|
2176 |
app bundle. |
|
64512 | 2177 |
|
2178 |
Also note that jEdit provides a heap space monitor in the status line |
|
2179 |
(bottom-right). Double-clicking on that causes full garbage-collection, |
|
2180 |
which sometimes helps in low-memory situations. |
|
58618 | 2181 |
\<close> |
53770
db362319d766
added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents:
53769
diff
changeset
|
2182 |
|
67399 | 2183 |
end |