doc-src/IsarRef/Thy/syntax.thy
author wenzelm
Fri, 02 May 2008 16:36:05 +0200
changeset 26767 cc127cc0951b
parent 26760 2de4ba348f06
child 26777 134529bc72db
permissions -rw-r--r--
converted pure.tex to Thy/pure.thy;
wenzelm@26767
     1
(* $Id$ *)
wenzelm@26754
     2
wenzelm@26754
     3
theory "syntax"
wenzelm@26754
     4
imports CPure
wenzelm@26754
     5
begin
wenzelm@26754
     6
wenzelm@26754
     7
chapter {* Syntax primitives *}
wenzelm@26754
     8
wenzelm@26754
     9
text {*
wenzelm@26754
    10
  The rather generic framework of Isabelle/Isar syntax emerges from
wenzelm@26754
    11
  three main syntactic categories: \emph{commands} of the top-level
wenzelm@26754
    12
  Isar engine (covering theory and proof elements), \emph{methods} for
wenzelm@26754
    13
  general goal refinements (analogous to traditional ``tactics''), and
wenzelm@26754
    14
  \emph{attributes} for operations on facts (within a certain
wenzelm@26754
    15
  context).  Subsequently we give a reference of basic syntactic
wenzelm@26754
    16
  entities underlying Isabelle/Isar syntax in a bottom-up manner.
wenzelm@26754
    17
  Concrete theory and proof language elements will be introduced later
wenzelm@26754
    18
  on.
wenzelm@26754
    19
wenzelm@26754
    20
  \medskip In order to get started with writing well-formed
wenzelm@26754
    21
  Isabelle/Isar documents, the most important aspect to be noted is
wenzelm@26754
    22
  the difference of \emph{inner} versus \emph{outer} syntax.  Inner
wenzelm@26754
    23
  syntax is that of Isabelle types and terms of the logic, while outer
wenzelm@26754
    24
  syntax is that of Isabelle/Isar theory sources (specifications and
wenzelm@26754
    25
  proofs).  As a general rule, inner syntax entities may occur only as
wenzelm@26754
    26
  \emph{atomic entities} within outer syntax.  For example, the string
wenzelm@26760
    27
  @{verbatim "\"x + y\""} and identifier @{verbatim z} are legal term
wenzelm@26760
    28
  specifications within a theory, while @{verbatim "x + y"} without
wenzelm@26760
    29
  quotes is not.
wenzelm@26754
    30
wenzelm@26754
    31
  Printed theory documents usually omit quotes to gain readability
wenzelm@26760
    32
  (this is a matter of {\LaTeX} macro setup, say via @{verbatim
wenzelm@26760
    33
  "\\isabellestyle"}, see also \cite{isabelle-sys}).  Experienced
wenzelm@26754
    34
  users of Isabelle/Isar may easily reconstruct the lost technical
wenzelm@26754
    35
  information, while mere readers need not care about quotes at all.
wenzelm@26754
    36
wenzelm@26754
    37
  \medskip Isabelle/Isar input may contain any number of input
wenzelm@26760
    38
  termination characters ``@{verbatim ";"}'' (semicolon) to separate
wenzelm@26754
    39
  commands explicitly.  This is particularly useful in interactive
wenzelm@26754
    40
  shell sessions to make clear where the current command is intended
wenzelm@26754
    41
  to end.  Otherwise, the interpreter loop will continue to issue a
wenzelm@26760
    42
  secondary prompt ``@{verbatim "#"}'' until an end-of-command is
wenzelm@26760
    43
  clearly recognized from the input syntax, e.g.\ encounter of the
wenzelm@26760
    44
  next command keyword.
wenzelm@26754
    45
wenzelm@26754
    46
  More advanced interfaces such as Proof~General \cite{proofgeneral}
wenzelm@26754
    47
  do not require explicit semicolons, the amount of input text is
wenzelm@26754
    48
  determined automatically by inspecting the present content of the
wenzelm@26754
    49
  Emacs text buffer.  In the printed presentation of Isabelle/Isar
wenzelm@26754
    50
  documents semicolons are omitted altogether for readability.
wenzelm@26754
    51
wenzelm@26754
    52
  \begin{warn}
wenzelm@26754
    53
    Proof~General requires certain syntax classification tables in
wenzelm@26754
    54
    order to achieve properly synchronized interaction with the
wenzelm@26754
    55
    Isabelle/Isar process.  These tables need to be consistent with
wenzelm@26754
    56
    the Isabelle version and particular logic image to be used in a
wenzelm@26754
    57
    running session (common object-logics may well change the outer
wenzelm@26754
    58
    syntax).  The standard setup should work correctly with any of the
wenzelm@26754
    59
    ``official'' logic images derived from Isabelle/HOL (including
wenzelm@26754
    60
    HOLCF etc.).  Users of alternative logics may need to tell
wenzelm@26760
    61
    Proof~General explicitly, e.g.\ by giving an option @{verbatim "-k ZF"}
wenzelm@26760
    62
    (in conjunction with @{verbatim "-l ZF"}, to specify the default
wenzelm@26760
    63
    logic image).  Note that option @{verbatim "-L"} does both
wenzelm@26760
    64
    of this at the same time.
wenzelm@26754
    65
  \end{warn}
wenzelm@26754
    66
*}
wenzelm@26754
    67
wenzelm@26754
    68
wenzelm@26754
    69
section {* Lexical matters \label{sec:lex-syntax} *}
wenzelm@26754
    70
wenzelm@26754
    71
