| author | blanchet | 
| Sun, 04 May 2014 19:01:36 +0200 | |
| changeset 56852 | b38c5b9cf590 | 
| parent 56466 | 08982abdcdad | 
| child 57310 | da107539996f | 
| permissions | -rw-r--r-- | 
| 53981 | 1 | (*:wrap=hard:maxLineLen=78:*) | 
| 2 | ||
| 53769 | 3 | theory JEdit | 
| 4 | imports Base | |
| 5 | begin | |
| 6 | ||
| 7 | chapter {* Introduction *}
 | |
| 8 | ||
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 9 | section {* Concepts and terminology *}
 | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 10 | |
| 54321 | 11 | text {* Isabelle/jEdit is a Prover IDE that integrates \emph{parallel
 | 
| 12 |   proof checking} \cite{Wenzel:2009,Wenzel:2013:ITP} with
 | |
| 13 |   \emph{asynchronous user interaction}
 | |
| 14 |   \cite{Wenzel:2010,Wenzel:2012:UITP-EPTCS}, based on a
 | |
| 15 |   document-oriented approach to \emph{continuous proof processing}
 | |
| 16 |   \cite{Wenzel:2011:CICM,Wenzel:2012}. Many concepts and system
 | |
| 17 | components are fit together in order to make this work. The main | |
| 54329 | 18 | building blocks are as follows. | 
| 53769 | 19 | |
| 54321 | 20 |   \begin{description}
 | 
| 53769 | 21 | |
| 54329 | 22 | \item [PIDE] is a general framework for Prover IDEs based on | 
| 23 | Isabelle/Scala. It is built around a concept of parallel and | |
| 24 | asynchronous document processing, which is supported natively by the | |
| 25 | parallel proof engine that is implemented in Isabelle/ML. The prover | |
| 26 | discontinues the traditional TTY-based command loop, and supports | |
| 27 | direct editing of formal source text with rich formal markup for GUI | |
| 28 | rendering. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 29 | |
| 54321 | 30 | \item [Isabelle/ML] is the implementation and extension language of | 
| 31 |   Isabelle, see also \cite{isabelle-implementation}. It is integrated
 | |
| 54329 | 32 | into the logical context of Isabelle/Isar and allows to manipulate | 
| 33 | logical entities directly. Arbitrary add-on tools may be implemented | |
| 34 | for object-logics such as Isabelle/HOL. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 35 | |
| 54321 | 36 | \item [Isabelle/Scala] is the system programming language of | 
| 37 | Isabelle. It extends the pure logical environment of Isabelle/ML | |
| 38 | towards the ``real world'' of graphical user interfaces, text | |
| 39 | editors, IDE frameworks, web services etc. Special infrastructure | |
| 54372 | 40 | allows to transfer algebraic datatypes and formatted text easily | 
| 41 | between ML and Scala, using asynchronous protocol commands. | |
| 54321 | 42 | |
| 54329 | 43 | \item [jEdit] is a sophisticated text editor implemented in | 
| 54703 | 44 |   Java.\footnote{@{url "http://www.jedit.org"}} It is easily extensible
 | 
| 54329 | 45 | by plugins written in languages that work on the JVM, e.g.\ | 
| 54703 | 46 |   Scala\footnote{@{url "http://www.scala-lang.org/"}}.
 | 
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 47 | |
| 54321 | 48 | \item [Isabelle/jEdit] is the main example application of the PIDE | 
| 54329 | 49 | framework and the default user-interface for Isabelle. It targets | 
| 50 | both beginners and experts. Technically, Isabelle/jEdit combines a | |
| 54372 | 51 | slightly modified version of the jEdit code base with a special | 
| 52 | plugin for Isabelle, integrated as standalone application for the | |
| 53 | main operating system platforms: Linux, Windows, Mac OS X. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 54 | |
| 54321 | 55 |   \end{description}
 | 
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 56 | |
| 54321 | 57 | The subtle differences of Isabelle/ML versus Standard ML, | 
| 54330 | 58 | Isabelle/Scala versus Scala, Isabelle/jEdit versus jEdit need to be | 
| 59 | taken into account when discussing any of these PIDE building blocks | |
| 60 | in public forums, mailing lists, or even scientific publications. | |
| 61 | *} | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 62 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 63 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 64 | section {* The Isabelle/jEdit Prover IDE *}
 | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 65 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 66 | text {*
 | 
| 54357 | 67 |   \begin{figure}[htb]
 | 
| 54331 | 68 |   \begin{center}
 | 
| 53773 | 69 |   \includegraphics[width=\textwidth]{isabelle-jedit}
 | 
| 54331 | 70 |   \end{center}
 | 
| 71 |   \caption{The Isabelle/jEdit Prover IDE}
 | |
| 54357 | 72 |   \label{fig:isabelle-jedit}
 | 
| 54331 | 73 |   \end{figure}
 | 
| 53773 | 74 | |
| 54357 | 75 |   Isabelle/jEdit (\figref{fig:isabelle-jedit}) consists of some
 | 
| 76 | plugins for the well-known jEdit text editor | |
| 54703 | 77 |   @{url "http://www.jedit.org"}, according to the following principles.
 | 
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 78 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 79 |   \begin{itemize}
 | 
| 53769 | 80 | |
| 54321 | 81 | \item The original jEdit look-and-feel is generally preserved, | 
| 54348 | 82 | although some default properties are changed to accommodate Isabelle | 
| 83 | (e.g.\ the text area font). | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 84 | |
| 54321 | 85 | \item Formal Isabelle/Isar text is checked asynchronously while | 
| 86 | editing. The user is in full command of the editor, and the prover | |
| 87 | refrains from locking portions of the buffer. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 88 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 89 | \item Prover feedback works via colors, boxes, squiggly underline, | 
| 54348 | 90 | hyperlinks, popup windows, icons, clickable output --- all based on | 
| 54321 | 91 | semantic markup produced by Isabelle in the background. | 
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 92 | |
| 54321 | 93 |   \item Using the mouse together with the modifier key @{verbatim
 | 
| 94 |   CONTROL} (Linux, Windows) or @{verbatim COMMAND} (Mac OS X) exposes
 | |
| 95 | additional formal content via tooltips and/or hyperlinks. | |
| 53769 | 96 | |
| 54321 | 97 | \item Formal output (in popups etc.) may be explored recursively, | 
| 98 | using the same techniques as in the editor source buffer. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 99 | |
| 54329 | 100 |   \item Additional panels (e.g.\ \emph{Output}, \emph{Symbols}) are
 | 
| 101 | organized by the Dockable Window Manager of jEdit, which also allows | |
| 102 | multiple floating instances of each window class. | |
| 103 | ||
| 54321 | 104 | \item The prover process and source files are managed on the editor | 
| 105 | side. The prover operates on timeless and stateless document | |
| 54348 | 106 | content as provided via Isabelle/Scala. | 
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 107 | |
| 54321 | 108 |   \item Plugin options of jEdit (for the \emph{Isabelle} plugin) give
 | 
| 54348 | 109 | access to a selection of Isabelle/Scala options and its persistent | 
| 54321 | 110 | preferences, usually with immediate effect on the prover back-end or | 
| 111 | editor front-end. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 112 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 113 | \item The logic image of the prover session may be specified within | 
| 54348 | 114 | Isabelle/jEdit. The new image is provided automatically by the | 
| 115 | Isabelle build tool after restart of the application. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 116 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 117 |   \end{itemize}
 | 
| 53769 | 118 | *} | 
| 119 | ||
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 120 | |
| 54321 | 121 | subsection {* Documentation *}
 | 
