| author | wenzelm | 
| Tue, 01 Mar 2016 10:32:55 +0100 | |
| changeset 62480 | f2e8984adef7 | 
| parent 62274 | 199f4d6dab0a | 
| child 62912 | 745d31e63c21 | 
| permissions | -rw-r--r-- | 
| 61656 | 1  | 
(*:maxLineLen=78:*)  | 
2  | 
||
| 27043 | 3  | 
theory Document_Preparation  | 
| 42651 | 4  | 
imports Base Main  | 
| 27043 | 5  | 
begin  | 
6  | 
||
| 58618 | 7  | 
chapter \<open>Document preparation \label{ch:document-prep}\<close>
 | 
| 27043 | 8  | 
|
| 61624 | 9  | 
text \<open>  | 
10  | 
Isabelle/Isar provides a simple document preparation system based on  | 
|
11  | 
  {PDF-\LaTeX}, with support for hyperlinks and bookmarks within that format.
 | 
|
12  | 
This allows to produce papers, books, theses etc.\ from Isabelle theory  | 
|
13  | 
sources.  | 
|
| 27043 | 14  | 
|
| 61624 | 15  | 
  {\LaTeX} output is generated while processing a \<^emph>\<open>session\<close> in batch mode, as
 | 
16  | 
  explained in the \<^emph>\<open>The Isabelle System Manual\<close> @{cite "isabelle-system"}.
 | 
|
17  | 
The main Isabelle tools to get started with document preparation are  | 
|
18  | 
  @{tool_ref mkroot} and @{tool_ref build}.
 | 
|
| 27043 | 19  | 
|
| 61624 | 20  | 
  The classic Isabelle/HOL tutorial @{cite "isabelle-hol-book"} also explains
 | 
21  | 
some aspects of theory presentation.  | 
|
22  | 
\<close>  | 
|
| 27043 | 23  | 
|
24  | 
||
| 58618 | 25  | 
section \<open>Markup commands \label{sec:markup}\<close>
 | 
| 27043 | 26  | 
|
| 58618 | 27  | 
text \<open>  | 
| 27043 | 28  | 
  \begin{matharray}{rcl}
 | 
| 61493 | 29  | 
    @{command_def "chapter"} & : & \<open>any \<rightarrow> any\<close> \\
 | 
30  | 
    @{command_def "section"} & : & \<open>any \<rightarrow> any\<close> \\
 | 
|
31  | 
    @{command_def "subsection"} & : & \<open>any \<rightarrow> any\<close> \\
 | 
|
32  | 
    @{command_def "subsubsection"} & : & \<open>any \<rightarrow> any\<close> \\
 | 
|
33  | 
    @{command_def "paragraph"} & : & \<open>any \<rightarrow> any\<close> \\
 | 
|
34  | 
    @{command_def "subparagraph"} & : & \<open>any \<rightarrow> any\<close> \\
 | 
|
35  | 
    @{command_def "text"} & : & \<open>any \<rightarrow> any\<close> \\
 | 
|
36  | 
    @{command_def "txt"} & : & \<open>any \<rightarrow> any\<close> \\
 | 
|
37  | 
    @{command_def "text_raw"} & : & \<open>any \<rightarrow> any\<close> \\
 | 
|
| 27043 | 38  | 
  \end{matharray}
 | 
39  | 
||
| 61624 | 40  | 
Markup commands provide a structured way to insert text into the document  | 
41  | 
  generated from a theory. Each markup command takes a single @{syntax text}
 | 
|
42  | 
  argument, which is passed as argument to a corresponding {\LaTeX} macro. The
 | 
|
43  | 
  default macros provided by @{file "~~/lib/texinputs/isabelle.sty"} can be
 | 
|
44  | 
  redefined according to the needs of the underlying document and {\LaTeX}
 | 
|
45  | 
styles.  | 
|
| 28747 | 46  | 
|
| 61624 | 47  | 
  Note that formal comments (\secref{sec:comments}) are similar to markup
 | 
48  | 
commands, but have a different status within Isabelle/Isar syntax.  | 
|
| 27043 | 49  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
50  | 
  @{rail \<open>
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
51  | 
    (@@{command chapter} | @@{command section} | @@{command subsection} |
 | 
| 61463 | 52  | 
      @@{command subsubsection} | @@{command paragraph} | @@{command subparagraph} |
 | 
53  | 
      @@{command text} | @@{command txt} | @@{command text_raw}) @{syntax text}
 | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
54  | 
\<close>}  | 
| 27043 | 55  | 
|
| 61624 | 56  | 
    \<^descr> @{command chapter}, @{command section}, @{command subsection} etc.\ mark
 | 
57  | 
section headings within the theory source. This works in any context, even  | 
|
58  | 
    before the initial @{command theory} command. The corresponding {\LaTeX}
 | 
|
59  | 
macros are \<^verbatim>\<open>\isamarkupchapter\<close>, \<^verbatim>\<open>\isamarkupsection\<close>,  | 
|
60  | 
\<^verbatim>\<open>\isamarkupsubsection\<close> etc.\  | 
|
| 27043 | 61  | 
|
| 61624 | 62  | 
    \<^descr> @{command text} and @{command txt} specify paragraphs of plain text.
 | 
63  | 
    This corresponds to a {\LaTeX} environment \<^verbatim>\<open>\begin{isamarkuptext}\<close> \<open>\<dots>\<close>
 | 
|
64  | 
    \<^verbatim>\<open>\end{isamarkuptext}\<close> etc.
 | 
|
| 27043 | 65  | 
|
| 61624 | 66  | 
    \<^descr> @{command text_raw} is similar to @{command text}, but without any
 | 
67  | 
    surrounding markup environment. This allows to inject arbitrary {\LaTeX}
 | 
|
68  | 
source into the generated document.  | 
|
| 27043 | 69  | 
|
| 61463 | 70  | 
All text passed to any of the above markup commands may refer to formal  | 
| 61624 | 71  | 
  entities via \<^emph>\<open>document antiquotations\<close>, see also \secref{sec:antiq}. These
 | 
72  | 
are interpreted in the present theory or proof context.  | 
|
| 27043 | 73  | 
|
| 61421 | 74  | 
\<^medskip>  | 
| 61624 | 75  | 
The proof markup commands closely resemble those for theory specifications,  | 
76  | 
  but have a different formal status and produce different {\LaTeX} macros.
 | 
|
| 58618 | 77  | 
\<close>  | 
| 27043 | 78  | 
|
79  | 
||
| 60286 | 80  | 
section \<open>Document antiquotations \label{sec:antiq}\<close>
 | 