text {*
wenzelm@26754
    72
  The Isabelle/Isar outer syntax provides token classes as presented
wenzelm@26754
    73
  below; most of these coincide with the inner lexical syntax as
wenzelm@26754
    74
  presented in \cite{isabelle-ref}.
wenzelm@26754
    75
wenzelm@26754
    76
  \begin{matharray}{rcl}
wenzelm@26754
    77
    @{syntax_def ident} & = & letter\,quasiletter^* \\
wenzelm@26754
    78
    @{syntax_def longident} & = & ident (\verb,.,ident)^+ \\
wenzelm@26754
    79
    @{syntax_def symident} & = & sym^+ ~|~ \verb,\,\verb,<,ident\verb,>, \\
wenzelm@26754
    80
    @{syntax_def nat} & = & digit^+ \\
wenzelm@26754
    81
    @{syntax_def var} & = & ident ~|~ \verb,?,ident ~|~ \verb,?,ident\verb,.,nat \\
wenzelm@26754
    82
    @{syntax_def typefree} & = & \verb,',ident \\
wenzelm@26754
    83
    @{syntax_def typevar} & = & typefree ~|~ \verb,?,typefree ~|~ \verb,?,typefree\verb,.,nat \\
wenzelm@26754
    84
    @{syntax_def string} & = & \verb,", ~\dots~ \verb,", \\
wenzelm@26754
    85
    @{syntax_def altstring} & = & \backquote ~\dots~ \backquote \\
wenzelm@26754
    86
    @{syntax_def verbatim} & = & \verb,{*, ~\dots~ \verb,*,\verb,}, \\[1ex]
wenzelm@26754
    87
wenzelm@26754
    88
    letter & = & latin ~|~ \verb,\,\verb,<,latin\verb,>, ~|~ \verb,\,\verb,<,latin\,latin\verb,>, ~|~ greek ~|~ \\
wenzelm@26754
    89
           &   & \verb,\<^isub>, ~|~ \verb,\<^isup>, \\
wenzelm@26754
    90
    quasiletter & = & letter ~|~ digit ~|~ \verb,_, ~|~ \verb,', \\
wenzelm@26754
    91
    latin & = & \verb,a, ~|~ \dots ~|~ \verb,z, ~|~ \verb,A, ~|~ \dots ~|~ \verb,Z, \\
wenzelm@26754
    92
    digit & = & \verb,0, ~|~ \dots ~|~ \verb,9, \\
wenzelm@26754
    93
    sym & = & \verb,!, ~|~ \verb,#, ~|~ \verb,$, ~|~ \verb,%, ~|~ \verb,&, ~|~
wenzelm@26754
    94
     \verb,*, ~|~ \verb,+, ~|~ \verb,-, ~|~ \verb,/, ~|~ \\
wenzelm@26754
    95
    & & \verb,<, ~|~ \verb,=, ~|~ \verb,>, ~|~ \verb,?, ~|~ \texttt{\at} ~|~
wenzelm@26754
    96
    \verb,^, ~|~ \verb,_, ~|~ \verb,|, ~|~ \verb,~, \\
wenzelm@26754
    97
    greek & = & \verb,\<alpha>, ~|~ \verb,\<beta>, ~|~ \verb,\<gamma>, ~|~ \verb,\<delta>, ~| \\
wenzelm@26754
    98
          &   & \verb,\<epsilon>, ~|~ \verb,\<zeta>, ~|~ \verb,\<eta>, ~|~ \verb,\<theta>, ~| \\
wenzelm@26754
    99
          &   & \verb,\<iota>, ~|~ \verb,\<kappa>, ~|~ \verb,\<mu>, ~|~ \verb,\<nu>, ~| \\
wenzelm@26754
   100
          &   & \verb,\<xi>, ~|~ \verb,\<pi>, ~|~ \verb,\<rho>, ~|~ \verb,\<sigma>, ~|~ \verb,\<tau>, ~| \\
wenzelm@26754
   101
          &   & \verb,\<upsilon>, ~|~ \verb,\<phi>, ~|~ \verb,\<chi>, ~|~ \verb,\<psi>, ~| \\
wenzelm@26754
   102
          &   & \verb,\<omega>, ~|~ \verb,\<Gamma>, ~|~ \verb,\<Delta>, ~|~ \verb,\<Theta>, ~| \\
wenzelm@26754
   103
          &   & \verb,\<Lambda>, ~|~ \verb,\<Xi>, ~|~ \verb,\<Pi>, ~|~ \verb,\<Sigma>, ~| \\
wenzelm@26754
   104
          &   & \verb,\<Upsilon>, ~|~ \verb,\<Phi>, ~|~ \verb,\<Psi>, ~|~ \verb,\<Omega>, \\
wenzelm@26754
   105
  \end{matharray}
wenzelm@26754
   106
wenzelm@26754
   107
  The syntax of @{syntax string} admits any characters, including
wenzelm@26760
   108
  newlines; ``@{verbatim "\""}'' (double-quote) and ``@{verbatim
wenzelm@26760
   109
  "\\"}'' (backslash) need to be escaped by a backslash; arbitrary
wenzelm@26760
   110
  character codes may be specified as ``@{verbatim "\\"}@{text ddd}'',
wenzelm@26760
   111
  with three decimal digits.  Alternative strings according to
wenzelm@26760
   112
  @{syntax altstring} are analogous, using single back-quotes instead.
wenzelm@26760
   113
  The body of @{syntax verbatim} may consist of any text not
wenzelm@26760
   114
  containing ``@{verbatim "*"}@{verbatim "}"}''; this allows
wenzelm@26760
   115
  convenient inclusion of quotes without further escapes.  The greek
wenzelm@26760
   116
  letters do \emph{not} include @{verbatim "\<lambda>"}, which is already used
wenzelm@26760
   117
  differently in the meta-logic.
wenzelm@26754
   118
wenzelm@26754
   119
  Common mathematical symbols such as @{text \<forall>} are represented in
wenzelm@26760
   120
  Isabelle as @{verbatim \<forall>}.  There are infinitely many Isabelle
wenzelm@26760
   121
  symbols like this, although proper presentation is left to front-end
wenzelm@26760
   122
  tools such as {\LaTeX} or Proof~General with the X-Symbol package.
wenzelm@26760
   123
  A list of standard Isabelle symbols that work well with these tools
wenzelm@26760
   124
  is given in \cite[appendix~A]{isabelle-sys}.
wenzelm@26754
   125
  
wenzelm@26760
   126
  Source comments take the form @{verbatim "(*"}~@{text
wenzelm@26760
   127
  "\<dots>"}~@{verbatim "*)"} and may be nested, although user-interface
wenzelm@26760
   128
  tools might prevent this.  Note that this form indicates source
wenzelm@26760
   129
  comments only, which are stripped after lexical analysis of the
wenzelm@26760
   130
  input.  The Isar document syntax also provides formal comments that
wenzelm@26760
   131
  are considered as part of the text (see \secref{sec:comments}).
wenzelm@26754
   132
*}
wenzelm@26754
   133
wenzelm@26754
   134
wenzelm@26754
   135
section {* Common syntax entities *}
wenzelm@26754
   136
wenzelm@26754
   137
text {*
wenzelm@26754
   138
  We now introduce several basic syntactic entities, such as names,
wenzelm@26754
   139
  terms, and theorem specifications, which are factored out of the
wenzelm@26754
   140
  actual Isar language elements to be described later.
wenzelm@26754
   141
*}
wenzelm@26754
   142
wenzelm@26754
   143
wenzelm@26754
   144
subsection {* Names *}
wenzelm@26754
   145
wenzelm@26754
   146
