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