doc-src/IsarRef/Thy/Outer_Syntax.thy
author wenzelm
Thu, 13 Nov 2008 21:59:02 +0100
changeset 28774 0e25ef17b06b
parent 28762 f5d79aeffd81
child 28775 d25fe9601dbd
permissions -rw-r--r--
more tuning of Pure grammer;
wenzelm@27037
     1
(* $Id$ *)
wenzelm@27037
     2
wenzelm@27037
     3
theory Outer_Syntax
wenzelm@27050
     4
imports Main
wenzelm@27037
     5
begin
wenzelm@27037
     6
wenzelm@27040
     7
chapter {* Outer syntax *}
wenzelm@27037
     8
wenzelm@27037
     9
text {*
wenzelm@27037
    10
  The rather generic framework of Isabelle/Isar syntax emerges from
wenzelm@27037
    11
  three main syntactic categories: \emph{commands} of the top-level
wenzelm@27037
    12
  Isar engine (covering theory and proof elements), \emph{methods} for
wenzelm@27037
    13
  general goal refinements (analogous to traditional ``tactics''), and
wenzelm@27037
    14
  \emph{attributes} for operations on facts (within a certain
wenzelm@27037
    15
  context).  Subsequently we give a reference of basic syntactic
wenzelm@27037
    16
  entities underlying Isabelle/Isar syntax in a bottom-up manner.
wenzelm@27037
    17
  Concrete theory and proof language elements will be introduced later
wenzelm@27037
    18
  on.
wenzelm@27037
    19
wenzelm@27037
    20
  \medskip In order to get started with writing well-formed
wenzelm@27037
    21
  Isabelle/Isar documents, the most important aspect to be noted is
wenzelm@27037
    22
  the difference of \emph{inner} versus \emph{outer} syntax.  Inner
wenzelm@27037
    23
  syntax is that of Isabelle types and terms of the logic, while outer
wenzelm@27037
    24
  syntax is that of Isabelle/Isar theory sources (specifications and
wenzelm@27037
    25
  proofs).  As a general rule, inner syntax entities may occur only as
wenzelm@27037
    26
  \emph{atomic entities} within outer syntax.  For example, the string
wenzelm@27037
    27
  @{verbatim "\"x + y\""} and identifier @{verbatim z} are legal term
wenzelm@27037
    28
  specifications within a theory, while @{verbatim "x + y"} without
wenzelm@27037
    29
  quotes is not.
wenzelm@27037
    30
wenzelm@27037
    31
  Printed theory documents usually omit quotes to gain readability
wenzelm@27037
    32
  (this is a matter of {\LaTeX} macro setup, say via @{verbatim
wenzelm@27037
    33
  "\\isabellestyle"}, see also \cite{isabelle-sys}).  Experienced
wenzelm@27037
    34
  users of Isabelle/Isar may easily reconstruct the lost technical
wenzelm@27037
    35
  information, while mere readers need not care about quotes at all.
wenzelm@27037
    36
wenzelm@27037
    37
  \medskip Isabelle/Isar input may contain any number of input
wenzelm@27037
    38
  termination characters ``@{verbatim ";"}'' (semicolon) to separate
wenzelm@27037
    39
  commands explicitly.  This is particularly useful in interactive
wenzelm@27037
    40
  shell sessions to make clear where the current command is intended
wenzelm@27037
    41
  to end.  Otherwise, the interpreter loop will continue to issue a
wenzelm@27037
    42
  secondary prompt ``@{verbatim "#"}'' until an end-of-command is
wenzelm@27037
    43
  clearly recognized from the input syntax, e.g.\ encounter of the
wenzelm@27037
    44
  next command keyword.
wenzelm@27037
    45
wenzelm@27037
    46
  More advanced interfaces such as Proof~General \cite{proofgeneral}
wenzelm@27037
    47
  do not require explicit semicolons, the amount of input text is
wenzelm@27037
    48
  determined automatically by inspecting the present content of the
wenzelm@27037
    49
  Emacs text buffer.  In the printed presentation of Isabelle/Isar
wenzelm@27037
    50
  documents semicolons are omitted altogether for readability.
wenzelm@27037
    51
wenzelm@27037
    52
  \begin{warn}
wenzelm@27037
    53
    Proof~General requires certain syntax classification tables in
wenzelm@27037
    54
    order to achieve properly synchronized interaction with the
wenzelm@27037
    55
    Isabelle/Isar process.  These tables need to be consistent with
wenzelm@27037
    56
    the Isabelle version and particular logic image to be used in a
wenzelm@27037
    57
    running session (common object-logics may well change the outer
wenzelm@27037
    58
    syntax).  The standard setup should work correctly with any of the
wenzelm@27037
    59
    ``official'' logic images derived from Isabelle/HOL (including
wenzelm@27037
    60
    HOLCF etc.).  Users of alternative logics may need to tell
wenzelm@27037
    61
    Proof~General explicitly, e.g.\ by giving an option @{verbatim "-k ZF"}
wenzelm@27037
    62
    (in conjunction with @{verbatim "-l ZF"}, to specify the default
wenzelm@27037
    63
    logic image).  Note that option @{verbatim "-L"} does both
wenzelm@27037
    64
    of this at the same time.
wenzelm@27037
    65
  \end{warn}
wenzelm@27037
    66
*}
wenzelm@27037
    67