text {*
wenzelm@26754
   147
  Entity \railqtok{name} usually refers to any name of types,
wenzelm@26754
   148
  constants, theorems etc.\ that are to be \emph{declared} or
wenzelm@26754
   149
  \emph{defined} (so qualified identifiers are excluded here).  Quoted
wenzelm@26754
   150
  strings provide an escape for non-identifier names or those ruled
wenzelm@26760
   151
  out by outer syntax keywords (e.g.\ quoted @{verbatim "\"let\""}).
wenzelm@26760
   152
  Already existing objects are usually referenced by
wenzelm@26760
   153
  \railqtok{nameref}.
wenzelm@26754
   154
wenzelm@26754
   155
  \indexoutertoken{name}\indexoutertoken{parname}\indexoutertoken{nameref}
wenzelm@26754
   156
  \indexoutertoken{int}
wenzelm@26754
   157
  \begin{rail}
wenzelm@26754
   158
    name: ident | symident | string | nat
wenzelm@26754
   159
    ;
wenzelm@26754
   160
    parname: '(' name ')'
wenzelm@26754
   161
    ;
wenzelm@26754
   162
    nameref: name | longident
wenzelm@26754
   163
    ;
wenzelm@26754
   164
    int: nat | '-' nat
wenzelm@26754
   165
    ;
wenzelm@26754
   166
  \end{rail}
wenzelm@26754
   167
*}
wenzelm@26754
   168
wenzelm@26754
   169
wenzelm@26754
   170
subsection {* Comments \label{sec:comments} *}
wenzelm@26754
   171
wenzelm@26754
   172
text {*
wenzelm@26754
   173
  Large chunks of plain \railqtok{text} are usually given
wenzelm@26760
   174
  \railtok{verbatim}, i.e.\ enclosed in @{verbatim "{"}@{verbatim
wenzelm@26760
   175
  "*"}~@{text "\<dots>"}~@{verbatim "*"}@{verbatim "}"}.  For convenience,
wenzelm@26760
   176
  any of the smaller text units conforming to \railqtok{nameref} are
wenzelm@26760
   177
  admitted as well.  A marginal \railnonterm{comment} is of the form
wenzelm@26760
   178
  @{verbatim "--"} \railqtok{text}.  Any number of these may occur
wenzelm@26760
   179
  within Isabelle/Isar commands.
wenzelm@26754
   180
wenzelm@26754
   181
  \indexoutertoken{text}\indexouternonterm{comment}
wenzelm@26754
   182
  \begin{rail}
wenzelm@26754
   183
    text: verbatim | nameref
wenzelm@26754
   184
    ;
wenzelm@26754
   185
    comment: '--' text
wenzelm@26754
   186
    ;
wenzelm@26754
   187
  \end{rail}
wenzelm@26754
   188
*}
wenzelm@26754
   189
wenzelm@26754
   190
wenzelm@26754
   191
subsection {* Type classes, sorts and arities *}
wenzelm@26754
   192
wenzelm@26754
   193
text {*
wenzelm@26754
   194
  Classes are specified by plain names.  Sorts have a very simple
wenzelm@26754
   195
  inner syntax, which is either a single class name @{text c} or a
wenzelm@26754
   196
  list @{text "{c\<^sub>1, \<dots>, c\<^sub>n}"} referring to the
wenzelm@26754
   197
  intersection of these classes.  The syntax of type arities is given
wenzelm@26754
   198
  directly at the outer level.
wenzelm@26754
   199
wenzelm@26754
   200
  \railalias{subseteq}{\isasymsubseteq}
wenzelm@26754
   201
  \railterm{subseteq}
wenzelm@26754
   202
wenzelm@26754
   203
  \indexouternonterm{sort}\indexouternonterm{arity}
wenzelm@26754
   204
  \indexouternonterm{classdecl}
wenzelm@26754
   205
  \begin{rail}
wenzelm@26754
   206
    classdecl: name (('<' | subseteq) (nameref + ','))?
wenzelm@26754
   207
    ;
wenzelm@26754
   208
    sort: nameref
wenzelm@26754
   209
    ;
wenzelm@26754
   210
    arity: ('(' (sort + ',') ')')? sort
wenzelm@26754
   211
    ;
wenzelm@26754
   212
  \end{rail}
wenzelm@26754
   213
*}
wenzelm@26754
   214
wenzelm@26754
   215
wenzelm@26754
   216
subsection {* Types and terms \label{sec:types-terms} *}
wenzelm@26754
   217
wenzelm@26754
   218
text {*
wenzelm@26754
   219
  The actual inner Isabelle syntax, that of types and terms of the
wenzelm@26754
   220
  logic, is far too sophisticated in order to be modelled explicitly
wenzelm@26754
   221
  at the outer theory level.  Basically, any such entity has to be
wenzelm@26754
   222
  quoted to turn it into a single token (the parsing and type-checking
wenzelm@26754
   223
  is performed internally later).  For convenience, a slightly more
wenzelm@26754
   224
  liberal convention is adopted: quotes may be omitted for any type or
wenzelm@26754
   225
  term that is already atomic at the outer level.  For example, one
wenzelm@26760
   226
  may just write @{verbatim x} instead of quoted @{verbatim "\"x\""}.
wenzelm@26760
   227
  Note that symbolic identifiers (e.g.\ @{verbatim "++"} or @{text
wenzelm@26760
   228
  "\<forall>"} are available as well, provided these have not been superseded
wenzelm@26760
   229
  by commands or other keywords already (such as @{verbatim "="} or
wenzelm@26760
   230
  @{verbatim "+"}).
wenzelm@26754
   231
wenzelm@26754
   232
  \indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}
wenzelm@26754
   233
  \begin{rail}
wenzelm@26754
   234
    type: nameref | typefree | typevar
wenzelm@26754
   235
    ;
wenzelm@26754
   236
    term: nameref | var
wenzelm@26754
   237
    ;
wenzelm@26754
   238
    prop: term
wenzelm@26754
   239
    ;
wenzelm@26754
   240
  \end{rail}
wenzelm@26754
   241
wenzelm@26754
   242
  Positional instantiations are indicated by giving a sequence of
wenzelm@26760
   243
  terms, or the placeholder ``@{verbatim _}'' (underscore), which
wenzelm@26760
   244
  means to skip a position.
wenzelm@26754
   245
wenzelm@26754
   246
  \indexoutertoken{inst}\indexoutertoken{insts}
wenzelm@26754
   247
  \begin{rail}
wenzelm@26754
   248
    inst: underscore | term
wenzelm@26754
   249
    ;
wenzelm@26754
   250
    insts: (inst *)
wenzelm@26754
   251
    ;
wenzelm@26754
   252
  \end{rail}
wenzelm@26754
   253
wenzelm@26754
   254
  Type declarations and definitions usually refer to
wenzelm@26754
   255
  \railnonterm{typespec} on the left-hand side.  This models basic
wenzelm@26754
   256
  type constructor application at the outer syntax level.  Note that
wenzelm@26754
   257
  only plain postfix notation is available here, but no infixes.
wenzelm@26754
   258
wenzelm@26754
   259
  \indexouternonterm{typespec}
wenzelm@26754
   260
  \begin{rail}
wenzelm@26754
   261
    typespec: (() | typefree | '(' ( typefree + ',' ) ')') name
wenzelm@26754
   262
    ;
wenzelm@26754
   263
  \end{rail}
wenzelm@26754
   264
*}
wenzelm@26754
   265