| 27043 | 81  | 
|
| 58618 | 82  | 
text \<open>  | 
| 27043 | 83  | 
  \begin{matharray}{rcl}
 | 
| 61493 | 84  | 
    @{antiquotation_def "theory"} & : & \<open>antiquotation\<close> \\
 | 
85  | 
    @{antiquotation_def "thm"} & : & \<open>antiquotation\<close> \\
 | 
|
86  | 
    @{antiquotation_def "lemma"} & : & \<open>antiquotation\<close> \\
 | 
|
87  | 
    @{antiquotation_def "prop"} & : & \<open>antiquotation\<close> \\
 | 
|
88  | 
    @{antiquotation_def "term"} & : & \<open>antiquotation\<close> \\
 | 
|
89  | 
    @{antiquotation_def term_type} & : & \<open>antiquotation\<close> \\
 | 
|
90  | 
    @{antiquotation_def typeof} & : & \<open>antiquotation\<close> \\
 | 
|
91  | 
    @{antiquotation_def const} & : & \<open>antiquotation\<close> \\
 | 
|
92  | 
    @{antiquotation_def abbrev} & : & \<open>antiquotation\<close> \\
 | 
|
93  | 
    @{antiquotation_def typ} & : & \<open>antiquotation\<close> \\
 | 
|
94  | 
    @{antiquotation_def type} & : & \<open>antiquotation\<close> \\
 | 
|
95  | 
    @{antiquotation_def class} & : & \<open>antiquotation\<close> \\
 | 
|
96  | 
    @{antiquotation_def "text"} & : & \<open>antiquotation\<close> \\
 | 
|
97  | 
    @{antiquotation_def goals} & : & \<open>antiquotation\<close> \\
 | 
|
98  | 
    @{antiquotation_def subgoals} & : & \<open>antiquotation\<close> \\
 | 
|
99  | 
    @{antiquotation_def prf} & : & \<open>antiquotation\<close> \\
 | 
|
100  | 
    @{antiquotation_def full_prf} & : & \<open>antiquotation\<close> \\
 | 
|
101  | 
    @{antiquotation_def ML} & : & \<open>antiquotation\<close> \\
 | 
|
102  | 
    @{antiquotation_def ML_op} & : & \<open>antiquotation\<close> \\
 | 
|
103  | 
    @{antiquotation_def ML_type} & : & \<open>antiquotation\<close> \\
 | 
|
104  | 
    @{antiquotation_def ML_structure} & : & \<open>antiquotation\<close> \\
 | 
|
105  | 
    @{antiquotation_def ML_functor} & : & \<open>antiquotation\<close> \\
 | 
|
106  | 
    @{antiquotation_def emph} & : & \<open>antiquotation\<close> \\
 | 
|
107  | 
    @{antiquotation_def bold} & : & \<open>antiquotation\<close> \\
 | 
|
108  | 
    @{antiquotation_def verbatim} & : & \<open>antiquotation\<close> \\
 | 
|
109  | 
    @{antiquotation_def "file"} & : & \<open>antiquotation\<close> \\
 | 
|
110  | 
    @{antiquotation_def "url"} & : & \<open>antiquotation\<close> \\
 | 
|
111  | 
    @{antiquotation_def "cite"} & : & \<open>antiquotation\<close> \\
 | 
|
| 61494 | 112  | 
    @{command_def "print_antiquotations"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | 
| 27043 | 113  | 
  \end{matharray}
 | 
114  | 
||
| 61614 | 115  | 
The overall content of an Isabelle/Isar theory may alternate between formal  | 
116  | 
and informal text. The main body consists of formal specification and proof  | 
|
117  | 
  commands, interspersed with markup commands (\secref{sec:markup}) or
 | 
|
118  | 
  document comments (\secref{sec:comments}). The argument of markup commands
 | 
|
119  | 
quotes informal text to be printed in the resulting document, but may again  | 
|
120  | 
refer to formal entities via \<^emph>\<open>document antiquotations\<close>.  | 
|
| 27043 | 121  | 
|
| 61503 | 122  | 
  For example, embedding \<^verbatim>\<open>@{term [show_types] "f x = a + x"}\<close>
 | 
| 
28749
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
123  | 
within a text block makes  | 
| 
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
124  | 
  \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} appear in the final {\LaTeX} document.
 | 
| 
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
125  | 
|
| 61614 | 126  | 
Antiquotations usually spare the author tedious typing of logical entities  | 
127  | 
in full detail. Even more importantly, some degree of consistency-checking  | 
|
128  | 
between the main body of formal text and its informal explanation is  | 
|
129  | 
achieved, since terms and types appearing in antiquotations are checked  | 
|
130  | 
within the current theory or proof context.  | 
|
| 27043 | 131  | 
|
| 61614 | 132  | 
\<^medskip>  | 
133  | 
Antiquotations are in general written as  | 
|
134  | 
  \<^verbatim>\<open>@{\<close>\<open>name\<close>~\<^verbatim>\<open>[\<close>\<open>options\<close>\<^verbatim>\<open>]\<close>~\<open>arguments\<close>\<^verbatim>\<open>}\<close>. The short form
 | 
