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