3 (*>*) |
3 (*>*) |
4 |
4 |
5 section {* Concrete Syntax \label{sec:concrete-syntax} *} |
5 section {* Concrete Syntax \label{sec:concrete-syntax} *} |
6 |
6 |
7 text {* |
7 text {* |
8 The core concept of Isabelle's framework for concrete |
8 The core concept of Isabelle's framework for concrete syntax is that |
9 syntax is that of \bfindex{mixfix annotations}. Associated |
9 of \bfindex{mixfix annotations}. Associated with any kind of |
10 with any kind of constant declaration, mixfixes affect both the |
10 constant declaration, mixfixes affect both the grammar productions |
11 grammar productions for the parser and output templates for the |
11 for the parser and output templates for the pretty printer. |
12 pretty printer. |
|
13 |
12 |
14 In full generality, parser and pretty printer configuration is a |
13 In full generality, parser and pretty printer configuration is a |
15 subtle affair \cite{isabelle-ref}. Your syntax |
14 subtle affair \cite{isabelle-ref}. Your syntax specifications need |
16 specifications need to interact properly with the |
15 to interact properly with the existing setup of Isabelle/Pure and |
17 existing setup of Isabelle/Pure and Isabelle/HOL\@. |
16 Isabelle/HOL\@. To avoid creating ambiguities with existing |
18 To avoid creating ambiguities with existing elements, it is |
17 elements, it is particularly important to give new syntactic |
19 particularly important to give new syntactic |
|
20 constructs the right precedence. |
18 constructs the right precedence. |
21 |
19 |
22 \medskip Subsequently we introduce a few simple syntax declaration |
20 \medskip Subsequently we introduce a few simple syntax declaration |
23 forms that already cover many common situations fairly well. |
21 forms that already cover many common situations fairly well. |
24 *} |
22 *} |
26 |
24 |
27 subsection {* Infix Annotations *} |
25 subsection {* Infix Annotations *} |
28 |
26 |
29 text {* |
27 text {* |
30 Syntax annotations may be included wherever constants are declared, |
28 Syntax annotations may be included wherever constants are declared, |
31 such as \isacommand{consts} and |
29 such as \isacommand{consts} and \isacommand{constdefs} --- and also |
32 \isacommand{constdefs} --- and also \isacommand{datatype}, which |
30 \isacommand{datatype}, which declares constructor operations. |
33 declares constructor operations. Type-constructors may be annotated as |
31 Type-constructors may be annotated as well, although this is less |
34 well, although this is less frequently encountered in practice (the |
32 frequently encountered in practice (the infix type @{text "\<times>"} comes |
35 infix type @{text "\<times>"} comes to mind). |
33 to mind). |
36 |
34 |
37 Infix declarations\index{infix annotations} provide a useful special |
35 Infix declarations\index{infix annotations} provide a useful special |
38 case of mixfixes. The following example of the |
36 case of mixfixes. The following example of the exclusive-or |
39 exclusive-or operation on boolean values illustrates typical infix |
37 operation on boolean values illustrates typical infix declarations. |
40 declarations. |
|
41 *} |
38 *} |
42 |
39 |
43 constdefs |
40 constdefs |
44 xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "[+]" 60) |
41 xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "[+]" 60) |
45 "A [+] B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)" |
42 "A [+] B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)" |
46 |
43 |
47 text {* |
44 text {* |
48 \noindent Now @{text "xor A B"} and @{text "A [+] B"} refer to the |
45 \noindent Now @{text "xor A B"} and @{text "A [+] B"} refer to the |
49 same expression internally. Any curried function with at least two |
46 same expression internally. Any curried function with at least two |
50 arguments may be given infix syntax. For partial |
47 arguments may be given infix syntax. For partial applications with |
51 applications with fewer than two operands, there is a notation |
48 fewer than two operands, there is a notation using the prefix~@{text |
52 using the prefix~\isa{op}. For instance, @{text xor} without arguments is represented |
49 op}. For instance, @{text xor} without arguments is represented as |
53 as @{text "op [+]"}; together with ordinary function application, this |
50 @{text "op [+]"}; together with ordinary function application, this |
54 turns @{text "xor A"} into @{text "op [+] A"}. |
51 turns @{text "xor A"} into @{text "op [+] A"}. |
55 |
52 |
56 \medskip The keyword \isakeyword{infixl} seen above specifies an |
53 \medskip The keyword \isakeyword{infixl} seen above specifies an |
57 infix operator that is nested to the \emph{left}: in iterated |
54 infix operator that is nested to the \emph{left}: in iterated |
58 applications the more complex expression appears on the left-hand |
55 applications the more complex expression appears on the left-hand |
59 side, and @{term "A [+] B [+] C"} stands for @{text "(A [+] B) [+] C"}. |
56 side, and @{term "A [+] B [+] C"} stands for @{text "(A [+] B) [+] |
60 Similarly, \isakeyword{infixr} specifies nesting to the |
57 C"}. Similarly, \isakeyword{infixr} means nesting to the |
61 \emph{right}, reading @{term "A [+] B [+] C"} as @{text "A [+] (B |
58 \emph{right}, reading @{term "A [+] B [+] C"} as @{text "A [+] (B |
62 [+] C)"}. In contrast, a \emph{non-oriented} declaration via |
59 [+] C)"}. A \emph{non-oriented} declaration via \isakeyword{infix} |
63 \isakeyword{infix} would render @{term "A [+] B [+] C"} illegal, but |
60 would render @{term "A [+] B [+] C"} illegal, but demand explicit |
64 demand explicit parentheses to indicate the intended grouping. |
61 parentheses to indicate the intended grouping. |
65 |
62 |
66 The string @{text [source] "[+]"} in our annotation refers to the |
63 The string @{text [source] "[+]"} in our annotation refers to the |
67 concrete syntax to represent the operator (a literal token), while |
64 concrete syntax to represent the operator (a literal token), while |
68 the number @{text 60} determines the precedence of the construct: |
65 the number @{text 60} determines the precedence of the construct: |
69 the syntactic priorities of the arguments and result. |
66 the syntactic priorities of the arguments and result. Isabelle/HOL |
70 Isabelle/HOL already uses up many popular combinations of |
67 already uses up many popular combinations of ASCII symbols for its |
71 ASCII symbols for its own use, including both @{text "+"} and @{text |
68 own use, including both @{text "+"} and @{text "++"}. Longer |
72 "++"}. Longer character combinations are |
69 character combinations are more likely to be still available for |
73 more likely to be still available for user extensions, such as our~@{text "[+]"}. |
70 user extensions, such as our~@{text "[+]"}. |
74 |
71 |
75 Operator precedences have a range of 0--1000. Very low or high priorities are |
72 Operator precedences have a range of 0--1000. Very low or high |
76 reserved for the meta-logic. HOL syntax |
73 priorities are reserved for the meta-logic. HOL syntax mainly uses |
77 mainly uses the range of 10--100: the equality infix @{text "="} is |
74 the range of 10--100: the equality infix @{text "="} is centered at |
78 centered at 50; logical connectives (like @{text "\<or>"} and @{text |
75 50; logical connectives (like @{text "\<or>"} and @{text "\<and>"}) are |
79 "\<and>"}) are below 50; algebraic ones (like @{text "+"} and @{text |
76 below 50; algebraic ones (like @{text "+"} and @{text "*"}) are |
80 "*"}) are above 50. User syntax should strive to coexist with common |
77 above 50. User syntax should strive to coexist with common HOL |
81 HOL forms, or use the mostly unused range 100--900. |
78 forms, or use the mostly unused range 100--900. |
82 |
|
83 *} |
79 *} |
84 |
80 |
85 |
81 |
86 subsection {* Mathematical Symbols \label{sec:syntax-symbols} *} |
82 subsection {* Mathematical Symbols \label{sec:syntax-symbols} *} |
87 |
83 |
88 text {* |
84 text {* |
89 Concrete syntax based on ASCII characters has inherent |
85 Concrete syntax based on ASCII characters has inherent limitations. |
90 limitations. Mathematical notation demands a larger repertoire |
86 Mathematical notation demands a larger repertoire of glyphs. |
91 of glyphs. Several standards of extended character sets have been |
87 Several standards of extended character sets have been proposed over |
92 proposed over decades, but none has become universally available so |
88 decades, but none has become universally available so far. Isabelle |
93 far. Isabelle supports a generic notion of \bfindex{symbols} as the |
89 has its own notion of \bfindex{symbols} as the smallest entities of |
94 smallest entities of source text, without referring to internal |
90 source text, without referring to internal encodings. There are |
95 encodings. There are three kinds of such ``generalized |
91 three kinds of such ``generalized characters'': |
96 characters'': |
|
97 |
92 |
98 \begin{enumerate} |
93 \begin{enumerate} |
99 |
94 |
100 \item 7-bit ASCII characters |
95 \item 7-bit ASCII characters |
101 |
96 |
105 |
100 |
106 \end{enumerate} |
101 \end{enumerate} |
107 |
102 |
108 Here $ident$ may be any identifier according to the usual Isabelle |
103 Here $ident$ may be any identifier according to the usual Isabelle |
109 conventions. This results in an infinite store of symbols, whose |
104 conventions. This results in an infinite store of symbols, whose |
110 interpretation is left to further front-end tools. For example, |
105 interpretation is left to further front-end tools. For example, the |
111 both the user-interface of Proof~General + X-Symbol and the Isabelle |
106 user-interface of Proof~General + X-Symbol and the Isabelle document |
112 document processor (see \S\ref{sec:document-preparation}) display |
107 processor (see \S\ref{sec:document-preparation}) display the |
113 the \verb,\,\verb,<forall>, symbol as~@{text \<forall>}. |
108 \verb,\,\verb,<forall>, symbol as~@{text \<forall>}. |
114 |
109 |
115 A list of standard Isabelle symbols is given in |
110 A list of standard Isabelle symbols is given in |
116 \cite[appendix~A]{isabelle-sys}. You may introduce their own |
111 \cite[appendix~A]{isabelle-sys}. You may introduce your own |
117 interpretation of further symbols by configuring the appropriate |
112 interpretation of further symbols by configuring the appropriate |
118 front-end tool accordingly, e.g.\ by defining certain {\LaTeX} |
113 front-end tool accordingly, e.g.\ by defining certain {\LaTeX} |
119 macros (see also \S\ref{sec:doc-prep-symbols}). There are also a |
114 macros (see also \S\ref{sec:doc-prep-symbols}). There are also a |
120 few predefined control symbols, such as \verb,\,\verb,<^sub>, and |
115 few predefined control symbols, such as \verb,\,\verb,<^sub>, and |
121 \verb,\,\verb,<^sup>, for sub- and superscript of the subsequent |
116 \verb,\,\verb,<^sup>, for sub- and superscript of the subsequent |
122 printable symbol, respectively. For example, \verb,A\<^sup>\<star>, is |
117 printable symbol, respectively. For example, \verb,A\<^sup>\<star>, is |
123 output as @{text "A\<^sup>\<star>"}. |
118 output as @{text "A\<^sup>\<star>"}. |
124 |
119 |
125 \medskip Replacing our definition of @{text xor} by the following |
120 \medskip Replacing our definition of @{text xor} by the following |
126 specifies a Isabelle symbol for the new operator: |
121 specifies a Isabelle symbol for the new operator: |
127 *} |
122 *} |
128 |
123 |
129 (*<*) |
124 (*<*) |
130 hide const xor |
125 hide const xor |
138 (*>*) |
133 (*>*) |
139 |
134 |
140 text {* |
135 text {* |
141 \noindent The X-Symbol package within Proof~General provides several |
136 \noindent The X-Symbol package within Proof~General provides several |
142 input methods to enter @{text \<oplus>} in the text. If all fails one may |
137 input methods to enter @{text \<oplus>} in the text. If all fails one may |
143 just type a named entity \verb,\,\verb,<oplus>, by hand; the |
138 just type a named entity \verb,\,\verb,<oplus>, by hand; the |
144 corresponding symbol will immediately be displayed. |
139 corresponding symbol will be displayed after further input. |
145 |
140 |
146 \medskip More flexible is to provide alternative |
141 \medskip More flexible is to provide alternative syntax forms |
147 syntax forms through the \bfindex{print mode} concept~\cite{isabelle-ref}. |
142 through the \bfindex{print mode} concept~\cite{isabelle-ref}. By |
148 By convention, the mode of |
143 convention, the mode of ``$xsymbols$'' is enabled whenever |
149 ``$xsymbols$'' is enabled whenever Proof~General's X-Symbol mode or |
144 Proof~General's X-Symbol mode or {\LaTeX} output is active. Now |
150 {\LaTeX} output is active. Now consider the following hybrid |
145 consider the following hybrid declaration of @{text xor}: |
151 declaration of @{text xor}: |
|
152 *} |
146 *} |
153 |
147 |
154 (*<*) |
148 (*<*) |
155 hide const xor |
149 hide const xor |
156 ML_setup {* Context.>> (Theory.add_path "version2") *} |
150 ML_setup {* Context.>> (Theory.add_path "version2") *} |
172 "(xsymbols)"}, is optional. Also note that its type merely serves |
166 "(xsymbols)"}, is optional. Also note that its type merely serves |
173 for syntactic purposes, and is \emph{not} checked for consistency |
167 for syntactic purposes, and is \emph{not} checked for consistency |
174 with the real constant. |
168 with the real constant. |
175 |
169 |
176 \medskip We may now write @{text "A [+] B"} or @{text "A \<oplus> B"} in |
170 \medskip We may now write @{text "A [+] B"} or @{text "A \<oplus> B"} in |
177 input, while output uses the nicer syntax of $xsymbols$, provided |
171 input, while output uses the nicer syntax of $xsymbols$ whenever |
178 that print mode is active. Such an arrangement is particularly |
172 that print mode is active. Such an arrangement is particularly |
179 useful for interactive development, where users may type ASCII |
173 useful for interactive development, where users may type ASCII text |
180 text and see mathematical symbols displayed during proofs. |
174 and see mathematical symbols displayed during proofs. |
181 *} |
175 *} |
182 |
176 |
183 |
177 |
184 subsection {* Prefix Annotations *} |
178 subsection {* Prefix Annotations *} |
185 |
179 |
186 text {* |
180 text {* |
187 Prefix syntax annotations\index{prefix annotation} are another |
181 Prefix syntax annotations\index{prefix annotation} are another form |
188 form of mixfixes \cite{isabelle-ref}, without any |
182 of mixfixes \cite{isabelle-ref}, without any template arguments or |
189 template arguments or priorities --- just some bits of literal |
183 priorities --- just some literal syntax. The following example |
190 syntax. The following example illustrates this idea idea by |
184 associates common symbols with the constructors of a datatype. |
191 associating common symbols with the constructors of a datatype. |
|
192 *} |
185 *} |
193 |
186 |
194 datatype currency = |
187 datatype currency = |
195 Euro nat ("\<euro>") |
188 Euro nat ("\<euro>") |
196 | Pounds nat ("\<pounds>") |
189 | Pounds nat ("\<pounds>") |
206 printed as @{term "\<euro> 10"}; only the head of the application is |
199 printed as @{term "\<euro> 10"}; only the head of the application is |
207 subject to our concrete syntax. This rather simple form already |
200 subject to our concrete syntax. This rather simple form already |
208 achieves conformance with notational standards of the European |
201 achieves conformance with notational standards of the European |
209 Commission. |
202 Commission. |
210 |
203 |
211 Prefix syntax also works for \isakeyword{consts} or |
204 Prefix syntax works the same way for \isakeyword{consts} or |
212 \isakeyword{constdefs}. |
205 \isakeyword{constdefs}. |
213 *} |
206 *} |
214 |
207 |
215 |
208 |
216 subsection {* Syntax Translations \label{sec:syntax-translations} *} |
209 subsection {* Syntax Translations \label{sec:syntax-translations} *} |
217 |
210 |
218 text{* |
211 text{* |
219 Mixfix syntax annotations merely decorate |
212 Mixfix syntax annotations merely decorate particular constant |
220 particular constant application forms with |
213 application forms with concrete syntax, for instance replacing \ |
221 concrete syntax, for instance replacing \ @{text "xor A B"} by @{text "A \<oplus> B"}. Occasionally, the relationship between some piece of |
214 @{text "xor A B"} by @{text "A \<oplus> B"}. Occasionally, the |
222 notation and its internal form is more complicated. Here we need |
215 relationship between some piece of notation and its internal form is |
223 \bfindex{syntax translations}. |
216 more complicated. Here we need \bfindex{syntax translations}. |
224 |
217 |
225 Using the \isakeyword{syntax}\index{syntax (command)}, command we |
218 Using the \isakeyword{syntax}\index{syntax (command)}, command we |
226 introduce uninterpreted notational elements. Then |
219 introduce uninterpreted notational elements. Then |
227 \commdx{translations} relate input forms to complex logical |
220 \commdx{translations} relate input forms to complex logical |
228 expressions. This provides a simple mechanism for |
221 expressions. This provides a simple mechanism for syntactic macros; |
229 syntactic macros; even heavier transformations may be written in ML |
222 even heavier transformations may be written in ML |
230 \cite{isabelle-ref}. |
223 \cite{isabelle-ref}. |
231 |
224 |
232 \medskip A typical use of syntax translations is to introduce |
225 \medskip A typical use of syntax translations is to introduce |
233 relational notation for membership in a set of pair, |
226 relational notation for membership in a set of pair, replacing \ |
234 replacing \ @{text "(x, y) \<in> sim"} by @{text "x \<approx> y"}. |
227 @{text "(x, y) \<in> sim"} by @{text "x \<approx> y"}. |
235 *} |
228 *} |
236 |
229 |
237 consts |
230 consts |
238 sim :: "('a \<times> 'a) set" |
231 sim :: "('a \<times> 'a) set" |
239 |
232 |
265 instead of \isakeyword{syntax} + \isakeyword{translations}. The |
258 instead of \isakeyword{syntax} + \isakeyword{translations}. The |
266 present formulation has the virtue that expressions are immediately |
259 present formulation has the virtue that expressions are immediately |
267 replaced by the ``definition'' upon parsing; the effect is reversed |
260 replaced by the ``definition'' upon parsing; the effect is reversed |
268 upon printing. |
261 upon printing. |
269 |
262 |
270 This sort of translation is appropriate when the |
263 This sort of translation is appropriate when the defined concept is |
271 defined concept is a trivial variation on an |
264 a trivial variation on an existing one. On the other hand, syntax |
272 existing one. On the other hand, syntax translations do not scale |
265 translations do not scale up well to large hierarchies of concepts. |
273 up well to large hierarchies of concepts. Translations |
266 Translations do not replace definitions! |
274 do not replace definitions! |
|
275 *} |
267 *} |
276 |
268 |
277 |
269 |
278 section {* Document Preparation \label{sec:document-preparation} *} |
270 section {* Document Preparation \label{sec:document-preparation} *} |
279 |
271 |
280 text {* |
272 text {* |
281 Isabelle/Isar is centered around the concept of \bfindex{formal |
273 Isabelle/Isar is centered around the concept of \bfindex{formal |
282 proof documents}\index{documents|bold}. The outcome of a |
274 proof documents}\index{documents|bold}. The outcome of a formal |
283 formal development effort is meant to be a human-readable record, |
275 development effort is meant to be a human-readable record, presented |
284 presented as browsable PDF file or printed on paper. The overall |
276 as browsable PDF file or printed on paper. The overall document |
285 document structure follows traditional mathematical articles, with |
277 structure follows traditional mathematical articles, with sections, |
286 sections, intermediate explanations, definitions, theorems and |
278 intermediate explanations, definitions, theorems and proofs. |
287 proofs. |
|
288 |
279 |
289 \medskip The Isabelle document preparation system essentially acts |
280 \medskip The Isabelle document preparation system essentially acts |
290 as a front-end to {\LaTeX}. After checking specifications and |
281 as a front-end to {\LaTeX}. After checking specifications and |
291 proofs formally, the theory sources are turned into typesetting |
282 proofs formally, the theory sources are turned into typesetting |
292 instructions in a schematic manner. This lets you write |
283 instructions in a schematic manner. This lets you write authentic |
293 authentic reports on theory developments with little effort: many |
284 reports on theory developments with little effort: many technical |
294 technical consistency checks are handled by the system. |
285 consistency checks are handled by the system. |
295 |
286 |
296 Here is an example to illustrate the idea of Isabelle document |
287 Here is an example to illustrate the idea of Isabelle document |
297 preparation. |
288 preparation. |
298 *} |
289 *} |
299 |
290 |
314 *} |
305 *} |
315 |
306 |
316 text_raw {* \end{quotation} *} |
307 text_raw {* \end{quotation} *} |
317 |
308 |
318 text {* |
309 text {* |
319 The above document output has been produced by the following theory |
310 \noindent The above document output has been produced as follows: |
320 specification: |
|
321 |
311 |
322 \begin{ttbox} |
312 \begin{ttbox} |
323 text {\ttlbrace}* |
313 text {\ttlbrace}* |
324 The following datatype definition of {\at}{\ttlbrace}text "'a bintree"{\ttrbrace} |
314 The following datatype definition of {\at}{\ttlbrace}text "'a bintree"{\ttrbrace} |
325 models binary trees with nodes being decorated by elements |
315 models binary trees with nodes being decorated by elements |
326 of type {\at}{\ttlbrace}typ 'a{\ttrbrace}. |
316 of type {\at}{\ttlbrace}typ 'a{\ttrbrace}. |
327 *{\ttrbrace} |
317 *{\ttrbrace} |
328 |
318 |
329 datatype 'a bintree = |
319 datatype 'a bintree = |
330 Leaf | Branch 'a "'a bintree" "'a bintree" |
320 Leaf | Branch 'a "'a bintree" "'a bintree" |
331 |
321 \end{ttbox} |
|
322 \begin{ttbox} |
332 text {\ttlbrace}* |
323 text {\ttlbrace}* |
333 {\ttback}noindent The datatype induction rule generated here is |
324 {\ttback}noindent The datatype induction rule generated here is |
334 of the form {\at}{\ttlbrace}thm [display] bintree.induct [no_vars]{\ttrbrace} |
325 of the form {\at}{\ttlbrace}thm [display] bintree.induct [no_vars]{\ttrbrace} |
335 *{\ttrbrace} |
326 *{\ttrbrace} |
336 \end{ttbox} |
327 \end{ttbox}\vspace{-\medskipamount} |
337 |
328 |
338 \noindent Here we have augmented the theory by formal comments |
329 \noindent Here we have augmented the theory by formal comments |
339 (using \isakeyword{text} blocks). The informal parts may again |
330 (using \isakeyword{text} blocks), the informal parts may again refer |
340 refer to formal entities by means of ``antiquotations'' (such as |
331 to formal entities by means of ``antiquotations'' (such as |
341 \texttt{\at}\verb,{text "'a bintree"}, or |
332 \texttt{\at}\verb,{text "'a bintree"}, or |
342 \texttt{\at}\verb,{typ 'a},), see also \S\ref{sec:doc-prep-text}. |
333 \texttt{\at}\verb,{typ 'a},), see also \S\ref{sec:doc-prep-text}. |
343 *} |
334 *} |
344 |
335 |
345 |
336 |
347 |
338 |
348 text {* |
339 text {* |
349 In contrast to the highly interactive mode of Isabelle/Isar theory |
340 In contrast to the highly interactive mode of Isabelle/Isar theory |
350 development, the document preparation stage essentially works in |
341 development, the document preparation stage essentially works in |
351 batch-mode. An Isabelle \bfindex{session} consists of a collection |
342 batch-mode. An Isabelle \bfindex{session} consists of a collection |
352 of source files that may contribute to an output document. |
343 of source files that may contribute to an output document. Each |
353 Each session is derived from a single parent, usually |
344 session is derived from a single parent, usually an object-logic |
354 an object-logic image like \texttt{HOL}. This results in an overall |
345 image like \texttt{HOL}. This results in an overall tree structure, |
355 tree structure, which is reflected by the output location in the |
346 which is reflected by the output location in the file system |
356 file system (usually rooted at \verb,~/isabelle/browser_info,). |
347 (usually rooted at \verb,~/isabelle/browser_info,). |
357 |
348 |
358 \medskip The easiest way to manage Isabelle sessions is via |
349 \medskip The easiest way to manage Isabelle sessions is via |
359 \texttt{isatool mkdir} (generates an initial session source setup) |
350 \texttt{isatool mkdir} (generates an initial session source setup) |
360 and \texttt{isatool make} (run sessions controlled by |
351 and \texttt{isatool make} (run sessions controlled by |
361 \texttt{IsaMakefile}). For example, a new session |
352 \texttt{IsaMakefile}). For example, a new session |
425 distributed with Isabelle. Further packages may be required in |
416 distributed with Isabelle. Further packages may be required in |
426 particular applications, say for unusual mathematical symbols. |
417 particular applications, say for unusual mathematical symbols. |
427 |
418 |
428 \medskip Any additional files for the {\LaTeX} stage go into the |
419 \medskip Any additional files for the {\LaTeX} stage go into the |
429 \texttt{MySession/document} directory as well. In particular, |
420 \texttt{MySession/document} directory as well. In particular, |
430 adding a file named \texttt{root.bib} causes an |
421 adding a file named \texttt{root.bib} causes an automatic run of |
431 automatic run of \texttt{bibtex} to process a bibliographic |
422 \texttt{bibtex} to process a bibliographic database; see also |
432 database; see also \texttt{isatool document} \cite{isabelle-sys}. |
423 \texttt{isatool document} \cite{isabelle-sys}. |
433 |
424 |
434 \medskip Any failure of the document preparation phase in an |
425 \medskip Any failure of the document preparation phase in an |
435 Isabelle batch session leaves the generated sources in their target |
426 Isabelle batch session leaves the generated sources in their target |
436 location, identified by the accompanying error message. This |
427 location, identified by the accompanying error message. This lets |
437 lets you trace {\LaTeX} problems with the generated files at |
428 you trace {\LaTeX} problems with the generated files at hand. |
438 hand. |
|
439 *} |
429 *} |
440 |
430 |
441 |
431 |
442 subsection {* Structure Markup *} |
432 subsection {* Structure Markup *} |
443 |
433 |
468 |
458 |
469 From the Isabelle perspective, each markup command takes a single |
459 From the Isabelle perspective, each markup command takes a single |
470 $text$ argument (delimited by \verb,",~@{text \<dots>}~\verb,", or |
460 $text$ argument (delimited by \verb,",~@{text \<dots>}~\verb,", or |
471 \verb,{,\verb,*,~@{text \<dots>}~\verb,*,\verb,},). After stripping any |
461 \verb,{,\verb,*,~@{text \<dots>}~\verb,*,\verb,},). After stripping any |
472 surrounding white space, the argument is passed to a {\LaTeX} macro |
462 surrounding white space, the argument is passed to a {\LaTeX} macro |
473 \verb,\isamarkupXYZ, for any command \isakeyword{XYZ}. These macros |
463 \verb,\isamarkupXYZ, for command \isakeyword{XYZ}. These macros are |
474 are defined in \verb,isabelle.sty, according to the meaning given in |
464 defined in \verb,isabelle.sty, according to the meaning given in the |
475 the rightmost column above. |
465 rightmost column above. |
476 |
466 |
477 \medskip The following source fragment illustrates structure markup |
467 \medskip The following source fragment illustrates structure markup |
478 of a theory. Note that {\LaTeX} labels may be included inside of |
468 of a theory. Note that {\LaTeX} labels may be included inside of |
479 section headings as well. |
469 section headings as well. |
480 |
470 |
499 subsection {\ttlbrace}* Main theorem {\ttback}label{\ttlbrace}sec:main-theorem{\ttrbrace} *{\ttrbrace} |
489 subsection {\ttlbrace}* Main theorem {\ttback}label{\ttlbrace}sec:main-theorem{\ttrbrace} *{\ttrbrace} |
500 |
490 |
501 theorem main: \dots |
491 theorem main: \dots |
502 |
492 |
503 end |
493 end |
504 \end{ttbox} |
494 \end{ttbox}\vspace{-\medskipamount} |
505 |
495 |
506 You may occasionally want to change the meaning of markup |
496 You may occasionally want to change the meaning of markup commands, |
507 commands, say via \verb,\renewcommand, in \texttt{root.tex}. For example, |
497 say via \verb,\renewcommand, in \texttt{root.tex}. For example, |
508 \verb,\isamarkupheader, is a good candidate for some tuning. |
498 \verb,\isamarkupheader, is a good candidate for some tuning. We |
509 We could |
499 could move it up in the hierarchy to become \verb,\chapter,. |
510 move it up in the hierarchy to become \verb,\chapter,. |
|
511 |
500 |
512 \begin{verbatim} |
501 \begin{verbatim} |
513 \renewcommand{\isamarkupheader}[1]{\chapter{#1}} |
502 \renewcommand{\isamarkupheader}[1]{\chapter{#1}} |
514 \end{verbatim} |
503 \end{verbatim} |
515 |
504 |
516 \noindent Now we must change the |
505 \noindent Now we must change the document class given in |
517 document class given in \texttt{root.tex} to something that supports |
506 \texttt{root.tex} to something that supports chapters. A suitable |
518 chapters. A suitable command is |
507 command is \verb,\documentclass{report},. |
519 \verb,\documentclass{report},. |
|
520 |
508 |
521 \medskip The {\LaTeX} macro \verb,\isabellecontext, is maintained to |
509 \medskip The {\LaTeX} macro \verb,\isabellecontext, is maintained to |
522 hold the name of the current theory context. This is particularly |
510 hold the name of the current theory context. This is particularly |
523 useful for document headings: |
511 useful for document headings: |
524 |
512 |
585 |
573 |
586 \begin{verbatim} |
574 \begin{verbatim} |
587 \renewcommand{\isastyletxt}{\isastyletext} |
575 \renewcommand{\isastyletxt}{\isastyletext} |
588 \end{verbatim} |
576 \end{verbatim} |
589 |
577 |
590 \medskip The $text$ part of these markup commands |
578 \medskip The $text$ part of Isabelle markup commands essentially |
591 essentially inserts \emph{quoted material} into a |
579 inserts \emph{quoted material} into a formal text, mainly for |
592 formal text, mainly for instruction of the reader. An |
580 instruction of the reader. An \bfindex{antiquotation} is again a |
593 \bfindex{antiquotation} is again a formal object embedded into such |
581 formal object embedded into such an informal portion. The |
594 an informal portion. The interpretation of antiquotations is |
582 interpretation of antiquotations is limited to some well-formedness |
595 limited to some well-formedness checks, with the result being pretty |
583 checks, with the result being pretty printed to the resulting |
596 printed to the resulting document. Quoted text blocks together with |
584 document. Quoted text blocks together with antiquotations provide |
597 antiquotations provide an attractive means of referring to formal |
585 an attractive means of referring to formal entities, with good |
598 entities, with good confidence in getting the technical details |
586 confidence in getting the technical details right (especially syntax |
599 right (especially syntax and types). |
587 and types). |
600 |
588 |
601 The general syntax of antiquotations is as follows: |
589 The general syntax of antiquotations is as follows: |
602 \texttt{{\at}{\ttlbrace}$name$ $arguments${\ttrbrace}}, or |
590 \texttt{{\at}{\ttlbrace}$name$ $arguments${\ttrbrace}}, or |
603 \texttt{{\at}{\ttlbrace}$name$ [$options$] $arguments${\ttrbrace}} |
591 \texttt{{\at}{\ttlbrace}$name$ [$options$] $arguments${\ttrbrace}} |
604 for a comma-separated list of options consisting of a $name$ or |
592 for a comma-separated list of options consisting of a $name$ or |
605 \texttt{$name$=$value$}. The syntax of $arguments$ depends on the |
593 \texttt{$name$=$value$} each. The syntax of $arguments$ depends on |
606 kind of antiquotation, it generally follows the same conventions for |
594 the kind of antiquotation, it generally follows the same conventions |
607 types, terms, or theorems as in the formal part of a theory. |
595 for types, terms, or theorems as in the formal part of a theory. |
608 |
596 |
609 \medskip This sentence demonstrates quotations and antiquotations: |
597 \medskip This sentence demonstrates quotations and antiquotations: |
610 @{term "%x y. x"} is a well-typed term. |
598 @{term "%x y. x"} is a well-typed term. |
611 |
599 |
612 \medskip\noindent The output above was produced as follows: |
600 \medskip\noindent The output above was produced as follows: |
613 \begin{ttbox} |
601 \begin{ttbox} |
614 text {\ttlbrace}* |
602 text {\ttlbrace}* |
615 This sentence demonstrates quotations and antiquotations: |
603 This sentence demonstrates quotations and antiquotations: |
616 {\at}{\ttlbrace}term "%x y. x"{\ttrbrace} is a well-typed term. |
604 {\at}{\ttlbrace}term "%x y. x"{\ttrbrace} is a well-typed term. |
617 *{\ttrbrace} |
605 *{\ttrbrace} |
618 \end{ttbox} |
606 \end{ttbox}\vspace{-\medskipamount} |
619 |
607 |
620 The notational change from the ASCII character~\verb,%, to the |
608 The notational change from the ASCII character~\verb,%, to the |
621 symbol~@{text \<lambda>} reveals that Isabelle printed this term, |
609 symbol~@{text \<lambda>} reveals that Isabelle printed this term, after |
622 after parsing and type-checking. Document preparation |
610 parsing and type-checking. Document preparation enables symbolic |
623 enables symbolic output by default. |
611 output by default. |
624 |
612 |
625 \medskip The next example includes an option to modify Isabelle's |
613 \medskip The next example includes an option to modify Isabelle's |
626 \verb,show_types, flag. The antiquotation |
614 \verb,show_types, flag. The antiquotation |
627 \texttt{{\at}}\verb,{term [show_types] "%x y. x"}, produces |
615 \texttt{{\at}}\verb,{term [show_types] "%x y. x"}, produces the |
628 the output @{term [show_types] "%x y. x"}. |
616 output @{term [show_types] "%x y. x"}. Type inference has figured |
629 Type inference has figured out the most |
617 out the most general typings in the present theory context. Terms |
630 general typings in the present theory context. Terms |
618 may acquire different typings due to constraints imposed by their |
631 may acquire different typings due to constraints imposed |
619 environment; within a proof, for example, variables are given the |
632 by their environment; within a proof, for example, variables are given |
620 same types as they have in the main goal statement. |
633 the same types as they have in the main goal statement. |
|
634 |
621 |
635 \medskip Several further kinds of antiquotations and options are |
622 \medskip Several further kinds of antiquotations and options are |
636 available \cite{isabelle-sys}. Here are a few commonly used |
623 available \cite{isabelle-sys}. Here are a few commonly used |
637 combinations: |
624 combinations: |
638 |
625 |
680 straightforward generalization of ASCII characters. While Isabelle |
668 straightforward generalization of ASCII characters. While Isabelle |
681 does not impose any interpretation of the infinite collection of |
669 does not impose any interpretation of the infinite collection of |
682 named symbols, {\LaTeX} documents use canonical glyphs for certain |
670 named symbols, {\LaTeX} documents use canonical glyphs for certain |
683 standard symbols \cite[appendix~A]{isabelle-sys}. |
671 standard symbols \cite[appendix~A]{isabelle-sys}. |
684 |
672 |
685 The {\LaTeX} code produced from Isabelle text follows a |
673 The {\LaTeX} code produced from Isabelle text follows a simple |
686 simple scheme. You can tune the final appearance by |
674 scheme. You can tune the final appearance by redefining certain |
687 redefining certain macros, say in \texttt{root.tex} of the document. |
675 macros, say in \texttt{root.tex} of the document. |
688 |
676 |
689 \begin{enumerate} |
677 \begin{enumerate} |
690 |
678 |
691 \item 7-bit ASCII characters: letters \texttt{A\dots Z} and |
679 \item 7-bit ASCII characters: letters \texttt{A\dots Z} and |
692 \texttt{a\dots z} are output directly, digits are passed as an |
680 \texttt{a\dots z} are output directly, digits are passed as an |
693 argument to the \verb,\isadigit, macro, other characters are |
681 argument to the \verb,\isadigit, macro, other characters are |
694 replaced by specifically named macros of the form |
682 replaced by specifically named macros of the form |
695 \verb,\isacharXYZ,. |
683 \verb,\isacharXYZ,. |
696 |
684 |
697 \item Named symbols: \verb,\,\verb,<,$XYZ$\verb,>, is turned into |
685 \item Named symbols: \verb,\,\verb,<XYZ>, is turned into |
698 \verb,{\isasym,$XYZ$\verb,},; note the additional braces. |
686 \verb,{\isasymXYZ},; note the additional braces. |
699 |
687 |
700 \item Named control symbols: \verb,\,\verb,<^,$XYZ$\verb,>, is |
688 \item Named control symbols: \verb,\,\verb,<^XYZ>, is turned into |
701 turned into \verb,\isactrl,$XYZ$; subsequent symbols may act as |
689 \verb,\isactrlXYZ,; subsequent symbols may act as arguments if the |
702 arguments if the corresponding macro is defined accordingly. |
690 control macro is defined accordingly. |
703 |
691 |
704 \end{enumerate} |
692 \end{enumerate} |
705 |
693 |
706 You may occasionally wish to give new {\LaTeX} interpretations of |
694 You may occasionally wish to give new {\LaTeX} interpretations of |
707 named symbols. This merely requires an appropriate definition of |
695 named symbols. This merely requires an appropriate definition of |
708 \verb,\isasym,$XYZ$\verb,, for \verb,\,\verb,<,$XYZ$\verb,>, (see |
696 \verb,\isasymXYZ,, for \verb,\,\verb,<XYZ>, (see |
709 \texttt{isabelle.sty} for working examples). Control symbols are |
697 \texttt{isabelle.sty} for working examples). Control symbols are |
710 slightly more difficult to get right, though. |
698 slightly more difficult to get right, though. |
711 |
699 |
712 \medskip The \verb,\isabellestyle, macro provides a high-level |
700 \medskip The \verb,\isabellestyle, macro provides a high-level |
713 interface to tune the general appearance of individual symbols. For |
701 interface to tune the general appearance of individual symbols. For |
735 Alternatively, one may tune the theory loading process in |
723 Alternatively, one may tune the theory loading process in |
736 \texttt{ROOT.ML} itself: traversal of the theory dependency graph |
724 \texttt{ROOT.ML} itself: traversal of the theory dependency graph |
737 may be fine-tuned by adding \verb,use_thy, invocations, although |
725 may be fine-tuned by adding \verb,use_thy, invocations, although |
738 topological sorting still has to be observed. Moreover, the ML |
726 topological sorting still has to be observed. Moreover, the ML |
739 operator \verb,no_document, temporarily disables document generation |
727 operator \verb,no_document, temporarily disables document generation |
740 while executing a theory loader command; its usage is like this: |
728 while executing a theory loader command. Its usage is like this: |
741 |
729 |
742 \begin{verbatim} |
730 \begin{verbatim} |
743 no_document use_thy "T"; |
731 no_document use_thy "T"; |
744 \end{verbatim} |
732 \end{verbatim} |
745 |
733 |
746 \medskip Theory output may be suppressed more selectively. |
734 \medskip Theory output may be suppressed more selectively. Research |
747 Research articles and slides usually do not include the |
735 articles and slides usually do not include the formal content in |
748 formal content in full. In order to delimit \bfindex{ignored |
736 full. Delimiting \bfindex{ignored material} by the special source |
749 material}, special source comments |
737 comments \verb,(,\verb,*,\verb,<,\verb,*,\verb,), and |
750 \verb,(,\verb,*,\verb,<,\verb,*,\verb,), and |
738 \verb,(,\verb,*,\verb,>,\verb,*,\verb,), tells the document |
751 \verb,(,\verb,*,\verb,>,\verb,*,\verb,), may be included in the |
739 preparation system to suppress these parts; the formal checking of |
752 text. Only document preparation is affected; the formal |
740 the theory is unchanged. |
753 checking of the theory is unchanged. |
741 |
754 |
742 In this example, we hide a theory's \isakeyword{theory} and |
755 In this example, we suppress a theory's uninteresting |
743 \isakeyword{end} brackets: |
756 \isakeyword{theory} and \isakeyword{end} brackets: |
|
757 |
744 |
758 \medskip |
745 \medskip |
759 |
746 |
760 \begin{tabular}{l} |
747 \begin{tabular}{l} |
761 \verb,(,\verb,*,\verb,<,\verb,*,\verb,), \\ |
748 \verb,(,\verb,*,\verb,<,\verb,*,\verb,), \\ |
784 \begin{verbatim} |
771 \begin{verbatim} |
785 by (auto(*<*)simp add: int_less_le(*>*)) |
772 by (auto(*<*)simp add: int_less_le(*>*)) |
786 \end{verbatim} |
773 \end{verbatim} |
787 %(* |
774 %(* |
788 |
775 |
789 \medskip Suppressing portions of printed text demands care. |
776 \medskip Suppressing portions of printed text demands care. You |
790 You should not misrepresent |
777 should not misrepresent the underlying theory development. It is |
791 the underlying theory development. It is |
778 easy to invalidate the visible text by hiding references to |
792 easy to invalidate the visible text by hiding |
779 questionable axioms. |
793 references to questionable axioms. |
|
794 |
780 |
795 Authentic reports of Isabelle/Isar theories, say as part of a |
781 Authentic reports of Isabelle/Isar theories, say as part of a |
796 library, should suppress nothing. |
782 library, should suppress nothing. Other users may need the full |
797 Other users may need the full information for their own derivative |
783 information for their own derivative work. If a particular |
798 work. If a particular formalization appears inadequate for general |
784 formalization appears inadequate for general public coverage, it is |
799 public coverage, it is often more appropriate to think of a better |
785 often more appropriate to think of a better way in the first place. |
800 way in the first place. |
|
801 |
786 |
802 \medskip Some technical subtleties of the |
787 \medskip Some technical subtleties of the |
803 \verb,(,\verb,*,\verb,<,\verb,*,\verb,),~\verb,(,\verb,*,\verb,>,\verb,*,\verb,), |
788 \verb,(,\verb,*,\verb,<,\verb,*,\verb,),~\verb,(,\verb,*,\verb,>,\verb,*,\verb,), |
804 elements need to be kept in mind, too --- the system performs few |
789 elements need to be kept in mind, too --- the system performs few |
805 sanity checks here. Arguments of markup commands and formal |
790 sanity checks here. Arguments of markup commands and formal |