|
135  | 
  \<^verbatim>\<open>\\<close>\<^verbatim>\<open><^\<close>\<open>name\<close>\<^verbatim>\<open>>\<close>\<open>\<open>argument_content\<close>\<close> (without surrounding \<^verbatim>\<open>@{\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close>)
 | 
|
136  | 
works for a single argument that is a cartouche. A cartouche without special  | 
|
137  | 
decoration is equivalent to \<^verbatim>\<open>\<^cartouche>\<close>\<open>\<open>argument_content\<close>\<close>, which is  | 
|
138  | 
  equivalent to \<^verbatim>\<open>@{cartouche\<close>~\<open>\<open>argument_content\<close>\<close>\<^verbatim>\<open>}\<close>. The special name
 | 
|
| 61595 | 139  | 
  @{antiquotation_def cartouche} is defined in the context: Isabelle/Pure
 | 
140  | 
  introduces that as an alias to @{antiquotation_ref text} (see below).
 | 
|
141  | 
Consequently, \<open>\<open>foo_bar + baz \<le> bazar\<close>\<close> prints literal quasi-formal text  | 
|
| 61614 | 142  | 
(unchecked). A control symbol \<^verbatim>\<open>\\<close>\<^verbatim>\<open><^\<close>\<open>name\<close>\<^verbatim>\<open>>\<close> within the body text, but  | 
143  | 
  without a subsequent cartouche, is equivalent to \<^verbatim>\<open>@{\<close>\<open>name\<close>\<^verbatim>\<open>}\<close>.
 | 
|
| 61472 | 144  | 
|
| 61474 | 145  | 
\begingroup  | 
146  | 
  \def\isasymcontrolstart{\isatt{\isacharbackslash\isacharless\isacharcircum}}
 | 
|
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59783 
diff
changeset
 | 
147  | 
  @{rail \<open>
 | 
| 61472 | 148  | 
    @{syntax_def antiquotation}:
 | 
| 61474 | 149  | 
      '@{' antiquotation_body '}' |
 | 
| 
61491
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
150  | 
      '\<controlstart>' @{syntax_ref name} '>' @{syntax_ref cartouche} |
 | 
| 
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
151  | 
      @{syntax_ref cartouche}
 | 
| 
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
152  | 
;  | 
| 
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
153  | 
options: '[' (option * ',') ']'  | 
| 
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
154  | 
;  | 
| 
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
155  | 
    option: @{syntax name} | @{syntax name} '=' @{syntax name}
 | 
| 
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
156  | 
;  | 
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59783 
diff
changeset
 | 
157  | 
\<close>}  | 
| 61474 | 158  | 
\endgroup  | 
| 
59917
 
9830c944670f
more uniform "verbose" option to print name space;
 
wenzelm 
parents: 
59783 
diff
changeset
 | 
159  | 
|
| 61614 | 160  | 
Note that the syntax of antiquotations may \<^emph>\<open>not\<close> include source comments  | 
161  | 
  \<^verbatim>\<open>(*\<close>~\<open>\<dots>\<close>~\<^verbatim>\<open>*)\<close> nor verbatim text \<^verbatim>\<open>{*\<close>~\<open>\<dots>\<close>~\<^verbatim>\<open>*}\<close>.
 | 
|
| 
61491
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
162  | 
|
| 43618 | 163  | 
%% FIXME less monolithic presentation, move to individual sections!?  | 
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
164  | 
  @{rail \<open>
 | 
| 61472 | 165  | 
    @{syntax_def antiquotation_body}:
 | 
| 61614 | 166  | 
      (@@{antiquotation text} | @@{antiquotation cartouche} | @@{antiquotation theory_text})
 | 
167  | 
        options @{syntax text} |
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
168  | 
      @@{antiquotation theory} options @{syntax name} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
169  | 
      @@{antiquotation thm} options styles @{syntax thmrefs} |
 | 
| 42617 | 170  | 
      @@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} @{syntax method}? |
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
171  | 
      @@{antiquotation prop} options styles @{syntax prop} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
172  | 
      @@{antiquotation term} options styles @{syntax term} |
 | 
| 43618 | 173  | 
      @@{antiquotation (HOL) value} options styles @{syntax term} |
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
174  | 
      @@{antiquotation term_type} options styles @{syntax term} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
175  | 
      @@{antiquotation typeof} options styles @{syntax term} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
176  | 
      @@{antiquotation const} options @{syntax term} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
177  | 
      @@{antiquotation abbrev} options @{syntax term} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
178  | 
      @@{antiquotation typ} options @{syntax type} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
179  | 
      @@{antiquotation type} options @{syntax name} |
 | 
| 61623 | 180  | 
      @@{antiquotation class} options @{syntax name} |
 | 
181  | 
      (@@{antiquotation command} | @@{antiquotation method} | @@{antiquotation attribute})
 | 
|
182  | 
        options @{syntax name}
 | 
|
| 46261 | 183  | 
;  | 
184  | 
    @{syntax antiquotation}:
 | 
|
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
185  | 
      @@{antiquotation goals} options |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
186  | 
      @@{antiquotation subgoals} options |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
187  | 
      @@{antiquotation prf} options @{syntax thmrefs} |
 | 
| 
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
188  | 
      @@{antiquotation full_prf} options @{syntax thmrefs} |
 | 
| 
58069
 
0255436b3d85
more liberal embedded "text", which includes cartouches;
 
wenzelm 
parents: 
56594 
diff
changeset
 | 
189  | 
      @@{antiquotation ML} options @{syntax text} |
 | 
| 
 
0255436b3d85
more liberal embedded "text", which includes cartouches;
 
wenzelm 
parents: 
56594 
diff
changeset
 | 
190  | 
      @@{antiquotation ML_op} options @{syntax text} |
 | 
| 
 
0255436b3d85
more liberal embedded "text", which includes cartouches;
 
wenzelm 
parents: 
56594 
diff
changeset
 | 
191  | 
      @@{antiquotation ML_type} options @{syntax text} |
 | 
| 
 
0255436b3d85
more liberal embedded "text", which includes cartouches;
 
wenzelm 
parents: 
56594 
diff
changeset
 | 
192  | 
      @@{antiquotation ML_structure} options @{syntax text} |
 | 
| 
 
0255436b3d85
more liberal embedded "text", which includes cartouches;
 
wenzelm 
parents: 
56594 
diff
changeset
 | 
193  | 
      @@{antiquotation ML_functor} options @{syntax text} |
 | 
| 
61473
 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 
wenzelm 
parents: 
61472 
diff
changeset
 | 
194  | 
      @@{antiquotation emph} options @{syntax text} |
 | 
| 
 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 
wenzelm 
parents: 
61472 
diff
changeset
 | 
195  | 
      @@{antiquotation bold} options @{syntax text} |
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
58618 
diff
changeset
 | 
196  | 
      @@{antiquotation verbatim} options @{syntax text} |
 | 
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
197  | 
      @@{antiquotation "file"} options @{syntax name} |
 | 
| 54705 | 198  | 
      @@{antiquotation file_unchecked} options @{syntax name} |
 | 
| 58593 | 199  | 
      @@{antiquotation url} options @{syntax name} |
 | 
200  | 
      @@{antiquotation cite} options @{syntax cartouche}? (@{syntax name} + @'and')
 | 
|
| 27043 | 201  | 
;  | 
| 32892 | 202  | 
    styles: '(' (style + ',') ')'
 | 
| 32891 | 203  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
204  | 
    style: (@{syntax name} +)
 | 
| 
61473
 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 
wenzelm 
parents: 
61472 
diff
changeset
 | 
205  | 
;  | 
| 
 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 
wenzelm 
parents: 
61472 
diff
changeset
 | 
206  | 
    @@{command print_antiquotations} ('!'?)
 | 
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
207  | 
\<close>}  | 
| 27043 | 208  | 
|
| 61614 | 209  | 
  \<^descr> \<open>@{text s}\<close> prints uninterpreted source text \<open>s\<close>, i.e.\ inner syntax. This
 | 