wenzelm@26754
   266
wenzelm@26754
   267
subsection {* Mixfix annotations *}
wenzelm@26754
   268
wenzelm@26754
   269
text {*
wenzelm@26754
   270
  Mixfix annotations specify concrete \emph{inner} syntax of Isabelle
wenzelm@26754
   271
  types and terms.  Some commands such as @{command "types"} (see
wenzelm@26760
   272
  \secref{sec:types-pure}) admit infixes only, while @{command
wenzelm@26760
   273
  "consts"} (see \secref{sec:consts}) and @{command "syntax"} (see
wenzelm@26760
   274
  \secref{sec:syn-trans}) support the full range of general mixfixes
wenzelm@26754
   275
  and binders.
wenzelm@26754
   276
wenzelm@26754
   277
  \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
wenzelm@26754
   278
  \begin{rail}
wenzelm@26754
   279
    infix: '(' ('infix' | 'infixl' | 'infixr') string? nat ')'
wenzelm@26754
   280
    ;
wenzelm@26754
   281
    mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')'
wenzelm@26754
   282
    ;
wenzelm@26754
   283
    structmixfix: mixfix | '(' 'structure' ')'
wenzelm@26754
   284
    ;
wenzelm@26754
   285
wenzelm@26754
   286
    prios: '[' (nat + ',') ']'
wenzelm@26754
   287
    ;
wenzelm@26754
   288
  \end{rail}
wenzelm@26754
   289
wenzelm@26754
   290
  Here the \railtok{string} specifications refer to the actual mixfix
wenzelm@26754
   291
  template (see also \cite{isabelle-ref}), which may include literal
wenzelm@26760
   292
  text, spacing, blocks, and arguments (denoted by ``@{verbatim _}'');
wenzelm@26760
   293
  the special symbol ``@{verbatim "\<index>"}'' (printed as ``@{text "\<index>"}'')
wenzelm@26760
   294
  represents an index argument that specifies an implicit structure
wenzelm@26760
   295
  reference (see also \secref{sec:locale}).  Infix and binder
wenzelm@26760
   296
  declarations provide common abbreviations for particular mixfix
wenzelm@26760
   297
  declarations.  So in practice, mixfix templates mostly degenerate to
wenzelm@26760
   298
  literal text for concrete syntax, such as ``@{verbatim "++"}'' for
wenzelm@26760
   299
  an infix symbol, or ``@{verbatim "++"}@{text "\<index>"}'' for an infix of
wenzelm@26760
   300
  an implicit structure.
wenzelm@26754
   301
*}
wenzelm@26754
   302
wenzelm@26754
   303
wenzelm@26754
   304
subsection {* Proof methods \label{sec:syn-meth} *}
wenzelm@26754
   305
wenzelm@26754
   306
text {*
wenzelm@26754
   307
  Proof methods are either basic ones, or expressions composed of
wenzelm@26760
   308
  methods via ``@{verbatim ","}'' (sequential composition),
wenzelm@26760
   309
  ``@{verbatim "|"}'' (alternative choices), ``@{verbatim "?"}'' 
wenzelm@26760
   310
  (try), ``@{verbatim "+"}'' (repeat at least once), ``@{verbatim
wenzelm@26760
   311
  "["}@{text n}@{verbatim "]"}'' (restriction to first @{text n}
wenzelm@26760
   312
  sub-goals, with default @{text "n = 1"}).  In practice, proof
wenzelm@26760
   313
  methods are usually just a comma separated list of
wenzelm@26760
   314
  \railqtok{nameref}~\railnonterm{args} specifications.  Note that
wenzelm@26760
   315
  parentheses may be dropped for single method specifications (with no
wenzelm@26760
   316
  arguments).
wenzelm@26754
   317
wenzelm@26754
   318
  \indexouternonterm{method}
wenzelm@26754
   319
  \begin{rail}
wenzelm@26754
   320
    method: (nameref | '(' methods ')') (() | '?' | '+' | '[' nat? ']')
wenzelm@26754
   321
    ;
wenzelm@26754
   322
    methods: (nameref args | method) + (',' | '|')
wenzelm@26754
   323
    ;
wenzelm@26754
   324
  \end{rail}
wenzelm@26754
   325
wenzelm@26754
   326
  Proper Isar proof methods do \emph{not} admit arbitrary goal
wenzelm@26754
   327
  addressing, but refer either to the first sub-goal or all sub-goals
wenzelm@26760
   328
  uniformly.  The goal restriction operator ``@{text "[n]"}''
wenzelm@26754
   329
  evaluates a method expression within a sandbox consisting of the
wenzelm@26760
   330
  first @{text n} sub-goals (which need to exist).  For example, the
wenzelm@26760
   331
  method ``@{text "simp_all[3]"}'' simplifies the first three
wenzelm@26760
   332
  sub-goals, while ``@{text "(rule foo, simp_all)[]"}'' simplifies all
wenzelm@26760
   333
  new goals that emerge from applying rule @{text "foo"} to the
wenzelm@26760
   334
  originally first one.
wenzelm@26754
   335
wenzelm@26754
   336
  Improper methods, notably tactic emulations, offer a separate
wenzelm@26754
   337
  low-level goal addressing scheme as explicit argument to the
wenzelm@26760
   338
  individual tactic being involved.  Here ``@{text "[!]"}'' refers to
wenzelm@26760
   339
  all goals, and ``@{text "[n-]"}'' to all goals starting from @{text
wenzelm@26760
   340
  "n"}.
wenzelm@26754
   341
wenzelm@26754
   342
  \indexouternonterm{goalspec}
wenzelm@26754
   343
  \begin{rail}
wenzelm@26754
   344
    goalspec: '[' (nat '-' nat | nat '-' | nat | '!' ) ']'
wenzelm@26754
   345
    ;
wenzelm@26754
   346
  \end{rail}
wenzelm@26754
   347
*}
wenzelm@26754
   348
wenzelm@26754
   349
wenzelm@26754
   350
subsection {* Attributes and theorems \label{sec:syn-att} *}
wenzelm@26754
   351
wenzelm@26754
   352
