src/Doc/JEdit/JEdit.thy
author wenzelm
Thu, 31 Oct 2013 16:10:35 +0100
changeset 54354 4e6defdc24ac
parent 54353 6692c355ebc1
child 54355 08cbb9461b48
permissions -rw-r--r--
more on automatically tried tools;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
53981
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
     1
(*:wrap=hard:maxLineLen=78:*)
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
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
     7
chapter {* Introduction *}
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
     8
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
     9
section {* Concepts and terminology *}
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    10
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    11
text {* Isabelle/jEdit is a Prover IDE that integrates \emph{parallel
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    12
  proof checking} \cite{Wenzel:2009,Wenzel:2013:ITP} with
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    13
  \emph{asynchronous user interaction}
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    14
  \cite{Wenzel:2010,Wenzel:2012:UITP-EPTCS}, based on a
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    15
  document-oriented approach to \emph{continuous proof processing}
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    16
  \cite{Wenzel:2011:CICM,Wenzel:2012}. Many concepts and system
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    17
  components are fit together in order to make this work. The main
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    18
  building blocks are as follows.
53769
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
    19
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    20
  \begin{description}
53769
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
    21
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    22
  \item [PIDE] is a general framework for Prover IDEs based on
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    23
  Isabelle/Scala. It is built around a concept of parallel and
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    24
  asynchronous document processing, which is supported natively by the
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    25
  parallel proof engine that is implemented in Isabelle/ML. The prover
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    26
  discontinues the traditional TTY-based command loop, and supports
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    27
  direct editing of formal source text with rich formal markup for GUI
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    28
  rendering.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    29
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    30
  \item [Isabelle/ML] is the implementation and extension language of
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    31
  Isabelle, see also \cite{isabelle-implementation}. It is integrated
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    32
  into the logical context of Isabelle/Isar and allows to manipulate
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    33
  logical entities directly. Arbitrary add-on tools may be implemented
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    34
  for object-logics such as Isabelle/HOL.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    35
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    36
  \item [Isabelle/Scala] is the system programming language of
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    37
  Isabelle. It extends the pure logical environment of Isabelle/ML
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    38
  towards the ``real world'' of graphical user interfaces, text
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    39
  editors, IDE frameworks, web services etc.  Special infrastructure
54348
wenzelm
parents: 54331
diff changeset
    40
  allows to transfer algebraic datatype values and formatted text
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    41
  easily between ML and Scala, using asynchronous protocol commands.
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    42
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    43
  \item [jEdit] is a sophisticated text editor implemented in
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    44
  Java.\footnote{\url{http://www.jedit.org}} It is easily extensible
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    45
  by plugins written in languages that work on the JVM, e.g.\
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    46
  Scala\footnote{\url{http://www.scala-lang.org/}}.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    47
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    48
  \item [Isabelle/jEdit] is the main example application of the PIDE
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    49
  framework and the default user-interface for Isabelle. It targets
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    50
  both beginners and experts. Technically, Isabelle/jEdit combines a
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    51
  slightly modified version of the official jEdit code base with a
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    52
  special plugin for Isabelle, integrated as standalone application
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    53
  for the main operating system platforms: Linux, Windows, Mac OS X.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    54
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    55
  \end{description}
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    56
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    57
  The subtle differences of Isabelle/ML versus Standard ML,
54330
wenzelm
parents: 54329
diff changeset
    58
  Isabelle/Scala versus Scala, Isabelle/jEdit versus jEdit need to be
wenzelm
parents: 54329
diff changeset
    59
  taken into account when discussing any of these PIDE building blocks
wenzelm
parents: 54329
diff changeset
    60
  in public forums, mailing lists, or even scientific publications.
wenzelm
parents: 54329
diff changeset
    61
  *}
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    62
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    63
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    64
section {* The Isabelle/jEdit Prover IDE *}
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    65
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    66
text {*
54331
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
    67
  \begin{figure}[htbp]
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
    68
  \begin{center}
53773
36703fcea740 added canonical screenshot;
wenzelm
parents: 53771
diff changeset
    69
  \includegraphics[width=\textwidth]{isabelle-jedit}
54331
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
    70
  \end{center}
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
    71
  \caption{The Isabelle/jEdit Prover IDE}
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
    72
  \end{figure}
53773
36703fcea740 added canonical screenshot;
wenzelm
parents: 53771
diff changeset
    73
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    74
  Isabelle/jEdit consists of some plugins for the well-known jEdit
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    75
  text editor \url{http://www.jedit.org}, according to the following
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    76
  principles.
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    77
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    78
  \begin{itemize}
53769
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
    79
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    80
  \item The original jEdit look-and-feel is generally preserved,
54348
wenzelm
parents: 54331
diff changeset
    81
  although some default properties are changed to accommodate Isabelle
wenzelm
parents: 54331
diff changeset
    82
  (e.g.\ the text area font).
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    83
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    84
  \item Formal Isabelle/Isar text is checked asynchronously while
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    85
  editing.  The user is in full command of the editor, and the prover
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    86
  refrains from locking portions of the buffer.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    87
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    88
  \item Prover feedback works via colors, boxes, squiggly underline,
54348
wenzelm
parents: 54331
diff changeset
    89
  hyperlinks, popup windows, icons, clickable output --- all based on
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    90
  semantic markup produced by Isabelle in the background.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    91
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    92
  \item Using the mouse together with the modifier key @{verbatim
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    93
  CONTROL} (Linux, Windows) or @{verbatim COMMAND} (Mac OS X) exposes
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    94
  additional formal content via tooltips and/or hyperlinks.
53769
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
    95
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    96
  \item Formal output (in popups etc.) may be explored recursively,
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
    97
  using the same techniques as in the editor source buffer.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
    98
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
    99
  \item Additional panels (e.g.\ \emph{Output}, \emph{Symbols}) are
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   100
  organized by the Dockable Window Manager of jEdit, which also allows
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   101
  multiple floating instances of each window class.
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   102
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   103
  \item The prover process and source files are managed on the editor
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   104
  side.  The prover operates on timeless and stateless document
54348
wenzelm
parents: 54331
diff changeset
   105
  content as provided via Isabelle/Scala.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   106
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   107
  \item Plugin options of jEdit (for the \emph{Isabelle} plugin) give
54348
wenzelm
parents: 54331
diff changeset
   108
  access to a selection of Isabelle/Scala options and its persistent
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   109
  preferences, usually with immediate effect on the prover back-end or
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   110
  editor front-end.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   111
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   112
  \item The logic image of the prover session may be specified within
54348
wenzelm
parents: 54331
diff changeset
   113
  Isabelle/jEdit.  The new image is provided automatically by the
wenzelm
parents: 54331
diff changeset
   114
  Isabelle build tool after restart of the application.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   115
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   116
  \end{itemize}
53769
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
   117
*}
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
   118
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   119
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   120
subsection {* Documentation *}
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   121
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   122
text {* Regular jEdit documentation is accessible via its @{verbatim
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   123
  Help} menu or @{verbatim F1} keyboard shortcut. This includes a full
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   124
  \emph{User's Guide} and \emph{Frequently Asked Questions} for this
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   125
  sophisticated text editor.
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   126
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   127
  Most of this information about jEdit is relevant for Isabelle/jEdit
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   128
  as well, but one needs to keep in mind that defaults sometimes
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   129
  differ, and the official jEdit documentation does not know about the
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   130
  Isabelle plugin with its special support for theory editing.
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   131
*}
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   132
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   133
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   134
subsection {* Plugins *}
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   135
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   136
text {* The \emph{Plugin Manager} of jEdit allows to augment editor
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   137
  functionality by JVM modules (jars) that are provided by the central
54348
wenzelm
parents: 54331
diff changeset
   138
  plugin repository, which is accessible by various mirror sites.
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   139
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   140
  Connecting to the plugin server infrastructure of the jEdit project
54348
wenzelm
parents: 54331
diff changeset
   141
  allows to update bundled plugins or to add further functionality.
wenzelm
parents: 54331
diff changeset
   142
  This needs to be done with the usual care for such an open bazaar of
wenzelm
parents: 54331
diff changeset
   143
  contributions. Arbitrary combinations of add-on features are apt to
wenzelm
parents: 54331
diff changeset
   144
  cause problems.  It is advisable to start with the default
wenzelm
parents: 54331
diff changeset
   145
  configuration of Isabelle/jEdit and develop some understanding how
wenzelm
parents: 54331
diff changeset
   146
  it is supposed to work, before loading additional plugins at a grand
wenzelm
parents: 54331
diff changeset
   147
  scale.
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   148
54348
wenzelm
parents: 54331
diff changeset
   149
  \medskip The main \emph{Isabelle} plugin is an integral part of
wenzelm
parents: 54331
diff changeset
   150
  Isabelle/jEdit and needs to remain active at all times! A few
wenzelm
parents: 54331
diff changeset
   151
  additional plugins are bundled with Isabelle/jEdit for convenience
wenzelm
parents: 54331
diff changeset
   152
  or out of necessity, notably \emph{Console} with its Isabelle/Scala
wenzelm
parents: 54331
diff changeset
   153
  sub-plugin and \emph{SideKick} with some Isabelle-specific parsers
wenzelm
parents: 54331
diff changeset
   154
  for document tree structure.  The \emph{ErrorList} plugin is
wenzelm
parents: 54331
diff changeset
   155
  required by \emph{SideKick}, but not used specifically in
wenzelm
parents: 54331
diff changeset
   156
  Isabelle/jEdit. *}
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   157
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   158
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   159
subsection {* Options *}
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   160
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   161
text {* Both jEdit and Isabelle have distinctive management of
54348
wenzelm
parents: 54331
diff changeset
   162
  persistent options.
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   163
54348
wenzelm
parents: 54331
diff changeset
   164
  Regular jEdit options are accessible via the dialogs for
wenzelm
parents: 54331
diff changeset
   165
  \emph{Global Options} and \emph{Plugin Options}.  Changed properties
wenzelm
parents: 54331
diff changeset
   166
  are stored eventually in @{verbatim
wenzelm
parents: 54331
diff changeset
   167
  "$ISABELLE_HOME_USER/jedit/properties"}.  In contrast, Isabelle
wenzelm
parents: 54331
diff changeset
   168
  system options are managed by Isabelle/Scala and changes stored in
wenzelm
parents: 54331
diff changeset
   169
  @{verbatim "$ISABELLE_HOME_USER/etc/preferences"}, independently of
wenzelm
parents: 54331
diff changeset
   170
  the jEdit properties.  See also \cite{isabelle-sys}, especially the
wenzelm
parents: 54331
diff changeset
   171
  coverage of sessions and command-line tools like @{tool build} or
wenzelm
parents: 54331
diff changeset
   172
  @{tool options}.
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   173
54348
wenzelm
parents: 54331
diff changeset
   174
  Those Isabelle options that are declared as \textbf{public} are
wenzelm
parents: 54331
diff changeset
   175
  configurable in jEdit via \emph{Plugin Options / Isabelle /
wenzelm
parents: 54331
diff changeset
   176
  General}.  Moreover, there are various options for rendering of
wenzelm
parents: 54331
diff changeset
   177
  document content, which are configurable via \emph{Plugin Options /
wenzelm
parents: 54331
diff changeset
   178
  Isabelle / Rendering}.  Thus \emph{Plugin Options / Isabelle} in
wenzelm
parents: 54331
diff changeset
   179
  jEdit provides a view on certain Isabelle options.  Note that some
wenzelm
parents: 54331
diff changeset
   180
  of these options affect general parameters that are relevant outside
wenzelm
parents: 54331
diff changeset
   181
  Isabelle/jEdit as well, e.g.\ @{system_option threads} or
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   182
  @{system_option parallel_proofs} for the Isabelle build tool
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   183
  \cite{isabelle-sys}.
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   184
54348
wenzelm
parents: 54331
diff changeset
   185
  \medskip All options are loaded on startup and saved on shutdown of
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   186
  Isabelle/jEdit.  Editing the machine-generated files @{verbatim
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   187
  "$ISABELLE_HOME_USER/jedit/properties"} or @{verbatim
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   188
  "$ISABELLE_HOME_USER/etc/preferences"} manually while the
54348
wenzelm
parents: 54331
diff changeset
   189
  application is running is likely to cause surprise due to lost
wenzelm
parents: 54331
diff changeset
   190
  update!  *}
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   191
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   192
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   193
subsection {* Keymaps *}
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   194
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   195
text {* Keyboard shortcuts used to be managed as jEdit properties in
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   196
  the past, but recent versions (2013) have a separate concept of
54348
wenzelm
parents: 54331
diff changeset
   197
  \emph{keymap} that is configurable via \emph{Global Options /
wenzelm
parents: 54331
diff changeset
   198
  Shortcuts}.  The @{verbatim imported} keymap is derived from the
54330
wenzelm
parents: 54329
diff changeset
   199
  initial environment of properties that is available at the first
54348
wenzelm
parents: 54331
diff changeset
   200
  start of the editor; afterwards the keymap file takes precedence.
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   201
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   202
  This is relevant for Isabelle/jEdit due to various fine-tuning of
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   203
  default properties, and additional keyboard shortcuts for Isabelle
54348
wenzelm
parents: 54331
diff changeset
   204
  specific functionality. Users may change their keymap, but need to
wenzelm
parents: 54331
diff changeset
   205
  copy Isabelle-specific key bindings manually.  *}
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   206
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   207
54322
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   208
subsection {* Look-and-feel *}
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   209
54330
wenzelm
parents: 54329
diff changeset
   210
text {* jEdit is a Java/AWT/Swing application with some ambition to
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   211
  support ``native'' look-and-feel on all platforms, within the limits
54330
wenzelm
parents: 54329
diff changeset
   212
  of what Oracle as Java provider and major operating system
wenzelm
parents: 54329
diff changeset
   213
  distributors allow (see also \secref{sec:problems}).
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   214
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   215
  Isabelle/jEdit enables platform-specific look-and-feel by default as
54330
wenzelm
parents: 54329
diff changeset
   216
  follows:
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   217
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   218
  \begin{description}
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   219
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   220
  \item[Linux] The platform-independent \emph{Nimbus} is used by
54322
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   221
  default, but the classic \emph{Metal} also works.  \emph{GTK+} works
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   222
  under the side-condition that the overall GTK theme is selected in a
54348
wenzelm
parents: 54331
diff changeset
   223
  Swing-friendly way.\footnote{GTK support in Java/Swing was once
wenzelm
parents: 54331
diff changeset
   224
  marketed aggressively by Sun, but never quite finished, and is today
wenzelm
parents: 54331
diff changeset
   225
  (2013) lagging a bit behind further development of Swing and GTK.}
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   226
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   227
  \item[Windows] Regular \emph{Windows} is used by default, but
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   228
  platform-independent \emph{Nimbus} and \emph{Metal} also work.
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   229
54348
wenzelm
parents: 54331
diff changeset
   230
  \item[Mac OS X] Regular \emph{Mac OS X} is used by default, but
wenzelm
parents: 54331
diff changeset
   231
  platform-independent \emph{Nimbus} and \emph{Metal} also work.
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   232
  Moreover the bundled \emph{MacOSX} plugin provides various functions
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   233
  that are expected from applications on that particular platform:
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   234
  quit from menu or dock, preferences menu, drag-and-drop of text
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   235
  files on the application, full-screen mode for main editor windows
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   236
  etc.
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   237
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   238
  \end{description}
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   239
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   240
  Users may experiment with different look-and-feels, but need to keep
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   241
  in mind that this extra variance of GUI functionality is unlikely to
54348
wenzelm
parents: 54331
diff changeset
   242
  work in arbitrary combinations.  The historic \emph{CDE/Motif} is
wenzelm
parents: 54331
diff changeset
   243
  better avoided.  After changing the look-and-feel in \emph{Global
wenzelm
parents: 54331
diff changeset
   244
  Options / Appearance}, it is advisable to restart Isabelle/jEdit in
wenzelm
parents: 54331
diff changeset
   245
  order to take full effect.  *}
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   246
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   247
54351
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   248
subsection {* File-system access *}
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   249
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   250
text {* File specifications in jEdit follow various formats and
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   251
  conventions according to \emph{Virtual File Systems}, which may be
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   252
  also provided by additional plugins.  This allows to access remote
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   253
  files via the @{verbatim "http:"} protocol prefix, for example.
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   254
  Isabelle/jEdit attempts to work with the file-system access model of
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   255
  jEdit as far as possible.  In particular, theory sources are passed
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   256
  directly from the editor to the prover, without indirection via
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   257
  files.
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   258
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   259
  Despite the flexibility of URLs in jEdit, local files are
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   260
  particularly important and are accessible without protocol prefix.
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   261
  Here the path notation is that of the Java Virtual Machine on the
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   262
  underlying platform.  On Windows the preferred form uses
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   263
  backslashes, but happens to accept Unix/POSIX forward slashes, too.
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   264
  Further differences arise due to drive letters and network shares.
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   265
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   266
  The Java notation for files needs to be distinguished from the one
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   267
  of Isabelle, which uses POSIX notation with forward slashes on
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   268
  \emph{all} platforms.\footnote{Isabelle on Windows uses Cygwin
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   269
  file-system access.}  Moreover, environment variables from the
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   270
  Isabelle process may be used freely, e.g.\ @{file
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   271
  "$ISABELLE_HOME/etc/symbols"} or @{file
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   272
  "$ISABELLE_JDK_HOME/README.html"}.  There are special shortcuts:
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   273
  @{verbatim "~"} for @{file "$USER_HOME"}, and @{verbatim "~~"} for
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   274
  @{file "$ISABELLE_HOME"}.
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   275
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   276
  \medskip Since jEdit happens to support environment variables within
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   277
  file specifications as well, it is natural to use similar notation
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   278
  within the editor, e.g.\ in the file-browser.  This does not work in
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   279
  full generality, though, due to the bias of jEdit towards
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   280
  platform-specific notation and of Isabelle towards POSIX.  Moreover,
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   281
  the Isabelle settings environment is not yet active when starting
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   282
  Isabelle/jEdit via its standard application wrapper (in contrast to
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   283
  @{verbatim "isabelle jedit"} run from the command line).
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   284
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   285
  For convenience, Isabelle/jEdit imitates at least @{verbatim
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   286
  "$ISABELLE_HOME"} and @{verbatim "$ISABELLE_HOME_USER"} within the
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   287
  Java process environment, in order to allow easy access to these
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   288
  important places from the editor.
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   289
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   290
  Moreover note that path specifications in prover input or output
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   291
  usually include formal markup that turns it into a hyperlink (see
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   292
  also \secref{sec:tooltips-hyperlinks}).  This allows to open the
54351
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   293
  corresponding file in the text editor, independently of the path
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   294
  notation.  *}
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   295
5cbe32533cdb more on file-system access;
wenzelm
parents: 54350
diff changeset
   296
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   297
chapter {* Prover IDE functionality *}
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   298
54353
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   299
section {* Text buffers and theories \label{sec:buffers-theories} *}
54322
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   300
54350
wenzelm
parents: 54349
diff changeset
   301
text {* As regular text editor, jEdit maintains a collection of open
wenzelm
parents: 54349
diff changeset
   302
  \emph{text buffers} to store source files; each buffer may be
wenzelm
parents: 54349
diff changeset
   303
  associated with any number of visible \emph{text areas}.  Buffers
wenzelm
parents: 54349
diff changeset
   304
  are subject to an \emph{edit mode} that is determined from the file
wenzelm
parents: 54349
diff changeset
   305
  type.  Files with extension \texttt{.thy} are assigned to the mode
wenzelm
parents: 54349
diff changeset
   306
  \emph{isabelle} and treated specifically.
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   307
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   308
  \medskip Isabelle theory files are automatically added to the formal
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   309
  document model of Isabelle/Scala, which maintains a family of
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   310
  versions of all sources for the prover.  The \emph{Theories} panel
54322
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   311
  provides an overview of the status of continuous checking of theory
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   312
  sources.  Unlike batch sessions \cite{isabelle-sys}, theory nodes
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   313
  are identified by full path names; this allows to work with multiple
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   314
  (disjoint) Isabelle sessions simultaneously within the same editor
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   315
  session.
54322
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   316
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   317
  Certain events to open or update buffers with theory files cause
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   318
  Isabelle/jEdit to resolve dependencies of \emph{theory imports}.
54350
wenzelm
parents: 54349
diff changeset
   319
  The system requests to load additional files into editor buffers, in
wenzelm
parents: 54349
diff changeset
   320
  order to be included in the theory document model for further
wenzelm
parents: 54349
diff changeset
   321
  checking.  It is also possible to resolve dependencies
wenzelm
parents: 54349
diff changeset
   322
  automatically, depending on \emph{Plugin Options / Isabelle /
wenzelm
parents: 54349
diff changeset
   323
  General / Auto load} (Isabelle system option @{system_option
wenzelm
parents: 54349
diff changeset
   324
  jedit_auto_load}).
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   325
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   326
  \medskip The open text area views on theory buffers define the
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   327
  visible \emph{perspective} of Isabelle/jEdit.  This is taken as a
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   328
  hint for document processing: the prover ensures that those parts of
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   329
  a theory where the user is looking are checked, while other parts
54350
wenzelm
parents: 54349
diff changeset
   330
  that are presently not required are ignored.  The perspective is
54330
wenzelm
parents: 54329
diff changeset
   331
  changed by opening or closing text area windows, or scrolling within
wenzelm
parents: 54329
diff changeset
   332
  some window.
54322
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   333
54330
wenzelm
parents: 54329
diff changeset
   334
  The \emph{Theories} panel provides some further options to influence
wenzelm
parents: 54329
diff changeset
   335
  the process of continuous checking: it may be switched off globally
54350
wenzelm
parents: 54349
diff changeset
   336
  to restrict the prover to superficial processing of command syntax.
wenzelm
parents: 54349
diff changeset
   337
  It is also possible to indicate theory nodes as \emph{required} for
54330
wenzelm
parents: 54329
diff changeset
   338
  continuous checking: this means such nodes and all their imports are
wenzelm
parents: 54329
diff changeset
   339
  always processed independently of the visibility status (if
wenzelm
parents: 54329
diff changeset
   340
  continuous checking is enabled).  Big theory libraries that are
wenzelm
parents: 54329
diff changeset
   341
  marked as required can have significant impact on performance,
wenzelm
parents: 54329
diff changeset
   342
  though.
54322
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   343
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   344
  \medskip Formal markup of checked theory content is turned into GUI
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   345
  rendering, based on a standard repertoire known from IDEs for
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   346
  programming languages: colors, icons, highlighting, squiggly
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   347
  underline, tooltips, hyperlinks etc.  For outer syntax of
54350
wenzelm
parents: 54349
diff changeset
   348
  Isabelle/Isar there is some traditional syntax-highlighting via
wenzelm
parents: 54349
diff changeset
   349
  static keyword tables and tokenization within the editor.  In
wenzelm
parents: 54349
diff changeset
   350
  contrast, the painting of inner syntax (term language etc.)\ uses
wenzelm
parents: 54349
diff changeset
   351
  semantic information that is reported dynamically from the logical
wenzelm
parents: 54349
diff changeset
   352
  context.  Thus the prover can provide additional markup to help the
wenzelm
parents: 54349
diff changeset
   353
  user to understand the meaning of formal text, and to produce more
wenzelm
parents: 54349
diff changeset
   354
  text with some add-on tools (e.g.\ information messages by automated
wenzelm
parents: 54349
diff changeset
   355
  provers or disprovers running in the background).
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   356
*}
54322
9d156ded3c2a more documentation;
wenzelm
parents: 54321
diff changeset
   357
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   358
54354
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   359
section {* Prover output \label{sec:prover-output} *}
54353
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   360
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   361
text {* Prover output consists of \emph{markup} and \emph{messages}.
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   362
  Both are directly attached to the corresponding positions in the
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   363
  original source text, and visualized in the text area, e.g.\ as text
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   364
  colours for free and bound variables, or as squiggly underline for
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   365
  warnings, errors etc.  In the latter case, the corresponding
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   366
  messages are shown by hovering with the mouse over the highlighted
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   367
  text --- although in many situations the user should already get
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   368
  some clue by looking at the text highlighting alone.
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   369
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   370
  The ``gutter area'' on the left-hand-side of the text area uses
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   371
  icons to provide a summary of the messages within the corresponding
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   372
  line of text.  Message priorities are used to prefer errors over
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   373
  warnings, warnings over information messages etc.  Plain output is
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   374
  ignored here.
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   375
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   376
  The ``overview area'' on the right-hand-side of the text area uses
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   377
  similar information to paint small rectangles for the overall status
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   378
  of the whole text buffer.  The graphics is scaled to fit the logical
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   379
  buffer length into the given window height.  Mouse clicks on the
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   380
  overview area position the cursor approximately to the corresponding
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   381
  line of text in the buffer.
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   382
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   383
  Another course-grained overview is provided by the \emph{Theories}
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   384
  panel (\secref{sec:buffers-theories}), but without direct
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   385
  correspondence to text positions.  A double-click on one of the
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   386
  theory entries with their status overview opens the corresponding
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   387
  text buffer, without changing the cursor position.
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   388
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   389
  \medskip In addition, the \emph{Output} panel displays prover
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   390
  messages that correspond to a given command, within a separate
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   391
  window.
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   392
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   393
  The cursor position in the presently active text area determines the
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   394
  prover commands whose cumulative message output is appended an shown
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   395
  in that window (in canonical order according to the processing of
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   396
  the command).  There are also control elements to modify the update
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   397
  policy of the output wrt.\ continued editor movements.  This is
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   398
  particularly useful with several independent instances of the
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   399
  \emph{Output} panel, which the Dockable Window Manager of jEdit can
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   400
  handle conveniently.
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   401
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   402
  Former users of the old TTY interaction model (e.g.\ Proof~General)
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   403
  might find a separate window for prover messages familiar, but it is
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   404
  important to understand that the main Prover IDE feedback happens
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   405
  elsewhere.  It is possible to do meaningful proof editing
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   406
  efficiently while using the secondary window only rarely.
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   407
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   408
  The main purpose of the output window is to ``debug'' unclear
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   409
  situations by inspecting internal state of the prover.\footnote{In
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   410
  that sense, unstructured tactic scripts depend on continous
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   411
  debugging with internal state inspection.} Consequently, some
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   412
  special messages for \emph{tracing} or \emph{proof state} only
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   413
  appear here, and are not attached to the original source.
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   414
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   415
  \medskip In any case, prover messages also contain markup that may
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   416
  be explored recursively via tooltips or hyperlinks (see
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   417
  \secref{sec:tooltips-hyperlinks}), or clicked directly to initiate
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   418
  certain actions (e.g.\ see \secref{sec:sledgehammer}).  *}
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   419
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   420
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   421
section {* Tooltips and hyperlinks \label{sec:tooltips-hyperlinks} *}
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   422
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   423
text {* Formally processed text (prover input or output) contains rich
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   424
  markup information that can be explored further by using the
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   425
  @{verbatim CONTROL} modifier key on Linux and Windows, or @{verbatim
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   426
  COMMAND} on Mac OS X.  Hovering with the mouse while the modifier is
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   427
  pressed reveals a \emph{tooltip} (grey box over the text with a
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   428
  yellow popup) and/or a \emph{hyperlink} (black rectangle over the
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   429
  text); see also \figref{fig:tooltip}.
54331
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   430
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   431
  \begin{figure}[htbp]
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   432
  \begin{center}
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   433
  \includegraphics[scale=0.6]{popup1}
54331
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   434
  \end{center}
54350
wenzelm
parents: 54349
diff changeset
   435
  \caption{Tooltip and hyperlink for some formal entity.}
wenzelm
parents: 54349
diff changeset
   436
  \label{fig:tooltip}
54331
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   437
  \end{figure}
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   438
54350
wenzelm
parents: 54349
diff changeset
   439
  Tooltip popups use the same rendering principles as the main text
wenzelm
parents: 54349
diff changeset
   440
  area, and further tooltips and/or hyperlinks may be exposed
wenzelm
parents: 54349
diff changeset
   441
  recursively by the same mechanism; see
wenzelm
parents: 54349
diff changeset
   442
  \figref{fig:nested-tooltips}.
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   443
54331
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   444
  \begin{figure}[htbp]
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   445
  \begin{center}
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   446
  \includegraphics[scale=0.6]{popup2}
54331
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   447
  \end{center}
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   448
  \caption{Nested tooltips over formal entities.}
54350
wenzelm
parents: 54349
diff changeset
   449
  \label{fig:nested-tooltips}
54331
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   450
  \end{figure}
54350
wenzelm
parents: 54349
diff changeset
   451
wenzelm
parents: 54349
diff changeset
   452
  The tooltip popup window provides some controls to \emph{close} or
wenzelm
parents: 54349
diff changeset
   453
  \emph{detach} the window, turning it into a separate \emph{Info}
wenzelm
parents: 54349
diff changeset
   454
  dockable window managed by jEdit.  The @{verbatim ESCAPE} key closes
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   455
  \emph{all} popups, which is particularly relevant when nested
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   456
  tooltips are stacking up.
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   457
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   458
  \medskip A black rectangle in the text indicates a hyperlink that
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   459
  may be followed by a mouse click (while the @{verbatim CONTROL} or
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   460
  @{verbatim COMMAND} modifier key is still pressed). Presently (2013)
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   461
  there is no systematic way to return to the original location within
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   462
  the editor.
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   463
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   464
  Also note that the link target may be a file that is itself not
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   465
  subject to formal document processing of the editor session and thus
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   466
  prevents further exploration: the chain of hyperlinks may end in
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   467
  some source file of the underlying logic image, even within the
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   468
  Isabelle/ML bootstrap sources of Isabelle/Pure, where the formal
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   469
  markup is less detailed. *}
54321
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   470
41951f427ac7 more documentation;
wenzelm
parents: 54320
diff changeset
   471
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   472
section {* Isabelle symbols *}
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   473
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   474
text {* Isabelle sources consist of \emph{symbols} that extend plain
54330
wenzelm
parents: 54329
diff changeset
   475
  ASCII to allow infinitely many mathematical symbols within the
wenzelm
parents: 54329
diff changeset
   476
  formal sources.  This works without depending on particular
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   477
  encodings or varying Unicode standards
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   478
  \cite{Wenzel:2011:CICM}.\footnote{Raw Unicode characters within
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   479
  formal sources would compromise portability and reliability in the
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   480
  face of changing interpretation of various unexpected features of
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   481
  Unicode.}
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   482
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   483
  For the prover back-end, formal text consists of ASCII characters
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   484
  that are grouped according to some simple rules, e.g.\ as plain
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   485
  ``@{verbatim a}'' or symbolic ``@{verbatim "\<alpha>"}''.
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   486
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   487
  For the editor front-end, a certain subset of symbols is rendered
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   488
  physically via Unicode glyphs, in order to show ``@{verbatim "\<alpha>"}''
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   489
  as ``@{text "\<alpha>"}'', for example.  This symbol interpretation is
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   490
  specified by the Isabelle system distribution in @{file
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   491
  "$ISABELLE_HOME/etc/symbols"} and may be augmented by the user in
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   492
  @{verbatim "$ISABELLE_HOME_USER/etc/symbols"}.
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   493
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   494
  The appendix of \cite{isabelle-isar-ref} gives an overview of the
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   495
  standard interpretation of finitely many symbols from the infinite
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   496
  collection.  Uninterpreted symbols are displayed literally, e.g.\
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   497
  ``@{verbatim "\<foobar>"}''.  Overlap of Unicode characters used in
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   498
  symbol interpretation with informal ones that might appear e.g.\ in
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   499
  comments needs to be avoided!
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   500
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   501
  \medskip \paragraph{Encoding.} Technically, the Unicode view on
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   502
  Isabelle symbols is an \emph{encoding} in jEdit (not in the
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   503
  underlying JVM) that is called @{verbatim "UTF-8-Isabelle"}.  It is
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   504
  provided by the Isabelle/jEdit plugin and enabled by default for all
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   505
  source files.  Sometimes such defaults are reset accidentally, or
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   506
  malformed UTF-8 sequences in the text force jEdit to fall back on a
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   507
  different encoding like @{verbatim "ISO-8859-15"}.  In that case,
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   508
  verbatim ``@{verbatim "\<alpha>"}'' will be shown in the text buffer
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   509
  instead of its Unicode rendering ``@{text "\<alpha>"}''.  The jEdit menu
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   510
  operation \emph{File / Reload with Encoding / UTF-8-Isabelle} helps
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   511
  to resolve such problems, potentially after repairing malformed
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   512
  parts of the text.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   513
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   514
  \medskip \paragraph{Font.} Correct rendering via Unicode requires a
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   515
  font that contains glyphs for the corresponding codepoints.  Most
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   516
  system fonts lack that, so Isabelle/jEdit prefers its own
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   517
  application font @{verbatim IsabelleText}, which ensures that
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   518
  standard collection of Isabelle symbols are actually seen on the
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   519
  screen (or printer).
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   520
54330
wenzelm
parents: 54329
diff changeset
   521
  Note that a Java/AWT/Swing application can load additional fonts
wenzelm
parents: 54329
diff changeset
   522
  only if they are not installed as system font already!  This means
wenzelm
parents: 54329
diff changeset
   523
  some old version of @{verbatim IsabelleText} that happens to be
wenzelm
parents: 54329
diff changeset
   524
  already present prevents Isabelle/jEdit from using its current
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   525
  bundled version.  This results in missing glyphs (black rectangles),
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   526
  when some obsolete version of @{verbatim IsabelleText} is used by
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   527
  accident.  This problem can be avoided by refraining to ``install''
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   528
  any version of @{verbatim IsabelleText} in the first place.
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   529
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   530
  \medskip \paragraph{Input methods.} In principle, Isabelle/jEdit
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   531
  could delegate the problem to produce Isabelle symbols in their
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   532
  Unicode rendering to the underlying operating system and its
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   533
  \emph{input methods}.  Regular jEdit also provides various ways to
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   534
  work with \emph{abbreviations} to produce certain non-ASCII
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   535
  characters.  Since none of these standard input methods work
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   536
  satisfactorily for the mathematical characters required for
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   537
  Isabelle, various specific Isabelle/jEdit mechanisms are provided.
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   538
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   539
  Here is a summary for practically relevant input methods for
54330
wenzelm
parents: 54329
diff changeset
   540
  Isabelle symbols:
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   541
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   542
  \begin{enumerate}
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   543
54330
wenzelm
parents: 54329
diff changeset
   544
  \item The \emph{Symbols} panel with some GUI buttons to insert
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   545
  certain symbols in the text buffer.  There are also tooltips to
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   546
  reveal to official Isabelle representation with some additional
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   547
  information about \emph{symbol abbreviations} (see below).
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   548
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   549
  \item Copy / paste from decoded source files: text that is rendered
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   550
  as Unicode already can be re-used to produce further text.  This
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   551
  also works between different applications, e.g.\ Isabelle/jEdit and
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   552
  some web browser or mail client, as long as the same Unicode view on
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   553
  Isabelle symbols is used uniformly.
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   554
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   555
  \item Copy / paste from prover output within Isabelle/jEdit.  The
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   556
  same principles as for text buffers apply, but note that \emph{copy}
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   557
  in secondary Isabelle/jEdit windows works via the keyboard shortcut
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   558
  @{verbatim "C+c"}.  The jEdit menu actions always refer to the
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   559
  primary text area!
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   560
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   561
  \item Completion provided by Isabelle plugin (see
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   562
  \secref{sec:completion}).  Isabelle symbols have a canonical name
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   563
  and optional abbreviations.  This can be used with the text
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   564
  completion mechanism of Isabelle/jEdit, to replace a prefix of the
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   565
  actual symbol like @{verbatim "\<lambda>"}, or its backslashed name
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   566
  @{verbatim "\\"}@{verbatim "lambda"}, or its ASCII abbreviation
54330
wenzelm
parents: 54329
diff changeset
   567
  @{verbatim "%"} by the Unicode rendering.
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   568
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   569
  The following table is an extract of the information provided by the
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   570
  standard @{file "$ISABELLE_HOME/etc/symbols"} file:
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   571
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   572
  \medskip
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   573
  \begin{tabular}{lll}
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   574
    \textbf{symbol} & \textbf{abbreviation} & \textbf{backslashed name} \\\hline
53771
wenzelm
parents: 53770
diff changeset
   575
    @{text "\<lambda>"}   & @{verbatim "%"}     &  @{verbatim "\\lambda"}         \\
wenzelm
parents: 53770
diff changeset
   576
    @{text "\<Rightarrow>"}  & @{verbatim "=>"}    &  @{verbatim "\\Rightarrow"}     \\
wenzelm
parents: 53770
diff changeset
   577
    @{text "\<Longrightarrow>"} & @{verbatim "==>"}   &  @{verbatim "\\Longrightarrow"} \\
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   578
53771
wenzelm
parents: 53770
diff changeset
   579
    @{text "\<And>"}  & @{verbatim "!!"}    &  @{verbatim "\\And"}            \\
wenzelm
parents: 53770
diff changeset
   580
    @{text "\<equiv>"}  & @{verbatim "=="}    &  @{verbatim "\\equiv"}          \\
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   581
53771
wenzelm
parents: 53770
diff changeset
   582
    @{text "\<forall>"}   & @{verbatim "!"}     &  @{verbatim "\\forall"}         \\
wenzelm
parents: 53770
diff changeset
   583
    @{text "\<exists>"}   & @{verbatim "?"}     &  @{verbatim "\\exists"}         \\
wenzelm
parents: 53770
diff changeset
   584
    @{text "\<longrightarrow>"} & @{verbatim "-->"}   &  @{verbatim "\\longrightarrow"} \\
wenzelm
parents: 53770
diff changeset
   585
    @{text "\<and>"}   & @{verbatim "&"}     &  @{verbatim "\\and"}            \\
wenzelm
parents: 53770
diff changeset
   586
    @{text "\<or>"}   & @{verbatim "|"}     &  @{verbatim "\\or"}             \\
wenzelm
parents: 53770
diff changeset
   587
    @{text "\<not>"}   & @{verbatim "~"}     &  @{verbatim "\\not"}            \\
wenzelm
parents: 53770
diff changeset
   588
    @{text "\<noteq>"}   & @{verbatim "~="}    &  @{verbatim "\\noteq"}          \\
wenzelm
parents: 53770
diff changeset
   589
    @{text "\<in>"}   & @{verbatim ":"}     &  @{verbatim "\\in"}             \\
wenzelm
parents: 53770
diff changeset
   590
    @{text "\<notin>"}   & @{verbatim "~:"}    &  @{verbatim "\\notin"}          \\
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   591
  \end{tabular}
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   592
  \medskip
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   593
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   594
  Note that the above abbreviations refer to the input method. The
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   595
  logical notation provides ASCII alternatives that often coincide,
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   596
  but deviate occasionally.  This occasionally causes user confusion
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   597
  with very old-fashioned Isabelle source that use ASCII replacement
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   598
  notation like @{verbatim "!"} or @{verbatim "ALL"} directly.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   599
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   600
  \end{enumerate}
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   601
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   602
  Raw Unicode characters within prover source files should be
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   603
  restricted to informal parts, e.g.\ to write text in non-latin
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   604
  alphabets.  Mathematical symbols should be defined via the official
54331
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   605
  rendering tables, to avoid problems with portability and long-term
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   606
  storage of formal text.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   607
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   608
  \paragraph{Control symbols.} There are some special control symbols
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   609
  to modify the style of a single symbol (without nesting). Control
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   610
  symbols may be applied to a region of selected text, either using
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   611
  the \emph{Symbols} panel or keyboard shortcuts or jEdit actions.
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   612
  These editor operations produce a separate control symbol for each
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   613
  symbol in the text, in order to make the whole text appear in a
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   614
  certain style.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   615
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   616
  \medskip
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   617
  \begin{tabular}{llll}
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   618
    style & \textbf{symbol} & shortcut & action \\\hline
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   619
    superscript & @{verbatim "\<^sup>"} & @{verbatim "C+e UP"} & @{verbatim "isabelle.control-sup"} \\
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   620
    subscript & @{verbatim "\<^sub>"} & @{verbatim "C+e DOWN"} & @{verbatim "isabelle.control-sub"} \\
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   621
    bold face & @{verbatim "\<^bold>"} & @{verbatim "C+e RIGHT"} & @{verbatim "isabelle.control-bold"} \\
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   622
    reset & & @{verbatim "C+e LEFT"} & @{verbatim "isabelle.control-reset"} \\
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   623
  \end{tabular}
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   624
  \medskip
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   625
54352
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   626
  To produce a single control symbol, it is also possible to complete
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   627
  on @{verbatim "\\"}@{verbatim sup}, @{verbatim "\\"}@{verbatim sub},
7557f9f1d4aa misc tuning and clarification;
wenzelm
parents: 54351
diff changeset
   628
  @{verbatim "\\"}@{verbatim bold} as for regular symbols.  *}
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   629
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   630
54323
d521407f8d0f more documentation;
wenzelm
parents: 54322
diff changeset
   631
section {* Text completion \label{sec:completion} *}
53769
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
   632
53981
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   633
text {*
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   634
  Text completion works via some light-weight GUI popup, which is triggered by
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   635
  keyboard events during the normal editing process in the main jEdit text
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   636
  area and a few additional text fields. The popup interprets special keys:
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   637
  @{verbatim TAB}, @{verbatim ESCAPE}, @{verbatim UP}, @{verbatim DOWN},
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   638
  @{verbatim PAGE_UP}, @{verbatim PAGE_DOWN}. All other key events are passed
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   639
  to the jEdit text area --- this allows to ignore unwanted completions most
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   640
  of the time and continue typing quickly.
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   641
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   642
  Various Isabelle plugin options control the popup behavior and immediate
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   643
  insertion into buffer.
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   644
54320
wenzelm
parents: 54037
diff changeset
   645
  Isabelle Symbols are completed in backslashed forms, e.g.\ @{verbatim
53981
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   646
  "\\"}@{verbatim "forall"} or @{verbatim "\<forall>"} that both produce the Isabelle
53982
wenzelm
parents: 53981
diff changeset
   647
  symbol @{text "\<forall>"} in its Unicode rendering. Alternatively, symbol
wenzelm
parents: 53981
diff changeset
   648
  abbreviations may be used as specified in @{file
wenzelm
parents: 53981
diff changeset
   649
  "$ISABELLE_HOME/etc/symbols"}.
53981
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   650
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   651
  \emph{Explicit completion} works via standard jEdit shortcut @{verbatim
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   652
  "C+b"}, which is remapped to action @{verbatim "isabelle.complete"}, with a
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   653
  fall-back on regular @{verbatim "complete-word"} for non-Isabelle buffers.
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   654
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   655
  \emph{Implicit completion} works via keyboard input on text area, with popup
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   656
  or immediate insertion into buffer. Plain words require at least 3
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   657
  characters to be completed.
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   658
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   659
  \emph{Immediate completion} means the (unique) replacement text is inserted
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   660
  into the buffer without popup. This mode ignores plain words and requires
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   661
  more than one characters for symbol abbreviations. Otherwise it falls back
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   662
  on completion popup.
1f4d6870b7b2 more on text completion;
wenzelm
parents: 53886
diff changeset
   663
*}
53769
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
   664
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   665
54353
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   666
section {* Automatically tried tools *}
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   667
54354
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   668
text {* Continuous document processing works asynchronously in the
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   669
  background.  Visible document source that has been evaluated already
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   670
  may get augmented by additional results of \emph{asynchronous print
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   671
  functions}.  The canonical example is proof state output, which is
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   672
  always enabled.  More heavy-weight print functions may be applied,
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   673
  in order to prove or disprove parts of the formal text by other
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   674
  means.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   675
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   676
  Isabelle/HOL provides various automatically tried tools that operate
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   677
  on outermost goal statements (e.g.\ @{command lemma}, @{command
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   678
  theorem}), independently of the state of the current proof attempt.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   679
  They work implicitly without any arguments.  Results are output as
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   680
  \emph{information messages}, which are indicated in the text area by
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   681
  blue squiggles and a blue information sign in the gutter.  The
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   682
  message content may be shown as for any other message, see also
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   683
  \secref{sec:prover-output}.  Some tools produce output with
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   684
  \emph{sendback} markup, which means that clicking on certain parts
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   685
  of the output inserts that text into the source in the proper place.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   686
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   687
  \medskip The following Isabelle system options control the behaviour
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   688
  of automatically tried tools (see also the jEdit dialog window
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   689
  \emph{Plugin Options / Isabelle / General / Automatically tried
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   690
  tools}):
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   691
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   692
  \begin{itemize}
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   693
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   694
  \item @{system_option auto_methods} controls automatic use of a
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   695
  combination of standard proof methods (@{method auto}, @{method
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   696
  simp}, @{method blast}, etc.).  This corresponds to the command
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   697
  @{command "try0"}.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   698
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   699
  The tool is disabled by default, since unparameterized invocation of
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   700
  standard proof methods often consumes substantial CPU resources
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   701
  without leading to success.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   702
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   703
  \item @{system_option auto_nitpick} controls a slightly reduced
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   704
  version of @{command nitpick}, which tests for counterexamples using
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   705
  first-order relational logic. See also the Nitpick manual
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   706
  \cite{isabelle-nitpick}.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   707
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   708
  This tool is disabled by default, due to the extra overhead of
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   709
  invoking an external Java process for each attempt to disprove a
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   710
  subgoal.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   711
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   712
  \item @{system_option auto_quickcheck} controls automatic use of
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   713
  @{command quickcheck}, which tests for counterexamples using a
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   714
  series of assignments for free variables of a subgoal.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   715
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   716
  This tool is \emph{enabled} by default.  It requires little
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   717
  overhead, but is a bit weaker than @{command nitpick}.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   718
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   719
  \item @{system_option auto_sledgehammer} controls a significantly
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   720
  reduced version of @{command sledgehammer}, which attempts to prove
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   721
  a subgoal using external automatic provers. See also the
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   722
  Sledgehammer manual \cite{isabelle-sledgehammer}.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   723
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   724
  This tool is disabled by default, due to the relatively heavy nature
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   725
  of Sledgehammer.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   726
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   727
  \item @{system_option auto_solve_direct} controls automatic use of
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   728
  @{command solve_direct}, which checks whether the current subgoals
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   729
  can be solved directly by an existing theorem.  This also helps to
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   730
  detect duplicate lemmas.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   731
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   732
  This tool is \emph{enabled} by default.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   733
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   734
  \end{itemize}
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   735
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   736
  Invocation of automatically tried tools is subject to some global
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   737
  policies of parallel execution, which may be configured as follows:
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   738
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   739
  \begin{itemize}
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   740
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   741
  \item @{system_option auto_time_limit} (default 2.0) determines the
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   742
  timeout (in seconds) for each tool execution individually.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   743
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   744
  \item @{system_option auto_time_start} (default 1.0) determines the
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   745
  start delay (in seconds) for automatically tried tools, after the
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   746
  main command evaluation is finished.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   747
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   748
  \end{itemize}
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   749
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   750
  Each tool is submitted independently to the pool of parallel
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   751
  execution tasks in Isabelle/ML, using hardwired priorities according
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   752
  to its relative ``heaviness''.  The main stages of evaluation and
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   753
  printing of proof states take precedence, but an already running
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   754
  tool is not canceled and may thus reduce reactivity of proof
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   755
  document processing.
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   756
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   757
  Users should experiment how the available CPU resources (number of
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   758
  cores) are best invested to get additional feedback from prover in
4e6defdc24ac more on automatically tried tools;
wenzelm
parents: 54353
diff changeset
   759
  the background, by using weaker or stronger tools.  *}
54353
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   760
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   761
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   762
section {* Sledgehammer \label{sec:sledgehammer} *}
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   763
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   764
text {*
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   765
  FIXME
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   766
*}
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   767
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   768
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   769
section {* Find theorems *}
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   770
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   771
text {*
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   772
  FIXME
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   773
*}
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   774
6692c355ebc1 more on prover output;
wenzelm
parents: 54352
diff changeset
   775
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   776
chapter {* Known problems and workarounds \label{sec:problems} *}
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   777
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   778
text {*
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   779
  \begin{itemize}
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   780
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   781
  \item \textbf{Problem:} Lack of dependency management for auxiliary files
53771
wenzelm
parents: 53770
diff changeset
   782
  that contribute to a theory (e.g.\ @{command ML_file}).
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   783
54330
wenzelm
parents: 54329
diff changeset
   784
  \textbf{Workaround:} Re-load files manually within the prover, by
wenzelm
parents: 54329
diff changeset
   785
  editing corresponding command in the text.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   786
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   787
  \item \textbf{Problem:} Odd behavior of some diagnostic commands with
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   788
  global side-effects, like writing a physical file.
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   789
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   790
  \textbf{Workaround:} Avoid such commands.
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   791
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   792
  \item \textbf{Problem:} No way to delete document nodes from the overall
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   793
  collection of theories.
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   794
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   795
  \textbf{Workaround:} Ignore unused files.  Restart whole
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   796
  Isabelle/jEdit session in worst-case situation.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   797
54330
wenzelm
parents: 54329
diff changeset
   798
  \item \textbf{Problem:} Keyboard shortcuts @{verbatim "C+PLUS"} and
wenzelm
parents: 54329
diff changeset
   799
  @{verbatim "C+MINUS"} for adjusting the editor font size depend on
wenzelm
parents: 54329
diff changeset
   800
  platform details and national keyboards.
wenzelm
parents: 54329
diff changeset
   801
wenzelm
parents: 54329
diff changeset
   802
  \textbf{Workaround:} Use numeric keypad or rebind keys in the
wenzelm
parents: 54329
diff changeset
   803
  jEdit Shortcuts options dialog.
wenzelm
parents: 54329
diff changeset
   804
53886
c83727c7a510 updated documentation concerning MacOSX plugin 1.3;
wenzelm
parents: 53773
diff changeset
   805
  \item \textbf{Problem:} The Mac OS X keyboard shortcut @{verbatim
54349
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   806
  "COMMAND+COMMA"} for application \emph{Preferences} is in conflict
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   807
  with the jEdit default shortcut for \emph{Incremental Search Bar}
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   808
  (action @{verbatim "quick-search"}).
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   809
53886
c83727c7a510 updated documentation concerning MacOSX plugin 1.3;
wenzelm
parents: 53773
diff changeset
   810
  \textbf{Workaround:} Remap in jEdit manually according to national
c83727c7a510 updated documentation concerning MacOSX plugin 1.3;
wenzelm
parents: 53773
diff changeset
   811
  keyboard, e.g.\ @{verbatim "COMMAND+SLASH"} on English ones.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   812
54349
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   813
  \item \textbf{Problem:} Mac OS X system fonts sometimes lead to
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   814
  character drop-outs in the main text area.
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   815
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   816
  \textbf{Workaround:} Use the default @{verbatim IsabelleText} font.
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   817
  (Do not install that font on the system.)
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   818
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   819
  \item \textbf{Problem:} Some Linux / X11 input methods such as IBus
54330
wenzelm
parents: 54329
diff changeset
   820
  tend to disrupt key event handling of Java/AWT/Swing.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   821
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   822
  \textbf{Workaround:} Do not use input methods, reset the environment
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   823
  variable @{verbatim XMODIFIERS} within Isabelle settings (default in
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   824
  Isabelle2013-1).
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   825
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   826
  \item \textbf{Problem:} Some Linux / X11 window managers that are
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   827
  not ``re-parenting'' cause problems with additional windows opened
54331
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   828
  by Java. This affects either historic or neo-minimalistic window
9e944630be0c more screenshots;
wenzelm
parents: 54330
diff changeset
   829
  managers like @{verbatim awesome} or @{verbatim xmonad}.
54329
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   830
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   831
  \textbf{Workaround:} Use regular re-parenting window manager.
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   832
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   833
  \item \textbf{Problem:} Recent forks of Linux / X11 window managers
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   834
  and desktop environments (variants of Gnome) disrupt the handling of
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   835
  menu popups and mouse positions of Java/AWT/Swing.
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   836
e0fa4bd16c80 misc tuning;
wenzelm
parents: 54323
diff changeset
   837
  \textbf{Workaround:} Use mainstream versions of Linux desktops.
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   838
54349
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   839
  \item \textbf{Problem:} Full-screen mode via jEdit action @{verbatim
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   840
  "toggle-full-screen"} (default shortcut @{verbatim F11}) works on
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   841
  Windows, but not on Mac OS X or various Linux / X11 window managers.
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   842
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   843
  \textbf{Workaround:} Use native full-screen control of the window
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   844
  manager, if available (notably on Mac OS X).
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   845
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   846
  \item \textbf{Problem:} Full-screen mode and dockable windows in
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   847
  \emph{floating} state may lead to confusion about window placement
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   848
  (depending on platform characteristics).
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   849
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   850
  \textbf{Workaround:} Avoid this combination.
fd5ddf2bce76 more on problems and workarounds;
wenzelm
parents: 54348
diff changeset
   851
53770
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   852
  \end{itemize}
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   853
*}
db362319d766 added/updated material from src/Tools/jEdit/README.html;
wenzelm
parents: 53769
diff changeset
   854
53769
036e80175bdd basic setup for Isabelle/jEdit documentation;
wenzelm
parents:
diff changeset
   855
end