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