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