| author | nipkow |
| Fri, 30 Oct 2015 20:01:05 +0100 | |
| changeset 61525 | 87244a9cfe40 |
| parent 61503 | 28e788ca2c5d |
| child 61579 | 634cd44bb1d3 |
| permissions | -rw-r--r-- |
| 27037 | 1 |
theory Outer_Syntax |
| 42651 | 2 |
imports Base Main |
| 27037 | 3 |
begin |
4 |
||
| 58618 | 5 |
chapter \<open>Outer syntax --- the theory language \label{ch:outer-syntax}\<close>
|
| 27037 | 6 |
|
| 58618 | 7 |
text \<open> |
| 27037 | 8 |
The rather generic framework of Isabelle/Isar syntax emerges from |
| 61477 | 9 |
three main syntactic categories: \<^emph>\<open>commands\<close> of the top-level |
10 |
Isar engine (covering theory and proof elements), \<^emph>\<open>methods\<close> for |
|
| 27037 | 11 |
general goal refinements (analogous to traditional ``tactics''), and |
| 61477 | 12 |
\<^emph>\<open>attributes\<close> for operations on facts (within a certain |
| 27037 | 13 |
context). Subsequently we give a reference of basic syntactic |
14 |
entities underlying Isabelle/Isar syntax in a bottom-up manner. |
|
15 |
Concrete theory and proof language elements will be introduced later |
|
16 |
on. |
|
17 |
||
| 61421 | 18 |
\<^medskip> |
19 |
In order to get started with writing well-formed |
|
| 27037 | 20 |
Isabelle/Isar documents, the most important aspect to be noted is |
| 61477 | 21 |
the difference of \<^emph>\<open>inner\<close> versus \<^emph>\<open>outer\<close> syntax. Inner |
| 27037 | 22 |
syntax is that of Isabelle types and terms of the logic, while outer |
23 |
syntax is that of Isabelle/Isar theory sources (specifications and |
|
24 |
proofs). As a general rule, inner syntax entities may occur only as |
|
| 61477 | 25 |
\<^emph>\<open>atomic entities\<close> within outer syntax. For example, the string |
| 61503 | 26 |
\<^verbatim>\<open>"x + y"\<close> and identifier \<^verbatim>\<open>z\<close> are legal term |
27 |
specifications within a theory, while \<^verbatim>\<open>x + y\<close> without |
|
| 27037 | 28 |
quotes is not. |
29 |
||
30 |
Printed theory documents usually omit quotes to gain readability |
|
| 61503 | 31 |
(this is a matter of {\LaTeX} macro setup, say via \<^verbatim>\<open>\isabellestyle\<close>,
|
32 |
see also @{cite "isabelle-system"}). Experienced
|
|
| 27037 | 33 |
users of Isabelle/Isar may easily reconstruct the lost technical |
34 |
information, while mere readers need not care about quotes at all. |
|
| 58618 | 35 |
\<close> |
| 27037 | 36 |
|
37 |
||
| 58618 | 38 |
section \<open>Commands\<close> |
| 50213 | 39 |
|
| 58618 | 40 |
text \<open> |
| 50213 | 41 |
\begin{matharray}{rcl}
|
| 61493 | 42 |
@{command_def "print_commands"}\<open>\<^sup>*\<close> & : & \<open>any \<rightarrow>\<close> \\
|
43 |
@{command_def "help"}\<open>\<^sup>*\<close> & : & \<open>any \<rightarrow>\<close> \\
|
|
| 50213 | 44 |
\end{matharray}
|
45 |
||
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
46 |
@{rail \<open>
|
| 50213 | 47 |
@@{command help} (@{syntax name} * )
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
48 |
\<close>} |
| 50213 | 49 |
|
| 61439 | 50 |
\<^descr> @{command "print_commands"} prints all outer syntax keywords
|
| 50213 | 51 |
and commands. |
52 |
||
| 61493 | 53 |
\<^descr> @{command "help"}~\<open>pats\<close> retrieves outer syntax
|
| 50213 | 54 |
commands according to the specified name patterns. |
| 58618 | 55 |
\<close> |
| 50213 | 56 |
|
57 |
||
| 58618 | 58 |
subsubsection \<open>Examples\<close> |
| 50213 | 59 |
|
| 58618 | 60 |
text \<open>Some common diagnostic commands are retrieved like this |
61 |
(according to usual naming conventions):\<close> |
|
| 50213 | 62 |
|
63 |
help "print" |
|
64 |
help "find" |
|
65 |
||
66 |
||
| 58618 | 67 |
section \<open>Lexical matters \label{sec:outer-lex}\<close>
|
| 27037 | 68 |
|
| 58618 | 69 |
text \<open>The outer lexical syntax consists of three main categories of |
| 28776 | 70 |
syntax tokens: |
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
71 |
|
| 61477 | 72 |
\<^enum> \<^emph>\<open>major keywords\<close> --- the command names that are available |
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
73 |
in the present logic session; |
|
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
74 |
|
| 61477 | 75 |
\<^enum> \<^emph>\<open>minor keywords\<close> --- additional literal tokens required |
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
76 |
by the syntax of commands; |
|
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
77 |
|
| 61477 | 78 |
\<^enum> \<^emph>\<open>named tokens\<close> --- various categories of identifiers etc. |
| 27037 | 79 |
|
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
80 |
|
| 28776 | 81 |
Major keywords and minor keywords are guaranteed to be disjoint. |
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
82 |
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:
28774
diff
changeset
|
83 |
theory text, without knowing the full details of command syntax. |
| 28776 | 84 |
Internally, there is some additional information about the kind of |
85 |
major keywords, which approximates the command type (theory command, |
|
86 |
proof command etc.). |
|
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
87 |
|
|
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
88 |
Keywords override named tokens. For example, the presence of a |
| 61503 | 89 |
command called \<^verbatim>\<open>term\<close> inhibits the identifier \<^verbatim>\<open>term\<close>, but the |
90 |
string \<^verbatim>\<open>"term"\<close> can be used instead. |
|
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
91 |
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:
28774
diff
changeset
|
92 |
addition to identifiers, wherever a named entity is expected. |
|
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
93 |
|
| 28776 | 94 |
When tokenizing a given input sequence, the lexer repeatedly takes |
95 |
the longest prefix of the input that forms a valid token. Spaces, |
|
96 |
tabs, newlines and formfeeds between tokens serve as explicit |
|
97 |
separators. |
|
98 |
||
| 61421 | 99 |
\<^medskip> |
100 |
The categories for named tokens are defined once and for all as follows. |
|
| 27037 | 101 |
|
| 28776 | 102 |
\begin{center}
|
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
103 |
\begin{supertabular}{rcl}
|
| 61493 | 104 |
@{syntax_def ident} & = & \<open>letter (subscript\<^sup>? quasiletter)\<^sup>*\<close> \\
|
| 61503 | 105 |
@{syntax_def longident} & = & \<open>ident(\<close>\<^verbatim>\<open>.\<close>\<open>ident)\<^sup>+\<close> \\
|
106 |
@{syntax_def symident} & = & \<open>sym\<^sup>+ |\<close>~~\<^verbatim>\<open>\\<close>\<^verbatim>\<open><\<close>\<open>ident\<close>\<^verbatim>\<open>>\<close> \\
|
|
| 61493 | 107 |
@{syntax_def nat} & = & \<open>digit\<^sup>+\<close> \\
|
| 61503 | 108 |
@{syntax_def float} & = & @{syntax_ref nat}\<^verbatim>\<open>.\<close>@{syntax_ref nat}~~\<open>|\<close>~~\<^verbatim>\<open>-\<close>@{syntax_ref nat}\<^verbatim>\<open>.\<close>@{syntax_ref nat} \\
|
109 |
@{syntax_def var} & = & \<^verbatim>\<open>?\<close>\<open>ident |\<close>~~\<^verbatim>\<open>?\<close>\<open>ident\<close>\<^verbatim>\<open>.\<close>\<open>nat\<close> \\
|
|
110 |
@{syntax_def typefree} & = & \<^verbatim>\<open>'\<close>\<open>ident\<close> \\
|
|
111 |
@{syntax_def typevar} & = & \<^verbatim>\<open>?\<close>\<open>typefree |\<close>~~\<^verbatim>\<open>?\<close>\<open>typefree\<close>\<^verbatim>\<open>.\<close>\<open>nat\<close> \\
|
|
112 |
@{syntax_def string} & = & \<^verbatim>\<open>"\<close> \<open>\<dots>\<close> \<^verbatim>\<open>"\<close> \\
|
|
113 |
@{syntax_def altstring} & = & \<^verbatim>\<open>`\<close> \<open>\<dots>\<close> \<^verbatim>\<open>`\<close> \\
|
|
| 61493 | 114 |
@{syntax_def cartouche} & = & @{verbatim "\<open>"} \<open>\<dots>\<close> @{verbatim "\<close>"} \\
|
| 61503 | 115 |
@{syntax_def verbatim} & = & \<^verbatim>\<open>{*\<close> \<open>\<dots>\<close> \<^verbatim>\<open>*}\<close> \\[1ex]
|
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
116 |
|
| 61503 | 117 |
\<open>letter\<close> & = & \<open>latin |\<close>~~\<^verbatim>\<open>\\<close>\<^verbatim>\<open><\<close>\<open>latin\<close>\<^verbatim>\<open>>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\\<close>\<^verbatim>\<open><\<close>\<open>latin latin\<close>\<^verbatim>\<open>>\<close>~~\<open>| greek |\<close> \\ |
118 |
\<open>subscript\<close> & = & \<^verbatim>\<open>\<^sub>\<close> \\ |
|
119 |
\<open>quasiletter\<close> & = & \<open>letter | digit |\<close>~~\<^verbatim>\<open>_\<close>~~\<open>|\<close>~~\<^verbatim>\<open>'\<close> \\ |
|
120 |
\<open>latin\<close> & = & \<^verbatim>\<open>a\<close>~~\<open>| \<dots> |\<close>~~\<^verbatim>\<open>z\<close>~~\<open>|\<close>~~\<^verbatim>\<open>A\<close>~~\<open>| \<dots> |\<close>~~\<^verbatim>\<open>Z\<close> \\ |
|
121 |
\<open>digit\<close> & = & \<^verbatim>\<open>0\<close>~~\<open>| \<dots> |\<close>~~\<^verbatim>\<open>9\<close> \\ |
|
122 |
\<open>sym\<close> & = & \<^verbatim>\<open>!\<close>~~\<open>|\<close>~~\<^verbatim>\<open>#\<close>~~\<open>|\<close>~~\<^verbatim>\<open>$\<close>~~\<open>|\<close>~~\<^verbatim>\<open>%\<close>~~\<open>|\<close>~~\<^verbatim>\<open>&\<close>~~\<open>|\<close>~~\<^verbatim>\<open>*\<close>~~\<open>|\<close>~~\<^verbatim>\<open>+\<close>~~\<open>|\<close>~~\<^verbatim>\<open>-\<close>~~\<open>|\<close>~~\<^verbatim>\<open>/\<close>~~\<open>|\<close> \\ |
|
123 |
& & \<^verbatim>\<open><\<close>~~\<open>|\<close>~~\<^verbatim>\<open>=\<close>~~\<open>|\<close>~~\<^verbatim>\<open>>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>?\<close>~~\<open>|\<close>~~\<^verbatim>\<open>@\<close>~~\<open>|\<close>~~\<^verbatim>\<open>^\<close>~~\<open>|\<close>~~\<^verbatim>\<open>_\<close>~~\<open>|\<close>~~\<^verbatim>\<open>|\<close>~~\<open>|\<close>~~\<^verbatim>\<open>~\<close> \\ |
|
124 |
\<open>greek\<close> & = & \<^verbatim>\<open>\<alpha>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<beta>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<gamma>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<delta>\<close>~~\<open>|\<close> \\ |
|
125 |
& & \<^verbatim>\<open>\<epsilon>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<zeta>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<eta>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<theta>\<close>~~\<open>|\<close> \\ |
|
126 |
& & \<^verbatim>\<open>\<iota>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<kappa>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<mu>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<nu>\<close>~~\<open>|\<close> \\ |
|
127 |
& & \<^verbatim>\<open>\<xi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<pi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<rho>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<sigma>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<tau>\<close>~~\<open>|\<close> \\ |
|
128 |
& & \<^verbatim>\<open>\<upsilon>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<phi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<chi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<psi>\<close>~~\<open>|\<close> \\ |
|
129 |
& & \<^verbatim>\<open>\<omega>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Gamma>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Delta>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Theta>\<close>~~\<open>|\<close> \\ |
|
130 |
& & \<^verbatim>\<open>\<Lambda>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Xi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Pi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Sigma>\<close>~~\<open>|\<close> \\ |
|
131 |
& & \<^verbatim>\<open>\<Upsilon>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Phi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Psi>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>\<Omega>\<close> \\ |
|
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
132 |
\end{supertabular}
|
| 28776 | 133 |
\end{center}
|
| 27037 | 134 |
|
| 28778 | 135 |
A @{syntax_ref var} or @{syntax_ref typevar} describes an unknown,
|
136 |
which is internally a pair of base name and index (ML type @{ML_type
|
|
137 |
indexname}). These components are either separated by a dot as in |
|
| 61503 | 138 |
\<open>?x.1\<close> or \<open>?x7.3\<close> or run together as in \<open>?x1\<close>. The latter form is possible |
139 |
if the base name does not end with digits. If the index is 0, it may be |
|
140 |
dropped altogether: \<open>?x\<close> and \<open>?x0\<close> and \<open>?x.0\<close> all refer to the |
|
| 61493 | 141 |
same unknown, with basename \<open>x\<close> and index 0. |
| 28778 | 142 |
|
143 |
The syntax of @{syntax_ref string} admits any characters, including
|
|
| 61503 | 144 |
newlines; ``\<^verbatim>\<open>"\<close>'' (double-quote) and ``\<^verbatim>\<open>\\<close>'' |
| 58724 | 145 |
(backslash) need to be escaped by a backslash; arbitrary |
| 61503 | 146 |
character codes may be specified as ``\<^verbatim>\<open>\\<close>\<open>ddd\<close>'', |
| 27037 | 147 |
with three decimal digits. Alternative strings according to |
| 28778 | 148 |
@{syntax_ref altstring} are analogous, using single back-quotes
|
149 |
instead. |
|
150 |
||
| 58725 | 151 |
The body of @{syntax_ref verbatim} may consist of any text not containing
|
| 61503 | 152 |
``\<^verbatim>\<open>*}\<close>''; this allows to include quotes without further |
153 |
escapes, but there is no way to escape ``\<^verbatim>\<open>*}\<close>''. Cartouches |
|
| 58725 | 154 |
do not have this limitation. |
| 28778 | 155 |
|
| 55033 | 156 |
A @{syntax_ref cartouche} consists of arbitrary text, with properly
|
| 61493 | 157 |
balanced blocks of ``@{verbatim "\<open>"}~\<open>\<dots>\<close>~@{verbatim
|
| 55033 | 158 |
"\<close>"}''. Note that the rendering of cartouche delimiters is |
| 61493 | 159 |
usually like this: ``\<open>\<open> \<dots> \<close>\<close>''. |
| 55033 | 160 |
|
| 61503 | 161 |
Source comments take the form \<^verbatim>\<open>(*\<close>~\<open>\<dots>\<close>~\<^verbatim>\<open>*)\<close> and may be nested, although the user-interface |
| 28778 | 162 |
might prevent this. Note that this form indicates source comments |
163 |
only, which are stripped after lexical analysis of the input. The |
|
| 61477 | 164 |
Isar syntax also provides proper \<^emph>\<open>document comments\<close> that are |
| 28778 | 165 |
considered as part of the text (see \secref{sec:comments}).
|
| 27037 | 166 |
|
| 61493 | 167 |
Common mathematical symbols such as \<open>\<forall>\<close> are represented in |
| 61503 | 168 |
Isabelle as \<^verbatim>\<open>\<forall>\<close>. There are infinitely many Isabelle |
| 27037 | 169 |
symbols like this, although proper presentation is left to front-end |
| 58842 | 170 |
tools such as {\LaTeX} or Isabelle/jEdit. A list of
|
|
47822
34b44d28fc4b
some updates concerning current Proof General 4.x, which lacks X-Symbol mode of 3.x;
wenzelm
parents:
47114
diff
changeset
|
171 |
predefined Isabelle symbols that work well with these tools is given |
| 61503 | 172 |
in \appref{app:symbols}. Note that \<^verbatim>\<open>\<lambda>\<close> does not belong
|
| 61493 | 173 |
to the \<open>letter\<close> category, since it is already used differently |
| 58618 | 174 |
in the Pure term language.\<close> |
| 27037 | 175 |
|
176 |
||
| 58618 | 177 |
section \<open>Common syntax entities\<close> |
| 27037 | 178 |
|
| 58618 | 179 |
text \<open> |
| 27037 | 180 |
We now introduce several basic syntactic entities, such as names, |
181 |
terms, and theorem specifications, which are factored out of the |
|
182 |
actual Isar language elements to be described later. |
|
| 58618 | 183 |
\<close> |
| 27037 | 184 |
|
185 |
||
| 58618 | 186 |
subsection \<open>Names\<close> |
| 27037 | 187 |
|
| 58618 | 188 |
text \<open>Entity @{syntax name} usually refers to any name of types,
|
| 61477 | 189 |
constants, theorems etc.\ that are to be \<^emph>\<open>declared\<close> or |
190 |
\<^emph>\<open>defined\<close> (so qualified identifiers are excluded here). Quoted |
|
| 27037 | 191 |
strings provide an escape for non-identifier names or those ruled |
| 61503 | 192 |
out by outer syntax keywords (e.g.\ quoted \<^verbatim>\<open>"let"\<close>). |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
193 |
Already existing objects are usually referenced by @{syntax
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
194 |
nameref}. |
| 27037 | 195 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
196 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
197 |
@{syntax_def name}: @{syntax ident} | @{syntax symident} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
198 |
@{syntax string} | @{syntax nat}
|
| 27037 | 199 |
; |
| 60131 | 200 |
@{syntax_def par_name}: '(' @{syntax name} ')'
|
| 27037 | 201 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
202 |
@{syntax_def nameref}: @{syntax name} | @{syntax longident}
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
203 |
\<close>} |
| 58618 | 204 |
\<close> |
| 40296 | 205 |
|
206 |
||
| 58618 | 207 |
subsection \<open>Numbers\<close> |
| 40296 | 208 |
|
| 58618 | 209 |
text \<open>The outer lexical syntax (\secref{sec:outer-lex}) admits
|
| 40296 | 210 |
natural numbers and floating point numbers. These are combined as |
211 |
@{syntax int} and @{syntax real} as follows.
|
|
212 |
||
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
213 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
214 |
@{syntax_def int}: @{syntax nat} | '-' @{syntax nat}
|
| 27037 | 215 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
216 |
@{syntax_def real}: @{syntax float} | @{syntax int}
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
217 |
\<close>} |
| 40296 | 218 |
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
219 |
Note that there is an overlap with the category @{syntax name},
|
| 40296 | 220 |
which also includes @{syntax nat}.
|
| 58618 | 221 |
\<close> |
| 27037 | 222 |
|
223 |
||
| 58618 | 224 |
subsection \<open>Comments \label{sec:comments}\<close>
|
| 27037 | 225 |
|
| 58725 | 226 |
text \<open>Large chunks of plain @{syntax text} are usually given @{syntax
|
| 61503 | 227 |
verbatim}, i.e.\ enclosed in \<^verbatim>\<open>{*\<close>~\<open>\<dots>\<close>~\<^verbatim>\<open>*}\<close>,
|
| 61493 | 228 |
or as @{syntax cartouche} \<open>\<open>\<dots>\<close>\<close>. For convenience, any of the
|
| 58725 | 229 |
smaller text units conforming to @{syntax nameref} are admitted as well. A
|
| 61503 | 230 |
marginal @{syntax comment} is of the form \<^verbatim>\<open>--\<close>~@{syntax text}.
|
| 58725 | 231 |
Any number of these may occur within Isabelle/Isar commands. |
| 27037 | 232 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
233 |
@{rail \<open>
|
|
56499
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56451
diff
changeset
|
234 |
@{syntax_def text}: @{syntax verbatim} | @{syntax cartouche} | @{syntax nameref}
|
| 27037 | 235 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
236 |
@{syntax_def comment}: '--' @{syntax text}
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
237 |
\<close>} |
| 58618 | 238 |
\<close> |
| 27037 | 239 |
|
240 |
||
| 58618 | 241 |
subsection \<open>Type classes, sorts and arities\<close> |
| 27037 | 242 |
|
| 58618 | 243 |
text \<open> |
| 27037 | 244 |
Classes are specified by plain names. Sorts have a very simple |
| 61493 | 245 |
inner syntax, which is either a single class name \<open>c\<close> or a |
246 |
list \<open>{c\<^sub>1, \<dots>, c\<^sub>n}\<close> referring to the
|
|
| 27037 | 247 |
intersection of these classes. The syntax of type arities is given |
248 |
directly at the outer level. |
|
249 |
||
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
250 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
251 |
@{syntax_def classdecl}: @{syntax name} (('<' | '\<subseteq>') (@{syntax nameref} + ','))?
|
| 27037 | 252 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
253 |
@{syntax_def sort}: @{syntax nameref}
|
| 27037 | 254 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
255 |
@{syntax_def arity}: ('(' (@{syntax sort} + ',') ')')? @{syntax sort}
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
256 |
\<close>} |
| 58618 | 257 |
\<close> |
| 27037 | 258 |
|
259 |
||
| 58618 | 260 |
subsection \<open>Types and terms \label{sec:types-terms}\<close>
|
| 27037 | 261 |
|
| 58618 | 262 |
text \<open> |
| 27037 | 263 |
The actual inner Isabelle syntax, that of types and terms of the |
264 |
logic, is far too sophisticated in order to be modelled explicitly |
|
265 |
at the outer theory level. Basically, any such entity has to be |
|
266 |
quoted to turn it into a single token (the parsing and type-checking |
|
267 |
is performed internally later). For convenience, a slightly more |
|
268 |
liberal convention is adopted: quotes may be omitted for any type or |
|
269 |
term that is already atomic at the outer level. For example, one |
|
| 61503 | 270 |
may just write \<^verbatim>\<open>x\<close> instead of quoted \<^verbatim>\<open>"x"\<close>. |
271 |
Note that symbolic identifiers (e.g.\ \<^verbatim>\<open>++\<close> or \<open>\<forall>\<close> are available as well, |
|
272 |
provided these have not been superseded |
|
273 |
by commands or other keywords already (such as \<^verbatim>\<open>=\<close> or |
|
274 |
\<^verbatim>\<open>+\<close>). |
|
| 27037 | 275 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
276 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
277 |
@{syntax_def type}: @{syntax nameref} | @{syntax typefree} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
278 |
@{syntax typevar}
|
| 27037 | 279 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
280 |
@{syntax_def term}: @{syntax nameref} | @{syntax var}
|
| 27037 | 281 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
282 |
@{syntax_def prop}: @{syntax term}
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
283 |
\<close>} |
| 27037 | 284 |
|
| 59853 | 285 |
Positional instantiations are specified as a sequence of terms, or the |
| 61493 | 286 |
placeholder ``\<open>_\<close>'' (underscore), which means to skip a position. |
| 27037 | 287 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
288 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
289 |
@{syntax_def inst}: '_' | @{syntax term}
|
| 27037 | 290 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
291 |
@{syntax_def insts}: (@{syntax inst} *)
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
292 |
\<close>} |
| 27037 | 293 |
|
| 61493 | 294 |
Named instantiations are specified as pairs of assignments \<open>v = |
295 |
t\<close>, which refer to schematic variables in some theorem that is |
|
| 59853 | 296 |
instantiated. Both type and terms instantiations are admitted, and |
297 |
distinguished by the usual syntax of variable names. |
|
298 |
||
299 |
@{rail \<open>
|
|
300 |
@{syntax_def named_inst}: variable '=' (type | term)
|
|
301 |
; |
|
302 |
@{syntax_def named_insts}: (named_inst @'and' +)
|
|
303 |
; |
|
304 |
variable: @{syntax name} | @{syntax var} | @{syntax typefree} | @{syntax typevar}
|
|
305 |
\<close>} |
|
306 |
||
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
307 |
Type declarations and definitions usually refer to @{syntax
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
308 |
typespec} on the left-hand side. This models basic type constructor |
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
309 |
application at the outer syntax level. Note that only plain postfix |
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
310 |
notation is available here, but no infixes. |
| 27037 | 311 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
312 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
313 |
@{syntax_def typespec}:
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
314 |
(() | @{syntax typefree} | '(' ( @{syntax typefree} + ',' ) ')') @{syntax name}
|
| 27037 | 315 |
; |
| 42705 | 316 |
@{syntax_def typespec_sorts}:
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
317 |
(() | (@{syntax typefree} ('::' @{syntax sort})?) |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
318 |
'(' ( (@{syntax typefree} ('::' @{syntax sort})?) + ',' ) ')') @{syntax name}
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
319 |
\<close>} |
| 58618 | 320 |
\<close> |
| 27037 | 321 |
|
322 |
||
| 58618 | 323 |
subsection \<open>Term patterns and declarations \label{sec:term-decls}\<close>
|
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
324 |
|
| 58618 | 325 |
text \<open>Wherever explicit propositions (or term fragments) occur in a |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
326 |
proof text, casual binding of schematic term variables may be given |
| 61493 | 327 |
specified via patterns of the form ``\<open>(\<IS> p\<^sub>1 \<dots> p\<^sub>n)\<close>''. |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
328 |
This works both for @{syntax term} and @{syntax prop}.
|
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
329 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
330 |
@{rail \<open>
|
| 42705 | 331 |
@{syntax_def term_pat}: '(' (@'is' @{syntax term} +) ')'
|
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
332 |
; |
| 42705 | 333 |
@{syntax_def prop_pat}: '(' (@'is' @{syntax prop} +) ')'
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
334 |
\<close>} |
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
335 |
|
| 61421 | 336 |
\<^medskip> |
| 61493 | 337 |
Declarations of local variables \<open>x :: \<tau>\<close> and |
338 |
logical propositions \<open>a : \<phi>\<close> represent different views on |
|
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
339 |
the same principle of introducing a local scope. In practice, one |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
340 |
may usually omit the typing of @{syntax vars} (due to
|
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
341 |
type-inference), and the naming of propositions (due to implicit |
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
342 |
references of current facts). In any case, Isar proof elements |
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
343 |
usually admit to introduce multiple such items simultaneously. |
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
344 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
345 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
346 |
@{syntax_def vars}: (@{syntax name} +) ('::' @{syntax type})?
|
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
347 |
; |
| 42705 | 348 |
@{syntax_def props}: @{syntax thmdecl}? (@{syntax prop} @{syntax prop_pat}? +)
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
349 |
\<close>} |
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
350 |
|
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
351 |
The treatment of multiple declarations corresponds to the |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
352 |
complementary focus of @{syntax vars} versus @{syntax props}. In
|
| 61493 | 353 |
``\<open>x\<^sub>1 \<dots> x\<^sub>n :: \<tau>\<close>'' the typing refers to all variables, while |
354 |
in \<open>a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n\<close> the naming refers to all propositions |
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
355 |
collectively. Isar language elements that refer to @{syntax vars}
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
356 |
or @{syntax props} typically admit separate typings or namings via
|
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
357 |
another level of iteration, with explicit @{keyword_ref "and"}
|
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
358 |
separators; e.g.\ see @{command "fix"} and @{command "assume"} in
|
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
359 |
\secref{sec:proof-context}.
|
| 59785 | 360 |
|
361 |
@{rail \<open>
|
|
362 |
@{syntax_def "fixes"}:
|
|
363 |
((@{syntax name} ('::' @{syntax type})? @{syntax mixfix}? | @{syntax vars}) + @'and')
|
|
364 |
; |
|
365 |
@{syntax_def "for_fixes"}: (@'for' @{syntax "fixes"})?
|
|
366 |
\<close>} |
|
367 |
||
368 |
The category @{syntax "fixes"} is a richer variant of @{syntax vars}: it
|
|
369 |
admits specification of mixfix syntax for the entities that are introduced |
|
| 61493 | 370 |
into the context. An optional suffix ``@{keyword "for"}~\<open>fixes\<close>''
|
| 59785 | 371 |
is admitted in many situations to indicate a so-called ``eigen-context'' |
372 |
of a formal element: the result will be exported and thus generalized over |
|
373 |
the given variables.\<close> |
|
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
374 |
|
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
375 |
|
| 58618 | 376 |
subsection \<open>Attributes and theorems \label{sec:syn-att}\<close>
|
| 27037 | 377 |
|
| 58618 | 378 |
text \<open>Attributes have their own ``semi-inner'' syntax, in the sense |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
379 |
that input conforming to @{syntax args} below is parsed by the
|
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
380 |
attribute a second time. The attribute argument specifications may |
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
381 |
be any sequence of atomic entities (identifiers, strings etc.), or |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
382 |
properly bracketed argument lists. Below @{syntax atom} refers to
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
383 |
any atomic entity, including any @{syntax keyword} conforming to
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
384 |
@{syntax symident}.
|
| 27037 | 385 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
386 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
387 |
@{syntax_def atom}: @{syntax nameref} | @{syntax typefree} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
388 |
@{syntax typevar} | @{syntax var} | @{syntax nat} | @{syntax float} |
|
|
55045
99056d23e05b
cartouche within nested args (attributes, methods, etc.);
wenzelm
parents:
55033
diff
changeset
|
389 |
@{syntax keyword} | @{syntax cartouche}
|
| 27037 | 390 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
391 |
arg: @{syntax atom} | '(' @{syntax args} ')' | '[' @{syntax args} ']'
|
| 27037 | 392 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
393 |
@{syntax_def args}: arg *
|
| 27037 | 394 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
395 |
@{syntax_def attributes}: '[' (@{syntax nameref} @{syntax args} * ',') ']'
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
396 |
\<close>} |
| 27037 | 397 |
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
398 |
Theorem specifications come in several flavors: @{syntax axmdecl}
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
399 |
and @{syntax thmdecl} usually refer to axioms, assumptions or
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
400 |
results of goal statements, while @{syntax thmdef} collects lists of
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
401 |
existing theorems. Existing theorems are given by @{syntax thmref}
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
402 |
and @{syntax thmrefs}, the former requires an actual singleton
|
| 27037 | 403 |
result. |
404 |
||
405 |
There are three forms of theorem references: |
|
| 60674 | 406 |
|
| 61493 | 407 |
\<^enum> named facts \<open>a\<close>, |
| 27037 | 408 |
|
| 61493 | 409 |
\<^enum> selections from named facts \<open>a(i)\<close> or \<open>a(j - k)\<close>, |
| 27037 | 410 |
|
| 61421 | 411 |
\<^enum> literal fact propositions using token syntax @{syntax_ref altstring}
|
| 61503 | 412 |
\<^verbatim>\<open>`\<close>\<open>\<phi>\<close>\<^verbatim>\<open>`\<close> or @{syntax_ref cartouche}
|
| 61493 | 413 |
\<open>\<open>\<phi>\<close>\<close> (see also method @{method_ref fact}).
|
| 27037 | 414 |
|
415 |
||
416 |
Any kind of theorem specification may include lists of attributes |
|
417 |
both on the left and right hand sides; attributes are applied to any |
|
418 |
immediately preceding fact. If names are omitted, the theorems are |
|
419 |
not stored within the theorem database of the theory or proof |
|
420 |
context, but any given attributes are applied nonetheless. |
|
421 |
||
| 61493 | 422 |
An extra pair of brackets around attributes (like ``\<open>[[simproc a]]\<close>'') abbreviates a theorem reference involving an |
| 27037 | 423 |
internal dummy fact, which will be ignored later on. So only the |
424 |
effect of the attribute on the background context will persist. |
|
425 |
This form of in-place declarations is particularly useful with |
|
426 |
commands like @{command "declare"} and @{command "using"}.
|
|
427 |
||
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
428 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
429 |
@{syntax_def axmdecl}: @{syntax name} @{syntax attributes}? ':'
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
430 |
; |
| 60631 | 431 |
@{syntax_def thmbind}:
|
432 |
@{syntax name} @{syntax attributes} | @{syntax name} | @{syntax attributes}
|
|
433 |
; |
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
434 |
@{syntax_def thmdecl}: thmbind ':'
|
| 27037 | 435 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
436 |
@{syntax_def thmdef}: thmbind '='
|
| 27037 | 437 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
438 |
@{syntax_def thmref}:
|
|
56499
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56451
diff
changeset
|
439 |
(@{syntax nameref} selection? | @{syntax altstring} | @{syntax cartouche})
|
|
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56451
diff
changeset
|
440 |
@{syntax attributes}? |
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
441 |
'[' @{syntax attributes} ']'
|
| 27037 | 442 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
443 |
@{syntax_def thmrefs}: @{syntax thmref} +
|
| 27037 | 444 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
445 |
selection: '(' ((@{syntax nat} | @{syntax nat} '-' @{syntax nat}?) + ',') ')'
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55045
diff
changeset
|
446 |
\<close>} |
| 58618 | 447 |
\<close> |
| 27037 | 448 |
|
| 60674 | 449 |
|
450 |
section \<open>Diagnostic commands\<close> |
|
451 |
||
452 |
text \<open> |
|
453 |
\begin{matharray}{rcl}
|
|
| 61493 | 454 |
@{command_def "print_theory"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
455 |
@{command_def "print_definitions"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
|
456 |
@{command_def "print_methods"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
|
457 |
@{command_def "print_attributes"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
|
458 |
@{command_def "print_theorems"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
|
459 |
@{command_def "find_theorems"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
|
460 |
@{command_def "find_consts"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
|
461 |
@{command_def "thm_deps"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
|
462 |
@{command_def "unused_thms"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
|
463 |
@{command_def "print_facts"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
|
464 |
@{command_def "print_term_bindings"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
|
| 60674 | 465 |
\end{matharray}
|
466 |
||
467 |
@{rail \<open>
|
|
468 |
(@@{command print_theory} |
|
|
| 61252 | 469 |
@@{command print_definitions} |
|
| 60674 | 470 |
@@{command print_methods} |
|
471 |
@@{command print_attributes} |
|
|
472 |
@@{command print_theorems} |
|
|
473 |
@@{command print_facts}) ('!'?)
|
|
474 |
; |
|
475 |
@@{command find_theorems} ('(' @{syntax nat}? 'with_dups'? ')')? \<newline> (thm_criterion*)
|
|
476 |
; |
|
477 |
thm_criterion: ('-'?) ('name' ':' @{syntax nameref} | 'intro' | 'elim' | 'dest' |
|
|
478 |
'solves' | 'simp' ':' @{syntax term} | @{syntax term})
|
|
479 |
; |
|
480 |
@@{command find_consts} (const_criterion*)
|
|
481 |
; |
|
482 |
const_criterion: ('-'?)
|
|
483 |
('name' ':' @{syntax nameref} | 'strict' ':' @{syntax type} | @{syntax type})
|
|
484 |
; |
|
485 |
@@{command thm_deps} @{syntax thmrefs}
|
|
486 |
; |
|
487 |
@@{command unused_thms} ((@{syntax name} +) '-' (@{syntax name} * ))?
|
|
488 |
\<close>} |
|
489 |
||
490 |
These commands print certain parts of the theory and proof context. |
|
491 |
Note that there are some further ones available, such as for the set |
|
492 |
of rules declared for simplifications. |
|
493 |
||
| 61439 | 494 |
\<^descr> @{command "print_theory"} prints the main logical content of the
|
| 61493 | 495 |
background theory; the ``\<open>!\<close>'' option indicates extra verbosity. |
| 60674 | 496 |
|
| 61439 | 497 |
\<^descr> @{command "print_definitions"} prints dependencies of definitional
|
| 61252 | 498 |
specifications within the background theory, which may be constants |
499 |
\secref{sec:consts} or types (\secref{sec:types-pure},
|
|
| 61493 | 500 |
\secref{sec:hol-typedef}); the ``\<open>!\<close>'' option indicates extra
|
| 61252 | 501 |
verbosity. |
502 |
||
| 61439 | 503 |
\<^descr> @{command "print_methods"} prints all proof methods available in the
|
| 61493 | 504 |
current theory context; the ``\<open>!\<close>'' option indicates extra |
| 60674 | 505 |
verbosity. |
506 |
||
| 61439 | 507 |
\<^descr> @{command "print_attributes"} prints all attributes available in the
|
| 61493 | 508 |
current theory context; the ``\<open>!\<close>'' option indicates extra |
| 60674 | 509 |
verbosity. |
510 |
||
| 61439 | 511 |
\<^descr> @{command "print_theorems"} prints theorems of the background theory
|
| 61493 | 512 |
resulting from the last command; the ``\<open>!\<close>'' option indicates |
| 60674 | 513 |
extra verbosity. |
514 |
||
| 61439 | 515 |
\<^descr> @{command "print_facts"} prints all local facts of the current
|
| 61493 | 516 |
context, both named and unnamed ones; the ``\<open>!\<close>'' option indicates |
| 60674 | 517 |
extra verbosity. |
518 |
||
| 61439 | 519 |
\<^descr> @{command "print_term_bindings"} prints all term bindings that
|
| 60674 | 520 |
are present in the context. |
521 |
||
| 61493 | 522 |
\<^descr> @{command "find_theorems"}~\<open>criteria\<close> retrieves facts
|
| 60674 | 523 |
from the theory or proof context matching all of given search |
| 61493 | 524 |
criteria. The criterion \<open>name: p\<close> selects all theorems |
525 |
whose fully qualified name matches pattern \<open>p\<close>, which may |
|
526 |
contain ``\<open>*\<close>'' wildcards. The criteria \<open>intro\<close>, |
|
527 |
\<open>elim\<close>, and \<open>dest\<close> select theorems that match the |
|
| 60674 | 528 |
current goal as introduction, elimination or destruction rules, |
| 61493 | 529 |
respectively. The criterion \<open>solves\<close> returns all rules |
| 60674 | 530 |
that would directly solve the current goal. The criterion |
| 61493 | 531 |
\<open>simp: t\<close> selects all rewrite rules whose left-hand side |
532 |
matches the given term. The criterion term \<open>t\<close> selects all |
|
533 |
theorems that contain the pattern \<open>t\<close> -- as usual, patterns |
|
534 |
may contain occurrences of the dummy ``\<open>_\<close>'', schematic |
|
| 60674 | 535 |
variables, and type constraints. |
536 |
||
| 61493 | 537 |
Criteria can be preceded by ``\<open>-\<close>'' to select theorems that |
| 61477 | 538 |
do \<^emph>\<open>not\<close> match. Note that giving the empty list of criteria |
539 |
yields \<^emph>\<open>all\<close> currently known facts. An optional limit for the |
|
| 60674 | 540 |
number of printed facts may be given; the default is 40. By |
541 |
default, duplicates are removed from the search result. Use |
|
| 61493 | 542 |
\<open>with_dups\<close> to display duplicates. |
| 60674 | 543 |
|
| 61493 | 544 |
\<^descr> @{command "find_consts"}~\<open>criteria\<close> prints all constants
|
545 |
whose type meets all of the given criteria. The criterion \<open>strict: ty\<close> is met by any type that matches the type pattern |
|
546 |
\<open>ty\<close>. Patterns may contain both the dummy type ``\<open>_\<close>'' |
|
547 |
and sort constraints. The criterion \<open>ty\<close> is similar, but it |
|
548 |
also matches against subtypes. The criterion \<open>name: p\<close> and |
|
549 |
the prefix ``\<open>-\<close>'' function as described for @{command
|
|
| 60674 | 550 |
"find_theorems"}. |
551 |
||
| 61493 | 552 |
\<^descr> @{command "thm_deps"}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close>
|
| 60674 | 553 |
visualizes dependencies of facts, using Isabelle's graph browser |
554 |
tool (see also @{cite "isabelle-system"}).
|
|
555 |
||
| 61493 | 556 |
\<^descr> @{command "unused_thms"}~\<open>A\<^sub>1 \<dots> A\<^sub>m - B\<^sub>1 \<dots> B\<^sub>n\<close>
|
557 |
displays all theorems that are proved in theories \<open>B\<^sub>1 \<dots> B\<^sub>n\<close> |
|
558 |
or their parents but not in \<open>A\<^sub>1 \<dots> A\<^sub>m\<close> or their parents and |
|
| 60674 | 559 |
that are never used. |
| 61493 | 560 |
If \<open>n\<close> is \<open>0\<close>, the end of the range of theories |
| 60674 | 561 |
defaults to the current theory. If no range is specified, |
562 |
only the unused theorems in the current theory are displayed. |
|
563 |
\<close> |
|
564 |
||
| 27037 | 565 |
end |