| author | haftmann | 
| Mon, 23 Feb 2009 21:38:36 +0100 | |
| changeset 30076 | f3043dafef5f | 
| parent 28761 | 9ec4482c9201 | 
| child 30168 | 9a20be5be90b | 
| permissions | -rw-r--r-- | 
| 27043 | 1 | (* $Id$ *) | 
| 2 | ||
| 3 | theory Document_Preparation | |
| 4 | imports Main | |
| 5 | begin | |
| 6 | ||
| 7 | chapter {* Document preparation \label{ch:document-prep} *}
 | |
| 8 | ||
| 28746 | 9 | text {* Isabelle/Isar provides a simple document preparation system
 | 
| 10 |   based on regular {PDF-\LaTeX} technology, with full support for
 | |
| 11 | hyper-links and bookmarks. Thus the results are well suited for WWW | |
| 12 | browsing and as printed copies. | |
| 27043 | 13 | |
| 28746 | 14 |   \medskip Isabelle generates {\LaTeX} output while running a
 | 
| 27043 | 15 |   \emph{logic session} (see also \cite{isabelle-sys}).  Getting
 | 
| 16 | started with a working configuration for common situations is quite | |
| 17 |   easy by using the Isabelle @{verbatim mkdir} and @{verbatim make}
 | |
| 18 | tools. First invoke | |
| 19 | \begin{ttbox}
 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
27881diff
changeset | 20 | isabelle mkdir Foo | 
| 27043 | 21 | \end{ttbox}
 | 
| 28746 | 22 |   to initialize a separate directory for session @{verbatim Foo} (it
 | 
| 23 |   is safe to experiment, since @{verbatim "isabelle mkdir"} never
 | |
| 24 |   overwrites existing files).  Ensure that @{verbatim "Foo/ROOT.ML"}
 | |
| 27043 | 25 | holds ML commands to load all theories required for this session; | 
| 26 |   furthermore @{verbatim "Foo/document/root.tex"} should include any
 | |
| 27 |   special {\LaTeX} macro packages required for your document (the
 | |
| 28 | default is usually sufficient as a start). | |
| 29 | ||
| 30 |   The session is controlled by a separate @{verbatim IsaMakefile}
 | |
| 31 | (with crude source dependencies by default). This file is located | |
| 32 |   one level up from the @{verbatim Foo} directory location.  Now
 | |
| 33 | invoke | |
| 34 | \begin{ttbox}
 | |
| 28504 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 wenzelm parents: 
27881diff
changeset | 35 | isabelle make Foo | 
| 27043 | 36 | \end{ttbox}
 | 
| 37 |   to run the @{verbatim Foo} session, with browser information and
 | |
| 38 | document preparation enabled. Unless any errors are reported by | |
| 39 |   Isabelle or {\LaTeX}, the output will appear inside the directory
 | |
| 28746 | 40 |   defined by the @{verbatim ISABELLE_BROWSER_INFO} setting (as
 | 
| 41 | reported by the batch job in verbose mode). | |
| 27043 | 42 | |
| 43 |   \medskip You may also consider to tune the @{verbatim usedir}
 | |
| 28746 | 44 |   options in @{verbatim IsaMakefile}, for example to switch the output
 | 
| 45 |   format between @{verbatim pdf} and @{verbatim dvi}, or activate the
 | |
| 27043 | 46 |   @{verbatim "-D"} option to retain a second copy of the generated
 | 
| 28746 | 47 |   {\LaTeX} sources (for manual inspection or separate runs of
 | 
| 48 |   @{executable latex}).
 | |
| 27043 | 49 | |
| 50 |   \medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys}
 | |
| 51 | for further details on Isabelle logic sessions and theory | |
| 52 |   presentation.  The Isabelle/HOL tutorial \cite{isabelle-hol-book}
 | |
| 28746 | 53 | also covers theory presentation to some extent. | 
| 27043 | 54 | *} | 
| 55 | ||
| 56 | ||
| 57 | section {* Markup commands \label{sec:markup} *}
 | |