| 122 | ||
| 123 | text {* Regular jEdit documentation is accessible via its @{verbatim
 | |
| 54655 
9e8189a841f7
further encouragement to read jEdit documentation;
 wenzelm parents: 
54643diff
changeset | 124 |   Help} menu or @{verbatim F1} keyboard shortcut.  This includes a
 | 
| 
9e8189a841f7
further encouragement to read jEdit documentation;
 wenzelm parents: 
54643diff
changeset | 125 |   full \emph{User's Guide} and \emph{Frequently Asked Questions} for
 | 
| 
9e8189a841f7
further encouragement to read jEdit documentation;
 wenzelm parents: 
54643diff
changeset | 126 | this sophisticated text editor. The same can be browsed without the | 
| 
9e8189a841f7
further encouragement to read jEdit documentation;
 wenzelm parents: 
54643diff
changeset | 127 | technical restrictions of the built-in Java HTML viewer here: | 
| 54703 | 128 |   @{url "http://www.jedit.org/index.php?page=docs"} (potentially for a
 | 
| 54655 
9e8189a841f7
further encouragement to read jEdit documentation;
 wenzelm parents: 
54643diff
changeset | 129 | different version of jEdit). | 
| 54321 | 130 | |
| 54329 | 131 | Most of this information about jEdit is relevant for Isabelle/jEdit | 
| 132 | as well, but one needs to keep in mind that defaults sometimes | |
| 133 | differ, and the official jEdit documentation does not know about the | |
| 134 | Isabelle plugin with its special support for theory editing. | |
| 54321 | 135 | *} | 
| 136 | ||
| 137 | ||
| 138 | subsection {* Plugins *}
 | |
| 139 | ||
| 140 | text {* The \emph{Plugin Manager} of jEdit allows to augment editor
 | |
| 141 | functionality by JVM modules (jars) that are provided by the central | |
| 54372 | 142 | plugin repository, which is accessible via various mirror sites. | 
| 54321 | 143 | |
| 144 | Connecting to the plugin server infrastructure of the jEdit project | |
| 54348 | 145 | allows to update bundled plugins or to add further functionality. | 
| 146 | This needs to be done with the usual care for such an open bazaar of | |
| 147 | contributions. Arbitrary combinations of add-on features are apt to | |
| 148 | cause problems. It is advisable to start with the default | |
| 149 | configuration of Isabelle/jEdit and develop some understanding how | |
| 150 | it is supposed to work, before loading additional plugins at a grand | |
| 151 | scale. | |
| 54321 | 152 | |
| 54348 | 153 |   \medskip The main \emph{Isabelle} plugin is an integral part of
 | 
| 154 | Isabelle/jEdit and needs to remain active at all times! A few | |
| 155 | additional plugins are bundled with Isabelle/jEdit for convenience | |
| 156 |   or out of necessity, notably \emph{Console} with its Isabelle/Scala
 | |
| 157 |   sub-plugin and \emph{SideKick} with some Isabelle-specific parsers
 | |
| 158 |   for document tree structure.  The \emph{ErrorList} plugin is
 | |
| 159 |   required by \emph{SideKick}, but not used specifically in
 | |
| 160 | Isabelle/jEdit. *} | |
| 54321 | 161 | |
| 162 | ||
| 163 | subsection {* Options *}
 | |
| 164 | ||
| 54329 | 165 | text {* Both jEdit and Isabelle have distinctive management of
 | 
| 54348 | 166 | persistent options. | 
| 54321 | 167 | |
| 54372 | 168 | Regular jEdit options are accessible via the dialog for | 
| 169 |   \emph{Plugins / Plugin Options}, which is also accessible via
 | |
| 170 |   \emph{Utilities / Global Options}.  Changed properties are stored in
 | |
| 171 |   @{file_unchecked "$ISABELLE_HOME_USER/jedit/properties"}.  In
 | |
| 172 | contrast, Isabelle system options are managed by Isabelle/Scala and | |
| 173 |   changed values stored in @{file_unchecked
 | |
| 174 | "$ISABELLE_HOME_USER/etc/preferences"}, independently of the jEdit | |
| 175 |   properties.  See also \cite{isabelle-sys}, especially the coverage
 | |
| 176 |   of sessions and command-line tools like @{tool build} or @{tool
 | |
| 177 | options}. | |
| 54321 | 178 | |
| 54348 | 179 |   Those Isabelle options that are declared as \textbf{public} are
 | 
| 180 |   configurable in jEdit via \emph{Plugin Options / Isabelle /
 | |
| 181 | General}. Moreover, there are various options for rendering of | |
| 182 |   document content, which are configurable via \emph{Plugin Options /
 | |
| 183 |   Isabelle / Rendering}.  Thus \emph{Plugin Options / Isabelle} in
 | |
| 54372 | 184 | jEdit provides a view on a subset of Isabelle system options. Note | 
| 185 | that some of these options affect general parameters that are | |
| 186 |   relevant outside Isabelle/jEdit as well, e.g.\ @{system_option
 | |
| 187 |   threads} or @{system_option parallel_proofs} for the Isabelle build
 | |
| 188 |   tool \cite{isabelle-sys}.
 | |
| 54329 | 189 | |
| 54348 | 190 | \medskip All options are loaded on startup and saved on shutdown of | 
| 54372 | 191 |   Isabelle/jEdit.  Editing the machine-generated @{file_unchecked
 | 
| 192 |   "$ISABELLE_HOME_USER/jedit/properties"} or @{file_unchecked
 | |
| 54329 | 193 | "$ISABELLE_HOME_USER/etc/preferences"} manually while the | 
| 54348 | 194 | application is running is likely to cause surprise due to lost | 
| 195 | update! *} | |
| 54321 | 196 | |
| 197 | ||
| 198 | subsection {* Keymaps *}
 | |
| 199 | ||
| 200 | text {* Keyboard shortcuts used to be managed as jEdit properties in
 | |
| 201 | the past, but recent versions (2013) have a separate concept of | |
| 54348 | 202 |   \emph{keymap} that is configurable via \emph{Global Options /
 | 
| 203 |   Shortcuts}.  The @{verbatim imported} keymap is derived from the
 | |
| 54330 | 204 | initial environment of properties that is available at the first | 
| 54348 | 205 | start of the editor; afterwards the keymap file takes precedence. | 
| 54321 | 206 | |
| 207 | This is relevant for Isabelle/jEdit due to various fine-tuning of | |
| 208 | default properties, and additional keyboard shortcuts for Isabelle | |
| 54372 | 209 | specific functionality. Users may change their keymap later, but | 
| 210 | need to copy Isabelle-specific key bindings manually (see also | |
| 211 |   @{file_unchecked "$ISABELLE_HOME_USER/jedit/keymaps"}).  *}
 | |
| 54321 | 212 | |
| 213 | ||
| 54322 | 214 | subsection {* Look-and-feel *}
 | 
| 54321 | 215 | |
| 54330 | 216 | text {* jEdit is a Java/AWT/Swing application with some ambition to
 | 
| 54329 | 217 | support ``native'' look-and-feel on all platforms, within the limits | 
| 54330 | 218 | of what Oracle as Java provider and major operating system | 
| 219 |   distributors allow (see also \secref{sec:problems}).
 | |
| 54321 | 220 | |
| 54329 | 221 | Isabelle/jEdit enables platform-specific look-and-feel by default as | 
| 54330 | 222 | follows: | 
| 54321 | 223 | |
| 224 |   \begin{description}
 | |
| 225 | ||
| 226 |   \item[Linux] The platform-independent \emph{Nimbus} is used by
 | |
| 54372 | 227 | default. | 
| 54321 | 228 | |
| 54372 | 229 |   \emph{GTK+} works under the side-condition that the overall GTK
 | 
| 230 |   theme is selected in a Swing-friendly way.\footnote{GTK support in
 | |
| 231 | Java/Swing was once marketed aggressively by Sun, but never quite | |
| 232 | finished, and is today (2013) lagging a bit behind further | |
| 54643 | 233 | development of Swing and GTK. The graphics rendering performance | 
| 234 | can be worse than for other Swing look-and-feels.} | |
| 54321 | 235 | |
| 54381 | 236 |   \item[Windows] Regular \emph{Windows} is used by default, but
 | 
| 237 |   \emph{Windows Classic} also works.
 | |
| 54372 | 238 | |
| 239 |   \item[Mac OS X] Regular \emph{Mac OS X} is used by default.
 | |
| 240 | ||
| 54321 | 241 |   Moreover the bundled \emph{MacOSX} plugin provides various functions
 | 
| 242 | that are expected from applications on that particular platform: | |
| 243 | quit from menu or dock, preferences menu, drag-and-drop of text | |
| 244 | files on the application, full-screen mode for main editor windows | |
| 245 | etc. | |
| 246 | ||
| 247 |   \end{description}
 | |
| 248 | ||
| 249 | Users may experiment with different look-and-feels, but need to keep | |
| 54329 | 250 | in mind that this extra variance of GUI functionality is unlikely to | 
| 54372 | 251 | work in arbitrary combinations. The platform-independent | 
| 252 |   \emph{Nimbus} and \emph{Metal} should always work.  The historic
 | |
| 253 |   \emph{CDE/Motif} is better avoided.
 | |
| 254 | ||
| 255 |   After changing the look-and-feel in \emph{Global Options /
 | |
| 256 | Appearance}, it is advisable to restart Isabelle/jEdit in order to | |
| 257 | take full effect. *} | |
| 54321 | 258 | |
| 259 | ||
| 54372 | 260 | chapter {* Prover IDE functionality *}
 | 
| 261 | ||
| 262 | section {* File-system access *}
 | |
| 54351 | 263 | |
| 264 | text {* File specifications in jEdit follow various formats and
 | |
| 265 |   conventions according to \emph{Virtual File Systems}, which may be
 | |
| 266 | also provided by additional plugins. This allows to access remote | |
| 267 |   files via the @{verbatim "http:"} protocol prefix, for example.
 | |
| 268 | Isabelle/jEdit attempts to work with the file-system access model of | |
| 269 | jEdit as far as possible. In particular, theory sources are passed | |
| 270 | directly from the editor to the prover, without indirection via | |
| 271 | files. | |
| 272 | ||
| 273 | Despite the flexibility of URLs in jEdit, local files are | |
| 274 | particularly important and are accessible without protocol prefix. | |
| 275 | Here the path notation is that of the Java Virtual Machine on the | |
| 276 | underlying platform. On Windows the preferred form uses | |
| 277 | backslashes, but happens to accept Unix/POSIX forward slashes, too. | |
| 278 | Further differences arise due to drive letters and network shares. | |
| 279 | ||
| 280 | The Java notation for files needs to be distinguished from the one | |
| 281 | of Isabelle, which uses POSIX notation with forward slashes on | |
| 282 |   \emph{all} platforms.\footnote{Isabelle on Windows uses Cygwin
 | |
| 283 | file-system access.} Moreover, environment variables from the | |
| 284 |   Isabelle process may be used freely, e.g.\ @{file
 | |
| 55880 | 285 |   "$ISABELLE_HOME/etc/symbols"} or @{file_unchecked "$POLYML_HOME/README"}.
 | 
| 54466 | 286 |   There are special shortcuts: @{file "~"} for @{file "$USER_HOME"}
 | 
| 287 |   and @{file "~~"} for @{file "$ISABELLE_HOME"}.
 | |
| 54351 | 288 | |
| 289 | \medskip Since jEdit happens to support environment variables within | |
| 290 | file specifications as well, it is natural to use similar notation | |
| 291 | within the editor, e.g.\ in the file-browser. This does not work in | |
| 292 | full generality, though, due to the bias of jEdit towards | |
| 293 | platform-specific notation and of Isabelle towards POSIX. Moreover, | |
| 294 | the Isabelle settings environment is not yet active when starting | |
| 295 | Isabelle/jEdit via its standard application wrapper (in contrast to | |
| 296 |   @{verbatim "isabelle jedit"} run from the command line).
 | |
| 297 | ||
| 298 |   For convenience, Isabelle/jEdit imitates at least @{verbatim
 | |
| 299 |   "$ISABELLE_HOME"} and @{verbatim "$ISABELLE_HOME_USER"} within the
 | |
| 300 | Java process environment, in order to allow easy access to these | |
| 301 | important places from the editor. | |
| 302 | ||
| 303 | Moreover note that path specifications in prover input or output | |
| 304 | usually include formal markup that turns it into a hyperlink (see | |
| 54352 | 305 |   also \secref{sec:tooltips-hyperlinks}).  This allows to open the
 | 
| 54351 | 306 | corresponding file in the text editor, independently of the path | 
| 307 | notation. *} | |
| 308 | ||
| 309 | ||
| 54353 | 310 | section {* Text buffers and theories \label{sec:buffers-theories} *}
 | 
| 54322 | 311 | |
| 54350 | 312 | text {* As regular text editor, jEdit maintains a collection of open
 | 
| 313 |   \emph{text buffers} to store source files; each buffer may be
 | |
| 314 |   associated with any number of visible \emph{text areas}.  Buffers
 | |
| 315 |   are subject to an \emph{edit mode} that is determined from the file
 | |
| 316 |   type.  Files with extension \texttt{.thy} are assigned to the mode
 | |
| 317 |   \emph{isabelle} and treated specifically.
 | |
| 54321 | 318 | |
| 54329 | 319 | \medskip Isabelle theory files are automatically added to the formal | 
| 320 | document model of Isabelle/Scala, which maintains a family of | |
| 321 |   versions of all sources for the prover.  The \emph{Theories} panel
 | |
| 54322 | 322 | provides an overview of the status of continuous checking of theory | 
| 54329 | 323 |   sources.  Unlike batch sessions \cite{isabelle-sys}, theory nodes
 | 
| 324 | are identified by full path names; this allows to work with multiple | |
| 325 | (disjoint) Isabelle sessions simultaneously within the same editor | |
| 326 | session. | |
| 54322 | 327 | |
| 54329 | 328 | Certain events to open or update buffers with theory files cause | 
| 329 |   Isabelle/jEdit to resolve dependencies of \emph{theory imports}.
 | |
| 54350 | 330 | The system requests to load additional files into editor buffers, in | 
| 331 | order to be included in the theory document model for further | |
| 332 | checking. It is also possible to resolve dependencies | |
| 333 |   automatically, depending on \emph{Plugin Options / Isabelle /
 | |
| 54372 | 334 | General / Auto load}. | 
| 54321 | 335 | |
| 54329 | 336 | \medskip The open text area views on theory buffers define the | 
| 337 |   visible \emph{perspective} of Isabelle/jEdit.  This is taken as a
 | |
| 338 | hint for document processing: the prover ensures that those parts of | |
| 339 | a theory where the user is looking are checked, while other parts | |
| 54350 | 340 | that are presently not required are ignored. The perspective is | 
| 54330 | 341 | changed by opening or closing text area windows, or scrolling within | 
| 54372 | 342 | an editor window. | 
| 54322 | 343 | |
| 54330 | 344 |   The \emph{Theories} panel provides some further options to influence
 | 
| 345 | the process of continuous checking: it may be switched off globally | |
| 54350 | 346 | to restrict the prover to superficial processing of command syntax. | 
| 347 |   It is also possible to indicate theory nodes as \emph{required} for
 | |
| 54330 | 348 | continuous checking: this means such nodes and all their imports are | 
| 349 | always processed independently of the visibility status (if | |
| 350 | continuous checking is enabled). Big theory libraries that are | |
| 351 | marked as required can have significant impact on performance, | |
| 352 | though. | |
| 54322 | 353 | |
| 354 | \medskip Formal markup of checked theory content is turned into GUI | |
| 355 | rendering, based on a standard repertoire known from IDEs for | |
| 356 | programming languages: colors, icons, highlighting, squiggly | |
| 54329 | 357 | underline, tooltips, hyperlinks etc. For outer syntax of | 
| 54350 | 358 | Isabelle/Isar there is some traditional syntax-highlighting via | 
| 359 | static keyword tables and tokenization within the editor. In | |
| 360 | contrast, the painting of inner syntax (term language etc.)\ uses | |
| 361 | semantic information that is reported dynamically from the logical | |
| 362 | context. Thus the prover can provide additional markup to help the | |
| 363 | user to understand the meaning of formal text, and to produce more | |
| 364 | text with some add-on tools (e.g.\ information messages by automated | |
| 365 | provers or disprovers running in the background). | |
| 54352 | 366 | *} | 
| 54322 | 367 | |
| 54352 | 368 | |
| 54354 | 369 | section {* Prover output \label{sec:prover-output} *}
 | 
| 54353 | 370 | |
| 371 | text {* Prover output consists of \emph{markup} and \emph{messages}.
 | |
| 372 | Both are directly attached to the corresponding positions in the | |
| 373 | original source text, and visualized in the text area, e.g.\ as text | |
| 374 | colours for free and bound variables, or as squiggly underline for | |
| 54357 | 375 |   warnings, errors etc.\ (see also \figref{fig:output}).  In the
 | 
| 376 | latter case, the corresponding messages are shown by hovering with | |
| 377 | the mouse over the highlighted text --- although in many situations | |
| 54372 | 378 | the user should already get some clue by looking at the position of | 
| 379 | the text highlighting. | |
| 54357 | 380 | |
| 381 |   \begin{figure}[htb]
 | |
| 382 |   \begin{center}
 | |
| 54373 | 383 |   \includegraphics[width=\textwidth]{output}
 | 
| 54357 | 384 |   \end{center}
 | 
| 54372 | 385 |   \caption{Multiple views on prover output: gutter area with icon,
 | 
| 386 | text area with popup, overview area, Theories panel, Output panel} | |
| 54357 | 387 |   \label{fig:output}
 | 
| 388 |   \end{figure}
 | |
| 54353 | 389 | |
| 390 | The ``gutter area'' on the left-hand-side of the text area uses | |
| 54372 | 391 | icons to provide a summary of the messages within the adjacent | 
| 54353 | 392 | line of text. Message priorities are used to prefer errors over | 
| 393 | warnings, warnings over information messages etc. Plain output is | |
| 394 | ignored here. | |
| 395 | ||
| 396 | The ``overview area'' on the right-hand-side of the text area uses | |
| 397 | similar information to paint small rectangles for the overall status | |
| 398 | of the whole text buffer. The graphics is scaled to fit the logical | |
| 399 | buffer length into the given window height. Mouse clicks on the | |
| 400 | overview area position the cursor approximately to the corresponding | |
| 54372 | 401 | line of text in the buffer. Repainting the overview in real-time | 
| 402 | causes problems with big theories, so it is restricted to part of | |
| 403 |   the text according to \emph{Plugin Options / Isabelle / General /
 | |
| 404 | Text Overview Limit} (in characters). | |
| 54353 | 405 | |
| 406 |   Another course-grained overview is provided by the \emph{Theories}
 | |
| 54372 | 407 | panel, but without direct correspondence to text positions. A | 
| 408 | double-click on one of the theory entries with their status overview | |
| 409 | opens the corresponding text buffer, without changing the cursor | |
| 410 | position. | |
| 54353 | 411 | |
| 412 |   \medskip In addition, the \emph{Output} panel displays prover
 | |
| 413 | messages that correspond to a given command, within a separate | |
| 414 | window. | |
| 415 | ||
| 416 | The cursor position in the presently active text area determines the | |
| 54372 | 417 | prover commands whose cumulative message output is appended and | 
| 418 | shown in that window (in canonical order according to the processing | |
| 419 | of the command). There are also control elements to modify the | |
| 420 | update policy of the output wrt.\ continued editor movements. This | |
| 421 | is particularly useful with several independent instances of the | |
| 54353 | 422 |   \emph{Output} panel, which the Dockable Window Manager of jEdit can
 | 
| 423 | handle conveniently. | |
| 424 | ||
| 425 | Former users of the old TTY interaction model (e.g.\ Proof~General) | |
| 426 | might find a separate window for prover messages familiar, but it is | |
| 427 | important to understand that the main Prover IDE feedback happens | |
| 428 | elsewhere. It is possible to do meaningful proof editing | |
| 54372 | 429 | efficiently, using secondary output windows only rarely. | 
| 54353 | 430 | |
| 431 | The main purpose of the output window is to ``debug'' unclear | |
| 432 |   situations by inspecting internal state of the prover.\footnote{In
 | |
| 54355 | 433 | that sense, unstructured tactic scripts depend on continuous | 
| 54353 | 434 | debugging with internal state inspection.} Consequently, some | 
| 435 |   special messages for \emph{tracing} or \emph{proof state} only
 | |
| 436 | appear here, and are not attached to the original source. | |
| 437 | ||
| 438 | \medskip In any case, prover messages also contain markup that may | |
| 439 | be explored recursively via tooltips or hyperlinks (see | |
| 440 |   \secref{sec:tooltips-hyperlinks}), or clicked directly to initiate
 | |
| 54355 | 441 |   certain actions (see \secref{sec:auto-tools} and
 | 
| 442 |   \secref{sec:sledgehammer}).  *}
 | |
| 54353 | 443 | |
| 444 | ||
| 54352 | 445 | section {* Tooltips and hyperlinks \label{sec:tooltips-hyperlinks} *}
 | 
| 446 | ||
| 447 | text {* Formally processed text (prover input or output) contains rich
 | |
| 448 | markup information that can be explored further by using the | |
| 449 |   @{verbatim CONTROL} modifier key on Linux and Windows, or @{verbatim
 | |
| 450 | COMMAND} on Mac OS X. Hovering with the mouse while the modifier is | |
| 451 |   pressed reveals a \emph{tooltip} (grey box over the text with a
 | |
| 452 |   yellow popup) and/or a \emph{hyperlink} (black rectangle over the
 | |
| 453 |   text); see also \figref{fig:tooltip}.
 | |
| 54331 | 454 | |
| 54357 | 455 |   \begin{figure}[htb]
 | 
| 54331 | 456 |   \begin{center}
 | 
| 54373 | 457 |   \includegraphics[scale=0.3]{popup1}
 | 
| 54331 | 458 |   \end{center}
 | 
| 54356 | 459 |   \caption{Tooltip and hyperlink for some formal entity}
 | 
| 54350 | 460 |   \label{fig:tooltip}
 | 
| 54331 | 461 |   \end{figure}
 | 
| 462 | ||
| 54350 | 463 | Tooltip popups use the same rendering principles as the main text | 
| 464 | area, and further tooltips and/or hyperlinks may be exposed | |
| 54357 | 465 |   recursively by the same mechanism; see \figref{fig:nested-tooltips}.
 | 
| 54323 | 466 | |
| 54357 | 467 |   \begin{figure}[htb]
 | 
| 54331 | 468 |   \begin{center}
 | 
| 54373 | 469 |   \includegraphics[scale=0.3]{popup2}
 | 
| 54331 | 470 |   \end{center}
 | 
| 54356 | 471 |   \caption{Nested tooltips over formal entities}
 | 
| 54350 | 472 |   \label{fig:nested-tooltips}
 | 
| 54331 | 473 |   \end{figure}
 | 
| 54350 | 474 | |
| 475 |   The tooltip popup window provides some controls to \emph{close} or
 | |
| 476 |   \emph{detach} the window, turning it into a separate \emph{Info}
 | |
| 54372 | 477 |   window managed by jEdit.  The @{verbatim ESCAPE} key closes
 | 
| 54352 | 478 |   \emph{all} popups, which is particularly relevant when nested
 | 
| 479 | tooltips are stacking up. | |
| 480 | ||
| 481 | \medskip A black rectangle in the text indicates a hyperlink that | |
| 482 |   may be followed by a mouse click (while the @{verbatim CONTROL} or
 | |
| 54361 | 483 |   @{verbatim COMMAND} modifier key is still pressed). Presently
 | 
| 54639 | 484 | (Isabelle2013-2) there is no systematic navigation within the | 
| 54372 | 485 | editor to return to the original location. | 
| 54352 | 486 | |
| 487 | Also note that the link target may be a file that is itself not | |
| 488 | subject to formal document processing of the editor session and thus | |
| 489 | prevents further exploration: the chain of hyperlinks may end in | |
| 54372 | 490 | some source file of the underlying logic image, or within the | 
| 491 | Isabelle/ML bootstrap sources of Isabelle/Pure. *} | |
| 54321 | 492 | |
| 493 | ||
| 54361 | 494 | section {* Text completion \label{sec:completion} *}
 | 
| 495 | ||
| 54362 | 496 | text {* \paragraph{Completion tables} are determined statically from
 | 
| 497 | the ``outer syntax'' of the underlying edit mode (for theory files | |
| 54372 | 498 | this is the syntax of Isar commands), and specifications of Isabelle | 
| 54362 | 499 |   symbols (see also \secref{sec:symbols}).
 | 
| 500 | ||
| 501 |   Symbols are completed in backslashed forms, e.g.\ @{verbatim
 | |
| 502 |   "\\"}@{verbatim "forall"} or @{verbatim "\<forall>"} that both produce the
 | |
| 503 |   Isabelle symbol @{text "\<forall>"} in its Unicode rendering.\footnote{The
 | |
| 504 | extra backslash avoids overlap with keywords of the buffer syntax, | |
| 54372 | 505 | and allows to produce Isabelle symbols robustly in most syntactic | 
| 54362 | 506 | contexts.} Alternatively, symbol abbreviations may be used as | 
| 507 |   specified in @{file "$ISABELLE_HOME/etc/symbols"}.
 | |
| 508 | ||
| 509 |   \paragraph{Completion popups} are required in situations of
 | |
| 510 | ambiguous completion results or where explicit confirmation is | |
| 511 | demanded before inserting completed text into the buffer. | |
| 512 | ||
| 513 | The popup is some minimally invasive GUI component over the text | |
| 514 |   area.  It interprets special keys @{verbatim TAB}, @{verbatim
 | |
| 515 |   ESCAPE}, @{verbatim UP}, @{verbatim DOWN}, @{verbatim PAGE_UP},
 | |
| 516 |   @{verbatim PAGE_DOWN}, but all other key events are passed to the
 | |
| 517 | underlying text area. This allows to ignore unwanted completions | |
| 518 | most of the time and continue typing quickly. | |
| 54361 | 519 | |
| 54362 | 520 | The meaning of special keys is as follows: | 
| 521 | ||
| 522 | \medskip | |
| 523 |   \begin{tabular}{ll}
 | |
| 54372 | 524 |   \textbf{key} & \textbf{action} \\\hline
 | 
| 54362 | 525 |   @{verbatim "TAB"} & select completion \\
 | 
| 526 |   @{verbatim "ESCAPE"} & dismiss popup \\
 | |
| 527 |   @{verbatim "UP"} & move up one item \\
 | |
| 528 |   @{verbatim "DOWN"} & move down one item \\
 | |
| 529 |   @{verbatim "PAGE_UP"} & move up one page of items \\
 | |
| 530 |   @{verbatim "PAGE_DOWN"} & move down one page of items \\
 | |
| 531 |   \end{tabular}
 | |
| 532 | \medskip | |
| 533 | ||
| 534 | Movement within the popup is only active for multiple items. | |
| 535 | Otherwise the corresponding key event retains its standard meaning | |
| 536 | within the underlying text area. | |
| 537 | ||
| 54372 | 538 |   \paragraph{Explicit completion} is triggered by the keyboard
 | 
| 56059 
2390391584c2
some document antiquotations for Isabelle/jEdit elements;
 wenzelm parents: 
55880diff
changeset | 539 |   shortcut @{verbatim "C+b"} (action @{action "isabelle.complete"}).
 | 
| 54372 | 540 |   This overrides the original jEdit binding for action @{verbatim
 | 
| 541 | "complete-word"}, but the latter is used as fall-back for | |
| 54362 | 542 | non-Isabelle edit modes. It is also possible to restore the | 
| 54372 | 543 |   original jEdit keyboard mapping of @{verbatim "complete-word"} via
 | 
| 544 |   \emph{Global Options / Shortcuts}.
 | |
| 54361 | 545 | |
| 54362 | 546 | Replacement text is inserted immediately into the buffer, unless | 
| 547 | ambiguous results demand an explicit popup. | |
| 548 | ||
| 549 |   \paragraph{Implicit completion} is triggered by regular keyboard
 | |
| 550 | input events during of the editing process in the main jEdit text | |
| 54372 | 551 | area (and a few additional text fields like the search criteria of | 
| 54362 | 552 |   the Find panel, see \secref{sec:find}).  Implicit completion depends
 | 
| 553 | on on further side-conditions: | |
| 554 | ||
| 555 |   \begin{enumerate}
 | |
| 556 | ||
| 557 |   \item The system option @{system_option jedit_completion} needs to
 | |
| 558 | be enabled (default). | |
| 559 | ||
| 54372 | 560 | \item Completion of syntax keywords requires at least 3 relevant | 
| 561 | characters in the text. | |
| 54362 | 562 | |
| 54372 | 563 |   \item The system option @{system_option jedit_completion_delay}
 | 
| 564 | determines an additional delay (0.0 by default), before opening a | |
| 565 | completion popup. | |
| 54361 | 566 | |
| 54362 | 567 |   \item The system option @{system_option jedit_completion_immediate}
 | 
| 568 | (disabled by default) controls whether replacement text should be | |
| 54372 | 569 | inserted immediately without popup. This is restricted to Isabelle | 
| 570 |   symbols and their abbreviations (\secref{sec:symbols}) --- plain
 | |
| 571 | keywords always demand a popup for clarity. | |
| 54362 | 572 | |
| 54372 | 573 | \item Completion of symbol abbreviations with only one relevant | 
| 574 | character in the text always enforces an explicit popup, | |
| 575 |   independently of @{system_option jedit_completion_immediate}.
 | |
| 54362 | 576 | |
| 577 |   \end{enumerate}
 | |
| 54361 | 578 | |
| 54362 | 579 |   These completion options may be configured in \emph{Plugin Options /
 | 
| 580 | Isabelle / General / Completion}. The default is quite moderate in | |
| 581 | showing occasional popups and refraining from immediate insertion. | |
| 54372 | 582 | An additional completion delay of 0.3 seconds will make it even less | 
| 583 | ambitious. | |
| 54361 | 584 | |
| 54372 | 585 |   In contrast, more aggressive completion works via @{system_option
 | 
| 54362 | 586 |   jedit_completion_delay}~@{verbatim "= 0.0"} and @{system_option
 | 
| 587 |   jedit_completion_immediate}~@{verbatim "= true"}.  Thus the editing
 | |
| 54372 | 588 | process becomes dependent on the system guessing correctly what the | 
| 589 | user had in mind. It requires some practice (and study of the | |
| 590 | symbol abbreviation tables) to become productive in this advanced | |
| 591 | mode. | |
| 54362 | 592 | |
| 54372 | 593 | In any case, unintended completions can be reverted by the regular | 
| 594 |   @{verbatim undo} operation of jEdit.  When editing embedded
 | |
| 595 |   languages such as ML, it is better to disable either @{system_option
 | |
| 54362 | 596 |   jedit_completion} or @{system_option jedit_completion_immediate}
 | 
| 54372 | 597 | temporarily. *} | 
| 54361 | 598 | |
| 599 | ||
| 54362 | 600 | section {* Isabelle symbols \label{sec:symbols} *}
 | 
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 601 | |
| 54323 | 602 | text {* Isabelle sources consist of \emph{symbols} that extend plain
 | 
| 54330 | 603 | ASCII to allow infinitely many mathematical symbols within the | 
| 604 | formal sources. This works without depending on particular | |
| 54372 | 605 | encodings and varying Unicode standards | 
| 54352 | 606 |   \cite{Wenzel:2011:CICM}.\footnote{Raw Unicode characters within
 | 
| 607 | formal sources would compromise portability and reliability in the | |
| 54372 | 608 | face of changing interpretation of special features of Unicode, such | 
| 609 | as Combining Characters or Bi-directional Text.} | |
| 54323 | 610 | |
| 611 | For the prover back-end, formal text consists of ASCII characters | |
| 612 | that are grouped according to some simple rules, e.g.\ as plain | |
| 613 |   ``@{verbatim a}'' or symbolic ``@{verbatim "\<alpha>"}''.
 | |
| 614 | ||
| 54352 | 615 | For the editor front-end, a certain subset of symbols is rendered | 
| 616 |   physically via Unicode glyphs, in order to show ``@{verbatim "\<alpha>"}''
 | |
| 617 |   as ``@{text "\<alpha>"}'', for example.  This symbol interpretation is
 | |
| 618 |   specified by the Isabelle system distribution in @{file
 | |
| 619 | "$ISABELLE_HOME/etc/symbols"} and may be augmented by the user in | |
| 54372 | 620 |   @{file_unchecked "$ISABELLE_HOME_USER/etc/symbols"}.
 | 
| 54323 | 621 | |
| 54329 | 622 |   The appendix of \cite{isabelle-isar-ref} gives an overview of the
 | 
| 623 | standard interpretation of finitely many symbols from the infinite | |
| 54352 | 624 | collection. Uninterpreted symbols are displayed literally, e.g.\ | 
| 625 |   ``@{verbatim "\<foobar>"}''.  Overlap of Unicode characters used in
 | |
| 54372 | 626 | symbol interpretation with informal ones (which might appear e.g.\ | 
| 627 | in comments) needs to be avoided! Raw Unicode characters within | |
| 628 | prover source files should be restricted to informal parts, e.g.\ to | |
| 629 | write text in non-latin alphabets in comments. | |
| 54323 | 630 | |
| 54352 | 631 |   \medskip \paragraph{Encoding.} Technically, the Unicode view on
 | 
| 632 |   Isabelle symbols is an \emph{encoding} in jEdit (not in the
 | |
| 633 |   underlying JVM) that is called @{verbatim "UTF-8-Isabelle"}.  It is
 | |
| 634 | provided by the Isabelle/jEdit plugin and enabled by default for all | |
| 635 | source files. Sometimes such defaults are reset accidentally, or | |
| 636 | malformed UTF-8 sequences in the text force jEdit to fall back on a | |
| 637 |   different encoding like @{verbatim "ISO-8859-15"}.  In that case,
 | |
| 638 |   verbatim ``@{verbatim "\<alpha>"}'' will be shown in the text buffer
 | |
| 639 |   instead of its Unicode rendering ``@{text "\<alpha>"}''.  The jEdit menu
 | |
| 640 |   operation \emph{File / Reload with Encoding / UTF-8-Isabelle} helps
 | |
| 641 | to resolve such problems, potentially after repairing malformed | |
| 642 | parts of the text. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 643 | |
| 54352 | 644 |   \medskip \paragraph{Font.} Correct rendering via Unicode requires a
 | 
| 645 | font that contains glyphs for the corresponding codepoints. Most | |
| 646 | system fonts lack that, so Isabelle/jEdit prefers its own | |
| 647 |   application font @{verbatim IsabelleText}, which ensures that
 | |
| 648 | standard collection of Isabelle symbols are actually seen on the | |
| 649 | screen (or printer). | |
| 54329 | 650 | |
| 54330 | 651 | Note that a Java/AWT/Swing application can load additional fonts | 
| 54372 | 652 | only if they are not installed on the operating system already! | 
| 653 |   Some old version of @{verbatim IsabelleText} that happens to be
 | |
| 654 | provided by the operating system would prevents Isabelle/jEdit from | |
| 655 | its bundled version. This could lead to missing glyphs (black | |
| 656 |   rectangles), when the system version of @{verbatim IsabelleText} is
 | |
| 657 | older than the application version. This problem can be avoided by | |
| 658 |   refraining to ``install'' any version of @{verbatim IsabelleText} in
 | |
| 659 | the first place (although it might be occasionally tempting to use | |
| 660 | the same font in other applications). | |
| 54329 | 661 | |
| 54323 | 662 |   \medskip \paragraph{Input methods.} In principle, Isabelle/jEdit
 | 
| 663 | could delegate the problem to produce Isabelle symbols in their | |
| 664 | Unicode rendering to the underlying operating system and its | |
| 54329 | 665 |   \emph{input methods}.  Regular jEdit also provides various ways to
 | 
| 666 |   work with \emph{abbreviations} to produce certain non-ASCII
 | |
| 667 | characters. Since none of these standard input methods work | |
| 668 | satisfactorily for the mathematical characters required for | |
| 669 | Isabelle, various specific Isabelle/jEdit mechanisms are provided. | |
| 54323 | 670 | |
| 54329 | 671 | Here is a summary for practically relevant input methods for | 
| 54330 | 672 | Isabelle symbols: | 
| 54323 | 673 | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 674 |   \begin{enumerate}
 | 
| 54323 | 675 | |
| 54330 | 676 |   \item The \emph{Symbols} panel with some GUI buttons to insert
 | 
| 54323 | 677 | certain symbols in the text buffer. There are also tooltips to | 
| 54372 | 678 | reveal the official Isabelle representation with some additional | 
| 54323 | 679 |   information about \emph{symbol abbreviations} (see below).
 | 
| 680 | ||
| 681 | \item Copy / paste from decoded source files: text that is rendered | |
| 54352 | 682 | as Unicode already can be re-used to produce further text. This | 
| 683 | also works between different applications, e.g.\ Isabelle/jEdit and | |
| 684 | some web browser or mail client, as long as the same Unicode view on | |
| 685 | Isabelle symbols is used uniformly. | |
| 54323 | 686 | |
| 54329 | 687 | \item Copy / paste from prover output within Isabelle/jEdit. The | 
| 688 |   same principles as for text buffers apply, but note that \emph{copy}
 | |
| 54352 | 689 | in secondary Isabelle/jEdit windows works via the keyboard shortcut | 
| 54372 | 690 |   @{verbatim "C+c"}, while jEdit menu actions always refer to the
 | 
| 54352 | 691 | primary text area! | 
| 54323 | 692 | |
| 693 | \item Completion provided by Isabelle plugin (see | |
| 694 |   \secref{sec:completion}).  Isabelle symbols have a canonical name
 | |
| 695 | and optional abbreviations. This can be used with the text | |
| 696 | completion mechanism of Isabelle/jEdit, to replace a prefix of the | |
| 54329 | 697 |   actual symbol like @{verbatim "\<lambda>"}, or its backslashed name
 | 
| 698 |   @{verbatim "\\"}@{verbatim "lambda"}, or its ASCII abbreviation
 | |
| 54330 | 699 |   @{verbatim "%"} by the Unicode rendering.
 | 
| 54323 | 700 | |
| 701 | The following table is an extract of the information provided by the | |
| 54329 | 702 |   standard @{file "$ISABELLE_HOME/etc/symbols"} file:
 | 
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 703 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 704 | \medskip | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 705 |   \begin{tabular}{lll}
 | 
| 54372 | 706 |     \textbf{symbol} & \textbf{backslashed name} & \textbf{abbreviation} \\\hline
 | 
| 707 |     @{text "\<lambda>"} & @{verbatim "\\lambda"} & @{verbatim "%"} \\
 | |
| 708 |     @{text "\<Rightarrow>"} & @{verbatim "\\Rightarrow"} & @{verbatim "=>"} \\
 | |
| 709 |     @{text "\<Longrightarrow>"} & @{verbatim "\\Longrightarrow"} & @{verbatim "==>"} \\[0.5ex]
 | |
| 710 |     @{text "\<And>"} & @{verbatim "\\And"} & @{verbatim "!!"} \\
 | |
| 711 |     @{text "\<equiv>"} & @{verbatim "\\equiv"} & @{verbatim "=="} \\[0.5ex]
 | |
| 712 |     @{text "\<forall>"} & @{verbatim "\\forall"} & @{verbatim "!"} \\
 | |
| 713 |     @{text "\<exists>"} & @{verbatim "\\exists"} & @{verbatim "?"} \\
 | |
| 714 |     @{text "\<longrightarrow>"} & @{verbatim "\\longrightarrow"} & @{verbatim "-->"} \\
 | |
| 715 |     @{text "\<and>"} & @{verbatim "\\and"} & @{verbatim "&"} \\
 | |
| 716 |     @{text "\<or>"} & @{verbatim "\\or"} & @{verbatim "|"} \\
 | |
| 717 |     @{text "\<not>"} & @{verbatim "\\not"} & @{verbatim "~"} \\
 | |
| 718 |     @{text "\<noteq>"} & @{verbatim "\\noteq"} & @{verbatim "~="} \\
 | |
| 719 |     @{text "\<in>"} & @{verbatim "\\in"} & @{verbatim ":"} \\
 | |
| 720 |     @{text "\<notin>"} & @{verbatim "\\notin"} & @{verbatim "~:"} \\
 | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 721 |   \end{tabular}
 | 
| 54329 | 722 | \medskip | 
| 723 | ||
| 724 | Note that the above abbreviations refer to the input method. The | |
| 725 | logical notation provides ASCII alternatives that often coincide, | |
| 54352 | 726 | but deviate occasionally. This occasionally causes user confusion | 
| 727 | with very old-fashioned Isabelle source that use ASCII replacement | |
| 54372 | 728 |   notation like @{verbatim "!"} or @{verbatim "ALL"} directly in the
 | 
| 729 | text. | |
| 730 | ||
| 731 | On the other hand, coincidence of symbol abbreviations with ASCII | |
| 732 | replacement syntax syntax helps to update old theory sources via | |
| 733 |   explicit completion (see also @{verbatim "C+b"} explained in
 | |
| 734 |   \secref{sec:completion}).
 | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 735 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 736 |   \end{enumerate}
 | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 737 | |
| 54323 | 738 |   \paragraph{Control symbols.} There are some special control symbols
 | 
| 54372 | 739 | to modify the display style of a single symbol (without | 
| 740 | nesting). Control symbols may be applied to a region of selected | |
| 741 |   text, either using the \emph{Symbols} panel or keyboard shortcuts or
 | |
| 742 | jEdit actions. These editor operations produce a separate control | |
| 743 | symbol for each symbol in the text, in order to make the whole text | |
| 744 | appear in a certain style. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 745 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 746 | \medskip | 
| 54352 | 747 |   \begin{tabular}{llll}
 | 
| 54372 | 748 |     \textbf{style} & \textbf{symbol} & \textbf{shortcut} & \textbf{action} \\\hline
 | 
| 56059 
2390391584c2
some document antiquotations for Isabelle/jEdit elements;
 wenzelm parents: 
55880diff
changeset | 749 |     superscript & @{verbatim "\<^sup>"} & @{verbatim "C+e UP"} & @{action "isabelle.control-sup"} \\
 | 
| 
2390391584c2
some document antiquotations for Isabelle/jEdit elements;
 wenzelm parents: 
55880diff
changeset | 750 |     subscript & @{verbatim "\<^sub>"} & @{verbatim "C+e DOWN"} & @{action "isabelle.control-sub"} \\
 | 
| 
2390391584c2
some document antiquotations for Isabelle/jEdit elements;
 wenzelm parents: 
55880diff
changeset | 751 |     bold face & @{verbatim "\<^bold>"} & @{verbatim "C+e RIGHT"} & @{action "isabelle.control-bold"} \\
 | 
| 
2390391584c2
some document antiquotations for Isabelle/jEdit elements;
 wenzelm parents: 
55880diff
changeset | 752 |     reset & & @{verbatim "C+e LEFT"} & @{action "isabelle.control-reset"} \\
 | 
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 753 |   \end{tabular}
 | 
| 54352 | 754 | \medskip | 
| 54323 | 755 | |
| 54352 | 756 | To produce a single control symbol, it is also possible to complete | 
| 757 |   on @{verbatim "\\"}@{verbatim sup}, @{verbatim "\\"}@{verbatim sub},
 | |
| 758 |   @{verbatim "\\"}@{verbatim bold} as for regular symbols.  *}
 | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 759 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 760 | |
| 54355 | 761 | section {* Automatically tried tools \label{sec:auto-tools} *}
 | 
| 54353 | 762 | |
| 54354 | 763 | text {* Continuous document processing works asynchronously in the
 | 
| 764 | background. Visible document source that has been evaluated already | |
| 765 |   may get augmented by additional results of \emph{asynchronous print
 | |
| 766 | functions}. The canonical example is proof state output, which is | |
| 767 | always enabled. More heavy-weight print functions may be applied, | |
| 768 | in order to prove or disprove parts of the formal text by other | |
| 769 | means. | |
| 770 | ||
| 771 | Isabelle/HOL provides various automatically tried tools that operate | |
| 772 |   on outermost goal statements (e.g.\ @{command lemma}, @{command
 | |
| 773 | theorem}), independently of the state of the current proof attempt. | |
| 774 | They work implicitly without any arguments. Results are output as | |
| 775 |   \emph{information messages}, which are indicated in the text area by
 | |
| 54356 | 776 | blue squiggles and a blue information sign in the gutter (see | 
| 777 |   \figref{fig:auto-tools}).  The message content may be shown as for
 | |
| 54372 | 778 |   other output (see also \secref{sec:prover-output}).  Some tools
 | 
| 54356 | 779 |   produce output with \emph{sendback} markup, which means that
 | 
| 780 | clicking on certain parts of the output inserts that text into the | |
| 781 | source in the proper place. | |
| 782 | ||
| 54357 | 783 |   \begin{figure}[htb]
 | 
| 54356 | 784 |   \begin{center}
 | 
| 54373 | 785 |   \includegraphics[scale=0.3]{auto-tools}
 | 
| 54356 | 786 |   \end{center}
 | 
| 787 |   \caption{Results of automatically tried tools}
 | |
| 788 |   \label{fig:auto-tools}
 | |
| 789 |   \end{figure}
 | |
| 54354 | 790 | |
| 54372 | 791 | \medskip The following Isabelle system options control the behavior | 
| 54354 | 792 | of automatically tried tools (see also the jEdit dialog window | 
| 793 |   \emph{Plugin Options / Isabelle / General / Automatically tried
 | |
| 794 | tools}): | |
| 795 | ||
| 796 |   \begin{itemize}
 | |
| 797 | ||
| 798 |   \item @{system_option auto_methods} controls automatic use of a
 | |
| 799 |   combination of standard proof methods (@{method auto}, @{method
 | |
| 54372 | 800 |   simp}, @{method blast}, etc.).  This corresponds to the Isar command
 | 
| 54354 | 801 |   @{command "try0"}.
 | 
| 802 | ||
| 803 | The tool is disabled by default, since unparameterized invocation of | |
| 804 | standard proof methods often consumes substantial CPU resources | |
| 805 | without leading to success. | |
| 806 | ||
| 807 |   \item @{system_option auto_nitpick} controls a slightly reduced
 | |
| 808 |   version of @{command nitpick}, which tests for counterexamples using
 | |
| 809 | first-order relational logic. See also the Nitpick manual | |
| 810 |   \cite{isabelle-nitpick}.
 | |
| 811 | ||
| 812 | This tool is disabled by default, due to the extra overhead of | |
| 813 | invoking an external Java process for each attempt to disprove a | |
| 814 | subgoal. | |
| 815 | ||
| 816 |   \item @{system_option auto_quickcheck} controls automatic use of
 | |
| 817 |   @{command quickcheck}, which tests for counterexamples using a
 | |
| 818 | series of assignments for free variables of a subgoal. | |
| 819 | ||
| 820 |   This tool is \emph{enabled} by default.  It requires little
 | |
| 821 |   overhead, but is a bit weaker than @{command nitpick}.
 | |
| 822 | ||
| 823 |   \item @{system_option auto_sledgehammer} controls a significantly
 | |
| 824 |   reduced version of @{command sledgehammer}, which attempts to prove
 | |
| 825 | a subgoal using external automatic provers. See also the | |
| 826 |   Sledgehammer manual \cite{isabelle-sledgehammer}.
 | |
| 827 | ||
| 828 | This tool is disabled by default, due to the relatively heavy nature | |
| 829 | of Sledgehammer. | |
| 830 | ||
| 831 |   \item @{system_option auto_solve_direct} controls automatic use of
 | |
| 832 |   @{command solve_direct}, which checks whether the current subgoals
 | |
| 833 | can be solved directly by an existing theorem. This also helps to | |
| 834 | detect duplicate lemmas. | |
| 835 | ||
| 836 |   This tool is \emph{enabled} by default.
 | |
| 837 | ||
| 838 |   \end{itemize}
 | |
| 839 | ||
| 840 | Invocation of automatically tried tools is subject to some global | |
| 841 | policies of parallel execution, which may be configured as follows: | |
| 842 | ||
| 843 |   \begin{itemize}
 | |
| 844 | ||
| 845 |   \item @{system_option auto_time_limit} (default 2.0) determines the
 | |
| 54372 | 846 | timeout (in seconds) for each tool execution. | 
| 54354 | 847 | |
| 848 |   \item @{system_option auto_time_start} (default 1.0) determines the
 | |
| 849 | start delay (in seconds) for automatically tried tools, after the | |
| 850 | main command evaluation is finished. | |
| 851 | ||
| 852 |   \end{itemize}
 | |
| 853 | ||
| 854 | Each tool is submitted independently to the pool of parallel | |
| 855 | execution tasks in Isabelle/ML, using hardwired priorities according | |
| 856 | to its relative ``heaviness''. The main stages of evaluation and | |
| 857 | printing of proof states take precedence, but an already running | |
| 858 | tool is not canceled and may thus reduce reactivity of proof | |
| 859 | document processing. | |
| 860 | ||
| 861 | Users should experiment how the available CPU resources (number of | |
| 862 | cores) are best invested to get additional feedback from prover in | |
| 54372 | 863 | the background, by using a selection of weaker or stronger tools. | 
| 864 | *} | |
| 54353 | 865 | |
| 866 | ||
| 867 | section {* Sledgehammer \label{sec:sledgehammer} *}
 | |
| 868 | ||
| 54356 | 869 | text {* The \emph{Sledgehammer} panel (\figref{fig:sledgehammer})
 | 
| 54372 | 870 | provides a view on some independent execution of the Isar command | 
| 871 |   @{command sledgehammer}, with process indicator (spinning wheel) and
 | |
| 872 | GUI elements for important Sledgehammer arguments and options. Any | |
| 54356 | 873 | number of Sledgehammer panels may be active, according to the | 
| 54372 | 874 | standard policies of Dockable Window Management in jEdit. Closing | 
| 875 | such windows also cancels the corresponding prover tasks. | |
| 54356 | 876 | |
| 54357 | 877 |   \begin{figure}[htb]
 | 
| 54356 | 878 |   \begin{center}
 | 
| 879 |   \includegraphics[scale=0.3]{sledgehammer}
 | |
| 880 |   \end{center}
 | |
| 881 |   \caption{An instance of the Sledgehammer panel}
 | |
| 882 |   \label{fig:sledgehammer}
 | |
| 883 |   \end{figure}
 | |
| 54355 | 884 | |
| 885 |   The \emph{Apply} button attaches a fresh invocation of @{command
 | |
| 886 | sledgehammer} to the command where the cursor is pointing in the | |
| 887 | text --- this should be some pending proof problem. Further buttons | |
| 888 |   like \emph{Cancel} and \emph{Locate} help to manage the running
 | |
| 889 | process. | |
| 890 | ||
| 891 | Results appear incrementally in the output window of the panel. | |
| 54372 | 892 |   Proposed proof snippets are marked-up as \emph{sendback}, which
 | 
| 54355 | 893 | means a single mouse click inserts the text into a suitable place of | 
| 894 | the original source. Some manual editing may be required | |
| 895 | nonetheless, say to remove earlier proof attempts. *} | |
| 54353 | 896 | |
| 897 | ||
| 54362 | 898 | section {* Find theorems \label{sec:find} *}
 | 
| 54353 | 899 | |
| 54356 | 900 | text {* The \emph{Find} panel (\figref{fig:find}) provides an
 | 
| 54372 | 901 |   independent view for the Isar command @{command find_theorems}.  The
 | 
| 902 | main text field accepts search criteria according to the syntax | |
| 903 |   @{syntax thmcriterium} given in \cite{isabelle-isar-ref}.  Further
 | |
| 904 |   options of @{command find_theorems} are available via GUI elements.
 | |
| 54356 | 905 | |
| 54357 | 906 |   \begin{figure}[htb]
 | 
| 54356 | 907 |   \begin{center}
 | 
| 908 |   \includegraphics[scale=0.3]{find}
 | |
| 909 |   \end{center}
 | |
| 910 |   \caption{An instance of the Find panel}
 | |
| 911 |   \label{fig:find}
 | |
| 912 |   \end{figure}
 | |
| 54355 | 913 | |
| 914 |   The \emph{Apply} button attaches a fresh invocation of @{command
 | |
| 915 | find_theorems} to the current context of the command where the | |
| 916 | cursor is pointing in the text, unless an alternative theory context | |
| 917 | (from the underlying logic image) is specified explicitly. *} | |
| 54353 | 918 | |
| 919 | ||
| 54358 | 920 | chapter {* Miscellaneous tools *}
 | 
| 921 | ||
| 922 | section {* SideKick *}
 | |
| 923 | ||
| 924 | text {* The \emph{SideKick} plugin of jEdit provides some general
 | |
| 925 | services to display buffer structure in a tree view. | |
| 926 | ||
| 927 | Isabelle/jEdit provides SideKick parsers for its main mode for | |
| 928 |   theory files, as well as some minor modes for the @{verbatim NEWS}
 | |
| 54372 | 929 |   file, session @{verbatim ROOT} files, and system @{verbatim
 | 
| 930 | options}. | |
| 54358 | 931 | |
| 932 |   Moreover, the special SideKick parser @{verbatim "isabelle-markup"}
 | |
| 933 | provides access to the full (uninterpreted) markup tree of the PIDE | |
| 934 | document model of the current buffer. This is occasionally useful | |
| 935 | for informative purposes, but the amount of displayed information | |
| 936 | might cause problems for large buffers, both for the human and the | |
| 937 | machine. | |
| 938 | *} | |
| 939 | ||
| 940 | ||
| 54359 | 941 | section {* Timing *}
 | 
| 942 | ||
| 943 | text {* Managed evaluation of commands within PIDE documents includes
 | |
| 944 | timing information, which consists of elapsed (wall-clock) time, CPU | |
| 945 | time, and GC (garbage collection) time. Note that in a | |
| 946 | multithreaded system it is difficult to measure execution time | |
| 947 | precisely: elapsed time is closer to the real requirements of | |
| 948 | runtime resources than CPU or GC time, which are both subject to | |
| 949 | influences from the parallel environment that are outside the scope | |
| 950 | of the current command transaction. | |
| 951 | ||
| 952 |   The \emph{Timing} panel provides an overview of cumulative command
 | |
| 953 | timings for each document node. Commands with elapsed time below | |
| 954 | the given threshold are ignored in the grand total. Nodes are | |
| 955 | sorted according to their overall timing. For the document node | |
| 956 | that corresponds to the current buffer, individual command timings | |
| 957 | are shown as well. A double-click on a theory node or command moves | |
| 958 | the editor focus to that particular source position. | |
| 959 | ||
| 960 | It is also possible to reveal individual timing information via some | |
| 961 | tooltip for the corresponding command keyword, using the technique | |
| 962 |   of mouse hovering with @{verbatim CONTROL}/@{verbatim COMMAND}
 | |
| 963 |   modifier key as explained in \secref{sec:tooltips-hyperlinks}.
 | |
| 964 | Actual display of timing depends on the global option | |
| 965 |   @{system_option jedit_timing_threshold}, which can be configured in
 | |
| 54360 | 966 | "Plugin Options / Isabelle / General". | 
| 967 | ||
| 968 |   \medskip The \emph{Monitor} panel provides a general impression of
 | |
| 969 | recent activity of the farm of worker threads in Isabelle/ML. Its | |
| 970 |   display is continuously updated according to @{system_option
 | |
| 971 | editor_chart_delay}. Note that the painting of the chart takes | |
| 972 | considerable runtime itself --- on the Java Virtual Machine that | |
| 973 | runs Isabelle/Scala, not Isabelle/ML. Internally, the | |
| 974 |   Isabelle/Scala module @{verbatim isabelle.ML_Statistics} provides
 | |
| 54372 | 975 | further access to statistics of Isabelle/ML. *} | 
| 54359 | 976 | |
| 977 | ||
| 54358 | 978 | section {* Isabelle/Scala console *}
 | 
| 979 | ||
| 980 | text {* The \emph{Console} plugin of jEdit manages various shells
 | |
| 981 |   (command interpreters), e.g.\ \emph{BeanShell}, which is the
 | |
| 54372 | 982 | official jEdit scripting language, and the cross-platform | 
| 983 |   \emph{System} shell.  Thus the console provides similar
 | |
| 984 |   functionality than the special buffers @{verbatim "*scratch*"} and
 | |
| 985 |   @{verbatim "*shell*"} in Emacs.
 | |
| 54358 | 986 | |
| 987 | Isabelle/jEdit extends the repertoire of the console by | |
| 988 |   \emph{Scala}, which is the regular Scala toplevel loop running
 | |
| 989 |   inside the \emph{same} JVM process as Isabelle/jEdit itself.  This
 | |
| 990 | means the Scala command interpreter has access to the JVM name space | |
| 991 | and state of the running Prover IDE application: the main entry | |
| 992 |   points are @{verbatim view} (the current editor view of jEdit) and
 | |
| 993 |   @{verbatim PIDE} (the Isabelle/jEdit plugin object).
 | |
| 994 | ||
| 995 | For example, the subsequent Scala snippet gets the PIDE document | |
| 996 | model of the current buffer within the current editor view: | |
| 997 | ||
| 998 |   \begin{center}
 | |
| 999 |   @{verbatim "PIDE.document_model(view.getBuffer).get"}
 | |
| 1000 |   \end{center}
 | |
| 1001 | ||
| 1002 | This helps to explore Isabelle/Scala functionality interactively. | |
| 1003 | Some care is required to avoid interference with the internals of | |
| 1004 | the running application, especially in production use. *} | |
| 1005 | ||
| 1006 | ||
| 1007 | section {* Low-level output *}
 | |
| 1008 | ||
| 1009 | text {* Prover output is normally shown directly in the main text area
 | |
| 54372 | 1010 |   or secondary \emph{Output} panels, as explained in
 | 
| 54358 | 1011 |   \secref{sec:prover-output}.
 | 
| 1012 | ||
| 1013 | Beyond this, it is occasionally useful to inspect low-level output | |
| 1014 | channels via some of the following additional panels: | |
| 1015 | ||
| 1016 |   \begin{itemize}
 | |
| 1017 | ||
| 1018 |   \item \emph{Protocol} shows internal messages between the
 | |
| 1019 | Isabelle/Scala and Isabelle/ML side of the PIDE editing protocol. | |
| 1020 | Recording of messages starts with the first activation of the | |
| 1021 | corresponding dockable window; earlier messages are lost. | |
| 1022 | ||
| 1023 | Actual display of protocol messages causes considerable slowdown, so | |
| 54372 | 1024 |   it is important to undock all \emph{Protocol} panels for production
 | 
| 1025 | work. | |
| 54358 | 1026 | |
| 1027 |   \item \emph{Raw Output} shows chunks of text from the @{verbatim
 | |
| 1028 |   stdout} and @{verbatim stderr} channels of the prover process.
 | |
| 1029 | Recording of output starts with the first activation of the | |
| 1030 | corresponding dockable window; earlier output is lost. | |
| 1031 | ||
| 1032 | The implicit stateful nature of physical I/O channels makes it | |
| 1033 | difficult to relate raw output to the actual command from where it | |
| 1034 | was originating. Parallel execution may add to the confusion. | |
| 1035 | Peeking at physical process I/O is only the last resort to diagnose | |
| 1036 | problems with tools that are not fully PIDE compliant. | |
| 1037 | ||
| 1038 | Under normal circumstances, prover output always works via managed | |
| 1039 |   message channels (corresponding to @{ML writeln}, @{ML warning},
 | |
| 1040 |   @{ML error} etc.\ in Isabelle/ML), which are displayed by regular
 | |
| 1041 |   means within the document model (\secref{sec:prover-output}).
 | |
| 1042 | ||
| 1043 |   \item \emph{Syslog} shows system messages that might be relevant to
 | |
| 1044 | diagnose problems with the startup or shutdown phase of the prover | |
| 1045 |   process; this also includes raw output on @{verbatim stderr}.
 | |
| 1046 | ||
| 1047 | A limited amount of syslog messages are buffered, independently of | |
| 1048 |   the docking state of the \emph{Syslog} panel.  This allows to
 | |
| 1049 | diagnose serious problems with Isabelle/PIDE process management, | |
| 1050 | outside of the actual protocol layer. | |
| 1051 | ||
| 1052 | Under normal situations, such low-level system output can be | |
| 1053 | ignored. | |
| 1054 | ||
| 1055 |   \end{itemize}
 | |
| 1056 | *} | |
| 1057 | ||
| 1058 | ||
| 54329 | 1059 | chapter {* Known problems and workarounds \label{sec:problems} *}
 | 
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1060 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1061 | text {*
 | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1062 |   \begin{itemize}
 | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1063 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1064 |   \item \textbf{Problem:} Odd behavior of some diagnostic commands with
 | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1065 | global side-effects, like writing a physical file. | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1066 | |
| 54372 | 1067 |   \textbf{Workaround:} Copy / paste complete command text from
 | 
| 1068 | elsewhere, or discontinue continuous checking temporarily. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1069 | |
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1070 |   \item \textbf{Problem:} No way to delete document nodes from the overall
 | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1071 | collection of theories. | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1072 | |
| 54329 | 1073 |   \textbf{Workaround:} Ignore unused files.  Restart whole
 | 
| 1074 | Isabelle/jEdit session in worst-case situation. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1075 | |
| 54330 | 1076 |   \item \textbf{Problem:} Keyboard shortcuts @{verbatim "C+PLUS"} and
 | 
| 1077 |   @{verbatim "C+MINUS"} for adjusting the editor font size depend on
 | |
| 1078 | platform details and national keyboards. | |
| 1079 | ||
| 54372 | 1080 |   \textbf{Workaround:} Rebind keys via \emph{Global Options /
 | 
| 1081 | Shortcuts}. | |
| 54330 | 1082 | |
| 53886 
c83727c7a510
updated documentation concerning MacOSX plugin 1.3;
 wenzelm parents: 
53773diff
changeset | 1083 |   \item \textbf{Problem:} The Mac OS X keyboard shortcut @{verbatim
 | 
| 54349 | 1084 |   "COMMAND+COMMA"} for application \emph{Preferences} is in conflict
 | 
| 1085 |   with the jEdit default shortcut for \emph{Incremental Search Bar}
 | |
| 1086 |   (action @{verbatim "quick-search"}).
 | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1087 | |
| 54372 | 1088 |   \textbf{Workaround:} Rebind key via \emph{Global Options /
 | 
| 1089 |   Shortcuts} according to national keyboard, e.g.\ @{verbatim
 | |
| 1090 | "COMMAND+SLASH"} on English ones. | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1091 | |
| 54349 | 1092 |   \item \textbf{Problem:} Mac OS X system fonts sometimes lead to
 | 
| 1093 | character drop-outs in the main text area. | |
| 1094 | ||
| 1095 |   \textbf{Workaround:} Use the default @{verbatim IsabelleText} font.
 | |
| 1096 | (Do not install that font on the system.) | |
| 1097 | ||
| 54329 | 1098 |   \item \textbf{Problem:} Some Linux / X11 input methods such as IBus
 | 
| 54330 | 1099 | tend to disrupt key event handling of Java/AWT/Swing. | 
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1100 | |
| 54329 | 1101 |   \textbf{Workaround:} Do not use input methods, reset the environment
 | 
| 1102 |   variable @{verbatim XMODIFIERS} within Isabelle settings (default in
 | |
| 54639 | 1103 | Isabelle2013-2). | 
| 54329 | 1104 | |
| 1105 |   \item \textbf{Problem:} Some Linux / X11 window managers that are
 | |
| 1106 | not ``re-parenting'' cause problems with additional windows opened | |
| 54331 | 1107 | by Java. This affects either historic or neo-minimalistic window | 
| 1108 |   managers like @{verbatim awesome} or @{verbatim xmonad}.
 | |
| 54329 | 1109 | |
| 1110 |   \textbf{Workaround:} Use regular re-parenting window manager.
 | |
| 1111 | ||
| 1112 |   \item \textbf{Problem:} Recent forks of Linux / X11 window managers
 | |
| 1113 | and desktop environments (variants of Gnome) disrupt the handling of | |
| 1114 | menu popups and mouse positions of Java/AWT/Swing. | |
| 1115 | ||
| 1116 |   \textbf{Workaround:} Use mainstream versions of Linux desktops.
 | |
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1117 | |
| 54349 | 1118 |   \item \textbf{Problem:} Full-screen mode via jEdit action @{verbatim
 | 
| 1119 |   "toggle-full-screen"} (default shortcut @{verbatim F11}) works on
 | |
| 1120 | Windows, but not on Mac OS X or various Linux / X11 window managers. | |
| 1121 | ||
| 1122 |   \textbf{Workaround:} Use native full-screen control of the window
 | |
| 54372 | 1123 | manager (notably on Mac OS X). | 
| 54349 | 1124 | |
| 1125 |   \item \textbf{Problem:} Full-screen mode and dockable windows in
 | |
| 1126 |   \emph{floating} state may lead to confusion about window placement
 | |
| 1127 | (depending on platform characteristics). | |
| 1128 | ||
| 1129 |   \textbf{Workaround:} Avoid this combination.
 | |
| 1130 | ||
| 53770 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1131 |   \end{itemize}
 | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1132 | *} | 
| 
db362319d766
added/updated material from src/Tools/jEdit/README.html;
 wenzelm parents: 
53769diff
changeset | 1133 | |
| 53769 | 1134 | end |