9 of \bfindex{mixfix annotations}. Associated with any kind of |
9 of \bfindex{mixfix annotations}. Associated with any kind of |
10 constant declaration, mixfixes affect both the grammar productions |
10 constant declaration, mixfixes affect both the grammar productions |
11 for the parser and output templates for the pretty printer. |
11 for the parser and output templates for the pretty printer. |
12 |
12 |
13 In full generality, parser and pretty printer configuration is a |
13 In full generality, parser and pretty printer configuration is a |
14 subtle affair~@{cite "isabelle-isar-ref"}. Your syntax specifications need |
14 subtle affair~\<^cite>\<open>"isabelle-isar-ref"\<close>. Your syntax specifications need |
15 to interact properly with the existing setup of Isabelle/Pure and |
15 to interact properly with the existing setup of Isabelle/Pure and |
16 Isabelle/HOL\@. To avoid creating ambiguities with existing |
16 Isabelle/HOL\@. To avoid creating ambiguities with existing |
17 elements, it is particularly important to give new syntactic |
17 elements, it is particularly important to give new syntactic |
18 constructs the right precedence. |
18 constructs the right precedence. |
19 |
19 |
104 interpretation is left to further front-end tools. For example, the |
104 interpretation is left to further front-end tools. For example, the |
105 Isabelle document processor (see \S\ref{sec:document-preparation}) |
105 Isabelle document processor (see \S\ref{sec:document-preparation}) |
106 display the \verb,\,\verb,<forall>, symbol as~\<open>\<forall>\<close>. |
106 display the \verb,\,\verb,<forall>, symbol as~\<open>\<forall>\<close>. |
107 |
107 |
108 A list of standard Isabelle symbols is given in |
108 A list of standard Isabelle symbols is given in |
109 @{cite "isabelle-isar-ref"}. You may introduce your own |
109 \<^cite>\<open>"isabelle-isar-ref"\<close>. You may introduce your own |
110 interpretation of further symbols by configuring the appropriate |
110 interpretation of further symbols by configuring the appropriate |
111 front-end tool accordingly, e.g.\ by defining certain {\LaTeX} |
111 front-end tool accordingly, e.g.\ by defining certain {\LaTeX} |
112 macros (see also \S\ref{sec:doc-prep-symbols}). There are also a |
112 macros (see also \S\ref{sec:doc-prep-symbols}). There are also a |
113 few predefined control symbols, such as \verb,\,\verb,<^sub>, and |
113 few predefined control symbols, such as \verb,\,\verb,<^sub>, and |
114 \verb,\,\verb,<^sup>, for sub- and superscript of the subsequent |
114 \verb,\,\verb,<^sup>, for sub- and superscript of the subsequent |
144 setup \<open>Sign.local_path\<close> |
144 setup \<open>Sign.local_path\<close> |
145 (*>*) |
145 (*>*) |
146 |
146 |
147 text \<open> |
147 text \<open> |
148 It is possible to provide alternative syntax forms |
148 It is possible to provide alternative syntax forms |
149 through the \bfindex{print mode} concept~@{cite "isabelle-isar-ref"}. By |
149 through the \bfindex{print mode} concept~\<^cite>\<open>"isabelle-isar-ref"\<close>. By |
150 convention, the mode of ``$xsymbols$'' is enabled whenever |
150 convention, the mode of ``$xsymbols$'' is enabled whenever |
151 Proof~General's X-Symbol mode or {\LaTeX} output is active. Now |
151 Proof~General's X-Symbol mode or {\LaTeX} output is active. Now |
152 consider the following hybrid declaration of \<open>xor\<close>: |
152 consider the following hybrid declaration of \<open>xor\<close>: |
153 \<close> |
153 \<close> |
154 |
154 |
178 |
178 |
179 subsection \<open>Prefix Annotations\<close> |
179 subsection \<open>Prefix Annotations\<close> |
180 |
180 |
181 text \<open> |
181 text \<open> |
182 Prefix syntax annotations\index{prefix annotation} are another form |
182 Prefix syntax annotations\index{prefix annotation} are another form |
183 of mixfixes @{cite "isabelle-isar-ref"}, without any template arguments or |
183 of mixfixes \<^cite>\<open>"isabelle-isar-ref"\<close>, without any template arguments or |
184 priorities --- just some literal syntax. The following example |
184 priorities --- just some literal syntax. The following example |
185 associates common symbols with the constructors of a datatype. |
185 associates common symbols with the constructors of a datatype. |
186 \<close> |
186 \<close> |
187 |
187 |
188 datatype currency = |
188 datatype currency = |
252 large hierarchies of concepts. Abbreviations do not replace |
252 large hierarchies of concepts. Abbreviations do not replace |
253 definitions. |
253 definitions. |
254 |
254 |
255 Abbreviations are a simplified form of the general concept of |
255 Abbreviations are a simplified form of the general concept of |
256 \emph{syntax translations}; even heavier transformations may be |
256 \emph{syntax translations}; even heavier transformations may be |
257 written in ML @{cite "isabelle-isar-ref"}. |
257 written in ML \<^cite>\<open>"isabelle-isar-ref"\<close>. |
258 \<close> |
258 \<close> |
259 |
259 |
260 |
260 |
261 section \<open>Document Preparation \label{sec:document-preparation}\<close> |
261 section \<open>Document Preparation \label{sec:document-preparation}\<close> |
262 |
262 |
340 \medskip The easiest way to manage Isabelle sessions is via |
340 \medskip The easiest way to manage Isabelle sessions is via |
341 \texttt{isabelle mkroot} (to generate an initial session source |
341 \texttt{isabelle mkroot} (to generate an initial session source |
342 setup) and \texttt{isabelle build} (to run sessions as specified in |
342 setup) and \texttt{isabelle build} (to run sessions as specified in |
343 the corresponding \texttt{ROOT} file). These Isabelle tools are |
343 the corresponding \texttt{ROOT} file). These Isabelle tools are |
344 described in further detail in the \emph{Isabelle System Manual} |
344 described in further detail in the \emph{Isabelle System Manual} |
345 @{cite "isabelle-system"}. |
345 \<^cite>\<open>"isabelle-system"\<close>. |
346 |
346 |
347 For example, a new session \texttt{MySession} (with document |
347 For example, a new session \texttt{MySession} (with document |
348 preparation) may be produced as follows: |
348 preparation) may be produced as follows: |
349 |
349 |
350 \begin{verbatim} |
350 \begin{verbatim} |
401 |
401 |
402 \medskip Any additional files for the {\LaTeX} stage go into the |
402 \medskip Any additional files for the {\LaTeX} stage go into the |
403 \texttt{MySession/document} directory as well. In particular, |
403 \texttt{MySession/document} directory as well. In particular, |
404 adding a file named \texttt{root.bib} causes an automatic run of |
404 adding a file named \texttt{root.bib} causes an automatic run of |
405 \texttt{bibtex} to process a bibliographic database; see also |
405 \texttt{bibtex} to process a bibliographic database; see also |
406 \texttt{isabelle document} @{cite "isabelle-system"}. |
406 \texttt{isabelle document} \<^cite>\<open>"isabelle-system"\<close>. |
407 |
407 |
408 \medskip Any failure of the document preparation phase in an |
408 \medskip Any failure of the document preparation phase in an |
409 Isabelle batch session leaves the generated sources in their target |
409 Isabelle batch session leaves the generated sources in their target |
410 location, identified by the accompanying error message. This lets |
410 location, identified by the accompanying error message. This lets |
411 you trace {\LaTeX} problems with the generated files at hand. |
411 you trace {\LaTeX} problems with the generated files at hand. |
471 text that are associated with formal Isabelle/Isar commands |
471 text that are associated with formal Isabelle/Isar commands |
472 (\bfindex{marginal comments}), or as standalone paragraphs within a |
472 (\bfindex{marginal comments}), or as standalone paragraphs within a |
473 theory or proof context (\bfindex{text blocks}). |
473 theory or proof context (\bfindex{text blocks}). |
474 |
474 |
475 \medskip Marginal comments are part of each command's concrete |
475 \medskip Marginal comments are part of each command's concrete |
476 syntax @{cite "isabelle-isar-ref"}; the common form is ``\verb,--,~$text$'' |
476 syntax \<^cite>\<open>"isabelle-isar-ref"\<close>; the common form is ``\verb,--,~$text$'' |
477 where $text$ is delimited by \verb,",\<open>\<dots>\<close>\verb,", or |
477 where $text$ is delimited by \verb,",\<open>\<dots>\<close>\verb,", or |
478 \verb,{,\verb,*,~\<open>\<dots>\<close>~\verb,*,\verb,}, as before. Multiple |
478 \verb,{,\verb,*,~\<open>\<dots>\<close>~\verb,*,\verb,}, as before. Multiple |
479 marginal comments may be given at the same time. Here is a simple |
479 marginal comments may be given at the same time. Here is a simple |
480 example: |
480 example: |
481 \<close> |
481 \<close> |
550 may acquire different typings due to constraints imposed by their |
550 may acquire different typings due to constraints imposed by their |
551 environment; within a proof, for example, variables are given the |
551 environment; within a proof, for example, variables are given the |
552 same types as they have in the main goal statement. |
552 same types as they have in the main goal statement. |
553 |
553 |
554 \medskip Several further kinds of antiquotations and options are |
554 \medskip Several further kinds of antiquotations and options are |
555 available @{cite "isabelle-isar-ref"}. Here are a few commonly used |
555 available \<^cite>\<open>"isabelle-isar-ref"\<close>. Here are a few commonly used |
556 combinations: |
556 combinations: |
557 |
557 |
558 \medskip |
558 \medskip |
559 |
559 |
560 \begin{tabular}{ll} |
560 \begin{tabular}{ll} |
598 As has been pointed out before (\S\ref{sec:syntax-symbols}), |
598 As has been pointed out before (\S\ref{sec:syntax-symbols}), |
599 Isabelle symbols are the smallest syntactic entities --- a |
599 Isabelle symbols are the smallest syntactic entities --- a |
600 straightforward generalization of ASCII characters. While Isabelle |
600 straightforward generalization of ASCII characters. While Isabelle |
601 does not impose any interpretation of the infinite collection of |
601 does not impose any interpretation of the infinite collection of |
602 named symbols, {\LaTeX} documents use canonical glyphs for certain |
602 named symbols, {\LaTeX} documents use canonical glyphs for certain |
603 standard symbols @{cite "isabelle-isar-ref"}. |
603 standard symbols \<^cite>\<open>"isabelle-isar-ref"\<close>. |
604 |
604 |
605 The {\LaTeX} code produced from Isabelle text follows a simple |
605 The {\LaTeX} code produced from Isabelle text follows a simple |
606 scheme. You can tune the final appearance by redefining certain |
606 scheme. You can tune the final appearance by redefining certain |
607 macros, say in \texttt{root.tex} of the document. |
607 macros, say in \texttt{root.tex} of the document. |
608 |
608 |
688 Tags observe the structure of proofs; adjacent commands with the |
688 Tags observe the structure of proofs; adjacent commands with the |
689 same tag are joined into a single region. The Isabelle document |
689 same tag are joined into a single region. The Isabelle document |
690 preparation system allows the user to specify how to interpret a |
690 preparation system allows the user to specify how to interpret a |
691 tagged region, in order to keep, drop, or fold the corresponding |
691 tagged region, in order to keep, drop, or fold the corresponding |
692 parts of the document. See the \emph{Isabelle System Manual} |
692 parts of the document. See the \emph{Isabelle System Manual} |
693 @{cite "isabelle-system"} for further details, especially on |
693 \<^cite>\<open>"isabelle-system"\<close> for further details, especially on |
694 \texttt{isabelle build} and \texttt{isabelle document}. |
694 \texttt{isabelle build} and \texttt{isabelle document}. |
695 |
695 |
696 Ignored material is specified by delimiting the original formal |
696 Ignored material is specified by delimiting the original formal |
697 source with special source comments |
697 source with special source comments |
698 \verb,(,\verb,*,\verb,<,\verb,*,\verb,), and |
698 \verb,(,\verb,*,\verb,<,\verb,*,\verb,), and |