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