wenzelm@27037
    68
wenzelm@28774
    69
section {* Lexical matters \label{sec:outer-lex} *}
wenzelm@27037
    70
wenzelm@28774
    71
text {* The Isabelle/Isar outer syntax provides token classes as
wenzelm@28774
    72
  presented below; most of these coincide with the inner lexical
wenzelm@28774
    73
  syntax as defined in \secref{sec:inner-lex}.
wenzelm@27037
    74
wenzelm@27037
    75
  \begin{matharray}{rcl}
wenzelm@27037
    76
    @{syntax_def ident} & = & letter\,quasiletter^* \\
wenzelm@27037
    77
    @{syntax_def longident} & = & ident (\verb,.,ident)^+ \\
wenzelm@27037
    78
    @{syntax_def symident} & = & sym^+ ~|~ \verb,\,\verb,<,ident\verb,>, \\
wenzelm@27037
    79
    @{syntax_def nat} & = & digit^+ \\
wenzelm@27037
    80
    @{syntax_def var} & = & ident ~|~ \verb,?,ident ~|~ \verb,?,ident\verb,.,nat \\
wenzelm@27037
    81
    @{syntax_def typefree} & = & \verb,',ident \\
wenzelm@27037
    82
    @{syntax_def typevar} & = & typefree ~|~ \verb,?,typefree ~|~ \verb,?,typefree\verb,.,nat \\
wenzelm@27037
    83
    @{syntax_def string} & = & \verb,", ~\dots~ \verb,", \\
wenzelm@27037
    84
    @{syntax_def altstring} & = & \backquote ~\dots~ \backquote \\
wenzelm@27037
    85
    @{syntax_def verbatim} & = & \verb,{*, ~\dots~ \verb,*,\verb,}, \\[1ex]
wenzelm@27037
    86
wenzelm@27037
    87
    letter & = & latin ~|~ \verb,\,\verb,<,latin\verb,>, ~|~ \verb,\,\verb,<,latin\,latin\verb,>, ~|~ greek ~|~ \\
wenzelm@27037
    88
           &   & \verb,\<^isub>, ~|~ \verb,\<^isup>, \\
wenzelm@27037
    89
    quasiletter & = & letter ~|~ digit ~|~ \verb,_, ~|~ \verb,', \\
wenzelm@27037
    90
    latin & = & \verb,a, ~|~ \dots ~|~ \verb,z, ~|~ \verb,A, ~|~ \dots ~|~ \verb,Z, \\
wenzelm@27037
    91
    digit & = & \verb,0, ~|~ \dots ~|~ \verb,9, \\
wenzelm@27037
    92
    sym & = & \verb,!, ~|~ \verb,#, ~|~ \verb,$, ~|~ \verb,%, ~|~ \verb,&, ~|~
wenzelm@27037
    93
     \verb,*, ~|~ \verb,+, ~|~ \verb,-, ~|~ \verb,/, ~|~ \\
wenzelm@27037
    94
    & & \verb,<, ~|~ \verb,=, ~|~ \verb,>, ~|~ \verb,?, ~|~ \texttt{\at} ~|~
wenzelm@27037
    95
    \verb,^, ~|~ \verb,_, ~|~ \verb,|, ~|~ \verb,~, \\
wenzelm@27037
    96
    greek & = & \verb,\<alpha>, ~|~ \verb,\<beta>, ~|~ \verb,\<gamma>, ~|~ \verb,\<delta>, ~| \\
wenzelm@27037
    97
          &   & \verb,\<epsilon>, ~|~ \verb,\<zeta>, ~|~ \verb,\<eta>, ~|~ \verb,\<theta>, ~| \\
wenzelm@27037
    98
          &   & \verb,\<iota>, ~|~ \verb,\<kappa>, ~|~ \verb,\<mu>, ~|~ \verb,\<nu>, ~| \\
wenzelm@27037
    99
          &   & \verb,\<xi>, ~|~ \verb,\<pi>, ~|~ \verb,\<rho>, ~|~ \verb,\<sigma>, ~|~ \verb,\<tau>, ~| \\
wenzelm@27037
   100
          &   & \verb,\<upsilon>, ~|~ \verb,\<phi>, ~|~ \verb,\<chi>, ~|~ \verb,\<psi>, ~| \\
wenzelm@27037
   101
          &   & \verb,\<omega>, ~|~ \verb,\<Gamma>, ~|~ \verb,\<Delta>, ~|~ \verb,\<Theta>, ~| \\
wenzelm@27037
   102
          &   & \verb,\<Lambda>, ~|~ \verb,\<Xi>, ~|~ \verb,\<Pi>, ~|~ \verb,\<Sigma>, ~| \\
wenzelm@27037
   103
          &   & \verb,\<Upsilon>, ~|~ \verb,\<Phi>, ~|~ \verb,\<Psi>, ~|~ \verb,\<Omega>, \\
wenzelm@27037
   104
  \end{matharray}
wenzelm@27037
   105
wenzelm@27037
   106
  The syntax of @{syntax string} admits any characters, including
wenzelm@27037
   107
  newlines; ``@{verbatim "\""}'' (double-quote) and ``@{verbatim
wenzelm@27037
   108
  "\\"}'' (backslash) need to be escaped by a backslash; arbitrary
wenzelm@27037
   109
  character codes may be specified as ``@{verbatim "\\"}@{text ddd}'',
wenzelm@27037
   110
  with three decimal digits.  Alternative strings according to
wenzelm@27037
   111
  @{syntax altstring} are analogous, using single back-quotes instead.
wenzelm@27037
   112
  The body of @{syntax verbatim} may consist of any text not
wenzelm@27037
   113
  containing ``@{verbatim "*"}@{verbatim "}"}''; this allows
wenzelm@27037
   114
  convenient inclusion of quotes without further escapes.  The greek
wenzelm@27037
   115
  letters do \emph{not} include @{verbatim "\<lambda>"}, which is already used
wenzelm@27037
   116
  differently in the meta-logic.
wenzelm@27037
   117
wenzelm@27037
   118
  Common mathematical symbols such as @{text \<forall>} are represented in
wenzelm@27037
   119
  Isabelle as @{verbatim \<forall>}.  There are infinitely many Isabelle
wenzelm@27037
   120
  symbols like this, although proper presentation is left to front-end
wenzelm@27037
   121
  tools such as {\LaTeX} or Proof~General with the X-Symbol package.
wenzelm@27037
   122
  A list of standard Isabelle symbols that work well with these tools
wenzelm@27037
   123
  is given in \cite[appendix~A]{isabelle-sys}.
wenzelm@27037
   124
  
wenzelm@27037
   125
  Source comments take the form @{verbatim "(*"}~@{text
wenzelm@27037
   126
  "\<dots>"}~@{verbatim "*)"} and may be nested, although user-interface
wenzelm@27037
   127
  tools might prevent this.  Note that this form indicates source
wenzelm@27037
   128
  comments only, which are stripped after lexical analysis of the
wenzelm@28748
   129
  input.  The Isar syntax also provides proper \emph{document
wenzelm@28748
   130
  comments} that are considered as part of the text (see
wenzelm@28748
   131
  \secref{sec:comments}).
wenzelm@27037
   132
*}
wenzelm@27037
   133