text {*
wenzelm@26760
   353
  Attributes (and proof methods, see \secref{sec:syn-meth}) have their
wenzelm@26754
   354
  own ``semi-inner'' syntax, in the sense that input conforming to
wenzelm@26754
   355
  \railnonterm{args} below is parsed by the attribute a second time.
wenzelm@26754
   356
  The attribute argument specifications may be any sequence of atomic
wenzelm@26754
   357
  entities (identifiers, strings etc.), or properly bracketed argument
wenzelm@26754
   358
  lists.  Below \railqtok{atom} refers to any atomic entity, including
wenzelm@26754
   359
  any \railtok{keyword} conforming to \railtok{symident}.
wenzelm@26754
   360
wenzelm@26754
   361
  \indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}
wenzelm@26754
   362
  \begin{rail}
wenzelm@26754
   363
    atom: nameref | typefree | typevar | var | nat | keyword
wenzelm@26754
   364
    ;
wenzelm@26754
   365
    arg: atom | '(' args ')' | '[' args ']'
wenzelm@26754
   366
    ;
wenzelm@26754
   367
    args: arg *
wenzelm@26754
   368
    ;
wenzelm@26754
   369
    attributes: '[' (nameref args * ',') ']'
wenzelm@26754
   370
    ;
wenzelm@26754
   371
  \end{rail}
wenzelm@26754
   372
wenzelm@26754
   373
  Theorem specifications come in several flavors:
wenzelm@26754
   374
  \railnonterm{axmdecl} and \railnonterm{thmdecl} usually refer to
wenzelm@26754
   375
  axioms, assumptions or results of goal statements, while
wenzelm@26754
   376
  \railnonterm{thmdef} collects lists of existing theorems.  Existing
wenzelm@26754
   377
  theorems are given by \railnonterm{thmref} and
wenzelm@26754
   378
  \railnonterm{thmrefs}, the former requires an actual singleton
wenzelm@26754
   379
  result.
wenzelm@26754
   380
wenzelm@26754
   381
  There are three forms of theorem references:
wenzelm@26754
   382
  \begin{enumerate}
wenzelm@26754
   383
  
wenzelm@26760
   384
  \item named facts @{text "a"},
wenzelm@26754
   385
wenzelm@26760
   386
  \item selections from named facts @{text "a(i)"} or @{text "a(j - k)"},
wenzelm@26754
   387
wenzelm@26754
   388
  \item literal fact propositions using @{syntax_ref altstring} syntax
wenzelm@26760
   389
  @{verbatim "`"}@{text "\<phi>"}@{verbatim "`"} (see also method
wenzelm@26760
   390
  @{method_ref fact} in \secref{sec:pure-meth-att}).
wenzelm@26754
   391
wenzelm@26754
   392
  \end{enumerate}
wenzelm@26754
   393
wenzelm@26754
   394
  Any kind of theorem specification may include lists of attributes
wenzelm@26754
   395
  both on the left and right hand sides; attributes are applied to any
wenzelm@26754
   396
  immediately preceding fact.  If names are omitted, the theorems are
wenzelm@26754
   397
  not stored within the theorem database of the theory or proof
wenzelm@26754
   398
  context, but any given attributes are applied nonetheless.
wenzelm@26754
   399
wenzelm@26760
   400
  An extra pair of brackets around attributes (like ``@{text
wenzelm@26760
   401
  "[[simproc a]]"}'') abbreviates a theorem reference involving an
wenzelm@26760
   402
  internal dummy fact, which will be ignored later on.  So only the
wenzelm@26760
   403
  effect of the attribute on the background context will persist.
wenzelm@26760
   404
  This form of in-place declarations is particularly useful with
wenzelm@26760
   405
  commands like @{command "declare"} and @{command "using"}.
wenzelm@26754
   406
wenzelm@26754
   407
  \indexouternonterm{axmdecl}\indexouternonterm{thmdecl}
wenzelm@26754
   408
  \indexouternonterm{thmdef}\indexouternonterm{thmref}
wenzelm@26754
   409
  \indexouternonterm{thmrefs}\indexouternonterm{selection}
wenzelm@26754
   410
  \begin{rail}
wenzelm@26754
   411
    axmdecl: name attributes? ':'
wenzelm@26754
   412
    ;
wenzelm@26754
   413
    thmdecl: thmbind ':'
wenzelm@26754
   414
    ;
wenzelm@26754
   415
    thmdef: thmbind '='
wenzelm@26754
   416
    ;
wenzelm@26754
   417
    thmref: (nameref selection? | altstring) attributes? | '[' attributes ']'
wenzelm@26754
   418
    ;
wenzelm@26754
   419
    thmrefs: thmref +
wenzelm@26754
   420
    ;
wenzelm@26754
   421
wenzelm@26754
   422
    thmbind: name attributes | name | attributes
wenzelm@26754
   423
    ;
wenzelm@26754
   424
    selection: '(' ((nat | nat '-' nat?) + ',') ')'
wenzelm@26754
   425
    ;
wenzelm@26754
   426
  \end{rail}
wenzelm@26754
   427
*}
wenzelm@26754
   428
wenzelm@26754
   429
wenzelm@26754
   430
subsection {* Term patterns and declarations \label{sec:term-decls} *}
wenzelm@26754
   431
wenzelm@26754
   432