210  | 
is particularly useful to print portions of text according to the Isabelle  | 
|
211  | 
document style, without demanding well-formedness, e.g.\ small pieces of  | 
|
212  | 
terms that should not be parsed or type-checked yet.  | 
|
| 
61491
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
213  | 
|
| 
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
214  | 
It is also possible to write this in the short form \<open>\<open>s\<close>\<close> without any  | 
| 
 
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
 
wenzelm 
parents: 
61479 
diff
changeset
 | 
215  | 
further decoration.  | 
| 27043 | 216  | 
|
| 61614 | 217  | 
  \<^descr> \<open>@{theory_text s}\<close> prints uninterpreted theory source text \<open>s\<close>, i.e.\
 | 
218  | 
outer syntax with command keywords and other tokens.  | 
|
| 27043 | 219  | 
|
| 61614 | 220  | 
  \<^descr> \<open>@{theory A}\<close> prints the name \<open>A\<close>, which is guaranteed to refer to a valid
 | 
221  | 
ancestor theory in the current context.  | 
|
222  | 
||
223  | 
  \<^descr> \<open>@{thm a\<^sub>1 \<dots> a\<^sub>n}\<close> prints theorems \<open>a\<^sub>1 \<dots> a\<^sub>n\<close>. Full fact expressions are
 | 
|
224  | 
  allowed here, including attributes (\secref{sec:syn-att}).
 | 
|
| 27043 | 225  | 
|
| 61493 | 226  | 
  \<^descr> \<open>@{prop \<phi>}\<close> prints a well-typed proposition \<open>\<phi>\<close>.
 | 
| 27043 | 227  | 
|
| 61614 | 228  | 
  \<^descr> \<open>@{lemma \<phi> by m}\<close> proves a well-typed proposition \<open>\<phi>\<close> by method \<open>m\<close> and
 | 
229  | 
prints the original \<open>\<phi>\<close>.  | 
|
| 27453 | 230  | 
|
| 61493 | 231  | 
  \<^descr> \<open>@{term t}\<close> prints a well-typed term \<open>t\<close>.
 | 
| 
43613
 
7afbaf5a406b
adding a minimalistic documentation of the value antiquotation in the Isar reference manual
 
bulwahn 
parents: 
42936 
diff
changeset
 | 
232  | 
|
| 61614 | 233  | 
  \<^descr> \<open>@{value t}\<close> evaluates a term \<open>t\<close> and prints its result, see also
 | 
234  | 
  @{command_ref (HOL) value}.
 | 
|
| 27043 | 235  | 
|
| 61614 | 236  | 
  \<^descr> \<open>@{term_type t}\<close> prints a well-typed term \<open>t\<close> annotated with its type.
 | 
| 
32898
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32892 
diff
changeset
 | 
237  | 
|
| 61614 | 238  | 
  \<^descr> \<open>@{typeof t}\<close> prints the type of a well-typed term \<open>t\<close>.
 | 
| 
32898
 
e871d897969c
term styles also cover antiquotations term_type and typeof
 
haftmann 
parents: 
32892 
diff
changeset
 | 
239  | 
|
| 61614 | 240  | 
  \<^descr> \<open>@{const c}\<close> prints a logical or syntactic constant \<open>c\<close>.
 | 
| 27043 | 241  | 
|
| 61614 | 242  | 
  \<^descr> \<open>@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}\<close> prints a constant abbreviation \<open>c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs\<close>
 | 
243  | 
as defined in the current context.  | 
|
| 
39305
 
d4fa19eb0822
'class' and 'type' are now antiquoations by default
 
haftmann 
parents: 
32898 
diff
changeset
 | 
244  | 
|
| 61493 | 245  | 
  \<^descr> \<open>@{typ \<tau>}\<close> prints a well-formed type \<open>\<tau>\<close>.
 | 
| 
39305
 
d4fa19eb0822
'class' and 'type' are now antiquoations by default
 
haftmann 
parents: 
32898 
diff
changeset
 | 
246  | 
|
| 61614 | 247  | 
  \<^descr> \<open>@{type \<kappa>}\<close> prints a (logical or syntactic) type constructor \<open>\<kappa>\<close>.
 | 
| 
39305
 
d4fa19eb0822
'class' and 'type' are now antiquoations by default
 
haftmann 
parents: 
32898 
diff
changeset
 | 
248  | 
|
| 61493 | 249  | 
  \<^descr> \<open>@{class c}\<close> prints a class \<open>c\<close>.
 | 
| 
39305
 
d4fa19eb0822
'class' and 'type' are now antiquoations by default
 
haftmann 
parents: 
32898 
diff
changeset
 | 
250  | 
|
| 61623 | 251  | 
  \<^descr> \<open>@{command name}\<close>, \<open>@{method name}\<close>, \<open>@{attribute name}\<close> print checked
 | 
252  | 
entities of the Isar language.  | 
|
253  | 
||
| 61614 | 254  | 
  \<^descr> \<open>@{goals}\<close> prints the current \<^emph>\<open>dynamic\<close> goal state. This is mainly for
 | 
255  | 
support of tactic-emulation scripts within Isar. Presentation of goal states  | 
|
256  | 
does not conform to the idea of human-readable proof documents!  | 
|
| 27043 | 257  | 
|
| 61614 | 258  | 
When explaining proofs in detail it is usually better to spell out the  | 
259  | 
reasoning via proper Isar proof commands, instead of peeking at the internal  | 
|
260  | 
machine configuration.  | 
|
| 27043 | 261  | 
|
| 61614 | 262  | 
  \<^descr> \<open>@{subgoals}\<close> is similar to \<open>@{goals}\<close>, but does not print the main goal.
 | 
| 27043 | 263  | 
|
| 61614 | 264  | 
  \<^descr> \<open>@{prf a\<^sub>1 \<dots> a\<^sub>n}\<close> prints the (compact) proof terms corresponding to the
 | 
265  | 
theorems \<open>a\<^sub>1 \<dots> a\<^sub>n\<close>. Note that this requires proof terms to be switched on  | 
|
266  | 
for the current logic session.  | 
|
| 27043 | 267  | 
|
| 61614 | 268  | 
  \<^descr> \<open>@{full_prf a\<^sub>1 \<dots> a\<^sub>n}\<close> is like \<open>@{prf a\<^sub>1 \<dots> a\<^sub>n}\<close>, but prints the full
 | 
