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