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