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