wenzelm@27037
   134
wenzelm@27037
   135
section {* Common syntax entities *}
wenzelm@27037
   136
wenzelm@27037
   137
text {*
wenzelm@27037
   138
  We now introduce several basic syntactic entities, such as names,
wenzelm@27037
   139
  terms, and theorem specifications, which are factored out of the
wenzelm@27037
   140
  actual Isar language elements to be described later.
wenzelm@27037
   141
*}
wenzelm@27037
   142
wenzelm@27037
   143
wenzelm@27037
   144
subsection {* Names *}
wenzelm@27037
   145
wenzelm@27037
   146
text {*
wenzelm@27037
   147
  Entity \railqtok{name} usually refers to any name of types,
wenzelm@27037
   148
  constants, theorems etc.\ that are to be \emph{declared} or
wenzelm@27037
   149
  \emph{defined} (so qualified identifiers are excluded here).  Quoted
wenzelm@27037
   150
  strings provide an escape for non-identifier names or those ruled
wenzelm@27037
   151
  out by outer syntax keywords (e.g.\ quoted @{verbatim "\"let\""}).
wenzelm@27037
   152
  Already existing objects are usually referenced by
wenzelm@27037
   153
  \railqtok{nameref}.
wenzelm@27037
   154
wenzelm@27037
   155
  \indexoutertoken{name}\indexoutertoken{parname}\indexoutertoken{nameref}
wenzelm@27037
   156
  \indexoutertoken{int}
wenzelm@27037
   157
  \begin{rail}
wenzelm@27037
   158
    name: ident | symident | string | nat
wenzelm@27037
   159
    ;
wenzelm@27037
   160
    parname: '(' name ')'
wenzelm@27037
   161
    ;
wenzelm@27037
   162
    nameref: name | longident
wenzelm@27037
   163
    ;
wenzelm@27037
   164
    int: nat | '-' nat
wenzelm@27037
   165
    ;
wenzelm@27037
   166
  \end{rail}
wenzelm@27037
   167
*}
wenzelm@27037
   168