269  | 
proof terms, i.e.\ also displays information omitted in the compact proof  | 
|
270  | 
term, which is denoted by ``\<open>_\<close>'' placeholders there.  | 
|
| 27043 | 271  | 
|
| 61614 | 272  | 
  \<^descr> \<open>@{ML s}\<close>, \<open>@{ML_op s}\<close>, \<open>@{ML_type s}\<close>, \<open>@{ML_structure s}\<close>, and
 | 
273  | 
  \<open>@{ML_functor s}\<close> check text \<open>s\<close> as ML value, infix operator, type,
 | 
|
274  | 
structure, and functor respectively. The source is printed verbatim.  | 
|
| 27043 | 275  | 
|
| 61614 | 276  | 
  \<^descr> \<open>@{emph s}\<close> prints document source recursively, with {\LaTeX} markup
 | 
277  | 
  \<^verbatim>\<open>\emph{\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close>.
 | 
|
| 
61473
 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 
wenzelm 
parents: 
61472 
diff
changeset
 | 
278  | 
|
| 61614 | 279  | 
  \<^descr> \<open>@{bold s}\<close> prints document source recursively, with {\LaTeX} markup
 | 
280  | 
  \<^verbatim>\<open>\textbf{\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close>.
 | 
|
| 
61473
 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 
wenzelm 
parents: 
61472 
diff
changeset
 | 
281  | 
|
| 61614 | 282  | 
  \<^descr> \<open>@{verbatim s}\<close> prints uninterpreted source text literally as ASCII
 | 
283  | 
characters, using some type-writer font style.  | 
|
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
58618 
diff
changeset
 | 
284  | 
|
| 61614 | 285  | 
  \<^descr> \<open>@{file path}\<close> checks that \<open>path\<close> refers to a file (or directory) and
 | 
286  | 
prints it verbatim.  | 
|
| 40801 | 287  | 
|
| 61614 | 288  | 
  \<^descr> \<open>@{file_unchecked path}\<close> is like \<open>@{file path}\<close>, but does not check the
 | 
289  | 
existence of the \<open>path\<close> within the file-system.  | 
|
| 54705 | 290  | 
|
| 61614 | 291  | 
  \<^descr> \<open>@{url name}\<close> produces markup for the given URL, which results in an
 | 
292  | 
active hyperlink within the text.  | 
|
| 
54702
 
3daeba5130f0
added document antiquotation @{url}, which produces formal markup for LaTeX and PIDE;
 
wenzelm 
parents: 
54346 
diff
changeset
 | 
293  | 
|
| 61614 | 294  | 
  \<^descr> \<open>@{cite name}\<close> produces a citation \<^verbatim>\<open>\cite{name}\<close> in {\LaTeX}, where the
 | 
295  | 
  name refers to some Bib{\TeX} database entry.
 | 
|
| 58593 | 296  | 
|
| 61614 | 297  | 
  The variant \<open>@{cite \<open>opt\<close> name}\<close> produces \<^verbatim>\<open>\cite[opt]{name}\<close> with some
 | 
298  | 
free-form optional argument. Multiple names are output with commas, e.g.  | 
|
299  | 
  \<open>@{cite foo \<AND> bar}\<close> becomes \<^verbatim>\<open>\cite{foo,bar}\<close>.
 | 
|
| 58593 | 300  | 
|
301  | 
  The {\LaTeX} macro name is determined by the antiquotation option
 | 
|
302  | 
  @{antiquotation_option_def cite_macro}, or the configuration option
 | 
|
| 61614 | 303  | 
  @{attribute cite_macro} in the context. For example, \<open>@{cite [cite_macro =
 | 
304  | 
  nocite] foobar}\<close> produces \<^verbatim>\<open>\nocite{foobar}\<close>.
 | 
|
| 
61473
 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 
wenzelm 
parents: 
61472 
diff
changeset
 | 
305  | 
|
| 61614 | 306  | 
  \<^descr> @{command "print_antiquotations"} prints all document antiquotations that
 | 
307  | 
are defined in the current context; the ``\<open>!\<close>'' option indicates extra  | 
|
308  | 
verbosity.  | 
|
| 58618 | 309  | 
\<close>  | 
| 27043 | 310  | 
|
| 
28749
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
311  | 
|
| 58618 | 312  | 
subsection \<open>Styled antiquotations\<close>  | 
| 
28749
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
313  | 
|
| 61624 | 314  | 
text \<open>  | 
315  | 
The antiquotations \<open>thm\<close>, \<open>prop\<close> and \<open>term\<close> admit an extra \<^emph>\<open>style\<close>  | 
|
316  | 
specification to modify the printed result. A style is specified by a name  | 
|
317  | 
with a possibly empty number of arguments; multiple styles can be sequenced  | 
|
318  | 
with commas. The following standard styles are available:  | 
|
| 27043 | 319  | 
|
| 61624 | 320  | 
\<^descr> \<open>lhs\<close> extracts the first argument of any application form with at least  | 
321  | 
two arguments --- typically meta-level or object-level equality, or any  | 
|
322  | 
other binary relation.  | 
|
| 27043 | 323  | 
|
| 61624 | 324  | 
\<^descr> \<open>rhs\<close> is like \<open>lhs\<close>, but extracts the second argument.  | 
| 27043 | 325  | 
|
| 61624 | 326  | 
\<^descr> \<open>concl\<close> extracts the conclusion \<open>C\<close> from a rule in Horn-clause normal form  | 
327  | 
\<open>A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C\<close>.  | 
|
| 27043 | 328  | 
|
| 61624 | 329  | 
\<^descr> \<open>prem\<close> \<open>n\<close> extract premise number \<open>n\<close> from from a rule in Horn-clause  | 
330  | 
normal form \<open>A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C\<close>.  | 
|
| 58618 | 331  | 
\<close>  | 
| 27043 | 332  | 
|
| 
28749
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
333  | 
|
| 58618 | 334  | 
subsection \<open>General options\<close>  | 
| 
28749
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
335  | 
|
| 61624 | 336  | 
text \<open>  | 
337  | 
The following options are available to tune the printed output of  | 
|
338  | 
antiquotations. Note that many of these coincide with system and  | 
|
| 54346 | 339  | 
configuration options of the same names.  | 
| 27043 | 340  | 
|
| 61624 | 341  | 
    \<^descr> @{antiquotation_option_def show_types}~\<open>= bool\<close> and
 | 
342  | 
    @{antiquotation_option_def show_sorts}~\<open>= bool\<close> control printing of
 | 
|
343  | 
explicit type and sort constraints.  | 
|
| 27043 | 344  | 
|
| 61624 | 345  | 
    \<^descr> @{antiquotation_option_def show_structs}~\<open>= bool\<close> controls printing of
 | 
346  | 
implicit structures.  | 
|
| 27043 | 347  | 
|
| 61624 | 348  | 
    \<^descr> @{antiquotation_option_def show_abbrevs}~\<open>= bool\<close> controls folding of
 | 
349  | 
abbreviations.  | 
|
| 
40879
 
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
 
wenzelm 
parents: 
40801 
diff
changeset
 | 
350  | 
|
| 61624 | 351  | 
    \<^descr> @{antiquotation_option_def names_long}~\<open>= bool\<close> forces names of types
 | 
352  | 
and constants etc.\ to be printed in their fully qualified internal form.  | 
|
| 27043 | 353  | 
|
| 61624 | 354  | 
    \<^descr> @{antiquotation_option_def names_short}~\<open>= bool\<close> forces names of types
 | 
355  | 
and constants etc.\ to be printed unqualified. Note that internalizing the  | 
|
356  | 
output again in the current context may well yield a different result.  | 
|
| 27043 | 357  | 
|
| 61624 | 358  | 
    \<^descr> @{antiquotation_option_def names_unique}~\<open>= bool\<close> determines whether the
 | 
359  | 
printed version of qualified names should be made sufficiently long to  | 
|
360  | 
avoid overlap with names declared further back. Set to \<open>false\<close> for more  | 
|
361  | 
concise output.  | 
|
| 27043 | 362  | 
|
| 61624 | 363  | 
    \<^descr> @{antiquotation_option_def eta_contract}~\<open>= bool\<close> prints terms in
 | 
364  | 
\<open>\<eta>\<close>-contracted form.  | 
|
| 27043 | 365  | 
|
| 61624 | 366  | 
    \<^descr> @{antiquotation_option_def display}~\<open>= bool\<close> indicates if the text is to
 | 
367  | 
be output as multi-line ``display material'', rather than a small piece of  | 
|
368  | 
text without line breaks (which is the default).  | 
|
| 27043 | 369  | 
|
| 61624 | 370  | 
In this mode the embedded entities are printed in the same style as the  | 
371  | 
main theory text.  | 
|
| 
28749
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
372  | 
|
| 61624 | 373  | 
    \<^descr> @{antiquotation_option_def break}~\<open>= bool\<close> controls line breaks in
 | 
374  | 
non-display material.  | 
|
| 27043 | 375  | 
|
| 61624 | 376  | 
    \<^descr> @{antiquotation_option_def quotes}~\<open>= bool\<close> indicates if the output
 | 
377  | 
should be enclosed in double quotes.  | 
|
| 27043 | 378  | 
|
| 61624 | 379  | 
    \<^descr> @{antiquotation_option_def mode}~\<open>= name\<close> adds \<open>name\<close> to the print mode
 | 
380  | 
    to be used for presentation. Note that the standard setup for {\LaTeX}
 | 
|
| 61997 | 381  | 
output is already present by default, with mode ``\<open>latex\<close>''.  | 
| 27043 | 382  | 
|
| 61624 | 383  | 
    \<^descr> @{antiquotation_option_def margin}~\<open>= nat\<close> and
 | 
384  | 
    @{antiquotation_option_def indent}~\<open>= nat\<close> change the margin or
 | 
|
385  | 
indentation for pretty printing of display material.  | 
|
386  | 
||
387  | 
    \<^descr> @{antiquotation_option_def goals_limit}~\<open>= nat\<close> determines the maximum
 | 
|
388  | 
number of subgoals to be printed (for goal-based antiquotation).  | 
|
| 27043 | 389  | 
|
| 61624 | 390  | 
    \<^descr> @{antiquotation_option_def source}~\<open>= bool\<close> prints the original source
 | 
391  | 
text of the antiquotation arguments, rather than its internal  | 
|
392  | 
    representation. Note that formal checking of @{antiquotation "thm"},
 | 
|
393  | 
    @{antiquotation "term"}, etc. is still enabled; use the @{antiquotation
 | 
|
394  | 
"text"} antiquotation for unchecked output.  | 
|
| 
28749
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
395  | 
|
| 61624 | 396  | 
Regular \<open>term\<close> and \<open>typ\<close> antiquotations with \<open>source = false\<close> involve a  | 
397  | 
full round-trip from the original source to an internalized logical entity  | 
|
398  | 
back to a source form, according to the syntax of the current context.  | 
|
399  | 
Thus the printed output is not under direct control of the author, it may  | 
|
400  | 
even fluctuate a bit as the underlying theory is changed later on.  | 
|
| 
28749
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
401  | 
|
| 61624 | 402  | 
    In contrast, @{antiquotation_option source}~\<open>= true\<close> admits direct
 | 
403  | 
printing of the given source text, with the desirable well-formedness  | 
|
404  | 
check in the background, but without modification of the printed text.  | 
|
| 
28749
 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 
wenzelm 
parents: 
28747 
diff
changeset
 | 
405  | 
|
| 61624 | 406  | 
For Boolean flags, ``\<open>name = true\<close>'' may be abbreviated as ``\<open>name\<close>''. All  | 
407  | 
of the above flags are disabled by default, unless changed specifically for  | 
|
408  | 
a logic session in the corresponding \<^verbatim>\<open>ROOT\<close> file.  | 
|
| 61458 | 409  | 
\<close>  | 
| 27043 | 410  | 
|
411  | 
||
| 62274 | 412  | 
section \<open>Markdown-like text structure\<close>  | 
413  | 
||
414  | 
text \<open>  | 
|
415  | 
  The markup commands @{command_ref text}, @{command_ref txt}, @{command_ref
 | 
|
416  | 
  text_raw} (\secref{sec:markup}) consist of plain text. Its internal
 | 
|
417  | 
structure consists of paragraphs and (nested) lists, using special Isabelle  | 
|
418  | 
symbols and some rules for indentation and blank lines. This quasi-visual  | 
|
419  | 
  format resembles \<^emph>\<open>Markdown\<close>\<^footnote>\<open>@{url "http://commonmark.org"}\<close>, but the
 | 
|
420  | 
full complexity of that notation is avoided.  | 
|
421  | 
||
422  | 
This is a summary of the main principles of minimal Markdown in Isabelle:  | 
|
423  | 
||
424  | 
\<^item> List items start with the following markers  | 
|
425  | 
\<^descr>[itemize:] \<^verbatim>\<open>\<^item>\<close>  | 
|
426  | 
\<^descr>[enumerate:] \<^verbatim>\<open>\<^enum>\<close>  | 
|
427  | 
\<^descr>[description:] \<^verbatim>\<open>\<^descr>\<close>  | 
|
428  | 
||
429  | 
\<^item> Adjacent list items with same indentation and same marker are grouped  | 
|
430  | 
into a single list.  | 
|
431  | 
||
432  | 
\<^item> Singleton blank lines separate paragraphs.  | 
|
433  | 
||
434  | 
\<^item> Multiple blank lines escape from the current list hierarchy.  | 
|
435  | 
||
436  | 
Notable differences to official Markdown:  | 
|
437  | 
||
438  | 
\<^item> Indentation of list items needs to match exactly.  | 
|
439  | 
||
440  | 
\<^item> Indentation is unlimited (official Markdown interprets four spaces as  | 
|
441  | 
block quote).  | 
|
442  | 
||
443  | 
\<^item> List items always consist of paragraphs --- there is no notion of  | 
|
444  | 
``tight'' list.  | 
|
445  | 
||
446  | 
\<^item> Section headings are expressed via Isar document markup commands  | 
|
447  | 
    (\secref{sec:markup}).
 | 
|
448  | 
||
449  | 
\<^item> URLs, font styles, other special content is expressed via antiquotations  | 
|
450  | 
    (\secref{sec:antiq}), usually with proper nesting of sub-languages via
 | 
|
451  | 
text cartouches.  | 
|
452  | 
\<close>  | 
|
453  | 
||
454  | 
||
| 58618 | 455  | 
section \<open>Markup via command tags \label{sec:tags}\<close>
 | 
| 27043 | 456  | 
|
| 61624 | 457  | 
text \<open>  | 
458  | 
Each Isabelle/Isar command may be decorated by additional presentation tags,  | 
|
459  | 
to indicate some modification in the way it is printed in the document.  | 
|
| 27043 | 460  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
461  | 
  @{rail \<open>
 | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
462  | 
    @{syntax_def tags}: ( tag * )
 | 
| 27043 | 463  | 
;  | 
| 
42596
 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 
wenzelm 
parents: 
40879 
diff
changeset
 | 
464  | 
    tag: '%' (@{syntax ident} | @{syntax string})
 | 
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
465  | 
\<close>}  | 
| 27043 | 466  | 
|
| 61624 | 467  | 
Some tags are pre-declared for certain classes of commands, serving as  | 
468  | 
default markup if no tags are given in the text:  | 
|
| 27043 | 469  | 
|
| 61421 | 470  | 
\<^medskip>  | 
| 27043 | 471  | 
  \begin{tabular}{ll}
 | 
| 61493 | 472  | 
\<open>theory\<close> & theory begin/end \\  | 
473  | 
\<open>proof\<close> & all proof commands \\  | 
|
474  | 
\<open>ML\<close> & all commands involving ML code \\  | 
|
| 27043 | 475  | 
  \end{tabular}
 | 
| 61421 | 476  | 
\<^medskip>  | 
| 27043 | 477  | 
|
| 61624 | 478  | 
  The Isabelle document preparation system @{cite "isabelle-system"} allows
 | 
479  | 
tagged command regions to be presented specifically, e.g.\ to fold proof  | 
|
480  | 
texts, or drop parts of the text completely.  | 
|
| 27043 | 481  | 
|
| 61624 | 482  | 
  For example ``@{command "by"}~\<open>%invisible auto\<close>'' causes that piece of proof
 | 
483  | 
to be treated as \<open>invisible\<close> instead of \<open>proof\<close> (the default), which may be  | 
|
484  | 
  shown or hidden depending on the document setup. In contrast, ``@{command
 | 
|
485  | 
"by"}~\<open>%visible auto\<close>'' forces this text to be shown invariably.  | 
|
| 27043 | 486  | 
|
| 61624 | 487  | 
Explicit tag specifications within a proof apply to all subsequent commands  | 
488  | 
  of the same level of nesting. For example, ``@{command "proof"}~\<open>%visible
 | 
|
489  | 
  \<dots>\<close>~@{command "qed"}'' forces the whole sub-proof to be typeset as \<open>visible\<close>
 | 
|
490  | 
(unless some of its parts are tagged differently).  | 
|
| 28750 | 491  | 
|
| 61421 | 492  | 
\<^medskip>  | 
| 61624 | 493  | 
Command tags merely produce certain markup environments for type-setting.  | 
494  | 
  The meaning of these is determined by {\LaTeX} macros, as defined in @{file
 | 
|
495  | 
"~~/lib/texinputs/isabelle.sty"} or by the document author. The Isabelle  | 
|
496  | 
document preparation tools also provide some high-level options to specify  | 
|
497  | 
the meaning of arbitrary tags to ``keep'', ``drop'', or ``fold'' the  | 
|
498  | 
corresponding parts of the text. Logic sessions may also specify ``document  | 
|
499  | 
versions'', where given tags are interpreted in some particular way. Again  | 
|
500  | 
  see @{cite "isabelle-system"} for further details.
 | 
|
| 58618 | 501  | 
\<close>  | 
| 27043 | 502  | 
|
503  | 
||
| 58618 | 504  | 
section \<open>Railroad diagrams\<close>  | 
| 42658 | 505  | 
|
| 58618 | 506  | 
text \<open>  | 
| 42658 | 507  | 
  \begin{matharray}{rcl}
 | 
| 61493 | 508  | 
    @{antiquotation_def "rail"} & : & \<open>antiquotation\<close> \\
 | 
| 42658 | 509  | 
  \end{matharray}
 | 
510  | 
||
| 55113 | 511  | 
  @{rail \<open>
 | 
| 59937 | 512  | 
    'rail' @{syntax text}
 | 
| 55113 | 513  | 
\<close>}  | 
| 42658 | 514  | 
|
| 61624 | 515  | 
  The @{antiquotation rail} antiquotation allows to include syntax diagrams
 | 
516  | 
  into Isabelle documents. {\LaTeX} requires the style file @{file
 | 
|
517  | 
"~~/lib/texinputs/railsetup.sty"}, which can be used via  | 
|
| 61503 | 518  | 
  \<^verbatim>\<open>\usepackage{railsetup}\<close> in \<^verbatim>\<open>root.tex\<close>, for example.
 | 
| 42658 | 519  | 
|
| 61624 | 520  | 
  The rail specification language is quoted here as Isabelle @{syntax string}
 | 
521  | 
  or text @{syntax "cartouche"}; it has its own grammar given below.
 | 
|
| 42658 | 522  | 
|
| 55120 | 523  | 
\begingroup  | 
| 61474 | 524  | 
  \def\isasymnewline{\isatt{\isacharbackslash\isacharless newline\isachargreater}}
 | 
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
525  | 
  @{rail \<open>
 | 
| 42658 | 526  | 
rule? + ';'  | 
527  | 
;  | 
|
528  | 
  rule: ((identifier | @{syntax antiquotation}) ':')? body
 | 
|
529  | 
;  | 
|
530  | 
body: concatenation + '|'  | 
|
531  | 
;  | 
|
532  | 
  concatenation: ((atom '?'?) +) (('*' | '+') atom?)?
 | 
|
533  | 
;  | 
|
534  | 
  atom: '(' body? ')' | identifier |
 | 
|
535  | 
    '@'? (string | @{syntax antiquotation}) |
 | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
536  | 
'\<newline>'  | 
| 
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
537  | 
\<close>}  | 
| 55120 | 538  | 
\endgroup  | 
| 42658 | 539  | 
|
| 61624 | 540  | 
  The lexical syntax of \<open>identifier\<close> coincides with that of @{syntax ident} in
 | 
541  | 
regular Isabelle syntax, but \<open>string\<close> uses single quotes instead of double  | 
|
542  | 
  quotes of the standard @{syntax string} category.
 | 
|
| 42658 | 543  | 
|
| 61624 | 544  | 
Each \<open>rule\<close> defines a formal language (with optional name), using a notation  | 
545  | 
that is similar to EBNF or regular expressions with recursion. The meaning  | 
|
546  | 
and visual appearance of these rail language elements is illustrated by the  | 
|
547  | 
following representative examples.  | 
|
| 42658 | 548  | 
|
| 61503 | 549  | 
\<^item> Empty \<^verbatim>\<open>()\<close>  | 
| 42658 | 550  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
551  | 
  @{rail \<open>()\<close>}
 | 
| 42658 | 552  | 
|
| 61503 | 553  | 
\<^item> Nonterminal \<^verbatim>\<open>A\<close>  | 
| 42658 | 554  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
555  | 
  @{rail \<open>A\<close>}
 | 
| 42658 | 556  | 
|
| 61503 | 557  | 
  \<^item> Nonterminal via Isabelle antiquotation \<^verbatim>\<open>@{syntax method}\<close>
 | 
| 42658 | 558  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
559  | 
  @{rail \<open>@{syntax method}\<close>}
 | 
| 42658 | 560  | 
|
| 61503 | 561  | 
\<^item> Terminal \<^verbatim>\<open>'xyz'\<close>  | 
| 42658 | 562  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
563  | 
  @{rail \<open>'xyz'\<close>}
 | 
| 42658 | 564  | 
|
| 61503 | 565  | 
\<^item> Terminal in keyword style \<^verbatim>\<open>@'xyz'\<close>  | 
| 42658 | 566  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
567  | 
  @{rail \<open>@'xyz'\<close>}
 | 
| 42658 | 568  | 
|
| 61503 | 569  | 
  \<^item> Terminal via Isabelle antiquotation \<^verbatim>\<open>@@{method rule}\<close>
 | 
| 42658 | 570  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
571  | 
  @{rail \<open>@@{method rule}\<close>}
 | 
| 42658 | 572  | 
|
| 61503 | 573  | 
\<^item> Concatenation \<^verbatim>\<open>A B C\<close>  | 
| 42658 | 574  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
575  | 
  @{rail \<open>A B C\<close>}
 | 
| 42658 | 576  | 
|
| 61503 | 577  | 
\<^item> Newline inside concatenation \<^verbatim>\<open>A B C \<newline> D E F\<close>  | 
| 42658 | 578  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
579  | 
  @{rail \<open>A B C \<newline> D E F\<close>}
 | 
| 42658 | 580  | 
|
| 61503 | 581  | 
\<^item> Variants \<^verbatim>\<open>A | B | C\<close>  | 
| 42658 | 582  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
583  | 
  @{rail \<open>A | B | C\<close>}
 | 
| 42658 | 584  | 
|
| 61503 | 585  | 
\<^item> Option \<^verbatim>\<open>A ?\<close>  | 
| 42658 | 586  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
587  | 
  @{rail \<open>A ?\<close>}
 | 
| 42658 | 588  | 
|
| 61503 | 589  | 
\<^item> Repetition \<^verbatim>\<open>A *\<close>  | 
| 42658 | 590  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
591  | 
  @{rail \<open>A *\<close>}
 | 
| 42658 | 592  | 
|
| 61503 | 593  | 
\<^item> Repetition with separator \<^verbatim>\<open>A * sep\<close>  | 
| 42658 | 594  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
595  | 
  @{rail \<open>A * sep\<close>}
 | 
| 42658 | 596  | 
|
| 61503 | 597  | 
\<^item> Strict repetition \<^verbatim>\<open>A +\<close>  | 
| 42658 | 598  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
599  | 
  @{rail \<open>A +\<close>}
 | 
| 42658 | 600  | 
|
| 61503 | 601  | 
\<^item> Strict repetition with separator \<^verbatim>\<open>A + sep\<close>  | 
| 42658 | 602  | 
|
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
603  | 
  @{rail \<open>A + sep\<close>}
 | 
| 58618 | 604  | 
\<close>  | 
| 42658 | 605  | 
|
606  | 
||
| 58618 | 607  | 
section \<open>Draft presentation\<close>  | 
| 27043 | 608  | 
|
| 58618 | 609  | 
text \<open>  | 
| 27043 | 610  | 
  \begin{matharray}{rcl}
 | 
| 61493 | 611  | 
    @{command_def "display_drafts"}\<open>\<^sup>*\<close> & : & \<open>any \<rightarrow>\<close> \\
 | 
| 27043 | 612  | 
  \end{matharray}
 | 
613  | 
||
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
614  | 
  @{rail \<open>
 | 
| 52549 | 615  | 
    @@{command display_drafts} (@{syntax name} +)
 | 
| 
55112
 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 
wenzelm 
parents: 
55029 
diff
changeset
 | 
616  | 
\<close>}  | 
| 27043 | 617  | 
|
| 61624 | 618  | 
  \<^descr> @{command "display_drafts"}~\<open>paths\<close> performs simple output of a given list
 | 
619  | 
of raw source files. Only those symbols that do not require additional  | 
|
620  | 
  {\LaTeX} packages are displayed properly, everything else is left verbatim.
 | 
|
| 58618 | 621  | 
\<close>  | 
| 27043 | 622  | 
|
623  | 
end  |