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