wenzelm@27037
   169
wenzelm@27037
   170
subsection {* Comments \label{sec:comments} *}
wenzelm@27037
   171
wenzelm@27037
   172
text {*
wenzelm@27037
   173
  Large chunks of plain \railqtok{text} are usually given
wenzelm@27037
   174
  \railtok{verbatim}, i.e.\ enclosed in @{verbatim "{"}@{verbatim
wenzelm@27037
   175
  "*"}~@{text "\<dots>"}~@{verbatim "*"}@{verbatim "}"}.  For convenience,
wenzelm@27037
   176
  any of the smaller text units conforming to \railqtok{nameref} are
wenzelm@27037
   177
  admitted as well.  A marginal \railnonterm{comment} is of the form
wenzelm@27037
   178
  @{verbatim "--"} \railqtok{text}.  Any number of these may occur
wenzelm@27037
   179
  within Isabelle/Isar commands.
wenzelm@27037
   180
wenzelm@27037
   181
  \indexoutertoken{text}\indexouternonterm{comment}
wenzelm@27037
   182
  \begin{rail}
wenzelm@27037
   183
    text: verbatim | nameref
wenzelm@27037
   184
    ;
wenzelm@27037
   185
    comment: '--' text
wenzelm@27037
   186
    ;
wenzelm@27037
   187
  \end{rail}
wenzelm@27037
   188
*}
wenzelm@27037
   189
wenzelm@27037
   190
wenzelm@27037
   191
subsection {* Type classes, sorts and arities *}
wenzelm@27037
   192
wenzelm@27037
   193