text {*
wenzelm@26754
   433
  Wherever explicit propositions (or term fragments) occur in a proof
wenzelm@26754
   434
  text, casual binding of schematic term variables may be given
wenzelm@26754
   435
  specified via patterns of the form ``@{text "(\<IS> p\<^sub>1 \<dots>
wenzelm@26754
   436
  p\<^sub>n)"}''.  This works both for \railqtok{term} and \railqtok{prop}.
wenzelm@26754
   437
wenzelm@26754
   438
  \indexouternonterm{termpat}\indexouternonterm{proppat}
wenzelm@26754
   439
  \begin{rail}
wenzelm@26754
   440
    termpat: '(' ('is' term +) ')'
wenzelm@26754
   441
    ;
wenzelm@26754
   442
    proppat: '(' ('is' prop +) ')'
wenzelm@26754
   443
    ;
wenzelm@26754
   444
  \end{rail}
wenzelm@26754
   445
wenzelm@26754
   446
  \medskip Declarations of local variables @{text "x :: \<tau>"} and
wenzelm@26754
   447
  logical propositions @{text "a : \<phi>"} represent different views on
wenzelm@26754
   448
  the same principle of introducing a local scope.  In practice, one
wenzelm@26754
   449
  may usually omit the typing of \railnonterm{vars} (due to
wenzelm@26754
   450
  type-inference), and the naming of propositions (due to implicit
wenzelm@26754
   451
  references of current facts).  In any case, Isar proof elements
wenzelm@26754
   452
  usually admit to introduce multiple such items simultaneously.
wenzelm@26754
   453
wenzelm@26754
   454
  \indexouternonterm{vars}\indexouternonterm{props}
wenzelm@26754
   455
  \begin{rail}
wenzelm@26754
   456
    vars: (name+) ('::' type)?
wenzelm@26754
   457
    ;
wenzelm@26754
   458
    props: thmdecl? (prop proppat? +)
wenzelm@26754
   459
    ;
wenzelm@26754
   460
  \end{rail}
wenzelm@26754
   461
wenzelm@26754
   462
  The treatment of multiple declarations corresponds to the
wenzelm@26754
   463
  complementary focus of \railnonterm{vars} versus
wenzelm@26754
   464
  \railnonterm{props}.  In ``@{text "x\<^sub>1 \<dots> x\<^sub>n :: \<tau>"}''
wenzelm@26754
   465
  the typing refers to all variables, while in @{text "a: \<phi>\<^sub>1 \<dots>
wenzelm@26754
   466
  \<phi>\<^sub>n"} the naming refers to all propositions collectively.
wenzelm@26754
   467
  Isar language elements that refer to \railnonterm{vars} or
wenzelm@26754
   468
  \railnonterm{props} typically admit separate typings or namings via
wenzelm@26754
   469
  another level of iteration, with explicit @{keyword_ref "and"}
wenzelm@26754
   470
  separators; e.g.\ see @{command "fix"} and @{command "assume"} in
wenzelm@26760
   471
  \secref{sec:proof-context}.
wenzelm@26754
   472
*}
wenzelm@26754
   473
wenzelm@26754
   474
wenzelm@26754
   475
subsection {* Antiquotations \label{sec:antiq} *}
wenzelm@26754
   476
wenzelm@26754
   477
text {*
wenzelm@26754
   478
  \begin{matharray}{rcl}
wenzelm@26754
   479
    @{antiquotation_def "theory"} & : & \isarantiq \\
wenzelm@26754
   480
    @{antiquotation_def "thm"} & : & \isarantiq \\
wenzelm@26754
   481
    @{antiquotation_def "prop"} & : & \isarantiq \\
wenzelm@26754
   482
    @{antiquotation_def "term"} & : & \isarantiq \\
wenzelm@26754
   483
    @{antiquotation_def const} & : & \isarantiq \\
wenzelm@26754
   484
    @{antiquotation_def abbrev} & : & \isarantiq \\
wenzelm@26754
   485
    @{antiquotation_def typeof} & : & \isarantiq \\
wenzelm@26754
   486
    @{antiquotation_def typ} & : & \isarantiq \\
wenzelm@26754
   487
    @{antiquotation_def thm_style} & : & \isarantiq \\
wenzelm@26754
   488
    @{antiquotation_def term_style} & : & \isarantiq \\
wenzelm@26754
   489
    @{antiquotation_def "text"} & : & \isarantiq \\
wenzelm@26754
   490
    @{antiquotation_def goals} & : & \isarantiq \\
wenzelm@26754
   491
    @{antiquotation_def subgoals} & : & \isarantiq \\
wenzelm@26754
   492
    @{antiquotation_def prf} & : & \isarantiq \\
wenzelm@26754
   493
    @{antiquotation_def full_prf} & : & \isarantiq \\
wenzelm@26754
   494
    @{antiquotation_def ML} & : & \isarantiq \\
wenzelm@26754
   495
    @{antiquotation_def ML_type} & : & \isarantiq \\
wenzelm@26754
   496
    @{antiquotation_def ML_struct} & : & \isarantiq \\
wenzelm@26754
   497
  \end{matharray}
wenzelm@26754
   498
wenzelm@26760
   499
  The text body of formal comments (see also \secref{sec:comments})
wenzelm@26760
   500
  may contain antiquotations of logical entities, such as theorems,
wenzelm@26760
   501
  terms and types, which are to be presented in the final output
wenzelm@26760
   502
  produced by the Isabelle document preparation system (see also
wenzelm@26760
   503
  \secref{sec:document-prep}).
wenzelm@26754
   504
wenzelm@26754
   505
  Thus embedding of ``@{text "@{term [show_types] \"f x = a + x\"}"}''
wenzelm@26754
   506
  within a text block would cause
wenzelm@26754
   507
  \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} to appear in the final {\LaTeX} document.  Also note that theorem
wenzelm@26754
   508
  antiquotations may involve attributes as well.  For example,
wenzelm@26760
   509
  @{text "@{thm sym [no_vars]}"} would print the theorem's
wenzelm@26760
   510
  statement where all schematic variables have been replaced by fixed
wenzelm@26760
   511
  ones, which are easier to read.
wenzelm@26754
   512
wenzelm@26754
   513
  \begin{rail}
wenzelm@26754
   514
    atsign lbrace antiquotation rbrace
wenzelm@26754
   515
    ;
wenzelm@26754
   516
wenzelm@26754
   517
    antiquotation:
wenzelm@26754
   518
      'theory' options name |
wenzelm@26754
   519
      'thm' options thmrefs |
wenzelm@26754
   520
      'prop' options prop |
wenzelm@26754
   521
      'term' options term |
wenzelm@26754
   522
      'const' options term |
wenzelm@26754
   523
      'abbrev' options term |
wenzelm@26754
   524
      'typeof' options term |
wenzelm@26754
   525
      'typ' options type |
wenzelm@26754
   526
      'thm\_style' options name thmref |
wenzelm@26754
   527
      'term\_style' options name term |
wenzelm@26754
   528
      'text' options name |
wenzelm@26754
   529
      'goals' options |
wenzelm@26754
   530
      'subgoals' options |
wenzelm@26754
   531
      'prf' options thmrefs |
wenzelm@26754
   532
      'full\_prf' options thmrefs |
wenzelm@26754
   533
      'ML' options name |
wenzelm@26754
   534
      'ML\_type' options name |
wenzelm@26754
   535
      'ML\_struct' options name
wenzelm@26754
   536
    ;
wenzelm@26754
   537
    options: '[' (option * ',') ']'
wenzelm@26754
   538
    ;
wenzelm@26754
   539
    option: name | name '=' name
wenzelm@26754
   540
    ;
wenzelm@26754
   541
  \end{rail}
wenzelm@26754
   542
wenzelm@26754
   543
  Note that the syntax of antiquotations may \emph{not} include source
wenzelm@26760
   544
  comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} or verbatim
wenzelm@26760
   545
  text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
wenzelm@26760
   546
  "*"}@{verbatim "}"}.
wenzelm@26754
   547
wenzelm@26754
   548
  \begin{descr}
wenzelm@26754
   549
  
wenzelm@26754
   550
  \item [@{text "@{theory A}"}] prints the name @{text "A"}, which is
wenzelm@26754
   551
  guaranteed to refer to a valid ancestor theory in the current