| 58 | ||
| 59 | text {*
 | |
| 60 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 61 |     @{command_def "header"} & : & @{text "toplevel \<rightarrow> toplevel"} \\[0.5ex]
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 62 |     @{command_def "chapter"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 63 |     @{command_def "section"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 64 |     @{command_def "subsection"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 65 |     @{command_def "subsubsection"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 66 |     @{command_def "text"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 67 |     @{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: 
28760diff
changeset | 68 |     @{command_def "sect"} & : & @{text "proof \<rightarrow> proof"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 69 |     @{command_def "subsect"} & : & @{text "proof \<rightarrow> proof"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 70 |     @{command_def "subsubsect"} & : & @{text "proof \<rightarrow> proof"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 71 |     @{command_def "txt"} & : & @{text "proof \<rightarrow> proof"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 72 |     @{command_def "txt_raw"} & : & @{text "proof \<rightarrow> proof"} \\
 | 
| 27043 | 73 |   \end{matharray}
 | 
| 74 | ||
| 28747 | 75 | Markup commands provide a structured way to insert text into the | 
| 76 | document generated from a theory. Each markup command takes a | |
| 77 |   single @{syntax text} argument, which is passed as argument to a
 | |
| 78 |   corresponding {\LaTeX} macro.  The default macros provided by
 | |
| 79 |   @{"file" "~~/lib/texinputs/isabelle.sty"} can be redefined according
 | |
| 80 |   to the needs of the underlying document and {\LaTeX} styles.
 | |
| 81 | ||
| 82 |   Note that formal comments (\secref{sec:comments}) are similar to
 | |
| 83 | markup commands, but have a different status within Isabelle/Isar | |
| 84 | syntax. | |
| 27043 | 85 | |
| 86 |   \begin{rail}
 | |
| 87 |     ('chapter' | 'section' | 'subsection' | 'subsubsection' | 'text') target? text
 | |
| 88 | ; | |
| 27049 | 89 |     ('header' | 'text\_raw' | 'sect' | 'subsect' | 'subsubsect' | 'txt' | 'txt\_raw') text
 | 
| 27043 | 90 | ; | 
| 91 |   \end{rail}
 | |
| 92 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 93 |   \begin{description}
 | 
| 27043 | 94 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 95 |   \item @{command header} provides plain text markup just preceding
 | 
| 28747 | 96 |   the formal beginning of a theory.  The corresponding {\LaTeX} macro
 | 
| 97 |   is @{verbatim "\\isamarkupheader"}, which acts like @{command
 | |
| 98 | section} by default. | |
| 27049 | 99 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 100 |   \item @{command chapter}, @{command section}, @{command subsection},
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 101 |   and @{command subsubsection} mark chapter and section headings
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 102 | within the main theory body or local theory targets. The | 
| 28747 | 103 |   corresponding {\LaTeX} macros are @{verbatim "\\isamarkupchapter"},
 | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 104 |   @{verbatim "\\isamarkupsection"}, @{verbatim
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 105 | "\\isamarkupsubsection"} etc. | 
| 28747 | 106 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 107 |   \item @{command sect}, @{command subsect}, and @{command subsubsect}
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 108 |   mark section headings within proofs.  The corresponding {\LaTeX}
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 109 |   macros are @{verbatim "\\isamarkupsect"}, @{verbatim
 | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 110 | "\\isamarkupsubsect"} etc. | 
| 27043 | 111 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 112 |   \item @{command text} and @{command txt} specify paragraphs of plain
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 113 |   text.  This corresponds to a {\LaTeX} environment @{verbatim
 | 
| 28747 | 114 |   "\\begin{isamarkuptext}"} @{text "\<dots>"} @{verbatim
 | 
| 115 |   "\\end{isamarkuptext}"} etc.
 | |
| 27043 | 116 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 117 |   \item @{command text_raw} and @{command txt_raw} insert {\LaTeX}
 | 
| 28747 | 118 | source into the output, without additional markup. Thus the full | 
| 119 | range of document manipulations becomes available, at the risk of | |
| 120 | messing up document output. | |
| 27043 | 121 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 122 |   \end{description}
 | 
| 27043 | 123 | |
| 28747 | 124 |   Except for @{command "text_raw"} and @{command "txt_raw"}, the text
 | 
| 125 | passed to any of the above markup commands may refer to formal | |
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 126 |   entities via \emph{document antiquotations}, see also
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 127 |   \secref{sec:antiq}.  These are interpreted in the present theory or
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 128 |   proof context, or the named @{text "target"}.
 | 
| 27043 | 129 | |
| 130 | \medskip The proof markup commands closely resemble those for theory | |
| 131 | specifications, but have a different formal status and produce | |
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 132 |   different {\LaTeX} macros.  The default definitions coincide for
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 133 |   analogous commands such as @{command section} and @{command sect}.
 | 
| 27043 | 134 | *} | 
| 135 | ||
| 136 | ||
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 137 | section {* Document Antiquotations \label{sec:antiq} *}
 | 
| 27043 | 138 | |
| 139 | text {*
 | |
| 140 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 141 |     @{antiquotation_def "theory"} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 142 |     @{antiquotation_def "thm"} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 143 |     @{antiquotation_def "lemma"} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 144 |     @{antiquotation_def "prop"} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 145 |     @{antiquotation_def "term"} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 146 |     @{antiquotation_def const} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 147 |     @{antiquotation_def abbrev} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 148 |     @{antiquotation_def typeof} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 149 |     @{antiquotation_def typ} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 150 |     @{antiquotation_def thm_style} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 151 |     @{antiquotation_def term_style} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 152 |     @{antiquotation_def "text"} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 153 |     @{antiquotation_def goals} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 154 |     @{antiquotation_def subgoals} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 155 |     @{antiquotation_def prf} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 156 |     @{antiquotation_def full_prf} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 157 |     @{antiquotation_def ML} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 158 |     @{antiquotation_def ML_type} & : & @{text antiquotation} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 159 |     @{antiquotation_def ML_struct} & : & @{text antiquotation} \\
 | 
| 27043 | 160 |   \end{matharray}
 | 
| 161 | ||
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 162 | The overall content of an Isabelle/Isar theory may alternate between | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 163 | formal and informal text. The main body consists of formal | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 164 | specification and proof commands, interspersed with markup commands | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 165 |   (\secref{sec:markup}) or document comments (\secref{sec:comments}).
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 166 | The argument of markup commands quotes informal text to be printed | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 167 | in the resulting document, but may again refer to formal entities | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 168 |   via \emph{document antiquotations}.
 | 
| 27043 | 169 | |
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 170 |   For example, embedding of ``@{text [source=false] "@{term [show_types] \"f x = a + x\"}"}''
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 171 | within a text block makes | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 172 |   \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: 
28747diff
changeset | 173 | |
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 174 | Antiquotations usually spare the author tedious typing of logical | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 175 | entities in full detail. Even more importantly, some degree of | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 176 | consistency-checking between the main body of formal text and its | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 177 | informal explanation is achieved, since terms and types appearing in | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 178 | antiquotations are checked within the current theory or proof | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 179 | context. | 
| 27043 | 180 | |
| 181 |   \begin{rail}
 | |
| 182 | atsign lbrace antiquotation rbrace | |
| 183 | ; | |
| 184 | ||
| 185 | antiquotation: | |
| 186 | 'theory' options name | | |
| 187 | 'thm' options thmrefs | | |
| 27453 | 188 | 'lemma' options prop 'by' method | | 
| 27043 | 189 | 'prop' options prop | | 
| 190 | 'term' options term | | |
| 191 | 'const' options term | | |
| 192 | 'abbrev' options term | | |
| 193 | 'typeof' options term | | |
| 194 | 'typ' options type | | |
| 195 | 'thm\_style' options name thmref | | |
| 196 | 'term\_style' options name term | | |
| 197 | 'text' options name | | |
| 198 | 'goals' options | | |
| 199 | 'subgoals' options | | |
| 200 | 'prf' options thmrefs | | |
| 201 | 'full\_prf' options thmrefs | | |
| 202 | 'ML' options name | | |
| 203 | 'ML\_type' options name | | |
| 204 | 'ML\_struct' options name | |
| 205 | ; | |
| 206 | options: '[' (option * ',') ']' | |
| 207 | ; | |
| 208 | option: name | name '=' name | |
| 209 | ; | |
| 210 |   \end{rail}
 | |
| 211 | ||
| 212 |   Note that the syntax of antiquotations may \emph{not} include source
 | |
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 213 |   comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} nor verbatim
 | 
| 27043 | 214 |   text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
 | 
| 215 |   "*"}@{verbatim "}"}.
 | |
| 216 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 217 |   \begin{description}
 | 
| 27043 | 218 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 219 |   \item @{text "@{theory A}"} prints the name @{text "A"}, which is
 | 
| 27043 | 220 | guaranteed to refer to a valid ancestor theory in the current | 
| 221 | context. | |
| 222 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 223 |   \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: 
28747diff
changeset | 224 | Full fact expressions are allowed here, including attributes | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 225 |   (\secref{sec:syn-att}).
 | 
| 27043 | 226 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 227 |   \item @{text "@{prop \<phi>}"} prints a well-typed proposition @{text
 | 
| 27043 | 228 | "\<phi>"}. | 
| 229 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 230 |   \item @{text "@{lemma \<phi> by m}"} proves a well-typed proposition
 | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 231 |   @{text "\<phi>"} by method @{text m} and prints the original @{text "\<phi>"}.
 | 
| 27453 | 232 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 233 |   \item @{text "@{term t}"} prints a well-typed term @{text "t"}.
 | 
| 27043 | 234 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 235 |   \item @{text "@{const c}"} prints a logical or syntactic constant
 | 
| 27043 | 236 |   @{text "c"}.
 | 
| 237 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 238 |   \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: 
28750diff
changeset | 239 |   @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in the current context.
 | 
| 27043 | 240 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 241 |   \item @{text "@{typeof t}"} prints the type of a well-typed term
 | 
| 27043 | 242 |   @{text "t"}.
 | 
| 243 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 244 |   \item @{text "@{typ \<tau>}"} prints a well-formed type @{text "\<tau>"}.
 | 
| 27043 | 245 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 246 |   \item @{text "@{thm_style s a}"} prints theorem @{text a},
 | 
| 27043 | 247 |   previously applying a style @{text s} to it (see below).
 | 
| 248 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 249 |   \item @{text "@{term_style s t}"} prints a well-typed term @{text t}
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 250 |   after applying a style @{text s} to it (see below).
 | 
| 27043 | 251 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 252 |   \item @{text "@{text s}"} prints uninterpreted source text @{text
 | 
| 27043 | 253 | s}. This is particularly useful to print portions of text according | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 254 | to the Isabelle document style, without demanding well-formedness, | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 255 | e.g.\ small pieces of terms that should not be parsed or | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 256 | type-checked yet. | 
| 27043 | 257 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 258 |   \item @{text "@{goals}"} prints the current \emph{dynamic} goal
 | 
| 27043 | 259 | state. This is mainly for support of tactic-emulation scripts | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 260 | within Isar. Presentation of goal states does not conform to the | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 261 | idea of human-readable proof documents! | 
| 27043 | 262 | |
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 263 | When explaining proofs in detail it is usually better to spell out | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 264 | the reasoning via proper Isar proof commands, instead of peeking at | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 265 | the internal machine configuration. | 
| 27043 | 266 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 267 |   \item @{text "@{subgoals}"} is similar to @{text "@{goals}"}, but
 | 
| 27043 | 268 | does not print the main goal. | 
| 269 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 270 |   \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: 
28747diff
changeset | 271 |   corresponding to the theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}. Note that this
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 272 | requires proof terms to be switched on for the current logic | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 273 | session. | 
| 27043 | 274 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 275 |   \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: 
28750diff
changeset | 276 | a\<^sub>n}"}, but prints the full proof terms, i.e.\ also displays | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 277 | information omitted in the compact proof term, which is denoted by | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 278 |   ``@{text _}'' placeholders there.
 | 
| 27043 | 279 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 280 |   \item @{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 281 |   "@{ML_struct s}"} check text @{text s} as ML value, type, and
 | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 282 | structure, respectively. The source is printed verbatim. | 
| 27043 | 283 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 284 |   \end{description}
 | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 285 | *} | 
| 27043 | 286 | |
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 287 | |
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 288 | subsubsection {* Styled antiquotations *}
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 289 | |
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 290 | text {* Some antiquotations like @{text thm_style} and @{text
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 291 |   term_style} admit an extra \emph{style} specification to modify the
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 292 | printed result. The following standard styles are available: | 
| 27043 | 293 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 294 |   \begin{description}
 | 
| 27043 | 295 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 296 |   \item @{text lhs} extracts the first argument of any application
 | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 297 | form with at least two arguments --- typically meta-level or | 
| 27043 | 298 | object-level equality, or any other binary relation. | 
| 299 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 300 |   \item @{text rhs} is like @{text lhs}, but extracts the second
 | 
| 27043 | 301 | argument. | 
| 302 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 303 |   \item @{text "concl"} extracts the conclusion @{text C} from a rule
 | 
| 27043 | 304 |   in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
 | 
| 305 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 306 |   \item @{text "prem1"}, \dots, @{text "prem9"} extract premise number
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 307 |   @{text "1, \<dots>, 9"}, respectively, from from a rule in Horn-clause
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 308 |   normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
 | 
| 27043 | 309 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 310 |   \end{description}
 | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 311 | *} | 
| 27043 | 312 | |
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 313 | |
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 314 | subsubsection {* General options *}
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 315 | |
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 316 | text {* The following options are available to tune the printed output
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 317 | of antiquotations. Note that many of these coincide with global ML | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 318 | flags of the same names. | 
| 27043 | 319 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 320 |   \begin{description}
 | 
| 27043 | 321 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 322 |   \item @{text "show_types = bool"} and @{text "show_sorts = bool"}
 | 
| 27043 | 323 | control printing of explicit type and sort constraints. | 
| 324 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 325 |   \item @{text "show_structs = bool"} controls printing of implicit
 | 
| 27043 | 326 | structures. | 
| 327 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 328 |   \item @{text "long_names = bool"} forces names of types and
 | 
| 27043 | 329 | constants etc.\ to be printed in their fully qualified internal | 
| 330 | form. | |
| 331 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 332 |   \item @{text "short_names = bool"} forces names of types and
 | 
| 27043 | 333 | constants etc.\ to be printed unqualified. Note that internalizing | 
| 334 | the output again in the current context may well yield a different | |
| 335 | result. | |
| 336 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 337 |   \item @{text "unique_names = bool"} determines whether the printed
 | 
| 27043 | 338 | version of qualified names should be made sufficiently long to avoid | 
| 339 |   overlap with names declared further back.  Set to @{text false} for
 | |
| 340 | more concise output. | |
| 341 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 342 |   \item @{text "eta_contract = bool"} prints terms in @{text
 | 
| 27043 | 343 | \<eta>}-contracted form. | 
| 344 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 345 |   \item @{text "display = bool"} indicates if the text is to be output
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 346 | as multi-line ``display material'', rather than a small piece of | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 347 | text without line breaks (which is the default). | 
| 27043 | 348 | |
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 349 | In this mode the embedded entities are printed in the same style as | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 350 | the main theory text. | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 351 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 352 |   \item @{text "break = bool"} controls line breaks in non-display
 | 
| 27043 | 353 | material. | 
| 354 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 355 |   \item @{text "quotes = bool"} indicates if the output should be
 | 
| 27043 | 356 | enclosed in double quotes. | 
| 357 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 358 |   \item @{text "mode = name"} adds @{text name} to the print mode to
 | 
| 27043 | 359 |   be used for presentation (see also \cite{isabelle-ref}).  Note that
 | 
| 360 |   the standard setup for {\LaTeX} output is already present by
 | |
| 361 |   default, including the modes @{text latex} and @{text xsymbols}.
 | |
| 362 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 363 |   \item @{text "margin = nat"} and @{text "indent = nat"} change the
 | 
| 27043 | 364 | margin or indentation for pretty printing of display material. | 
| 365 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 366 |   \item @{text "goals_limit = nat"} determines the maximum number of
 | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 367 | goals to be printed (for goal-based antiquotation). | 
| 27043 | 368 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 369 |   \item @{text "locale = name"} specifies an alternative locale
 | 
| 27043 | 370 | context used for evaluating and printing the subsequent argument. | 
| 371 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 372 |   \item @{text "source = bool"} prints the original source text of the
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 373 | antiquotation arguments, rather than its internal representation. | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 374 |   Note that formal checking of @{antiquotation "thm"}, @{antiquotation
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 375 |   "term"}, etc. is still enabled; use the @{antiquotation "text"}
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 376 | antiquotation for unchecked output. | 
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 377 | |
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 378 |   Regular @{text "term"} and @{text "typ"} antiquotations with @{text
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 379 | "source = false"} involve a full round-trip from the original source | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 380 | to an internalized logical entity back to a source form, according | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 381 | to the syntax of the current context. Thus the printed output is | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 382 | not under direct control of the author, it may even fluctuate a bit | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 383 | as the underlying theory is changed later on. | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 384 | |
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 385 |   In contrast, @{text "source = true"} admits direct printing of the
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 386 | given source text, with the desirable well-formedness check in the | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 387 | background, but without modification of the printed text. | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 388 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 389 |   \end{description}
 | 
| 27043 | 390 | |
| 391 |   For boolean flags, ``@{text "name = true"}'' may be abbreviated as
 | |
| 392 |   ``@{text name}''.  All of the above flags are disabled by default,
 | |
| 28749 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 393 |   unless changed from ML, say in the @{verbatim "ROOT.ML"} of the
 | 
| 
99f6da3bbbf7
renamed "formal comments" to "document comments";
 wenzelm parents: 
28747diff
changeset | 394 | logic session. | 
| 27043 | 395 | *} | 
| 396 | ||
| 397 | ||
| 28750 | 398 | section {* Markup via command tags \label{sec:tags} *}
 | 
| 27043 | 399 | |
| 28750 | 400 | text {* Each Isabelle/Isar command may be decorated by additional
 | 
| 401 | presentation tags, to indicate some modification in the way it is | |
| 402 | printed in the document. | |
| 27043 | 403 | |
| 404 |   \indexouternonterm{tags}
 | |
| 405 |   \begin{rail}
 | |
| 406 | tags: ( tag * ) | |
| 407 | ; | |
| 408 | tag: '\%' (ident | string) | |
| 409 |   \end{rail}
 | |
| 410 | ||
| 28750 | 411 | Some tags are pre-declared for certain classes of commands, serving | 
| 412 | as default markup if no tags are given in the text: | |
| 27043 | 413 | |
| 28750 | 414 | \medskip | 
| 27043 | 415 |   \begin{tabular}{ll}
 | 
| 416 |     @{text "theory"} & theory begin/end \\
 | |
| 417 |     @{text "proof"} & all proof commands \\
 | |
| 418 |     @{text "ML"} & all commands involving ML code \\
 | |
| 419 |   \end{tabular}
 | |
| 420 | ||
| 28750 | 421 | \medskip The Isabelle document preparation system | 
| 422 |   \cite{isabelle-sys} allows tagged command regions to be presented
 | |
| 27043 | 423 | specifically, e.g.\ to fold proof texts, or drop parts of the text | 
| 424 | completely. | |
| 425 | ||
| 28750 | 426 |   For example ``@{command "by"}~@{text "%invisible auto"}'' causes
 | 
| 427 |   that piece of proof to be treated as @{text invisible} instead of
 | |
| 428 |   @{text "proof"} (the default), which may be shown or hidden
 | |
| 27043 | 429 |   depending on the document setup.  In contrast, ``@{command
 | 
| 28750 | 430 |   "by"}~@{text "%visible auto"}'' forces this text to be shown
 | 
| 27043 | 431 | invariably. | 
| 432 | ||
| 433 | Explicit tag specifications within a proof apply to all subsequent | |
| 434 |   commands of the same level of nesting.  For example, ``@{command
 | |
| 28750 | 435 |   "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' forces the whole
 | 
| 436 |   sub-proof to be typeset as @{text visible} (unless some of its parts
 | |
| 437 | are tagged differently). | |
| 438 | ||
| 439 | \medskip Command tags merely produce certain markup environments for | |
| 440 |   type-setting.  The meaning of these is determined by {\LaTeX}
 | |
| 441 |   macros, as defined in @{"file" "~~/lib/texinputs/isabelle.sty"} or
 | |
| 442 | by the document author. The Isabelle document preparation tools | |
| 443 | also provide some high-level options to specify the meaning of | |
| 444 | arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding | |
| 445 | parts of the text. Logic sessions may also specify ``document | |
| 446 | versions'', where given tags are interpreted in some particular way. | |
| 447 |   Again see \cite{isabelle-sys} for further details.
 | |
| 27043 | 448 | *} | 
| 449 | ||
| 450 | ||
| 451 | section {* Draft presentation *}
 | |
| 452 | ||
| 453 | text {*
 | |
| 454 |   \begin{matharray}{rcl}
 | |
| 28761 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 455 |     @{command_def "display_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
 | 
| 
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
 wenzelm parents: 
28760diff
changeset | 456 |     @{command_def "print_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
 | 
| 27043 | 457 |   \end{matharray}
 | 
| 458 | ||
| 459 |   \begin{rail}
 | |
| 460 |     ('display\_drafts' | 'print\_drafts') (name +)
 | |
| 461 | ; | |
| 462 |   \end{rail}
 | |
| 463 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 464 |   \begin{description}
 | 
| 27043 | 465 | |
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 466 |   \item @{command "display_drafts"}~@{text paths} and @{command
 | 
| 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 467 |   "print_drafts"}~@{text paths} perform simple output of a given list
 | 
| 27043 | 468 | of raw source files. Only those symbols that do not require | 
| 469 |   additional {\LaTeX} packages are displayed properly, everything else
 | |
| 470 | is left verbatim. | |
| 471 | ||
| 28760 
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
 wenzelm parents: 
28750diff
changeset | 472 |   \end{description}
 | 
| 27043 | 473 | *} | 
| 474 | ||
| 475 | end |