text {*
wenzelm@27037
   194
  Classes are specified by plain names.  Sorts have a very simple
wenzelm@27037
   195
  inner syntax, which is either a single class name @{text c} or a
wenzelm@27037
   196
  list @{text "{c\<^sub>1, \<dots>, c\<^sub>n}"} referring to the
wenzelm@27037
   197
  intersection of these classes.  The syntax of type arities is given
wenzelm@27037
   198
  directly at the outer level.
wenzelm@27037
   199
wenzelm@27037
   200
  \indexouternonterm{sort}\indexouternonterm{arity}
wenzelm@27037
   201
  \indexouternonterm{classdecl}
wenzelm@27037
   202
  \begin{rail}
wenzelm@27037
   203
    classdecl: name (('<' | subseteq) (nameref + ','))?
wenzelm@27037
   204
    ;
wenzelm@27037
   205
    sort: nameref
wenzelm@27037
   206
    ;
wenzelm@27037
   207
    arity: ('(' (sort + ',') ')')? sort
wenzelm@27037
   208
    ;
wenzelm@27037
   209
  \end{rail}
wenzelm@27037
   210
*}
wenzelm@27037
   211
wenzelm@27037
   212
wenzelm@27037
   213
subsection {* Types and terms \label{sec:types-terms} *}
wenzelm@27037
   214
wenzelm@27037
   215
text {*
wenzelm@27037
   216
  The actual inner Isabelle syntax, that of types and terms of the
wenzelm@27037
   217
  logic, is far too sophisticated in order to be modelled explicitly
wenzelm@27037
   218
  at the outer theory level.  Basically, any such entity has to be
wenzelm@27037
   219
  quoted to turn it into a single token (the parsing and type-checking
wenzelm@27037
   220
  is performed internally later).  For convenience, a slightly more
wenzelm@27037
   221
  liberal convention is adopted: quotes may be omitted for any type or
wenzelm@27037
   222
  term that is already atomic at the outer level.  For example, one
wenzelm@27037
   223
  may just write @{verbatim x} instead of quoted @{verbatim "\"x\""}.
wenzelm@27037
   224
  Note that symbolic identifiers (e.g.\ @{verbatim "++"} or @{text
wenzelm@27037
   225
  "\<forall>"} are available as well, provided these have not been superseded
wenzelm@27037
   226
  by commands or other keywords already (such as @{verbatim "="} or
wenzelm@27037
   227
  @{verbatim "+"}).
wenzelm@27037
   228
wenzelm@27037
   229
  \indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop}
wenzelm@27037
   230
  \begin{rail}
wenzelm@27037
   231
    type: nameref | typefree | typevar
wenzelm@27037
   232
    ;
wenzelm@27037
   233
    term: nameref | var
wenzelm@27037
   234
    ;
wenzelm@27037
   235
    prop: term
wenzelm@27037
   236
    ;
wenzelm@27037
   237
  \end{rail}
wenzelm@27037
   238
wenzelm@27037
   239
  Positional instantiations are indicated by giving a sequence of
wenzelm@27037
   240
  terms, or the placeholder ``@{text _}'' (underscore), which means to
wenzelm@27037
   241
  skip a position.
wenzelm@27037
   242
wenzelm@27037
   243
  \indexoutertoken{inst}\indexoutertoken{insts}
wenzelm@27037
   244
  \begin{rail}
wenzelm@27037
   245
    inst: underscore | term
wenzelm@27037
   246
    ;
wenzelm@27037
   247
    insts: (inst *)
wenzelm@27037
   248
    ;
wenzelm@27037
   249
  \end{rail}
wenzelm@27037
   250
wenzelm@27037
   251
  Type declarations and definitions usually refer to
wenzelm@27037
   252
  \railnonterm{typespec} on the left-hand side.  This models basic
wenzelm@27037
   253
  type constructor application at the outer syntax level.  Note that
wenzelm@27037
   254
  only plain postfix notation is available here, but no infixes.
wenzelm@27037
   255
wenzelm@27037
   256
  \indexouternonterm{typespec}
wenzelm@27037
   257
  \begin{rail}
wenzelm@27037
   258
    typespec: (() | typefree | '(' ( typefree + ',' ) ')') name
wenzelm@27037
   259
    ;
wenzelm@27037
   260
  \end{rail}
wenzelm@27037
   261
*}
wenzelm@27037
   262
wenzelm@27037
   263
wenzelm@28754
   264
subsection {* Term patterns and declarations \label{sec:term-decls} *}
wenzelm@28754
   265
wenzelm@28754
   266