wenzelm@26754
   552
  context.
wenzelm@26754
   553
wenzelm@26760
   554
  \item [@{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"}] prints theorems
wenzelm@26760
   555
  @{text "a\<^sub>1 \<dots> a\<^sub>n"}.  Note that attribute specifications
wenzelm@26760
   556
  may be included as well (see also \secref{sec:syn-att}); the
wenzelm@26760
   557
  @{attribute_ref no_vars} rule (see \secref{sec:misc-meth-att}) would
wenzelm@26760
   558
  be particularly useful to suppress printing of schematic variables.
wenzelm@26754
   559
wenzelm@26754
   560
  \item [@{text "@{prop \<phi>}"}] prints a well-typed proposition @{text
wenzelm@26754
   561
  "\<phi>"}.
wenzelm@26754
   562
wenzelm@26754
   563
  \item [@{text "@{term t}"}] prints a well-typed term @{text "t"}.
wenzelm@26754
   564
wenzelm@26754
   565
  \item [@{text "@{const c}"}] prints a logical or syntactic constant
wenzelm@26754
   566
  @{text "c"}.
wenzelm@26754
   567
  
wenzelm@26754
   568
  \item [@{text "@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}"}] prints a constant
wenzelm@26754
   569
  abbreviation @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in
wenzelm@26754
   570
  the current context.
wenzelm@26754
   571
wenzelm@26754
   572
  \item [@{text "@{typeof t}"}] prints the type of a well-typed term
wenzelm@26754
   573
  @{text "t"}.
wenzelm@26754
   574
wenzelm@26754
   575
  \item [@{text "@{typ \<tau>}"}] prints a well-formed type @{text "\<tau>"}.
wenzelm@26754
   576
  
wenzelm@26754
   577
  \item [@{text "@{thm_style s a}"}] prints theorem @{text a},
wenzelm@26754
   578
  previously applying a style @{text s} to it (see below).
wenzelm@26754
   579
  
wenzelm@26754
   580
  \item [@{text "@{term_style s t}"}] prints a well-typed term @{text
wenzelm@26754
   581
  t} after applying a style @{text s} to it (see below).
wenzelm@26754
   582
wenzelm@26754
   583
  \item [@{text "@{text s}"}] prints uninterpreted source text @{text
wenzelm@26754
   584
  s}.  This is particularly useful to print portions of text according
wenzelm@26754
   585
  to the Isabelle {\LaTeX} output style, without demanding
wenzelm@26754
   586
  well-formedness (e.g.\ small pieces of terms that should not be
wenzelm@26754
   587
  parsed or type-checked yet).
wenzelm@26754
   588
wenzelm@26754
   589
  \item [@{text "@{goals}"}] prints the current \emph{dynamic} goal
wenzelm@26754
   590
  state.  This is mainly for support of tactic-emulation scripts
wenzelm@26754
   591
  within Isar --- presentation of goal states does not conform to
wenzelm@26754
   592
  actual human-readable proof documents.
wenzelm@26754
   593
wenzelm@26754
   594
  Please do not include goal states into document output unless you
wenzelm@26754
   595
  really know what you are doing!
wenzelm@26754
   596
  
wenzelm@26754
   597
  \item [@{text "@{subgoals}"}] is similar to @{text "@{goals}"}, but
wenzelm@26754
   598
  does not print the main goal.
wenzelm@26754
   599
  
wenzelm@26754
   600
  \item [@{text "@{prf a\<^sub>1 \<dots> a\<^sub>n}"}] prints the (compact)
wenzelm@26754
   601
  proof terms corresponding to the theorems @{text "a\<^sub>1 \<dots>
wenzelm@26754
   602
  a\<^sub>n"}. Note that this requires proof terms to be switched on
wenzelm@26754
   603
  for the current object logic (see the ``Proof terms'' section of the
wenzelm@26754
   604
  Isabelle reference manual for information on how to do this).
wenzelm@26754
   605
  
wenzelm@26754
   606
  \item [@{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"}] is like @{text
wenzelm@26754
   607
  "@{prf a\<^sub>1 \<dots> a\<^sub>n}"}, but displays the full proof terms,
wenzelm@26754
   608
  i.e.\ also displays information omitted in the compact proof term,
wenzelm@26760
   609
  which is denoted by ``@{verbatim _}'' placeholders there.
wenzelm@26754
   610
  
wenzelm@26754
   611
  \item [@{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
wenzelm@26754
   612
  "@{ML_struct s}"}] check text @{text s} as ML value, type, and
wenzelm@26754
   613
  structure, respectively.  The source is displayed verbatim.
wenzelm@26754
   614
wenzelm@26754
   615
  \end{descr}
wenzelm@26754
   616
wenzelm@26754
   617
  \medskip The following standard styles for use with @{text
wenzelm@26754
   618
  thm_style} and @{text term_style} are available:
wenzelm@26754
   619
wenzelm@26754
   620
  \begin{descr}
wenzelm@26754
   621
  
wenzelm@26754
   622
  \item [@{text lhs}] extracts the first argument of any application
wenzelm@26754
   623
  form with at least two arguments -- typically meta-level or
wenzelm@26754
   624
  object-level equality, or any other binary relation.
wenzelm@26754
   625
  
wenzelm@26754
   626
  \item [@{text rhs}] is like @{text lhs}, but extracts the second
wenzelm@26754
   627
  argument.
wenzelm@26754
   628
  
wenzelm@26754
   629
  \item [@{text "concl"}] extracts the conclusion @{text C} from a rule
wenzelm@26754
   630
  in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
wenzelm@26754
   631
  
wenzelm@26754
   632
  \item [@{text "prem1"}, \dots, @{text "prem9"}] extract premise
wenzelm@26760
   633
  number @{text "1, \<dots>, 9"}, respectively, from from a rule in
wenzelm@26754
   634
  Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
wenzelm@26754
   635
wenzelm@26754
   636
  \end{descr}
wenzelm@26754
   637
wenzelm@26754
   638
  \medskip
wenzelm@26754
   639
  The following options are available to tune the output.  Note that most of
wenzelm@26754
   640
  these coincide with ML flags of the same names (see also \cite{isabelle-ref}).
wenzelm@26754
   641
wenzelm@26754
   642
  \begin{descr}
wenzelm@26754
   643
wenzelm@26754
   644
  \item[@{text "show_types = bool"} and @{text "show_sorts = bool"}]
wenzelm@26754
   645
  control printing of explicit type and sort constraints.
wenzelm@26754
   646
wenzelm@26754
   647
  \item[@{text "show_structs = bool"}] controls printing of implicit
wenzelm@26754
   648
  structures.
wenzelm@26754
   649
wenzelm@26754
   650
  \item[@{text "long_names = bool"}] forces names of types and
wenzelm@26754
   651
  constants etc.\ to be printed in their fully qualified internal
wenzelm@26754
   652
  form.
wenzelm@26754
   653
wenzelm@26754
   654
  \item[@{text "short_names = bool"}] forces names of types and
wenzelm@26754
   655
  constants etc.\ to be printed unqualified.  Note that internalizing
wenzelm@26754
   656
  the output again in the current context may well yield a different
wenzelm@26754
   657
  result.
wenzelm@26754
   658
wenzelm@26754
   659
  \item[@{text "unique_names = bool"}] determines whether the printed
wenzelm@26754
   660
  version of qualified names should be made sufficiently long to avoid
wenzelm@26754
   661
  overlap with names declared further back.  Set to @{text false} for
wenzelm@26754
   662
  more concise output.
wenzelm@26754
   663
wenzelm@26754
   664
  \item[@{text "eta_contract = bool"}] prints terms in @{text
wenzelm@26754
   665
  \<eta>}-contracted form.
wenzelm@26754
   666
wenzelm@26754
   667
  \item[@{text "display = bool"}] indicates if the text is to be
wenzelm@26754
   668
  output as multi-line ``display material'', rather than a small piece
wenzelm@26754
   669
  of text without line breaks (which is the default).
wenzelm@26754
   670
wenzelm@26754
   671
  \item[@{text "break = bool"}] controls line breaks in non-display
wenzelm@26754
   672
  material.
wenzelm@26754
   673
wenzelm@26754
   674
  \item[@{text "quotes = bool"}] indicates if the output should be
wenzelm@26754
   675
  enclosed in double quotes.
wenzelm@26754
   676
wenzelm@26754
   677
  \item[@{text "mode = name"}] adds @{text name} to the print mode to
wenzelm@26754
   678
  be used for presentation (see also \cite{isabelle-ref}).  Note that
wenzelm@26754
   679
  the standard setup for {\LaTeX} output is already present by
wenzelm@26754
   680
  default, including the modes @{text latex} and @{text xsymbols}.
wenzelm@26754
   681
wenzelm@26754
   682
  \item[@{text "margin = nat"} and @{text "indent = nat"}] change the
wenzelm@26754
   683
  margin or indentation for pretty printing of display material.
wenzelm@26754
   684
wenzelm@26754
   685
  \item[@{text "source = bool"}] prints the source text of the
wenzelm@26754
   686
  antiquotation arguments, rather than the actual value.  Note that
wenzelm@26754
   687
  this does not affect well-formedness checks of @{antiquotation
wenzelm@26754
   688
  "thm"}, @{antiquotation "term"}, etc. (only the @{antiquotation
wenzelm@26754
   689
  "text"} antiquotation admits arbitrary output).
wenzelm@26754
   690
wenzelm@26754
   691
  \item[@{text "goals_limit = nat"}] determines the maximum number of
wenzelm@26754
   692
  goals to be printed.
wenzelm@26754
   693
wenzelm@26754
   694
  \item[@{text "locale = name"}] specifies an alternative locale
wenzelm@26754
   695
  context used for evaluating and printing the subsequent argument.
wenzelm@26754
   696
wenzelm@26754
   697
  \end{descr}
wenzelm@26754
   698
wenzelm@26754
   699
  For boolean flags, ``@{text "name = true"}'' may be abbreviated as
wenzelm@26754
   700
  ``@{text name}''.  All of the above flags are disabled by default,
wenzelm@26754
   701
  unless changed from ML.
wenzelm@26754
   702
wenzelm@26754
   703
  \medskip Note that antiquotations do not only spare the author from
wenzelm@26754
   704
  tedious typing of logical entities, but also achieve some degree of
wenzelm@26754
   705
  consistency-checking of informal explanations with formal
wenzelm@26754
   706
  developments: well-formedness of terms and types with respect to the
wenzelm@26754
   707
  current theory or proof context is ensured here.
wenzelm@26754
   708
*}
wenzelm@26754
   709
wenzelm@26754
   710
wenzelm@26754
   711
subsection {* Tagged commands \label{sec:tags} *}
wenzelm@26754
   712
wenzelm@26754
   713
text {*
wenzelm@26754
   714
  Each Isabelle/Isar command may be decorated by presentation tags:
wenzelm@26754
   715
wenzelm@26754
   716
  \indexouternonterm{tags}
wenzelm@26754
   717
  \begin{rail}
wenzelm@26754
   718
    tags: ( tag * )
wenzelm@26754
   719
    ;
wenzelm@26754
   720
    tag: '\%' (ident | string)
wenzelm@26754
   721
  \end{rail}
wenzelm@26754
   722
wenzelm@26754
   723
  The tags @{text "theory"}, @{text "proof"}, @{text "ML"} are already
wenzelm@26754
   724
  pre-declared for certain classes of commands:
wenzelm@26754
   725
wenzelm@26754
   726
 \medskip
wenzelm@26754
   727
wenzelm@26754
   728
  \begin{tabular}{ll}
wenzelm@26754
   729
    @{text "theory"} & theory begin/end \\
wenzelm@26754
   730
    @{text "proof"} & all proof commands \\
wenzelm@26754
   731
    @{text "ML"} & all commands involving ML code \\
wenzelm@26754
   732
  \end{tabular}
wenzelm@26754
   733
wenzelm@26754
   734
  \medskip The Isabelle document preparation system (see also
wenzelm@26754
   735
  \cite{isabelle-sys}) allows tagged command regions to be presented
wenzelm@26754
   736
  specifically, e.g.\ to fold proof texts, or drop parts of the text
wenzelm@26754
   737
  completely.
wenzelm@26754
   738
wenzelm@26754
   739
  For example ``@{command "by"}~@{text "%invisible auto"}'' would
wenzelm@26754
   740
  cause that piece of proof to be treated as @{text invisible} instead
wenzelm@26754
   741
  of @{text "proof"} (the default), which may be either show or hidden
wenzelm@26754
   742
  depending on the document setup.  In contrast, ``@{command
wenzelm@26754
   743
  "by"}~@{text "%visible auto"}'' would force this text to be shown
wenzelm@26754
   744
  invariably.
wenzelm@26754
   745
wenzelm@26754
   746
  Explicit tag specifications within a proof apply to all subsequent
wenzelm@26754
   747
  commands of the same level of nesting.  For example, ``@{command
wenzelm@26754
   748
  "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' would force the
wenzelm@26754
   749
  whole sub-proof to be typeset as @{text visible} (unless some of its
wenzelm@26754
   750
  parts are tagged differently).
wenzelm@26754
   751
*}
wenzelm@26754
   752
wenzelm@26754
   753
end