text {*
wenzelm@28754
   267
  Wherever explicit propositions (or term fragments) occur in a proof
wenzelm@28754
   268
  text, casual binding of schematic term variables may be given
wenzelm@28754
   269
  specified via patterns of the form ``@{text "(\<IS> p\<^sub>1 \<dots>
wenzelm@28754
   270
  p\<^sub>n)"}''.  This works both for \railqtok{term} and \railqtok{prop}.
wenzelm@28754
   271
wenzelm@28754
   272
  \indexouternonterm{termpat}\indexouternonterm{proppat}
wenzelm@28754
   273
  \begin{rail}
wenzelm@28754
   274
    termpat: '(' ('is' term +) ')'
wenzelm@28754
   275
    ;
wenzelm@28754
   276
    proppat: '(' ('is' prop +) ')'
wenzelm@28754
   277
    ;
wenzelm@28754
   278
  \end{rail}
wenzelm@28754
   279
wenzelm@28754
   280
  \medskip Declarations of local variables @{text "x :: \<tau>"} and
wenzelm@28754
   281
  logical propositions @{text "a : \<phi>"} represent different views on
wenzelm@28754
   282
  the same principle of introducing a local scope.  In practice, one
wenzelm@28754
   283
  may usually omit the typing of \railnonterm{vars} (due to
wenzelm@28754
   284
  type-inference), and the naming of propositions (due to implicit
wenzelm@28754
   285
  references of current facts).  In any case, Isar proof elements
wenzelm@28754
   286
  usually admit to introduce multiple such items simultaneously.
wenzelm@28754
   287
wenzelm@28754
   288
  \indexouternonterm{vars}\indexouternonterm{props}
wenzelm@28754
   289
  \begin{rail}
wenzelm@28754
   290
    vars: (name+) ('::' type)?
wenzelm@28754
   291
    ;
wenzelm@28754
   292
    props: thmdecl? (prop proppat? +)
wenzelm@28754
   293
    ;
wenzelm@28754
   294
  \end{rail}
wenzelm@28754
   295
wenzelm@28754
   296
  The treatment of multiple declarations corresponds to the
wenzelm@28754
   297
  complementary focus of \railnonterm{vars} versus
wenzelm@28754
   298
  \railnonterm{props}.  In ``@{text "x\<^sub>1 \<dots> x\<^sub>n :: \<tau>"}''
wenzelm@28754
   299
  the typing refers to all variables, while in @{text "a: \<phi>\<^sub>1 \<dots>
wenzelm@28754
   300
  \<phi>\<^sub>n"} the naming refers to all propositions collectively.
wenzelm@28754
   301
  Isar language elements that refer to \railnonterm{vars} or
wenzelm@28754
   302
  \railnonterm{props} typically admit separate typings or namings via
wenzelm@28754
   303
  another level of iteration, with explicit @{keyword_ref "and"}
wenzelm@28754
   304
  separators; e.g.\ see @{command "fix"} and @{command "assume"} in
wenzelm@28754
   305
  \secref{sec:proof-context}.
wenzelm@28754
   306
*}
wenzelm@28754
   307
wenzelm@28754
   308
wenzelm@27037
   309
subsection {* Attributes and theorems \label{sec:syn-att} *}
wenzelm@27037
   310
wenzelm@28754
   311
text {* Attributes have their own ``semi-inner'' syntax, in the sense
wenzelm@28754
   312
  that input conforming to \railnonterm{args} below is parsed by the
wenzelm@28754
   313
  attribute a second time.  The attribute argument specifications may
wenzelm@28754
   314
  be any sequence of atomic entities (identifiers, strings etc.), or
wenzelm@28754
   315
  properly bracketed argument lists.  Below \railqtok{atom} refers to
wenzelm@28754
   316
  any atomic entity, including any \railtok{keyword} conforming to
wenzelm@28754
   317
  \railtok{symident}.
wenzelm@27037
   318
wenzelm@27037
   319
  \indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}
wenzelm@27037
   320
  \begin{rail}
wenzelm@27037
   321
    atom: nameref | typefree | typevar | var | nat | keyword
wenzelm@27037
   322
    ;
wenzelm@27037
   323
    arg: atom | '(' args ')' | '[' args ']'
wenzelm@27037
   324
    ;
wenzelm@27037
   325
    args: arg *
wenzelm@27037
   326
    ;
wenzelm@27037
   327
    attributes: '[' (nameref args * ',') ']'
wenzelm@27037
   328
    ;
wenzelm@27037
   329
  \end{rail}
wenzelm@27037
   330
wenzelm@27037
   331
  Theorem specifications come in several flavors:
wenzelm@27037
   332
  \railnonterm{axmdecl} and \railnonterm{thmdecl} usually refer to
wenzelm@27037
   333
  axioms, assumptions or results of goal statements, while
wenzelm@27037
   334
  \railnonterm{thmdef} collects lists of existing theorems.  Existing
wenzelm@27037
   335
  theorems are given by \railnonterm{thmref} and
wenzelm@27037
   336
  \railnonterm{thmrefs}, the former requires an actual singleton
wenzelm@27037
   337
  result.
wenzelm@27037
   338
wenzelm@27037
   339
  There are three forms of theorem references:
wenzelm@27037
   340
  \begin{enumerate}
wenzelm@27037
   341
  
wenzelm@27037
   342
  \item named facts @{text "a"},
wenzelm@27037
   343
wenzelm@27037
   344
  \item selections from named facts @{text "a(i)"} or @{text "a(j - k)"},
wenzelm@27037
   345
wenzelm@27037
   346
  \item literal fact propositions using @{syntax_ref altstring} syntax
wenzelm@27037
   347
  @{verbatim "`"}@{text "\<phi>"}@{verbatim "`"} (see also method
wenzelm@28754
   348
  @{method_ref fact}).
wenzelm@27037
   349
wenzelm@27037
   350
  \end{enumerate}
wenzelm@27037
   351
wenzelm@27037
   352
  Any kind of theorem specification may include lists of attributes
wenzelm@27037
   353
  both on the left and right hand sides; attributes are applied to any
wenzelm@27037
   354
  immediately preceding fact.  If names are omitted, the theorems are
wenzelm@27037
   355
  not stored within the theorem database of the theory or proof
wenzelm@27037
   356
  context, but any given attributes are applied nonetheless.
wenzelm@27037
   357
wenzelm@27037
   358
  An extra pair of brackets around attributes (like ``@{text
wenzelm@27037
   359
  "[[simproc a]]"}'') abbreviates a theorem reference involving an
wenzelm@27037
   360
  internal dummy fact, which will be ignored later on.  So only the
wenzelm@27037
   361
  effect of the attribute on the background context will persist.
wenzelm@27037
   362
  This form of in-place declarations is particularly useful with
wenzelm@27037
   363
  commands like @{command "declare"} and @{command "using"}.
wenzelm@27037
   364
wenzelm@27037
   365
  \indexouternonterm{axmdecl}\indexouternonterm{thmdecl}
wenzelm@27037
   366
  \indexouternonterm{thmdef}\indexouternonterm{thmref}
wenzelm@27037
   367
  \indexouternonterm{thmrefs}\indexouternonterm{selection}
wenzelm@27037
   368
  \begin{rail}
wenzelm@27037
   369
    axmdecl: name attributes? ':'
wenzelm@27037
   370
    ;
wenzelm@27037
   371
    thmdecl: thmbind ':'
wenzelm@27037
   372
    ;
wenzelm@27037
   373
    thmdef: thmbind '='
wenzelm@27037
   374
    ;
wenzelm@27037
   375
    thmref: (nameref selection? | altstring) attributes? | '[' attributes ']'
wenzelm@27037
   376
    ;
wenzelm@27037
   377
    thmrefs: thmref +
wenzelm@27037
   378
    ;
wenzelm@27037
   379
wenzelm@27037
   380
    thmbind: name attributes | name | attributes
wenzelm@27037
   381
    ;
wenzelm@27037
   382
    selection: '(' ((nat | nat '-' nat?) + ',') ')'
wenzelm@27037
   383
    ;
wenzelm@27037
   384
  \end{rail}
wenzelm@27037
   385
*}
wenzelm@27037
   386
wenzelm@27037
   387
end