| author | blanchet | 
| Wed, 04 Nov 2015 15:07:23 +0100 | |
| changeset 61569 | 947ce60a06e1 | 
| parent 61503 | 28e788ca2c5d | 
| child 61572 | ddb3ac3fef45 | 
| permissions | -rw-r--r-- | 
| 28762 | 1 | theory Inner_Syntax | 
| 42651 | 2 | imports Base Main | 
| 28762 | 3 | begin | 
| 4 | ||
| 58618 | 5 | chapter \<open>Inner syntax --- the term language \label{ch:inner-syntax}\<close>
 | 
| 28762 | 6 | |
| 58618 | 7 | text \<open>The inner syntax of Isabelle provides concrete notation for | 
| 61493 | 8 | the main entities of the logical framework, notably \<open>\<lambda>\<close>-terms with types and type classes. Applications may either | 
| 46282 | 9 | extend existing syntactic categories by additional notation, or | 
| 10 | define new sub-languages that are linked to the standard term | |
| 61503 | 11 | language via some explicit markers. For example \<^verbatim>\<open>FOO\<close>~\<open>foo\<close> could | 
| 12 | embed the syntax corresponding for some | |
| 61493 | 13 | user-defined nonterminal \<open>foo\<close> --- within the bounds of the | 
| 46282 | 14 | given lexical syntax of Isabelle/Pure. | 
| 15 | ||
| 16 | The most basic way to specify concrete syntax for logical entities | |
| 17 |   works via mixfix annotations (\secref{sec:mixfix}), which may be
 | |
| 18 | usually given as part of the original declaration or via explicit | |
| 19 |   notation commands later on (\secref{sec:notation}).  This already
 | |
| 20 | covers many needs of concrete syntax without having to understand | |
| 21 | the full complexity of inner syntax layers. | |
| 22 | ||
| 23 | Further details of the syntax engine involves the classical | |
| 24 | distinction of lexical language versus context-free grammar (see | |
| 61477 | 25 |   \secref{sec:pure-syntax}), and various mechanisms for \<^emph>\<open>syntax
 | 
| 26 |   transformations\<close> (see \secref{sec:syntax-transformations}).
 | |
| 58618 | 27 | \<close> | 
| 46282 | 28 | |
| 29 | ||
| 58618 | 30 | section \<open>Printing logical entities\<close> | 
| 28762 | 31 | |
| 58618 | 32 | subsection \<open>Diagnostic commands \label{sec:print-diag}\<close>
 | 
| 28762 | 33 | |
| 58618 | 34 | text \<open> | 
| 28762 | 35 |   \begin{matharray}{rcl}
 | 
| 61493 | 36 |     @{command_def "typ"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | 
| 37 |     @{command_def "term"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 38 |     @{command_def "prop"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 39 |     @{command_def "thm"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 40 |     @{command_def "prf"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 41 |     @{command_def "full_prf"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | |
| 42 |     @{command_def "print_state"}\<open>\<^sup>*\<close> & : & \<open>any \<rightarrow>\<close> \\
 | |
| 28762 | 43 |   \end{matharray}
 | 
| 44 | ||
| 45 | These diagnostic commands assist interactive development by printing | |
| 46 | internal logical entities in a human-readable fashion. | |
| 47 | ||
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55108diff
changeset | 48 |   @{rail \<open>
 | 
| 48792 | 49 |     @@{command typ} @{syntax modes}? @{syntax type} ('::' @{syntax sort})?
 | 
| 28762 | 50 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 51 |     @@{command term} @{syntax modes}? @{syntax term}
 | 
| 28762 | 52 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 53 |     @@{command prop} @{syntax modes}? @{syntax prop}
 | 
| 28762 | 54 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 55 |     @@{command thm} @{syntax modes}? @{syntax thmrefs}
 | 
| 28762 | 56 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 57 |     ( @@{command prf} | @@{command full_prf} ) @{syntax modes}? @{syntax thmrefs}?
 | 
| 28762 | 58 | ; | 
| 52430 | 59 |     @@{command print_state} @{syntax modes}?
 | 
| 28762 | 60 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 61 |     @{syntax_def modes}: '(' (@{syntax name} + ) ')'
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55108diff
changeset | 62 | \<close>} | 
| 28762 | 63 | |
| 61493 | 64 |   \<^descr> @{command "typ"}~\<open>\<tau>\<close> reads and prints a type expression
 | 
| 48792 | 65 | according to the current context. | 
| 66 | ||
| 61493 | 67 |   \<^descr> @{command "typ"}~\<open>\<tau> :: s\<close> uses type-inference to
 | 
| 68 | determine the most general way to make \<open>\<tau>\<close> conform to sort | |
| 69 | \<open>s\<close>. For concrete \<open>\<tau>\<close> this checks if the type | |
| 70 | belongs to that sort. Dummy type parameters ``\<open>_\<close>'' | |
| 48792 | 71 | (underscore) are assigned to fresh type variables with most general | 
| 72 | sorts, according the the principles of type-inference. | |
| 28766 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28765diff
changeset | 73 | |
| 61493 | 74 |   \<^descr> @{command "term"}~\<open>t\<close> and @{command "prop"}~\<open>\<phi>\<close>
 | 
| 28766 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28765diff
changeset | 75 | read, type-check and print terms or propositions according to the | 
| 61493 | 76 | current theory or proof context; the inferred type of \<open>t\<close> is | 
| 28766 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28765diff
changeset | 77 | output as well. Note that these commands are also useful in | 
| 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28765diff
changeset | 78 | inspecting the current environment of term abbreviations. | 
| 28762 | 79 | |
| 61493 | 80 |   \<^descr> @{command "thm"}~\<open>a\<^sub>1 \<dots> a\<^sub>n\<close> retrieves
 | 
| 28762 | 81 | theorems from the current theory or proof context. Note that any | 
| 82 | attributes included in the theorem specifications are applied to a | |
| 83 | temporary context derived from the current theory or proof; the | |
| 61493 | 84 | result is discarded, i.e.\ attributes involved in \<open>a\<^sub>1, | 
| 85 | \<dots>, a\<^sub>n\<close> do not have any permanent effect. | |
| 28762 | 86 | |
| 61439 | 87 |   \<^descr> @{command "prf"} displays the (compact) proof term of the
 | 
| 28762 | 88 | current proof state (if present), or of the given theorems. Note | 
| 89 | that this requires proof terms to be switched on for the current | |
| 90 | object logic (see the ``Proof terms'' section of the Isabelle | |
| 91 | reference manual for information on how to do this). | |
| 92 | ||
| 61439 | 93 |   \<^descr> @{command "full_prf"} is like @{command "prf"}, but displays
 | 
| 28762 | 94 | the full proof term, i.e.\ also displays information omitted in the | 
| 61493 | 95 | compact proof term, which is denoted by ``\<open>_\<close>'' placeholders | 
| 28762 | 96 | there. | 
| 97 | ||
| 61439 | 98 |   \<^descr> @{command "print_state"} prints the current proof state (if
 | 
| 52430 | 99 | present), including current facts and goals. | 
| 28766 
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28765diff
changeset | 100 | |
| 28762 | 101 | |
| 61493 | 102 | All of the diagnostic commands above admit a list of \<open>modes\<close> | 
| 42926 | 103 | to be specified, which is appended to the current print mode; see | 
| 46284 | 104 |   also \secref{sec:print-modes}.  Thus the output behavior may be
 | 
| 105 | modified according particular print mode features. For example, | |
| 61493 | 106 |   @{command "print_state"}~\<open>(latex xsymbols)\<close> prints the
 | 
| 52430 | 107 | current proof state with mathematical symbols and special characters | 
| 46284 | 108 |   represented in {\LaTeX} source, according to the Isabelle style
 | 
| 60270 | 109 |   @{cite "isabelle-system"}.
 | 
| 28762 | 110 | |
| 111 |   Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more
 | |
| 112 | systematic way to include formal items into the printed text | |
| 113 | document. | |
| 58618 | 114 | \<close> | 
| 28762 | 115 | |
| 116 | ||
| 58618 | 117 | subsection \<open>Details of printed content\<close> | 
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 118 | |
| 58618 | 119 | text \<open> | 
| 42655 | 120 |   \begin{tabular}{rcll}
 | 
| 61493 | 121 |     @{attribute_def show_markup} & : & \<open>attribute\<close> \\
 | 
| 122 |     @{attribute_def show_types} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 123 |     @{attribute_def show_sorts} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 124 |     @{attribute_def show_consts} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 125 |     @{attribute_def show_abbrevs} & : & \<open>attribute\<close> & default \<open>true\<close> \\
 | |
| 126 |     @{attribute_def show_brackets} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 127 |     @{attribute_def names_long} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 128 |     @{attribute_def names_short} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 129 |     @{attribute_def names_unique} & : & \<open>attribute\<close> & default \<open>true\<close> \\
 | |
| 130 |     @{attribute_def eta_contract} & : & \<open>attribute\<close> & default \<open>true\<close> \\
 | |
| 131 |     @{attribute_def goals_limit} & : & \<open>attribute\<close> & default \<open>10\<close> \\
 | |
| 132 |     @{attribute_def show_main_goal} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 133 |     @{attribute_def show_hyps} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 134 |     @{attribute_def show_tags} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 135 |     @{attribute_def show_question_marks} & : & \<open>attribute\<close> & default \<open>true\<close> \\
 | |
| 42655 | 136 |   \end{tabular}
 | 
| 61421 | 137 | \<^medskip> | 
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 138 | |
| 42655 | 139 | These configuration options control the detail of information that | 
| 140 | is displayed for types, terms, theorems, goals etc. See also | |
| 141 |   \secref{sec:config}.
 | |
| 28765 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 142 | |
| 61439 | 143 |   \<^descr> @{attribute show_markup} controls direct inlining of markup
 | 
| 49699 | 144 | into the printed representation of formal entities --- notably type | 
| 145 | and sort constraints. This enables Prover IDE users to retrieve | |
| 146 | that information via tooltips or popups while hovering with the | |
| 147 | mouse over the output window, for example. Consequently, this | |
| 58842 | 148 | option is enabled by default for Isabelle/jEdit. | 
| 49699 | 149 | |
| 61439 | 150 |   \<^descr> @{attribute show_types} and @{attribute show_sorts} control
 | 
| 42655 | 151 | printing of type constraints for term variables, and sort | 
| 152 | constraints for type variables. By default, neither of these are | |
| 153 |   shown in output.  If @{attribute show_sorts} is enabled, types are
 | |
| 49699 | 154 | always shown as well. In Isabelle/jEdit, manual setting of these | 
| 155 |   options is normally not required thanks to @{attribute show_markup}
 | |
| 156 | above. | |
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 157 | |
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 158 | Note that displaying types and sorts may explain why a polymorphic | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 159 | inference rule fails to resolve with some goal, or why a rewrite | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 160 | rule does not apply as expected. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 161 | |
| 61439 | 162 |   \<^descr> @{attribute show_consts} controls printing of types of
 | 
| 42655 | 163 | constants when displaying a goal state. | 
| 28765 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 164 | |
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 165 | Note that the output can be enormous, because polymorphic constants | 
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 166 | often occur at several different type instances. | 
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 167 | |
| 61439 | 168 |   \<^descr> @{attribute show_abbrevs} controls folding of constant
 | 
| 42655 | 169 | abbreviations. | 
| 40879 
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
 wenzelm parents: 
40255diff
changeset | 170 | |
| 61439 | 171 |   \<^descr> @{attribute show_brackets} controls bracketing in pretty
 | 
| 42655 | 172 | printed output. If enabled, all sub-expressions of the pretty | 
| 28765 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 173 | printing tree will be parenthesized, even if this produces malformed | 
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 174 | term syntax! This crude way of showing the internal structure of | 
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 175 | pretty printed entities may occasionally help to diagnose problems | 
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 176 | with operator priorities, for example. | 
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 177 | |
| 61439 | 178 |   \<^descr> @{attribute names_long}, @{attribute names_short}, and
 | 
| 42669 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42655diff
changeset | 179 |   @{attribute names_unique} control the way of printing fully
 | 
| 42358 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
42279diff
changeset | 180 | qualified internal names in external form. See also | 
| 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
42279diff
changeset | 181 |   \secref{sec:antiq} for the document antiquotation options of the
 | 
| 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
42279diff
changeset | 182 | same names. | 
| 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
42279diff
changeset | 183 | |
| 61493 | 184 |   \<^descr> @{attribute eta_contract} controls \<open>\<eta>\<close>-contracted
 | 
| 42655 | 185 | printing of terms. | 
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 186 | |
| 61493 | 187 |   The \<open>\<eta>\<close>-contraction law asserts @{prop "(\<lambda>x. f x) \<equiv> f"},
 | 
| 188 | provided \<open>x\<close> is not free in \<open>f\<close>. It asserts | |
| 61477 | 189 |   \<^emph>\<open>extensionality\<close> of functions: @{prop "f \<equiv> g"} if @{prop "f x \<equiv>
 | 
| 61493 | 190 | g x"} for all \<open>x\<close>. Higher-order unification frequently puts | 
| 191 |   terms into a fully \<open>\<eta>\<close>-expanded form.  For example, if \<open>F\<close> has type \<open>(\<tau> \<Rightarrow> \<tau>) \<Rightarrow> \<tau>\<close> then its expanded form is @{term
 | |
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 192 | "\<lambda>h. F (\<lambda>x. h x)"}. | 
| 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 193 | |
| 61493 | 194 |   Enabling @{attribute eta_contract} makes Isabelle perform \<open>\<eta>\<close>-contractions before printing, so that @{term "\<lambda>h. F (\<lambda>x. h x)"}
 | 
| 195 | appears simply as \<open>F\<close>. | |
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 196 | |
| 61493 | 197 | Note that the distinction between a term and its \<open>\<eta>\<close>-expanded | 
| 28765 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 198 | form occasionally matters. While higher-order resolution and | 
| 61493 | 199 | rewriting operate modulo \<open>\<alpha>\<beta>\<eta>\<close>-conversion, some other tools | 
| 28765 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 200 | might look at terms more discretely. | 
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 201 | |
| 61439 | 202 |   \<^descr> @{attribute goals_limit} controls the maximum number of
 | 
| 51960 
61ac1efe02c3
option "goals_limit", with more uniform description;
 wenzelm parents: 
51657diff
changeset | 203 | subgoals to be printed. | 
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 204 | |
| 61439 | 205 |   \<^descr> @{attribute show_main_goal} controls whether the main result
 | 
| 42655 | 206 | to be proven should be displayed. This information might be | 
| 39130 | 207 | relevant for schematic goals, to inspect the current claim that has | 
| 208 | been synthesized so far. | |
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 209 | |
| 61439 | 210 |   \<^descr> @{attribute show_hyps} controls printing of implicit
 | 
| 42655 | 211 | hypotheses of local facts. Normally, only those hypotheses are | 
| 61477 | 212 | displayed that are \<^emph>\<open>not\<close> covered by the assumptions of the | 
| 42655 | 213 | current context: this situation indicates a fault in some tool being | 
| 214 | used. | |
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 215 | |
| 61477 | 216 |   By enabling @{attribute show_hyps}, output of \<^emph>\<open>all\<close> hypotheses
 | 
| 42655 | 217 | can be enforced, which is occasionally useful for diagnostic | 
| 218 | purposes. | |
| 28763 
b5e6122ff575
added pretty printing options (from old ref manual);
 wenzelm parents: 
28762diff
changeset | 219 | |
| 61439 | 220 |   \<^descr> @{attribute show_tags} controls printing of extra annotations
 | 
| 42655 | 221 | within theorems, such as internal position information, or the case | 
| 222 |   names being attached by the attribute @{attribute case_names}.
 | |
| 28765 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 223 | |
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 224 |   Note that the @{attribute tagged} and @{attribute untagged}
 | 
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 225 | attributes provide low-level access to the collection of tags | 
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 226 | associated with a theorem. | 
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 227 | |
| 61439 | 228 |   \<^descr> @{attribute show_question_marks} controls printing of question
 | 
| 61493 | 229 | marks for schematic variables, such as \<open>?x\<close>. Only the leading | 
| 28765 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 230 | question mark is affected, the remaining text is unchanged | 
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 231 | (including proper markup for schematic variables that might be | 
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 232 | relevant for user interfaces). | 
| 58618 | 233 | \<close> | 
| 28765 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 234 | |
| 
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
 wenzelm parents: 
28763diff
changeset | 235 | |
| 58618 | 236 | subsection \<open>Alternative print modes \label{sec:print-modes}\<close>
 | 
| 46284 | 237 | |
| 58618 | 238 | text \<open> | 
| 46284 | 239 |   \begin{mldecls}
 | 
| 240 |     @{index_ML print_mode_value: "unit -> string list"} \\
 | |
| 241 |     @{index_ML Print_Mode.with_modes: "string list -> ('a -> 'b) -> 'a -> 'b"} \\
 | |
| 242 |   \end{mldecls}
 | |
| 243 | ||
| 61477 | 244 | The \<^emph>\<open>print mode\<close> facility allows to modify various operations | 
| 46284 | 245 |   for printing.  Commands like @{command typ}, @{command term},
 | 
| 246 |   @{command thm} (see \secref{sec:print-diag}) take additional print
 | |
| 247 | modes as optional argument. The underlying ML operations are as | |
| 248 | follows. | |
| 249 | ||
| 61439 | 250 |   \<^descr> @{ML "print_mode_value ()"} yields the list of currently
 | 
| 46284 | 251 | active print mode names. This should be understood as symbolic | 
| 252 | representation of certain individual features for printing (with | |
| 253 | precedence from left to right). | |
| 254 | ||
| 61493 | 255 |   \<^descr> @{ML Print_Mode.with_modes}~\<open>modes f x\<close> evaluates
 | 
| 256 | \<open>f x\<close> in an execution context where the print mode is | |
| 257 | prepended by the given \<open>modes\<close>. This provides a thread-safe | |
| 46284 | 258 | way to augment print modes. It is also monotonic in the set of mode | 
| 259 | names: it retains the default print mode that certain | |
| 260 | user-interfaces might have installed for their proper functioning! | |
| 261 | ||
| 262 | ||
| 61421 | 263 | \<^medskip> | 
| 264 | The pretty printer for inner syntax maintains alternative | |
| 46284 | 265 | mixfix productions for any print mode name invented by the user, say | 
| 266 |   in commands like @{command notation} or @{command abbreviation}.
 | |
| 267 | Mode names can be arbitrary, but the following ones have a specific | |
| 268 | meaning by convention: | |
| 269 | ||
| 61503 | 270 | \<^item> \<^verbatim>\<open>""\<close> (the empty string): default mode; | 
| 46284 | 271 | implicitly active as last element in the list of modes. | 
| 272 | ||
| 61503 | 273 | \<^item> \<^verbatim>\<open>input\<close>: dummy print mode that is never active; may | 
| 46284 | 274 | be used to specify notation that is only available for input. | 
| 275 | ||
| 61503 | 276 | \<^item> \<^verbatim>\<open>internal\<close> dummy print mode that is never active; | 
| 46284 | 277 | used internally in Isabelle/Pure. | 
| 278 | ||
| 61503 | 279 | \<^item> \<^verbatim>\<open>xsymbols\<close>: enable proper mathematical symbols | 
| 46284 | 280 |   instead of ASCII art.\footnote{This traditional mode name stems from
 | 
| 58842 | 281 | the ``X-Symbol'' package for classic Proof~General with XEmacs.} | 
| 46284 | 282 | |
| 61503 | 283 |   \<^item> \<^verbatim>\<open>latex\<close>: additional mode that is active in {\LaTeX}
 | 
| 46284 | 284 | document preparation of Isabelle theory sources; allows to provide | 
| 285 | alternative output notation. | |
| 58618 | 286 | \<close> | 
| 46284 | 287 | |
| 288 | ||
| 58618 | 289 | section \<open>Mixfix annotations \label{sec:mixfix}\<close>
 | 
| 28762 | 290 | |
| 61477 | 291 | text \<open>Mixfix annotations specify concrete \<^emph>\<open>inner syntax\<close> of | 
| 35351 
7425aece4ee3
allow general mixfix syntax for type constructors;
 wenzelm parents: 
32833diff
changeset | 292 | Isabelle types and terms. Locally fixed parameters in toplevel | 
| 46290 | 293 | theorem statements, locale and class specifications also admit | 
| 294 | mixfix annotations in a fairly uniform manner. A mixfix annotation | |
| 50635 | 295 | describes the concrete syntax, the translation to abstract | 
| 46290 | 296 | syntax, and the pretty printing. Special case annotations provide a | 
| 297 | simple means of specifying infix operators and binders. | |
| 298 | ||
| 58552 | 299 |   Isabelle mixfix syntax is inspired by {\OBJ} @{cite OBJ}.  It allows
 | 
| 46290 | 300 | to specify any context-free priority grammar, which is more general | 
| 301 | than the fixity declarations of ML and Prolog. | |
| 28762 | 302 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55108diff
changeset | 303 |   @{rail \<open>
 | 
| 51654 
8450b944e58a
just one syntax category "mixfix" -- check structure annotation semantically;
 wenzelm parents: 
50636diff
changeset | 304 |     @{syntax_def mixfix}: '('
 | 
| 58761 | 305 |       (@{syntax template} prios? @{syntax nat}? |
 | 
| 306 |         (@'infix' | @'infixl' | @'infixr') @{syntax template} @{syntax nat} |
 | |
| 307 |         @'binder' @{syntax template} prios? @{syntax nat} |
 | |
| 308 | @'structure') ')' | |
| 46290 | 309 | ; | 
| 310 | template: string | |
| 46289 | 311 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 312 |     prios: '[' (@{syntax nat} + ',') ']'
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55108diff
changeset | 313 | \<close>} | 
| 28762 | 314 | |
| 61493 | 315 | The string given as \<open>template\<close> may include literal text, | 
| 316 | spacing, blocks, and arguments (denoted by ``\<open>_\<close>''); the | |
| 61503 | 317 | special symbol ``\<^verbatim>\<open>\<index>\<close>'' (printed as ``\<open>\<index>\<close>'') | 
| 51657 
3db1bbc82d8d
more accurate documentation of "(structure)" mixfix;
 wenzelm parents: 
51654diff
changeset | 318 |   represents an index argument that specifies an implicit @{keyword
 | 
| 
3db1bbc82d8d
more accurate documentation of "(structure)" mixfix;
 wenzelm parents: 
51654diff
changeset | 319 |   "structure"} reference (see also \secref{sec:locale}).  Only locally
 | 
| 
3db1bbc82d8d
more accurate documentation of "(structure)" mixfix;
 wenzelm parents: 
51654diff
changeset | 320 |   fixed variables may be declared as @{keyword "structure"}.
 | 
| 
3db1bbc82d8d
more accurate documentation of "(structure)" mixfix;
 wenzelm parents: 
51654diff
changeset | 321 | |
| 
3db1bbc82d8d
more accurate documentation of "(structure)" mixfix;
 wenzelm parents: 
51654diff
changeset | 322 | Infix and binder declarations provide common abbreviations for | 
| 
3db1bbc82d8d
more accurate documentation of "(structure)" mixfix;
 wenzelm parents: 
51654diff
changeset | 323 | particular mixfix declarations. So in practice, mixfix templates | 
| 
3db1bbc82d8d
more accurate documentation of "(structure)" mixfix;
 wenzelm parents: 
51654diff
changeset | 324 | mostly degenerate to literal text for concrete syntax, such as | 
| 61503 | 325 | ``\<^verbatim>\<open>++\<close>'' for an infix symbol. | 
| 326 | \<close> | |
| 28762 | 327 | |
| 46290 | 328 | |
| 58618 | 329 | subsection \<open>The general mixfix form\<close> | 
| 46290 | 330 | |
| 58618 | 331 | text \<open>In full generality, mixfix declarations work as follows. | 
| 61493 | 332 | Suppose a constant \<open>c :: \<tau>\<^sub>1 \<Rightarrow> \<dots> \<tau>\<^sub>n \<Rightarrow> \<tau>\<close> is annotated by | 
| 333 | \<open>(mixfix [p\<^sub>1, \<dots>, p\<^sub>n] p)\<close>, where \<open>mixfix\<close> is a string | |
| 334 | \<open>d\<^sub>0 _ d\<^sub>1 _ \<dots> _ d\<^sub>n\<close> consisting of delimiters that surround | |
| 46290 | 335 | argument positions as indicated by underscores. | 
| 28762 | 336 | |
| 337 | Altogether this determines a production for a context-free priority | |
| 61493 | 338 | grammar, where for each argument \<open>i\<close> the syntactic category | 
| 339 | is determined by \<open>\<tau>\<^sub>i\<close> (with priority \<open>p\<^sub>i\<close>), and the | |
| 340 | result category is determined from \<open>\<tau>\<close> (with priority \<open>p\<close>). Priority specifications are optional, with default 0 for | |
| 46292 | 341 |   arguments and 1000 for the result.\footnote{Omitting priorities is
 | 
| 342 | prone to syntactic ambiguities unless the delimiter tokens determine | |
| 61493 | 343 | fully bracketed notation, as in \<open>if _ then _ else _ fi\<close>.} | 
| 28762 | 344 | |
| 61493 | 345 | Since \<open>\<tau>\<close> may be again a function type, the constant | 
| 28762 | 346 | type scheme may have more argument positions than the mixfix | 
| 61493 | 347 | pattern. Printing a nested application \<open>c t\<^sub>1 \<dots> t\<^sub>m\<close> for | 
| 348 | \<open>m > n\<close> works by attaching concrete notation only to the | |
| 349 | innermost part, essentially by printing \<open>(c t\<^sub>1 \<dots> t\<^sub>n) \<dots> t\<^sub>m\<close> | |
| 28762 | 350 | instead. If a term has fewer arguments than specified in the mixfix | 
| 351 | template, the concrete syntax is ignored. | |
| 352 | ||
| 61421 | 353 | \<^medskip> | 
| 354 | A mixfix template may also contain additional directives | |
| 28762 | 355 | for pretty printing, notably spaces, blocks, and breaks. The | 
| 356 | general template format is a sequence over any of the following | |
| 357 | entities. | |
| 358 | ||
| 61493 | 359 | \<^descr> \<open>d\<close> is a delimiter, namely a non-empty sequence of | 
| 28771 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 360 | characters other than the following special characters: | 
| 28762 | 361 | |
| 61421 | 362 | \<^medskip> | 
| 28771 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 363 |   \begin{tabular}{ll}
 | 
| 61503 | 364 | \<^verbatim>\<open>'\<close> & single quote \\ | 
| 365 | \<^verbatim>\<open>_\<close> & underscore \\ | |
| 61493 | 366 | \<open>\<index>\<close> & index symbol \\ | 
| 61503 | 367 | \<^verbatim>\<open>(\<close> & open parenthesis \\ | 
| 368 | \<^verbatim>\<open>)\<close> & close parenthesis \\ | |
| 369 | \<^verbatim>\<open>/\<close> & slash \\ | |
| 28771 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 370 |   \end{tabular}
 | 
| 61421 | 371 | \<^medskip> | 
| 28762 | 372 | |
| 61503 | 373 | \<^descr> \<^verbatim>\<open>'\<close> escapes the special meaning of these | 
| 28771 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 374 | meta-characters, producing a literal version of the following | 
| 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 375 | character, unless that is a blank. | 
| 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 376 | |
| 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 377 | A single quote followed by a blank separates delimiters, without | 
| 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 378 | affecting printing, but input tokens may have additional white space | 
| 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 379 | here. | 
| 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 380 | |
| 61503 | 381 | \<^descr> \<^verbatim>\<open>_\<close> is an argument position, which stands for a | 
| 28762 | 382 | certain syntactic category in the underlying grammar. | 
| 383 | ||
| 61493 | 384 | \<^descr> \<open>\<index>\<close> is an indexed argument position; this is the place | 
| 28771 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 385 | where implicit structure arguments can be attached. | 
| 28762 | 386 | |
| 61493 | 387 | \<^descr> \<open>s\<close> is a non-empty sequence of spaces for printing. | 
| 28771 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 388 | This and the following specifications do not affect parsing at all. | 
| 28762 | 389 | |
| 61503 | 390 | \<^descr> \<^verbatim>\<open>(\<close>\<open>n\<close> opens a pretty printing block. The | 
| 28762 | 391 | optional number specifies how much indentation to add when a line | 
| 392 | break occurs within the block. If the parenthesis is not followed | |
| 393 | by digits, the indentation defaults to 0. A block specified via | |
| 61503 | 394 | \<^verbatim>\<open>(00\<close> is unbreakable. | 
| 28762 | 395 | |
| 61503 | 396 | \<^descr> \<^verbatim>\<open>)\<close> closes a pretty printing block. | 
| 28762 | 397 | |
| 61503 | 398 | \<^descr> \<^verbatim>\<open>//\<close> forces a line break. | 
| 28762 | 399 | |
| 61503 | 400 | \<^descr> \<^verbatim>\<open>/\<close>\<open>s\<close> allows a line break. Here \<open>s\<close> | 
| 28771 
4510201c6aaf
mixfix annotations: verbatim for special symbols;
 wenzelm parents: 
28770diff
changeset | 401 | stands for the string of spaces (zero or more) right after the | 
| 61477 | 402 | slash. These spaces are printed if the break is \<^emph>\<open>not\<close> taken. | 
| 28762 | 403 | |
| 404 | ||
| 405 | The general idea of pretty printing with blocks and breaks is also | |
| 58552 | 406 |   described in @{cite "paulson-ml2"}; it goes back to @{cite "Oppen:1980"}.
 | 
| 58618 | 407 | \<close> | 
| 28762 | 408 | |
| 409 | ||
| 58618 | 410 | subsection \<open>Infixes\<close> | 
| 46290 | 411 | |
| 58618 | 412 | text \<open>Infix operators are specified by convenient short forms that | 
| 46290 | 413 | abbreviate general mixfix annotations as follows: | 
| 414 | ||
| 415 |   \begin{center}
 | |
| 416 |   \begin{tabular}{lll}
 | |
| 417 | ||
| 61503 | 418 |   \<^verbatim>\<open>(\<close>@{keyword_def "infix"}~\<^verbatim>\<open>"\<close>\<open>sy\<close>\<^verbatim>\<open>"\<close> \<open>p\<close>\<^verbatim>\<open>)\<close>
 | 
| 61493 | 419 | & \<open>\<mapsto>\<close> & | 
| 61503 | 420 |   \<^verbatim>\<open>("(_\<close>~\<open>sy\<close>\<^verbatim>\<open>/ _)" [\<close>\<open>p + 1\<close>\<^verbatim>\<open>,\<close>~\<open>p + 1\<close>\<^verbatim>\<open>]\<close>~\<open>p\<close>\<^verbatim>\<open>)\<close> \\
 | 
| 421 |   \<^verbatim>\<open>(\<close>@{keyword_def "infixl"}~\<^verbatim>\<open>"\<close>\<open>sy\<close>\<^verbatim>\<open>"\<close> \<open>p\<close>\<^verbatim>\<open>)\<close>
 | |
| 61493 | 422 | & \<open>\<mapsto>\<close> & | 
| 61503 | 423 |   \<^verbatim>\<open>("(_\<close>~\<open>sy\<close>\<^verbatim>\<open>/ _)" [\<close>\<open>p\<close>\<^verbatim>\<open>,\<close>~\<open>p + 1\<close>\<^verbatim>\<open>]\<close>~\<open>p\<close>\<^verbatim>\<open>)\<close> \\
 | 
| 424 |   \<^verbatim>\<open>(\<close>@{keyword_def "infixr"}~\<^verbatim>\<open>"\<close>\<open>sy\<close>\<^verbatim>\<open>"\<close>~\<open>p\<close>\<^verbatim>\<open>)\<close>
 | |
| 61493 | 425 | & \<open>\<mapsto>\<close> & | 
| 61503 | 426 |   \<^verbatim>\<open>("(_\<close>~\<open>sy\<close>\<^verbatim>\<open>/ _)" [\<close>\<open>p + 1\<close>\<^verbatim>\<open>,\<close>~\<open>p\<close>\<^verbatim>\<open>]\<close>~\<open>p\<close>\<^verbatim>\<open>)\<close> \\
 | 
| 46290 | 427 | |
| 428 |   \end{tabular}
 | |
| 429 |   \end{center}
 | |
| 430 | ||
| 61503 | 431 | The mixfix template \<^verbatim>\<open>"(_\<close>~\<open>sy\<close>\<^verbatim>\<open>/ _)"\<close> | 
| 46292 | 432 | specifies two argument positions; the delimiter is preceded by a | 
| 433 | space and followed by a space or line break; the entire phrase is a | |
| 434 | pretty printing block. | |
| 46290 | 435 | |
| 61503 | 436 | The alternative notation \<^verbatim>\<open>op\<close>~\<open>sy\<close> is introduced | 
| 46290 | 437 | in addition. Thus any infix operator may be written in prefix form | 
| 438 | (as in ML), independently of the number of arguments in the term. | |
| 58618 | 439 | \<close> | 
| 46290 | 440 | |
| 441 | ||
| 58618 | 442 | subsection \<open>Binders\<close> | 
| 46290 | 443 | |
| 61477 | 444 | text \<open>A \<^emph>\<open>binder\<close> is a variable-binding construct such as a | 
| 61493 | 445 | quantifier. The idea to formalize \<open>\<forall>x. b\<close> as \<open>All | 
| 446 |   (\<lambda>x. b)\<close> for \<open>All :: ('a \<Rightarrow> bool) \<Rightarrow> bool\<close> already goes back
 | |
| 58552 | 447 |   to @{cite church40}.  Isabelle declarations of certain higher-order
 | 
| 46292 | 448 |   operators may be annotated with @{keyword_def "binder"} annotations
 | 
| 449 | as follows: | |
| 46290 | 450 | |
| 451 |   \begin{center}
 | |
| 61503 | 452 |   \<open>c ::\<close>~\<^verbatim>\<open>"\<close>\<open>(\<tau>\<^sub>1 \<Rightarrow> \<tau>\<^sub>2) \<Rightarrow> \<tau>\<^sub>3\<close>\<^verbatim>\<open>"  (\<close>@{keyword "binder"}~\<^verbatim>\<open>"\<close>\<open>sy\<close>\<^verbatim>\<open>" [\<close>\<open>p\<close>\<^verbatim>\<open>]\<close>~\<open>q\<close>\<^verbatim>\<open>)\<close>
 | 
| 46290 | 453 |   \end{center}
 | 
| 454 | ||
| 61493 | 455 | This introduces concrete binder syntax \<open>sy x. b\<close>, where | 
| 456 | \<open>x\<close> is a bound variable of type \<open>\<tau>\<^sub>1\<close>, the body \<open>b\<close> has type \<open>\<tau>\<^sub>2\<close> and the whole term has type \<open>\<tau>\<^sub>3\<close>. | |
| 457 | The optional integer \<open>p\<close> specifies the syntactic priority of | |
| 458 | the body; the default is \<open>q\<close>, which is also the priority of | |
| 46290 | 459 | the whole construct. | 
| 460 | ||
| 461 | Internally, the binder syntax is expanded to something like this: | |
| 462 |   \begin{center}
 | |
| 61503 | 463 |   \<open>c_binder ::\<close>~\<^verbatim>\<open>"\<close>\<open>idts \<Rightarrow> \<tau>\<^sub>2 \<Rightarrow> \<tau>\<^sub>3\<close>\<^verbatim>\<open>"  ("(3\<close>\<open>sy\<close>\<^verbatim>\<open>_./ _)" [0,\<close>~\<open>p\<close>\<^verbatim>\<open>]\<close>~\<open>q\<close>\<^verbatim>\<open>)\<close>
 | 
| 46290 | 464 |   \end{center}
 | 
| 465 | ||
| 466 |   Here @{syntax (inner) idts} is the nonterminal symbol for a list of
 | |
| 467 | identifiers with optional type constraints (see also | |
| 61503 | 468 |   \secref{sec:pure-grammar}).  The mixfix template \<^verbatim>\<open>"(3\<close>\<open>sy\<close>\<^verbatim>\<open>_./ _)"\<close>
 | 
| 469 | defines argument positions | |
| 46290 | 470 | for the bound identifiers and the body, separated by a dot with | 
| 471 | optional line break; the entire phrase is a pretty printing block of | |
| 61493 | 472 | indentation level 3. Note that there is no extra space after \<open>sy\<close>, so it needs to be included user specification if the binder | 
| 46290 | 473 | syntax ends with a token that may be continued by an identifier | 
| 474 |   token at the start of @{syntax (inner) idts}.
 | |
| 475 | ||
| 61493 | 476 | Furthermore, a syntax translation to transforms \<open>c_binder x\<^sub>1 | 
| 477 | \<dots> x\<^sub>n b\<close> into iterated application \<open>c (\<lambda>x\<^sub>1. \<dots> c (\<lambda>x\<^sub>n. b)\<dots>)\<close>. | |
| 58618 | 478 | This works in both directions, for parsing and printing.\<close> | 
| 46290 | 479 | |
| 480 | ||
| 58618 | 481 | section \<open>Explicit notation \label{sec:notation}\<close>
 | 
| 28762 | 482 | |
| 58618 | 483 | text \<open> | 
| 28762 | 484 |   \begin{matharray}{rcll}
 | 
| 61493 | 485 |     @{command_def "type_notation"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
 | 
| 486 |     @{command_def "no_type_notation"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
 | |
| 487 |     @{command_def "notation"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
 | |
| 488 |     @{command_def "no_notation"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
 | |
| 489 |     @{command_def "write"} & : & \<open>proof(state) \<rightarrow> proof(state)\<close> \\
 | |
| 28762 | 490 |   \end{matharray}
 | 
| 491 | ||
| 46288 | 492 | Commands that introduce new logical entities (terms or types) | 
| 493 | usually allow to provide mixfix annotations on the spot, which is | |
| 494 | convenient for default notation. Nonetheless, the syntax may be | |
| 495 | modified later on by declarations for explicit notation. This | |
| 496 | allows to add or delete mixfix annotations for of existing logical | |
| 497 | entities within the current context. | |
| 498 | ||
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55108diff
changeset | 499 |   @{rail \<open>
 | 
| 59783 
00b62aa9f430
tuned syntax diagrams -- no duplication of "target";
 wenzelm parents: 
58842diff
changeset | 500 |     (@@{command type_notation} | @@{command no_type_notation}) @{syntax mode}? \<newline>
 | 
| 
00b62aa9f430
tuned syntax diagrams -- no duplication of "target";
 wenzelm parents: 
58842diff
changeset | 501 |       (@{syntax nameref} @{syntax mixfix} + @'and')
 | 
| 35413 | 502 | ; | 
| 59783 
00b62aa9f430
tuned syntax diagrams -- no duplication of "target";
 wenzelm parents: 
58842diff
changeset | 503 |     (@@{command notation} | @@{command no_notation}) @{syntax mode}? \<newline>
 | 
| 51654 
8450b944e58a
just one syntax category "mixfix" -- check structure annotation semantically;
 wenzelm parents: 
50636diff
changeset | 504 |       (@{syntax nameref} @{syntax mixfix} + @'and')
 | 
| 28762 | 505 | ; | 
| 51654 
8450b944e58a
just one syntax category "mixfix" -- check structure annotation semantically;
 wenzelm parents: 
50636diff
changeset | 506 |     @@{command write} @{syntax mode}? (@{syntax nameref} @{syntax mixfix} + @'and')
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55108diff
changeset | 507 | \<close>} | 
| 28762 | 508 | |
| 61493 | 509 |   \<^descr> @{command "type_notation"}~\<open>c (mx)\<close> associates mixfix
 | 
| 35413 | 510 | syntax with an existing type constructor. The arity of the | 
| 511 | constructor is retrieved from the context. | |
| 46282 | 512 | |
| 61439 | 513 |   \<^descr> @{command "no_type_notation"} is similar to @{command
 | 
| 35413 | 514 | "type_notation"}, but removes the specified syntax annotation from | 
| 515 | the present context. | |
| 516 | ||
| 61493 | 517 |   \<^descr> @{command "notation"}~\<open>c (mx)\<close> associates mixfix
 | 
| 35413 | 518 | syntax with an existing constant or fixed variable. The type | 
| 519 | declaration of the given entity is retrieved from the context. | |
| 46282 | 520 | |
| 61439 | 521 |   \<^descr> @{command "no_notation"} is similar to @{command "notation"},
 | 
| 28762 | 522 | but removes the specified syntax annotation from the present | 
| 523 | context. | |
| 524 | ||
| 61439 | 525 |   \<^descr> @{command "write"} is similar to @{command "notation"}, but
 | 
| 36508 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 wenzelm parents: 
35413diff
changeset | 526 | works within an Isar proof body. | 
| 58618 | 527 | \<close> | 
| 28762 | 528 | |
| 28778 | 529 | |
| 58618 | 530 | section \<open>The Pure syntax \label{sec:pure-syntax}\<close>
 | 
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 531 | |
| 58618 | 532 | subsection \<open>Lexical matters \label{sec:inner-lex}\<close>
 | 
| 46282 | 533 | |
| 58618 | 534 | text \<open>The inner lexical syntax vaguely resembles the outer one | 
| 46282 | 535 |   (\secref{sec:outer-lex}), but some details are different.  There are
 | 
| 536 | two main categories of inner syntax tokens: | |
| 537 | ||
| 61477 | 538 | \<^enum> \<^emph>\<open>delimiters\<close> --- the literal tokens occurring in | 
| 46282 | 539 | productions of the given priority grammar (cf.\ | 
| 540 |   \secref{sec:priority-grammar});
 | |
| 541 | ||
| 61477 | 542 | \<^enum> \<^emph>\<open>named tokens\<close> --- various categories of identifiers etc. | 
| 46282 | 543 | |
| 544 | ||
| 545 | Delimiters override named tokens and may thus render certain | |
| 546 | identifiers inaccessible. Sometimes the logical context admits | |
| 547 | alternative ways to refer to the same entity, potentially via | |
| 548 | qualified names. | |
| 549 | ||
| 61421 | 550 | \<^medskip> | 
| 551 | The categories for named tokens are defined once and for | |
| 46282 | 552 | all as follows, reusing some categories of the outer token syntax | 
| 553 |   (\secref{sec:outer-lex}).
 | |
| 554 | ||
| 555 |   \begin{center}
 | |
| 556 |   \begin{supertabular}{rcl}
 | |
| 557 |     @{syntax_def (inner) id} & = & @{syntax_ref ident} \\
 | |
| 558 |     @{syntax_def (inner) longid} & = & @{syntax_ref longident} \\
 | |
| 559 |     @{syntax_def (inner) var} & = & @{syntax_ref var} \\
 | |
| 560 |     @{syntax_def (inner) tid} & = & @{syntax_ref typefree} \\
 | |
| 561 |     @{syntax_def (inner) tvar} & = & @{syntax_ref typevar} \\
 | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58409diff
changeset | 562 |     @{syntax_def (inner) num_token} & = & @{syntax_ref nat} \\
 | 
| 61503 | 563 |     @{syntax_def (inner) float_token} & = & @{syntax_ref nat}\<^verbatim>\<open>.\<close>@{syntax_ref nat} \\
 | 
| 564 |     @{syntax_def (inner) str_token} & = & \<^verbatim>\<open>''\<close> \<open>\<dots>\<close> \<^verbatim>\<open>''\<close> \\
 | |
| 565 |     @{syntax_def (inner) string_token} & = & \<^verbatim>\<open>"\<close> \<open>\<dots>\<close> \<^verbatim>\<open>"\<close> \\
 | |
| 61493 | 566 |     @{syntax_def (inner) cartouche} & = & @{verbatim "\<open>"} \<open>\<dots>\<close> @{verbatim "\<close>"} \\
 | 
| 46282 | 567 |   \end{supertabular}
 | 
| 568 |   \end{center}
 | |
| 569 | ||
| 570 |   The token categories @{syntax (inner) num_token}, @{syntax (inner)
 | |
| 58421 | 571 |   float_token}, @{syntax (inner) str_token}, @{syntax (inner) string_token},
 | 
| 572 |   and @{syntax (inner) cartouche} are not used in Pure. Object-logics may
 | |
| 573 | implement numerals and string literals by adding appropriate syntax | |
| 574 |   declarations, together with some translation functions (e.g.\ see @{file
 | |
| 575 | "~~/src/HOL/Tools/string_syntax.ML"}). | |
| 46282 | 576 | |
| 58421 | 577 |   The derived categories @{syntax_def (inner) num_const}, and @{syntax_def
 | 
| 578 | (inner) float_const}, provide robust access to the respective tokens: the | |
| 579 | syntax tree holds a syntactic constant instead of a free variable. | |
| 58618 | 580 | \<close> | 
| 46282 | 581 | |
| 582 | ||
| 58618 | 583 | subsection \<open>Priority grammars \label{sec:priority-grammar}\<close>
 | 
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 584 | |
| 61477 | 585 | text \<open>A context-free grammar consists of a set of \<^emph>\<open>terminal | 
| 586 | symbols\<close>, a set of \<^emph>\<open>nonterminal symbols\<close> and a set of | |
| 61493 | 587 | \<^emph>\<open>productions\<close>. Productions have the form \<open>A = \<gamma>\<close>, | 
| 588 | where \<open>A\<close> is a nonterminal and \<open>\<gamma>\<close> is a string of | |
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 589 | terminals and nonterminals. One designated nonterminal is called | 
| 61477 | 590 | the \<^emph>\<open>root symbol\<close>. The language defined by the grammar | 
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 591 | consists of all strings of terminals that can be derived from the | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 592 | root symbol by applying productions as rewrite rules. | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 593 | |
| 61477 | 594 | The standard Isabelle parser for inner syntax uses a \<^emph>\<open>priority | 
| 595 | grammar\<close>. Each nonterminal is decorated by an integer priority: | |
| 61493 | 596 | \<open>A\<^sup>(\<^sup>p\<^sup>)\<close>. In a derivation, \<open>A\<^sup>(\<^sup>p\<^sup>)\<close> may be rewritten | 
| 597 | using a production \<open>A\<^sup>(\<^sup>q\<^sup>) = \<gamma>\<close> only if \<open>p \<le> q\<close>. Any | |
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 598 | priority grammar can be translated into a normal context-free | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 599 | grammar by introducing new nonterminals and productions. | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 600 | |
| 61421 | 601 | \<^medskip> | 
| 61493 | 602 | Formally, a set of context free productions \<open>G\<close> | 
| 603 | induces a derivation relation \<open>\<longrightarrow>\<^sub>G\<close> as follows. Let \<open>\<alpha>\<close> and \<open>\<beta>\<close> denote strings of terminal or nonterminal symbols. | |
| 604 | Then \<open>\<alpha> A\<^sup>(\<^sup>p\<^sup>) \<beta> \<longrightarrow>\<^sub>G \<alpha> \<gamma> \<beta>\<close> holds if and only if \<open>G\<close> | |
| 605 | contains some production \<open>A\<^sup>(\<^sup>q\<^sup>) = \<gamma>\<close> for \<open>p \<le> q\<close>. | |
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 606 | |
| 61421 | 607 | \<^medskip> | 
| 608 | The following grammar for arithmetic expressions | |
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 609 | demonstrates how binding power and associativity of operators can be | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 610 | enforced by priorities. | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 611 | |
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 612 |   \begin{center}
 | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 613 |   \begin{tabular}{rclr}
 | 
| 61503 | 614 | \<open>A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\<close> & \<open>=\<close> & \<^verbatim>\<open>(\<close> \<open>A\<^sup>(\<^sup>0\<^sup>)\<close> \<^verbatim>\<open>)\<close> \\ | 
| 615 | \<open>A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\<close> & \<open>=\<close> & \<^verbatim>\<open>0\<close> \\ | |
| 616 | \<open>A\<^sup>(\<^sup>0\<^sup>)\<close> & \<open>=\<close> & \<open>A\<^sup>(\<^sup>0\<^sup>)\<close> \<^verbatim>\<open>+\<close> \<open>A\<^sup>(\<^sup>1\<^sup>)\<close> \\ | |
| 617 | \<open>A\<^sup>(\<^sup>2\<^sup>)\<close> & \<open>=\<close> & \<open>A\<^sup>(\<^sup>3\<^sup>)\<close> \<^verbatim>\<open>*\<close> \<open>A\<^sup>(\<^sup>2\<^sup>)\<close> \\ | |
| 618 | \<open>A\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>=\<close> & \<^verbatim>\<open>-\<close> \<open>A\<^sup>(\<^sup>3\<^sup>)\<close> \\ | |
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 619 |   \end{tabular}
 | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 620 |   \end{center}
 | 
| 61503 | 621 | The choice of priorities determines that \<^verbatim>\<open>-\<close> binds | 
| 622 | tighter than \<^verbatim>\<open>*\<close>, which binds tighter than \<^verbatim>\<open>+\<close>. | |
| 623 | Furthermore \<^verbatim>\<open>+\<close> associates to the left and | |
| 624 | \<^verbatim>\<open>*\<close> to the right. | |
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 625 | |
| 61421 | 626 | \<^medskip> | 
| 627 | For clarity, grammars obey these conventions: | |
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 628 | |
| 61421 | 629 | \<^item> All priorities must lie between 0 and 1000. | 
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 630 | |
| 61421 | 631 | \<^item> Priority 0 on the right-hand side and priority 1000 on the | 
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 632 | left-hand side may be omitted. | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 633 | |
| 61493 | 634 | \<^item> The production \<open>A\<^sup>(\<^sup>p\<^sup>) = \<alpha>\<close> is written as \<open>A = \<alpha> | 
| 635 | (p)\<close>, i.e.\ the priority of the left-hand side actually appears in | |
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 636 | a column on the far right. | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 637 | |
| 61493 | 638 | \<^item> Alternatives are separated by \<open>|\<close>. | 
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 639 | |
| 61493 | 640 | \<^item> Repetition is indicated by dots \<open>(\<dots>)\<close> in an informal | 
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 641 | but obvious way. | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 642 | |
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 643 | |
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 644 | Using these conventions, the example grammar specification above | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 645 | takes the form: | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 646 |   \begin{center}
 | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 647 |   \begin{tabular}{rclc}
 | 
| 61503 | 648 | \<open>A\<close> & \<open>=\<close> & \<^verbatim>\<open>(\<close> \<open>A\<close> \<^verbatim>\<open>)\<close> \\ | 
| 649 | & \<open>|\<close> & \<^verbatim>\<open>0\<close> & \qquad\qquad \\ | |
| 650 | & \<open>|\<close> & \<open>A\<close> \<^verbatim>\<open>+\<close> \<open>A\<^sup>(\<^sup>1\<^sup>)\<close> & \<open>(0)\<close> \\ | |
| 651 | & \<open>|\<close> & \<open>A\<^sup>(\<^sup>3\<^sup>)\<close> \<^verbatim>\<open>*\<close> \<open>A\<^sup>(\<^sup>2\<^sup>)\<close> & \<open>(2)\<close> \\ | |
| 652 | & \<open>|\<close> & \<^verbatim>\<open>-\<close> \<open>A\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>(3)\<close> \\ | |
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 653 |   \end{tabular}
 | 
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 654 |   \end{center}
 | 
| 58618 | 655 | \<close> | 
| 28769 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 656 | |
| 
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
 wenzelm parents: 
28767diff
changeset | 657 | |
| 58618 | 658 | subsection \<open>The Pure grammar \label{sec:pure-grammar}\<close>
 | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 659 | |
| 61493 | 660 | text \<open>The priority grammar of the \<open>Pure\<close> theory is defined | 
| 46287 | 661 | approximately like this: | 
| 28774 | 662 | |
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 663 |   \begin{center}
 | 
| 28773 | 664 |   \begin{supertabular}{rclr}
 | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 665 | |
| 61493 | 666 |   @{syntax_def (inner) any} & = & \<open>prop  |  logic\<close> \\\\
 | 
| 28772 | 667 | |
| 61503 | 668 |   @{syntax_def (inner) prop} & = & \<^verbatim>\<open>(\<close> \<open>prop\<close> \<^verbatim>\<open>)\<close> \\
 | 
| 669 | & \<open>|\<close> & \<open>prop\<^sup>(\<^sup>4\<^sup>)\<close> \<^verbatim>\<open>::\<close> \<open>type\<close> & \<open>(3)\<close> \\ | |
| 670 | & \<open>|\<close> & \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> \<^verbatim>\<open>==\<close> \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>(2)\<close> \\ | |
| 61493 | 671 | & \<open>|\<close> & \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> \<open>\<equiv>\<close> \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>(2)\<close> \\ | 
| 61503 | 672 | & \<open>|\<close> & \<open>prop\<^sup>(\<^sup>3\<^sup>)\<close> \<^verbatim>\<open>&&&\<close> \<open>prop\<^sup>(\<^sup>2\<^sup>)\<close> & \<open>(2)\<close> \\ | 
| 673 | & \<open>|\<close> & \<open>prop\<^sup>(\<^sup>2\<^sup>)\<close> \<^verbatim>\<open>==>\<close> \<open>prop\<^sup>(\<^sup>1\<^sup>)\<close> & \<open>(1)\<close> \\ | |
| 61493 | 674 | & \<open>|\<close> & \<open>prop\<^sup>(\<^sup>2\<^sup>)\<close> \<open>\<Longrightarrow>\<close> \<open>prop\<^sup>(\<^sup>1\<^sup>)\<close> & \<open>(1)\<close> \\ | 
| 61503 | 675 | & \<open>|\<close> & \<^verbatim>\<open>[|\<close> \<open>prop\<close> \<^verbatim>\<open>;\<close> \<open>\<dots>\<close> \<^verbatim>\<open>;\<close> \<open>prop\<close> \<^verbatim>\<open>|]\<close> \<^verbatim>\<open>==>\<close> \<open>prop\<^sup>(\<^sup>1\<^sup>)\<close> & \<open>(1)\<close> \\ | 
| 676 | & \<open>|\<close> & \<open>\<lbrakk>\<close> \<open>prop\<close> \<^verbatim>\<open>;\<close> \<open>\<dots>\<close> \<^verbatim>\<open>;\<close> \<open>prop\<close> \<open>\<rbrakk>\<close> \<open>\<Longrightarrow>\<close> \<open>prop\<^sup>(\<^sup>1\<^sup>)\<close> & \<open>(1)\<close> \\ | |
| 677 | & \<open>|\<close> & \<^verbatim>\<open>!!\<close> \<open>idts\<close> \<^verbatim>\<open>.\<close> \<open>prop\<close> & \<open>(0)\<close> \\ | |
| 678 | & \<open>|\<close> & \<open>\<And>\<close> \<open>idts\<close> \<^verbatim>\<open>.\<close> \<open>prop\<close> & \<open>(0)\<close> \\ | |
| 679 | & \<open>|\<close> & \<^verbatim>\<open>OFCLASS\<close> \<^verbatim>\<open>(\<close> \<open>type\<close> \<^verbatim>\<open>,\<close> \<open>logic\<close> \<^verbatim>\<open>)\<close> \\ | |
| 680 | & \<open>|\<close> & \<^verbatim>\<open>SORT_CONSTRAINT\<close> \<^verbatim>\<open>(\<close> \<open>type\<close> \<^verbatim>\<open>)\<close> \\ | |
| 681 | & \<open>|\<close> & \<^verbatim>\<open>TERM\<close> \<open>logic\<close> \\ | |
| 682 | & \<open>|\<close> & \<^verbatim>\<open>PROP\<close> \<open>aprop\<close> \\\\ | |
| 28772 | 683 | |
| 61503 | 684 |   @{syntax_def (inner) aprop} & = & \<^verbatim>\<open>(\<close> \<open>aprop\<close> \<^verbatim>\<open>)\<close> \\
 | 
| 685 | & \<open>|\<close> & \<open>id | longid | var |\<close>~~\<^verbatim>\<open>_\<close>~~\<open>|\<close>~~\<^verbatim>\<open>...\<close> \\ | |
| 686 | & \<open>|\<close> & \<^verbatim>\<open>CONST\<close> \<open>id |\<close>~~\<^verbatim>\<open>CONST\<close> \<open>longid\<close> \\ | |
| 687 | & \<open>|\<close> & \<^verbatim>\<open>XCONST\<close> \<open>id |\<close>~~\<^verbatim>\<open>XCONST\<close> \<open>longid\<close> \\ | |
| 61493 | 688 | & \<open>|\<close> & \<open>logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \<dots> any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\<close> & \<open>(999)\<close> \\\\ | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 689 | |
| 61503 | 690 |   @{syntax_def (inner) logic} & = & \<^verbatim>\<open>(\<close> \<open>logic\<close> \<^verbatim>\<open>)\<close> \\
 | 
| 691 | & \<open>|\<close> & \<open>logic\<^sup>(\<^sup>4\<^sup>)\<close> \<^verbatim>\<open>::\<close> \<open>type\<close> & \<open>(3)\<close> \\ | |
| 692 | & \<open>|\<close> & \<open>id | longid | var |\<close>~~\<^verbatim>\<open>_\<close>~~\<open>|\<close>~~\<^verbatim>\<open>...\<close> \\ | |
| 693 | & \<open>|\<close> & \<^verbatim>\<open>CONST\<close> \<open>id |\<close>~~\<^verbatim>\<open>CONST\<close> \<open>longid\<close> \\ | |
| 694 | & \<open>|\<close> & \<^verbatim>\<open>XCONST\<close> \<open>id |\<close>~~\<^verbatim>\<open>XCONST\<close> \<open>longid\<close> \\ | |
| 61493 | 695 | & \<open>|\<close> & \<open>logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \<dots> any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\<close> & \<open>(999)\<close> \\ | 
| 696 | & \<open>|\<close> & \<open>\<struct> index\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\<close> \\ | |
| 61503 | 697 | & \<open>|\<close> & \<^verbatim>\<open>%\<close> \<open>pttrns\<close> \<^verbatim>\<open>.\<close> \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>(3)\<close> \\ | 
| 698 | & \<open>|\<close> & \<open>\<lambda>\<close> \<open>pttrns\<close> \<^verbatim>\<open>.\<close> \<open>any\<^sup>(\<^sup>3\<^sup>)\<close> & \<open>(3)\<close> \\ | |
| 699 | & \<open>|\<close> & \<^verbatim>\<open>op\<close> \<^verbatim>\<open>==\<close>~~\<open>|\<close>~~\<^verbatim>\<open>op\<close> \<open>\<equiv>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>op\<close> \<^verbatim>\<open>&&&\<close> \\ | |
| 700 | & \<open>|\<close> & \<^verbatim>\<open>op\<close> \<^verbatim>\<open>==>\<close>~~\<open>|\<close>~~\<^verbatim>\<open>op\<close> \<open>\<Longrightarrow>\<close> \\ | |
| 701 | & \<open>|\<close> & \<^verbatim>\<open>TYPE\<close> \<^verbatim>\<open>(\<close> \<open>type\<close> \<^verbatim>\<open>)\<close> \\\\ | |
| 28772 | 702 | |
| 61503 | 703 |   @{syntax_def (inner) idt} & = & \<^verbatim>\<open>(\<close> \<open>idt\<close> \<^verbatim>\<open>)\<close>~~\<open>|  id  |\<close>~~\<^verbatim>\<open>_\<close> \\
 | 
| 704 | & \<open>|\<close> & \<open>id\<close> \<^verbatim>\<open>::\<close> \<open>type\<close> & \<open>(0)\<close> \\ | |
| 705 | & \<open>|\<close> & \<^verbatim>\<open>_\<close> \<^verbatim>\<open>::\<close> \<open>type\<close> & \<open>(0)\<close> \\\\ | |
| 28772 | 706 | |
| 61503 | 707 |   @{syntax_def (inner) index} & = & \<^verbatim>\<open>\<^bsub>\<close> \<open>logic\<^sup>(\<^sup>0\<^sup>)\<close> \<^verbatim>\<open>\<^esub>\<close>~~\<open>|  |  \<index>\<close> \\\\
 | 
| 46287 | 708 | |
| 61493 | 709 |   @{syntax_def (inner) idts} & = & \<open>idt  |  idt\<^sup>(\<^sup>1\<^sup>) idts\<close> & \<open>(0)\<close> \\\\
 | 
| 28772 | 710 | |
| 61493 | 711 |   @{syntax_def (inner) pttrn} & = & \<open>idt\<close> \\\\
 | 
| 28772 | 712 | |
| 61493 | 713 |   @{syntax_def (inner) pttrns} & = & \<open>pttrn  |  pttrn\<^sup>(\<^sup>1\<^sup>) pttrns\<close> & \<open>(0)\<close> \\\\
 | 
| 28774 | 714 | |
| 61503 | 715 |   @{syntax_def (inner) type} & = & \<^verbatim>\<open>(\<close> \<open>type\<close> \<^verbatim>\<open>)\<close> \\
 | 
| 716 | & \<open>|\<close> & \<open>tid | tvar |\<close>~~\<^verbatim>\<open>_\<close> \\ | |
| 717 | & \<open>|\<close> & \<open>tid\<close> \<^verbatim>\<open>::\<close> \<open>sort | tvar\<close>~~\<^verbatim>\<open>::\<close> \<open>sort |\<close>~~\<^verbatim>\<open>_\<close> \<^verbatim>\<open>::\<close> \<open>sort\<close> \\ | |
| 61493 | 718 | & \<open>|\<close> & \<open>type_name | type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) type_name\<close> \\ | 
| 61503 | 719 | & \<open>|\<close> & \<^verbatim>\<open>(\<close> \<open>type\<close> \<^verbatim>\<open>,\<close> \<open>\<dots>\<close> \<^verbatim>\<open>,\<close> \<open>type\<close> \<^verbatim>\<open>)\<close> \<open>type_name\<close> \\ | 
| 720 | & \<open>|\<close> & \<open>type\<^sup>(\<^sup>1\<^sup>)\<close> \<^verbatim>\<open>=>\<close> \<open>type\<close> & \<open>(0)\<close> \\ | |
| 61493 | 721 | & \<open>|\<close> & \<open>type\<^sup>(\<^sup>1\<^sup>)\<close> \<open>\<Rightarrow>\<close> \<open>type\<close> & \<open>(0)\<close> \\ | 
| 61503 | 722 | & \<open>|\<close> & \<^verbatim>\<open>[\<close> \<open>type\<close> \<^verbatim>\<open>,\<close> \<open>\<dots>\<close> \<^verbatim>\<open>,\<close> \<open>type\<close> \<^verbatim>\<open>]\<close> \<^verbatim>\<open>=>\<close> \<open>type\<close> & \<open>(0)\<close> \\ | 
| 723 | & \<open>|\<close> & \<^verbatim>\<open>[\<close> \<open>type\<close> \<^verbatim>\<open>,\<close> \<open>\<dots>\<close> \<^verbatim>\<open>,\<close> \<open>type\<close> \<^verbatim>\<open>]\<close> \<open>\<Rightarrow>\<close> \<open>type\<close> & \<open>(0)\<close> \\ | |
| 61493 | 724 |   @{syntax_def (inner) type_name} & = & \<open>id  |  longid\<close> \\\\
 | 
| 28772 | 725 | |
| 61503 | 726 |   @{syntax_def (inner) sort} & = & @{syntax class_name}~~\<open>|\<close>~~\<^verbatim>\<open>{}\<close> \\
 | 
| 727 |     & \<open>|\<close> & \<^verbatim>\<open>{\<close> @{syntax class_name} \<^verbatim>\<open>,\<close> \<open>\<dots>\<close> \<^verbatim>\<open>,\<close> @{syntax class_name} \<^verbatim>\<open>}\<close> \\
 | |
| 61493 | 728 |   @{syntax_def (inner) class_name} & = & \<open>id  |  longid\<close> \\
 | 
| 28773 | 729 |   \end{supertabular}
 | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 730 |   \end{center}
 | 
| 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 731 | |
| 61421 | 732 | \<^medskip> | 
| 61503 | 733 | Here literal terminals are printed \<^verbatim>\<open>verbatim\<close>; | 
| 28774 | 734 |   see also \secref{sec:inner-lex} for further token categories of the
 | 
| 735 | inner syntax. The meaning of the nonterminals defined by the above | |
| 736 | grammar is as follows: | |
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 737 | |
| 61439 | 738 |   \<^descr> @{syntax_ref (inner) any} denotes any term.
 | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 739 | |
| 61439 | 740 |   \<^descr> @{syntax_ref (inner) prop} denotes meta-level propositions,
 | 
| 28778 | 741 |   which are terms of type @{typ prop}.  The syntax of such formulae of
 | 
| 742 | the meta-logic is carefully distinguished from usual conventions for | |
| 61493 | 743 | object-logics. In particular, plain \<open>\<lambda>\<close>-term notation is | 
| 61477 | 744 |   \<^emph>\<open>not\<close> recognized as @{syntax (inner) prop}.
 | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 745 | |
| 61439 | 746 |   \<^descr> @{syntax_ref (inner) aprop} denotes atomic propositions, which
 | 
| 28778 | 747 |   are embedded into regular @{syntax (inner) prop} by means of an
 | 
| 61503 | 748 | explicit \<^verbatim>\<open>PROP\<close> token. | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 749 | |
| 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 750 |   Terms of type @{typ prop} with non-constant head, e.g.\ a plain
 | 
| 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 751 |   variable, are printed in this form.  Constants that yield type @{typ
 | 
| 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 752 | prop} are expected to provide their own concrete syntax; otherwise | 
| 28778 | 753 |   the printed version will appear like @{syntax (inner) logic} and
 | 
| 754 |   cannot be parsed again as @{syntax (inner) prop}.
 | |
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 755 | |
| 61439 | 756 |   \<^descr> @{syntax_ref (inner) logic} denotes arbitrary terms of a
 | 
| 28778 | 757 |   logical type, excluding type @{typ prop}.  This is the main
 | 
| 61493 | 758 | syntactic category of object-logic entities, covering plain \<open>\<lambda>\<close>-term notation (variables, abstraction, application), plus | 
| 28778 | 759 | anything defined by the user. | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 760 | |
| 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 761 | When specifying notation for logical entities, all logical types | 
| 61477 | 762 |   (excluding @{typ prop}) are \<^emph>\<open>collapsed\<close> to this single category
 | 
| 28778 | 763 |   of @{syntax (inner) logic}.
 | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 764 | |
| 61439 | 765 |   \<^descr> @{syntax_ref (inner) index} denotes an optional index term for
 | 
| 51657 
3db1bbc82d8d
more accurate documentation of "(structure)" mixfix;
 wenzelm parents: 
51654diff
changeset | 766 |   indexed syntax.  If omitted, it refers to the first @{keyword_ref
 | 
| 61493 | 767 | "structure"} variable in the context. The special dummy ``\<open>\<index>\<close>'' serves as pattern variable in mixfix annotations that | 
| 46287 | 768 | introduce indexed notation. | 
| 769 | ||
| 61439 | 770 |   \<^descr> @{syntax_ref (inner) idt} denotes identifiers, possibly
 | 
| 28778 | 771 | constrained by types. | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 772 | |
| 61439 | 773 |   \<^descr> @{syntax_ref (inner) idts} denotes a sequence of @{syntax_ref
 | 
| 28778 | 774 | (inner) idt}. This is the most basic category for variables in | 
| 61493 | 775 | iterated binders, such as \<open>\<lambda>\<close> or \<open>\<And>\<close>. | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 776 | |
| 61439 | 777 |   \<^descr> @{syntax_ref (inner) pttrn} and @{syntax_ref (inner) pttrns}
 | 
| 28778 | 778 | denote patterns for abstraction, cases bindings etc. In Pure, these | 
| 779 |   categories start as a merely copy of @{syntax (inner) idt} and
 | |
| 780 |   @{syntax (inner) idts}, respectively.  Object-logics may add
 | |
| 781 | additional productions for binding forms. | |
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 782 | |
| 61439 | 783 |   \<^descr> @{syntax_ref (inner) type} denotes types of the meta-logic.
 | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 784 | |
| 61439 | 785 |   \<^descr> @{syntax_ref (inner) sort} denotes meta-level sorts.
 | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 786 | |
| 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 787 | |
| 28774 | 788 | Here are some further explanations of certain syntax features. | 
| 28773 | 789 | |
| 61493 | 790 |   \<^item> In @{syntax (inner) idts}, note that \<open>x :: nat y\<close> is
 | 
| 791 | parsed as \<open>x :: (nat y)\<close>, treating \<open>y\<close> like a type | |
| 792 | constructor applied to \<open>nat\<close>. To avoid this interpretation, | |
| 793 | write \<open>(x :: nat) y\<close> with explicit parentheses. | |
| 28773 | 794 | |
| 61493 | 795 | \<^item> Similarly, \<open>x :: nat y :: nat\<close> is parsed as \<open>x :: | 
| 796 | (nat y :: nat)\<close>. The correct form is \<open>(x :: nat) (y :: | |
| 797 | nat)\<close>, or \<open>(x :: nat) y :: nat\<close> if \<open>y\<close> is last in the | |
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 798 | sequence of identifiers. | 
| 28773 | 799 | |
| 61421 | 800 | \<^item> Type constraints for terms bind very weakly. For example, | 
| 61493 | 801 | \<open>x < y :: nat\<close> is normally parsed as \<open>(x < y) :: | 
| 802 | nat\<close>, unless \<open><\<close> has a very low priority, in which case the | |
| 803 | input is likely to be ambiguous. The correct form is \<open>x < (y | |
| 804 | :: nat)\<close>. | |
| 28773 | 805 | |
| 61421 | 806 | \<^item> Dummy variables (written as underscore) may occur in different | 
| 28774 | 807 | roles. | 
| 28773 | 808 | |
| 61493 | 809 | \<^descr> A type ``\<open>_\<close>'' or ``\<open>_ :: sort\<close>'' acts like an | 
| 61458 | 810 | anonymous inference parameter, which is filled-in according to the | 
| 811 | most general type produced by the type-checking phase. | |
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 812 | |
| 61493 | 813 | \<^descr> A bound ``\<open>_\<close>'' refers to a vacuous abstraction, where | 
| 61458 | 814 | the body does not refer to the binding introduced here. As in the | 
| 61493 | 815 |     term @{term "\<lambda>x _. x"}, which is \<open>\<alpha>\<close>-equivalent to \<open>\<lambda>x y. x\<close>.
 | 
| 28773 | 816 | |
| 61493 | 817 | \<^descr> A free ``\<open>_\<close>'' refers to an implicit outer binding. | 
| 818 | Higher definitional packages usually allow forms like \<open>f x _ | |
| 819 | = x\<close>. | |
| 28773 | 820 | |
| 61493 | 821 | \<^descr> A schematic ``\<open>_\<close>'' (within a term pattern, see | 
| 61458 | 822 |     \secref{sec:term-decls}) refers to an anonymous variable that is
 | 
| 823 | implicitly abstracted over its context of locally bound variables. | |
| 61493 | 824 |     For example, this allows pattern matching of \<open>{x. f x = g
 | 
| 825 |     x}\<close> against \<open>{x. _ = _}\<close>, or even \<open>{_. _ = _}\<close> by
 | |
| 61458 | 826 | using both bound and schematic dummies. | 
| 28773 | 827 | |
| 61503 | 828 | \<^descr> The three literal dots ``\<^verbatim>\<open>...\<close>'' may be also | 
| 829 | written as ellipsis symbol \<^verbatim>\<open>\<dots>\<close>. In both cases this | |
| 28774 | 830 | refers to a special schematic variable, which is bound in the | 
| 831 | context. This special term abbreviation works nicely with | |
| 832 |   calculational reasoning (\secref{sec:calculation}).
 | |
| 833 | ||
| 61503 | 834 | \<^descr> \<^verbatim>\<open>CONST\<close> ensures that the given identifier is treated | 
| 46287 | 835 | as constant term, and passed through the parse tree in fully | 
| 836 | internalized form. This is particularly relevant for translation | |
| 837 |   rules (\secref{sec:syn-trans}), notably on the RHS.
 | |
| 838 | ||
| 61503 | 839 | \<^descr> \<^verbatim>\<open>XCONST\<close> is similar to \<^verbatim>\<open>CONST\<close>, but | 
| 46287 | 840 | retains the constant name as given. This is only relevant to | 
| 841 |   translation rules (\secref{sec:syn-trans}), notably on the LHS.
 | |
| 58618 | 842 | \<close> | 
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 843 | |
| 28777 | 844 | |
| 58618 | 845 | subsection \<open>Inspecting the syntax\<close> | 
| 28777 | 846 | |
| 58618 | 847 | text \<open> | 
| 46282 | 848 |   \begin{matharray}{rcl}
 | 
| 61493 | 849 |     @{command_def "print_syntax"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
 | 
| 46282 | 850 |   \end{matharray}
 | 
| 28777 | 851 | |
| 61439 | 852 |   \<^descr> @{command "print_syntax"} prints the inner syntax of the
 | 
| 46282 | 853 | current context. The output can be quite large; the most important | 
| 854 | sections are explained below. | |
| 28777 | 855 | |
| 61493 | 856 | \<^descr> \<open>lexicon\<close> lists the delimiters of the inner token | 
| 61458 | 857 |     language; see \secref{sec:inner-lex}.
 | 
| 28777 | 858 | |
| 61493 | 859 | \<^descr> \<open>prods\<close> lists the productions of the underlying | 
| 61458 | 860 |     priority grammar; see \secref{sec:priority-grammar}.
 | 
| 28777 | 861 | |
| 61493 | 862 | The nonterminal \<open>A\<^sup>(\<^sup>p\<^sup>)\<close> is rendered in plain text as \<open>A[p]\<close>; delimiters are quoted. Many productions have an extra | 
| 863 | \<open>\<dots> => name\<close>. These names later become the heads of parse | |
| 61458 | 864 | trees; they also guide the pretty printer. | 
| 28777 | 865 | |
| 61477 | 866 | Productions without such parse tree names are called \<^emph>\<open>copy | 
| 867 | productions\<close>. Their right-hand side must have exactly one | |
| 61458 | 868 | nonterminal symbol (or named token). The parser does not create a | 
| 869 | new parse tree node for copy productions, but simply returns the | |
| 870 | parse tree of the right-hand symbol. | |
| 46282 | 871 | |
| 61458 | 872 | If the right-hand side of a copy production consists of a single | 
| 61477 | 873 | nonterminal without any delimiters, then it is called a \<^emph>\<open>chain | 
| 874 | production\<close>. Chain productions act as abbreviations: conceptually, | |
| 61458 | 875 | they are removed from the grammar by adding new productions. | 
| 876 | Priority information attached to chain productions is ignored; only | |
| 61493 | 877 | the dummy value \<open>-1\<close> is displayed. | 
| 46282 | 878 | |
| 61493 | 879 | \<^descr> \<open>print modes\<close> lists the alternative print modes | 
| 61458 | 880 |     provided by this grammar; see \secref{sec:print-modes}.
 | 
| 28777 | 881 | |
| 61493 | 882 | \<^descr> \<open>parse_rules\<close> and \<open>print_rules\<close> relate to | 
| 61458 | 883 |     syntax translations (macros); see \secref{sec:syn-trans}.
 | 
| 46282 | 884 | |
| 61493 | 885 | \<^descr> \<open>parse_ast_translation\<close> and \<open>print_ast_translation\<close> list sets of constants that invoke | 
| 61458 | 886 | translation functions for abstract syntax trees, which are only | 
| 887 |     required in very special situations; see \secref{sec:tr-funs}.
 | |
| 28777 | 888 | |
| 61493 | 889 | \<^descr> \<open>parse_translation\<close> and \<open>print_translation\<close> | 
| 61458 | 890 | list the sets of constants that invoke regular translation | 
| 891 |     functions; see \secref{sec:tr-funs}.
 | |
| 58618 | 892 | \<close> | 
| 28774 | 893 | |
| 28770 
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
 wenzelm parents: 
28769diff
changeset | 894 | |
| 58618 | 895 | subsection \<open>Ambiguity of parsed expressions\<close> | 
| 46291 | 896 | |
| 58618 | 897 | text \<open> | 
| 46291 | 898 |   \begin{tabular}{rcll}
 | 
| 61493 | 899 |     @{attribute_def syntax_ambiguity_warning} & : & \<open>attribute\<close> & default \<open>true\<close> \\
 | 
| 900 |     @{attribute_def syntax_ambiguity_limit} & : & \<open>attribute\<close> & default \<open>10\<close> \\
 | |
| 46291 | 901 |   \end{tabular}
 | 
| 902 | ||
| 903 | Depending on the grammar and the given input, parsing may be | |
| 904 | ambiguous. Isabelle lets the Earley parser enumerate all possible | |
| 905 | parse trees, and then tries to make the best out of the situation. | |
| 906 | Terms that cannot be type-checked are filtered out, which often | |
| 907 | leads to a unique result in the end. Unlike regular type | |
| 908 | reconstruction, which is applied to the whole collection of input | |
| 909 | terms simultaneously, the filtering stage only treats each given | |
| 910 | term in isolation. Filtering is also not attempted for individual | |
| 911 |   types or raw ASTs (as required for @{command translations}).
 | |
| 912 | ||
| 913 | Certain warning or error messages are printed, depending on the | |
| 914 | situation and the given configuration options. Parsing ultimately | |
| 915 | fails, if multiple results remain after the filtering phase. | |
| 916 | ||
| 61439 | 917 |   \<^descr> @{attribute syntax_ambiguity_warning} controls output of
 | 
| 46512 
4f9f61f9b535
simplified configuration options for syntax ambiguity;
 wenzelm parents: 
46506diff
changeset | 918 | explicit warning messages about syntax ambiguity. | 
| 46291 | 919 | |
| 61439 | 920 |   \<^descr> @{attribute syntax_ambiguity_limit} determines the number of
 | 
| 46291 | 921 | resulting parse trees that are shown as part of the printed message | 
| 922 | in case of an ambiguity. | |
| 58618 | 923 | \<close> | 
| 46291 | 924 | |
| 925 | ||
| 58618 | 926 | section \<open>Syntax transformations \label{sec:syntax-transformations}\<close>
 | 
| 48113 | 927 | |
| 58618 | 928 | text \<open>The inner syntax engine of Isabelle provides separate | 
| 52413 | 929 | mechanisms to transform parse trees either via rewrite systems on | 
| 48113 | 930 |   first-order ASTs (\secref{sec:syn-trans}), or ML functions on ASTs
 | 
| 61493 | 931 |   or syntactic \<open>\<lambda>\<close>-terms (\secref{sec:tr-funs}).  This works
 | 
| 48113 | 932 | both for parsing and printing, as outlined in | 
| 933 |   \figref{fig:parse-print}.
 | |
| 934 | ||
| 935 |   \begin{figure}[htbp]
 | |
| 936 |   \begin{center}
 | |
| 937 |   \begin{tabular}{cl}
 | |
| 938 | string & \\ | |
| 61493 | 939 | \<open>\<down>\<close> & lexer + parser \\ | 
| 48113 | 940 | parse tree & \\ | 
| 61493 | 941 | \<open>\<down>\<close> & parse AST translation \\ | 
| 48113 | 942 | AST & \\ | 
| 61493 | 943 | \<open>\<down>\<close> & AST rewriting (macros) \\ | 
| 48113 | 944 | AST & \\ | 
| 61493 | 945 | \<open>\<down>\<close> & parse translation \\ | 
| 48113 | 946 | --- pre-term --- & \\ | 
| 61493 | 947 | \<open>\<down>\<close> & print translation \\ | 
| 48113 | 948 | AST & \\ | 
| 61493 | 949 | \<open>\<down>\<close> & AST rewriting (macros) \\ | 
| 48113 | 950 | AST & \\ | 
| 61493 | 951 | \<open>\<down>\<close> & print AST translation \\ | 
| 48113 | 952 | string & | 
| 953 |   \end{tabular}
 | |
| 954 |   \end{center}
 | |
| 955 |   \caption{Parsing and printing with translations}\label{fig:parse-print}
 | |
| 956 |   \end{figure}
 | |
| 957 | ||
| 958 | These intermediate syntax tree formats eventually lead to a pre-term | |
| 959 | with all names and binding scopes resolved, but most type | |
| 960 | information still missing. Explicit type constraints might be given by | |
| 961 | the user, or implicit position information by the system --- both | |
| 48816 | 962 | need to be passed-through carefully by syntax transformations. | 
| 48113 | 963 | |
| 61477 | 964 | Pre-terms are further processed by the so-called \<^emph>\<open>check\<close> and | 
| 965 | \<^emph>\<open>uncheck\<close> phases that are intertwined with type-inference (see | |
| 58552 | 966 |   also @{cite "isabelle-implementation"}).  The latter allows to operate
 | 
| 48113 | 967 | on higher-order abstract syntax with proper binding and type | 
| 968 | information already available. | |
| 969 | ||
| 970 | As a rule of thumb, anything that manipulates bindings of variables | |
| 971 | or constants needs to be implemented as syntax transformation (see | |
| 972 | below). Anything else is better done via check/uncheck: a prominent | |
| 973 |   example application is the @{command abbreviation} concept of
 | |
| 58618 | 974 | Isabelle/Pure.\<close> | 
| 48113 | 975 | |
| 976 | ||
| 58618 | 977 | subsection \<open>Abstract syntax trees \label{sec:ast}\<close>
 | 
| 48113 | 978 | |
| 58618 | 979 | text \<open>The ML datatype @{ML_type Ast.ast} explicitly represents the
 | 
| 48114 | 980 | intermediate AST format that is used for syntax rewriting | 
| 981 |   (\secref{sec:syn-trans}).  It is defined in ML as follows:
 | |
| 61408 
9020a3ba6c9a
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61143diff
changeset | 982 |   @{verbatim [display]
 | 
| 
9020a3ba6c9a
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61143diff
changeset | 983 | \<open>datatype ast = | 
| 
9020a3ba6c9a
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61143diff
changeset | 984 | Constant of string | | 
| 
9020a3ba6c9a
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61143diff
changeset | 985 | Variable of string | | 
| 
9020a3ba6c9a
@{verbatim [display]} supersedes old alltt/ttbox;
 wenzelm parents: 
61143diff
changeset | 986 | Appl of ast list\<close>} | 
| 48114 | 987 | |
| 988 | An AST is either an atom (constant or variable) or a list of (at | |
| 989 | least two) subtrees. Occasional diagnostic output of ASTs uses | |
| 990 | notation that resembles S-expression of LISP. Constant atoms are | |
| 991 | shown as quoted strings, variable atoms as non-quoted strings and | |
| 992 | applications as a parenthesized list of subtrees. For example, the | |
| 993 | AST | |
| 58724 | 994 |   @{ML [display] \<open>Ast.Appl [Ast.Constant "_abs", Ast.Variable "x", Ast.Variable "t"]\<close>}
 | 
| 61503 | 995 |   is pretty-printed as \<^verbatim>\<open>("_abs" x t)\<close>.  Note that
 | 
| 996 | \<^verbatim>\<open>()\<close> and \<^verbatim>\<open>(x)\<close> are excluded as ASTs, because | |
| 48114 | 997 | they have too few subtrees. | 
| 998 | ||
| 61421 | 999 | \<^medskip> | 
| 1000 | AST application is merely a pro-forma mechanism to indicate | |
| 61503 | 1001 | certain syntactic structures. Thus \<^verbatim>\<open>(c a b)\<close> could mean | 
| 48114 | 1002 | either term application or type application, depending on the | 
| 1003 | syntactic context. | |
| 1004 | ||
| 61503 | 1005 |   Nested application like \<^verbatim>\<open>(("_abs" x t) u)\<close> is also
 | 
| 48114 | 1006 | possible, but ASTs are definitely first-order: the syntax constant | 
| 61503 | 1007 | \<^verbatim>\<open>"_abs"\<close> does not bind the \<^verbatim>\<open>x\<close> in any way. | 
| 48114 | 1008 | Proper bindings are introduced in later stages of the term syntax, | 
| 61503 | 1009 |   where \<^verbatim>\<open>("_abs" x t)\<close> becomes an @{ML Abs} node and
 | 
| 1010 | occurrences of \<^verbatim>\<open>x\<close> in \<^verbatim>\<open>t\<close> are replaced by bound | |
| 48114 | 1011 | variables (represented as de-Bruijn indices). | 
| 58618 | 1012 | \<close> | 
| 48113 | 1013 | |
| 1014 | ||
| 58618 | 1015 | subsubsection \<open>AST constants versus variables\<close> | 
| 48114 | 1016 | |
| 58618 | 1017 | text \<open>Depending on the situation --- input syntax, output syntax, | 
| 56582 | 1018 |   translation patterns --- the distinction of atomic ASTs as @{ML
 | 
| 48114 | 1019 |   Ast.Constant} versus @{ML Ast.Variable} serves slightly different
 | 
| 1020 | purposes. | |
| 1021 | ||
| 61493 | 1022 | Input syntax of a term such as \<open>f a b = c\<close> does not yet | 
| 1023 | indicate the scopes of atomic entities \<open>f, a, b, c\<close>: they | |
| 48114 | 1024 | could be global constants or local variables, even bound ones | 
| 1025 |   depending on the context of the term.  @{ML Ast.Variable} leaves
 | |
| 1026 | this choice still open: later syntax layers (or translation | |
| 1027 | functions) may capture such a variable to determine its role | |
| 1028 | specifically, to make it a constant, bound variable, free variable | |
| 1029 | etc. In contrast, syntax translations that introduce already known | |
| 1030 |   constants would rather do it via @{ML Ast.Constant} to prevent
 | |
| 1031 | accidental re-interpretation later on. | |
| 1032 | ||
| 1033 |   Output syntax turns term constants into @{ML Ast.Constant} and
 | |
| 1034 |   variables (free or schematic) into @{ML Ast.Variable}.  This
 | |
| 61493 | 1035 | information is precise when printing fully formal \<open>\<lambda>\<close>-terms. | 
| 48114 | 1036 | |
| 61421 | 1037 | \<^medskip> | 
| 1038 |   AST translation patterns (\secref{sec:syn-trans}) that
 | |
| 52413 | 1039 | represent terms cannot distinguish constants and variables | 
| 61493 | 1040 | syntactically. Explicit indication of \<open>CONST c\<close> inside the | 
| 1041 | term language is required, unless \<open>c\<close> is known as special | |
| 61477 | 1042 |   \<^emph>\<open>syntax constant\<close> (see also @{command syntax}).  It is also
 | 
| 52413 | 1043 |   possible to use @{command syntax} declarations (without mixfix
 | 
| 1044 | annotation) to enforce that certain unqualified names are always | |
| 1045 | treated as constant within the syntax machinery. | |
| 48114 | 1046 | |
| 52413 | 1047 | The situation is simpler for ASTs that represent types or sorts, | 
| 1048 | since the concrete syntax already distinguishes type variables from | |
| 61493 | 1049 |   type constants (constructors).  So \<open>('a, 'b) foo\<close>
 | 
| 1050 | corresponds to an AST application of some constant for \<open>foo\<close> | |
| 1051 | and variable arguments for \<open>'a\<close> and \<open>'b\<close>. Note that | |
| 52413 | 1052 | the postfix application is merely a feature of the concrete syntax, | 
| 58618 | 1053 | while in the AST the constructor occurs in head position.\<close> | 
| 48114 | 1054 | |
| 1055 | ||
| 58618 | 1056 | subsubsection \<open>Authentic syntax names\<close> | 
| 48114 | 1057 | |
| 58618 | 1058 | text \<open>Naming constant entities within ASTs is another delicate | 
| 52413 | 1059 | issue. Unqualified names are resolved in the name space tables in | 
| 48114 | 1060 | the last stage of parsing, after all translations have been applied. | 
| 1061 | Since syntax transformations do not know about this later name | |
| 52413 | 1062 | resolution, there can be surprises in boundary cases. | 
| 48114 | 1063 | |
| 61477 | 1064 |   \<^emph>\<open>Authentic syntax names\<close> for @{ML Ast.Constant} avoid this
 | 
| 48114 | 1065 | problem: the fully-qualified constant name with a special prefix for | 
| 61493 | 1066 | its formal category (\<open>class\<close>, \<open>type\<close>, \<open>const\<close>, \<open>fixed\<close>) represents the information faithfully | 
| 48114 | 1067 | within the untyped AST format. Accidental overlap with free or | 
| 1068 | bound variables is excluded as well. Authentic syntax names work | |
| 1069 | implicitly in the following situations: | |
| 1070 | ||
| 61421 | 1071 | \<^item> Input of term constants (or fixed variables) that are | 
| 48114 | 1072 |   introduced by concrete syntax via @{command notation}: the
 | 
| 1073 | correspondence of a particular grammar production to some known term | |
| 1074 | entity is preserved. | |
| 1075 | ||
| 61421 | 1076 | \<^item> Input of type constants (constructors) and type classes --- | 
| 48114 | 1077 | thanks to explicit syntactic distinction independently on the | 
| 1078 | context. | |
| 1079 | ||
| 61421 | 1080 | \<^item> Output of term constants, type constants, type classes --- | 
| 48114 | 1081 | this information is already available from the internal term to be | 
| 1082 | printed. | |
| 1083 | ||
| 1084 | ||
| 1085 | In other words, syntax transformations that operate on input terms | |
| 48816 | 1086 | written as prefix applications are difficult to make robust. | 
| 1087 | Luckily, this case rarely occurs in practice, because syntax forms | |
| 58618 | 1088 | to be translated usually correspond to some concrete notation.\<close> | 
| 48114 | 1089 | |
| 1090 | ||
| 58618 | 1091 | subsection \<open>Raw syntax and translations \label{sec:syn-trans}\<close>
 | 
| 28762 | 1092 | |
| 58618 | 1093 | text \<open> | 
| 48117 | 1094 |   \begin{tabular}{rcll}
 | 
| 61493 | 1095 |     @{command_def "nonterminal"} & : & \<open>theory \<rightarrow> theory\<close> \\
 | 
| 1096 |     @{command_def "syntax"} & : & \<open>theory \<rightarrow> theory\<close> \\
 | |
| 1097 |     @{command_def "no_syntax"} & : & \<open>theory \<rightarrow> theory\<close> \\
 | |
| 1098 |     @{command_def "translations"} & : & \<open>theory \<rightarrow> theory\<close> \\
 | |
| 1099 |     @{command_def "no_translations"} & : & \<open>theory \<rightarrow> theory\<close> \\
 | |
| 1100 |     @{attribute_def syntax_ast_trace} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 1101 |     @{attribute_def syntax_ast_stats} & : & \<open>attribute\<close> & default \<open>false\<close> \\
 | |
| 48117 | 1102 |   \end{tabular}
 | 
| 61421 | 1103 | \<^medskip> | 
| 59783 
00b62aa9f430
tuned syntax diagrams -- no duplication of "target";
 wenzelm parents: 
58842diff
changeset | 1104 | |
| 46292 | 1105 | Unlike mixfix notation for existing formal entities | 
| 1106 |   (\secref{sec:notation}), raw syntax declarations provide full access
 | |
| 48115 | 1107 | to the priority grammar of the inner syntax, without any sanity | 
| 1108 | checks. This includes additional syntactic categories (via | |
| 1109 |   @{command nonterminal}) and free-form grammar productions (via
 | |
| 1110 |   @{command syntax}).  Additional syntax translations (or macros, via
 | |
| 1111 |   @{command translations}) are required to turn resulting parse trees
 | |
| 1112 | into proper representations of formal entities again. | |
| 46292 | 1113 | |
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55108diff
changeset | 1114 |   @{rail \<open>
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 1115 |     @@{command nonterminal} (@{syntax name} + @'and')
 | 
| 28762 | 1116 | ; | 
| 46494 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 wenzelm parents: 
46483diff
changeset | 1117 |     (@@{command syntax} | @@{command no_syntax}) @{syntax mode}? (constdecl +)
 | 
| 28762 | 1118 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 1119 |     (@@{command translations} | @@{command no_translations})
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 1120 |       (transpat ('==' | '=>' | '<=' | '\<rightleftharpoons>' | '\<rightharpoonup>' | '\<leftharpoondown>') transpat +)
 | 
| 28762 | 1121 | ; | 
| 1122 | ||
| 46494 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 wenzelm parents: 
46483diff
changeset | 1123 |     constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}?
 | 
| 
ea2ae63336f3
clarified outer syntax "constdecl", which is only local to some rail diagrams;
 wenzelm parents: 
46483diff
changeset | 1124 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 1125 |     mode: ('(' ( @{syntax name} | @'output' | @{syntax name} @'output' ) ')')
 | 
| 28762 | 1126 | ; | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 1127 |     transpat: ('(' @{syntax nameref} ')')? @{syntax string}
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55108diff
changeset | 1128 | \<close>} | 
| 28762 | 1129 | |
| 61493 | 1130 |   \<^descr> @{command "nonterminal"}~\<open>c\<close> declares a type
 | 
| 1131 | constructor \<open>c\<close> (without arguments) to act as purely syntactic | |
| 28762 | 1132 | type: a nonterminal symbol of the inner syntax. | 
| 1133 | ||
| 61493 | 1134 |   \<^descr> @{command "syntax"}~\<open>(mode) c :: \<sigma> (mx)\<close> augments the
 | 
| 46292 | 1135 | priority grammar and the pretty printer table for the given print | 
| 61503 | 1136 |   mode (default \<^verbatim>\<open>""\<close>). An optional keyword @{keyword_ref
 | 
| 46292 | 1137 | "output"} means that only the pretty printer table is affected. | 
| 1138 | ||
| 61493 | 1139 |   Following \secref{sec:mixfix}, the mixfix annotation \<open>mx =
 | 
| 1140 | template ps q\<close> together with type \<open>\<sigma> = \<tau>\<^sub>1 \<Rightarrow> \<dots> \<tau>\<^sub>n \<Rightarrow> \<tau>\<close> and | |
| 1141 | specify a grammar production. The \<open>template\<close> contains | |
| 1142 | delimiter tokens that surround \<open>n\<close> argument positions | |
| 61503 | 1143 | (\<^verbatim>\<open>_\<close>). The latter correspond to nonterminal symbols | 
| 61493 | 1144 | \<open>A\<^sub>i\<close> derived from the argument types \<open>\<tau>\<^sub>i\<close> as | 
| 46292 | 1145 | follows: | 
| 1146 | ||
| 61493 | 1147 | \<^item> \<open>prop\<close> if \<open>\<tau>\<^sub>i = prop\<close> | 
| 46292 | 1148 | |
| 61493 | 1149 | \<^item> \<open>logic\<close> if \<open>\<tau>\<^sub>i = (\<dots>)\<kappa>\<close> for logical type | 
| 1150 | constructor \<open>\<kappa> \<noteq> prop\<close> | |
| 46292 | 1151 | |
| 61493 | 1152 | \<^item> \<open>any\<close> if \<open>\<tau>\<^sub>i = \<alpha>\<close> for type variables | 
| 46292 | 1153 | |
| 61493 | 1154 | \<^item> \<open>\<kappa>\<close> if \<open>\<tau>\<^sub>i = \<kappa>\<close> for nonterminal \<open>\<kappa>\<close> | 
| 61458 | 1155 | (syntactic type constructor) | 
| 46292 | 1156 | |
| 61493 | 1157 | Each \<open>A\<^sub>i\<close> is decorated by priority \<open>p\<^sub>i\<close> from the | 
| 1158 | given list \<open>ps\<close>; missing priorities default to 0. | |
| 46292 | 1159 | |
| 1160 | The resulting nonterminal of the production is determined similarly | |
| 61493 | 1161 | from type \<open>\<tau>\<close>, with priority \<open>q\<close> and default 1000. | 
| 46292 | 1162 | |
| 61421 | 1163 | \<^medskip> | 
| 61493 | 1164 | Parsing via this production produces parse trees \<open>t\<^sub>1, \<dots>, t\<^sub>n\<close> for the argument slots. The resulting parse tree is | 
| 1165 | composed as \<open>c t\<^sub>1 \<dots> t\<^sub>n\<close>, by using the syntax constant \<open>c\<close> of the syntax declaration. | |
| 46292 | 1166 | |
| 1167 | Such syntactic constants are invented on the spot, without formal | |
| 1168 | check wrt.\ existing declarations. It is conventional to use plain | |
| 61493 | 1169 | identifiers prefixed by a single underscore (e.g.\ \<open>_foobar\<close>). Names should be chosen with care, to avoid clashes | 
| 48816 | 1170 | with other syntax declarations. | 
| 46292 | 1171 | |
| 61421 | 1172 | \<^medskip> | 
| 61503 | 1173 | The special case of copy production is specified by \<open>c =\<close>~\<^verbatim>\<open>""\<close> (empty string). | 
| 1174 | It means that the resulting parse tree \<open>t\<close> is copied directly, without any | |
| 46292 | 1175 | further decoration. | 
| 46282 | 1176 | |
| 61493 | 1177 |   \<^descr> @{command "no_syntax"}~\<open>(mode) decls\<close> removes grammar
 | 
| 1178 | declarations (and translations) resulting from \<open>decls\<close>, which | |
| 28762 | 1179 |   are interpreted in the same manner as for @{command "syntax"} above.
 | 
| 46282 | 1180 | |
| 61493 | 1181 |   \<^descr> @{command "translations"}~\<open>rules\<close> specifies syntactic
 | 
| 48115 | 1182 | translation rules (i.e.\ macros) as first-order rewrite rules on | 
| 48816 | 1183 |   ASTs (\secref{sec:ast}).  The theory context maintains two
 | 
| 61503 | 1184 | independent lists translation rules: parse rules (\<^verbatim>\<open>=>\<close> | 
| 1185 | or \<open>\<rightharpoonup>\<close>) and print rules (\<^verbatim>\<open><=\<close> or \<open>\<leftharpoondown>\<close>). | |
| 48115 | 1186 | For convenience, both can be specified simultaneously as parse~/ | 
| 61503 | 1187 | print rules (\<^verbatim>\<open>==\<close> or \<open>\<rightleftharpoons>\<close>). | 
| 48115 | 1188 | |
| 28762 | 1189 | Translation patterns may be prefixed by the syntactic category to be | 
| 61493 | 1190 | used for parsing; the default is \<open>logic\<close> which means that | 
| 48115 | 1191 | regular term syntax is used. Both sides of the syntax translation | 
| 1192 | rule undergo parsing and parse AST translations | |
| 1193 |   \secref{sec:tr-funs}, in order to perform some fundamental
 | |
| 61493 | 1194 | normalization like \<open>\<lambda>x y. b \<leadsto> \<lambda>x. \<lambda>y. b\<close>, but other AST | 
| 61477 | 1195 | translation rules are \<^emph>\<open>not\<close> applied recursively here. | 
| 48115 | 1196 | |
| 1197 | When processing AST patterns, the inner syntax lexer runs in a | |
| 1198 | different mode that allows identifiers to start with underscore. | |
| 1199 | This accommodates the usual naming convention for auxiliary syntax | |
| 1200 | constants --- those that do not have a logical counter part --- by | |
| 1201 | allowing to specify arbitrary AST applications within the term | |
| 1202 | syntax, independently of the corresponding concrete syntax. | |
| 1203 | ||
| 1204 |   Atomic ASTs are distinguished as @{ML Ast.Constant} versus @{ML
 | |
| 1205 | Ast.Variable} as follows: a qualified name or syntax constant | |
| 1206 |   declared via @{command syntax}, or parse tree head of concrete
 | |
| 1207 |   notation becomes @{ML Ast.Constant}, anything else @{ML
 | |
| 61493 | 1208 | Ast.Variable}. Note that \<open>CONST\<close> and \<open>XCONST\<close> within | 
| 48115 | 1209 |   the term language (\secref{sec:pure-grammar}) allow to enforce
 | 
| 1210 | treatment as constants. | |
| 1211 | ||
| 61493 | 1212 | AST rewrite rules \<open>(lhs, rhs)\<close> need to obey the following | 
| 48115 | 1213 | side-conditions: | 
| 1214 | ||
| 61493 | 1215 | \<^item> Rules must be left linear: \<open>lhs\<close> must not contain | 
| 61458 | 1216 |     repeated variables.\footnote{The deeper reason for this is that AST
 | 
| 1217 | equality is not well-defined: different occurrences of the ``same'' | |
| 1218 | AST could be decorated differently by accidental type-constraints or | |
| 1219 | source position information, for example.} | |
| 48115 | 1220 | |
| 61493 | 1221 | \<^item> Every variable in \<open>rhs\<close> must also occur in \<open>lhs\<close>. | 
| 48115 | 1222 | |
| 61493 | 1223 |   \<^descr> @{command "no_translations"}~\<open>rules\<close> removes syntactic
 | 
| 28762 | 1224 | translation rules, which are interpreted in the same manner as for | 
| 1225 |   @{command "translations"} above.
 | |
| 1226 | ||
| 61439 | 1227 |   \<^descr> @{attribute syntax_ast_trace} and @{attribute
 | 
| 48117 | 1228 | syntax_ast_stats} control diagnostic output in the AST normalization | 
| 1229 | process, when translation rules are applied to concrete input or | |
| 1230 | output. | |
| 1231 | ||
| 46293 | 1232 | |
| 1233 | Raw syntax and translations provides a slightly more low-level | |
| 1234 | access to the grammar and the form of resulting parse trees. It is | |
| 1235 | often possible to avoid this untyped macro mechanism, and use | |
| 1236 |   type-safe @{command abbreviation} or @{command notation} instead.
 | |
| 1237 |   Some important situations where @{command syntax} and @{command
 | |
| 1238 | translations} are really need are as follows: | |
| 1239 | ||
| 61421 | 1240 |   \<^item> Iterated replacement via recursive @{command translations}.
 | 
| 46293 | 1241 |   For example, consider list enumeration @{term "[a, b, c, d]"} as
 | 
| 1242 |   defined in theory @{theory List} in Isabelle/HOL.
 | |
| 1243 | ||
| 61421 | 1244 | \<^item> Change of binding status of variables: anything beyond the | 
| 46293 | 1245 |   built-in @{keyword "binder"} mixfix annotation requires explicit
 | 
| 1246 | syntax translations. For example, consider list filter | |
| 1247 |   comprehension @{term "[x \<leftarrow> xs . P]"} as defined in theory @{theory
 | |
| 1248 | List} in Isabelle/HOL. | |
| 61458 | 1249 | \<close> | 
| 46293 | 1250 | |
| 28762 | 1251 | |
| 58618 | 1252 | subsubsection \<open>Applying translation rules\<close> | 
| 48117 | 1253 | |
| 58618 | 1254 | text \<open>As a term is being parsed or printed, an AST is generated as | 
| 48117 | 1255 |   an intermediate form according to \figref{fig:parse-print}.  The AST
 | 
| 1256 | is normalized by applying translation rules in the manner of a | |
| 1257 | first-order term rewriting system. We first examine how a single | |
| 1258 | rule is applied. | |
| 1259 | ||
| 61493 | 1260 | Let \<open>t\<close> be the abstract syntax tree to be normalized and | 
| 1261 | \<open>(lhs, rhs)\<close> some translation rule. A subtree \<open>u\<close> | |
| 1262 | of \<open>t\<close> is called \<^emph>\<open>redex\<close> if it is an instance of \<open>lhs\<close>; in this case the pattern \<open>lhs\<close> is said to match the | |
| 1263 | object \<open>u\<close>. A redex matched by \<open>lhs\<close> may be | |
| 1264 | replaced by the corresponding instance of \<open>rhs\<close>, thus | |
| 1265 | \<^emph>\<open>rewriting\<close> the AST \<open>t\<close>. Matching requires some notion | |
| 61477 | 1266 |   of \<^emph>\<open>place-holders\<close> in rule patterns: @{ML Ast.Variable} serves
 | 
| 48117 | 1267 | this purpose. | 
| 1268 | ||
| 61493 | 1269 | More precisely, the matching of the object \<open>u\<close> against the | 
| 1270 | pattern \<open>lhs\<close> is performed as follows: | |
| 48117 | 1271 | |
| 61493 | 1272 |   \<^item> Objects of the form @{ML Ast.Variable}~\<open>x\<close> or @{ML
 | 
| 1273 |   Ast.Constant}~\<open>x\<close> are matched by pattern @{ML
 | |
| 1274 | Ast.Constant}~\<open>x\<close>. Thus all atomic ASTs in the object are | |
| 48117 | 1275 | treated as (potential) constants, and a successful match makes them | 
| 1276 | actual constants even before name space resolution (see also | |
| 1277 |   \secref{sec:ast}).
 | |
| 1278 | ||
| 61493 | 1279 |   \<^item> Object \<open>u\<close> is matched by pattern @{ML
 | 
| 1280 | Ast.Variable}~\<open>x\<close>, binding \<open>x\<close> to \<open>u\<close>. | |
| 48117 | 1281 | |
| 61493 | 1282 |   \<^item> Object @{ML Ast.Appl}~\<open>us\<close> is matched by @{ML
 | 
| 1283 | Ast.Appl}~\<open>ts\<close> if \<open>us\<close> and \<open>ts\<close> have the | |
| 48117 | 1284 | same length and each corresponding subtree matches. | 
| 1285 | ||
| 61421 | 1286 | \<^item> In every other case, matching fails. | 
| 48117 | 1287 | |
| 1288 | ||
| 61493 | 1289 | A successful match yields a substitution that is applied to \<open>rhs\<close>, generating the instance that replaces \<open>u\<close>. | 
| 48117 | 1290 | |
| 1291 | Normalizing an AST involves repeatedly applying translation rules | |
| 1292 | until none are applicable. This works yoyo-like: top-down, | |
| 1293 | bottom-up, top-down, etc. At each subtree position, rules are | |
| 1294 | chosen in order of appearance in the theory definitions. | |
| 1295 | ||
| 1296 |   The configuration options @{attribute syntax_ast_trace} and
 | |
| 48816 | 1297 |   @{attribute syntax_ast_stats} might help to understand this process
 | 
| 48117 | 1298 | and diagnose problems. | 
| 1299 | ||
| 1300 |   \begin{warn}
 | |
| 1301 | If syntax translation rules work incorrectly, the output of | |
| 61477 | 1302 |   @{command_ref print_syntax} with its \<^emph>\<open>rules\<close> sections reveals the
 | 
| 48117 | 1303 | actual internal forms of AST pattern, without potentially confusing | 
| 1304 | concrete syntax. Recall that AST constants appear as quoted strings | |
| 1305 | and variables without quotes. | |
| 1306 |   \end{warn}
 | |
| 1307 | ||
| 1308 |   \begin{warn}
 | |
| 61493 | 1309 |   If @{attribute_ref eta_contract} is set to \<open>true\<close>, terms
 | 
| 1310 | will be \<open>\<eta>\<close>-contracted \<^emph>\<open>before\<close> the AST rewriter sees | |
| 48117 | 1311 | them. Thus some abstraction nodes needed for print rules to match | 
| 61493 | 1312 | may vanish. For example, \<open>Ball A (\<lambda>x. P x)\<close> would contract | 
| 1313 | to \<open>Ball A P\<close> and the standard print rule would fail to | |
| 48117 | 1314 | apply. This problem can be avoided by hand-written ML translation | 
| 1315 |   functions (see also \secref{sec:tr-funs}), which is in fact the same
 | |
| 1316 |   mechanism used in built-in @{keyword "binder"} declarations.
 | |
| 1317 |   \end{warn}
 | |
| 58618 | 1318 | \<close> | 
| 48117 | 1319 | |
| 28762 | 1320 | |
| 58618 | 1321 | subsection \<open>Syntax translation functions \label{sec:tr-funs}\<close>
 | 
| 28762 | 1322 | |
| 58618 | 1323 | text \<open> | 
| 28762 | 1324 |   \begin{matharray}{rcl}
 | 
| 61493 | 1325 |     @{command_def "parse_ast_translation"} & : & \<open>theory \<rightarrow> theory\<close> \\
 | 
| 1326 |     @{command_def "parse_translation"} & : & \<open>theory \<rightarrow> theory\<close> \\
 | |
| 1327 |     @{command_def "print_translation"} & : & \<open>theory \<rightarrow> theory\<close> \\
 | |
| 1328 |     @{command_def "typed_print_translation"} & : & \<open>theory \<rightarrow> theory\<close> \\
 | |
| 1329 |     @{command_def "print_ast_translation"} & : & \<open>theory \<rightarrow> theory\<close> \\
 | |
| 1330 |     @{ML_antiquotation_def "class_syntax"} & : & \<open>ML antiquotation\<close> \\
 | |
| 1331 |     @{ML_antiquotation_def "type_syntax"} & : & \<open>ML antiquotation\<close> \\
 | |
| 1332 |     @{ML_antiquotation_def "const_syntax"} & : & \<open>ML antiquotation\<close> \\
 | |
| 1333 |     @{ML_antiquotation_def "syntax_const"} & : & \<open>ML antiquotation\<close> \\
 | |
| 28762 | 1334 |   \end{matharray}
 | 
| 1335 | ||
| 48118 | 1336 | Syntax translation functions written in ML admit almost arbitrary | 
| 1337 | manipulations of inner syntax, at the expense of some complexity and | |
| 1338 | obscurity in the implementation. | |
| 1339 | ||
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55108diff
changeset | 1340 |   @{rail \<open>
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 1341 |   ( @@{command parse_ast_translation} | @@{command parse_translation} |
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
42358diff
changeset | 1342 |     @@{command print_translation} | @@{command typed_print_translation} |
 | 
| 52143 | 1343 |     @@{command print_ast_translation}) @{syntax text}
 | 
| 48119 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1344 | ; | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1345 |   (@@{ML_antiquotation class_syntax} |
 | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1346 |    @@{ML_antiquotation type_syntax} |
 | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1347 |    @@{ML_antiquotation const_syntax} |
 | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1348 |    @@{ML_antiquotation syntax_const}) name
 | 
| 55112 
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
 wenzelm parents: 
55108diff
changeset | 1349 | \<close>} | 
| 28762 | 1350 | |
| 61439 | 1351 |   \<^descr> @{command parse_translation} etc. declare syntax translation
 | 
| 48119 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1352 | functions to the theory. Any of these commands have a single | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1353 |   @{syntax text} argument that refers to an ML expression of
 | 
| 52413 | 1354 | appropriate type as follows: | 
| 48118 | 1355 | |
| 61421 | 1356 | \<^medskip> | 
| 48119 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1357 |   {\footnotesize
 | 
| 52143 | 1358 |   \begin{tabular}{l}
 | 
| 1359 |   @{command parse_ast_translation} : \\
 | |
| 1360 |   \quad @{ML_type "(string * (Proof.context -> Ast.ast list -> Ast.ast)) list"} \\
 | |
| 1361 |   @{command parse_translation} : \\
 | |
| 1362 |   \quad @{ML_type "(string * (Proof.context -> term list -> term)) list"} \\
 | |
| 1363 |   @{command print_translation} : \\
 | |
| 1364 |   \quad @{ML_type "(string * (Proof.context -> term list -> term)) list"} \\
 | |
| 1365 |   @{command typed_print_translation} : \\
 | |
| 1366 |   \quad @{ML_type "(string * (Proof.context -> typ -> term list -> term)) list"} \\
 | |
| 1367 |   @{command print_ast_translation} : \\
 | |
| 1368 |   \quad @{ML_type "(string * (Proof.context -> Ast.ast list -> Ast.ast)) list"} \\
 | |
| 48118 | 1369 |   \end{tabular}}
 | 
| 61421 | 1370 | \<^medskip> | 
| 28762 | 1371 | |
| 61493 | 1372 | The argument list consists of \<open>(c, tr)\<close> pairs, where \<open>c\<close> is the syntax name of the formal entity involved, and \<open>tr\<close> a function that translates a syntax form \<open>c args\<close> into | 
| 1373 | \<open>tr ctxt args\<close> (depending on the context). The Isabelle/ML | |
| 1374 | naming convention for parse translations is \<open>c_tr\<close> and for | |
| 1375 | print translations \<open>c_tr'\<close>. | |
| 48118 | 1376 | |
| 1377 |   The @{command_ref print_syntax} command displays the sets of names
 | |
| 61493 | 1378 | associated with the translation functions of a theory under \<open>parse_ast_translation\<close> etc. | 
| 48118 | 1379 | |
| 61493 | 1380 |   \<^descr> \<open>@{class_syntax c}\<close>, \<open>@{type_syntax c}\<close>,
 | 
| 1381 |   \<open>@{const_syntax c}\<close> inline the authentic syntax name of the
 | |
| 48119 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1382 | given formal entities into the ML source. This is the | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1383 | fully-qualified logical name prefixed by a special marker to | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1384 | indicate its kind: thus different logical name spaces are properly | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1385 | distinguished within parse trees. | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1386 | |
| 61493 | 1387 |   \<^descr> \<open>@{const_syntax c}\<close> inlines the name \<open>c\<close> of
 | 
| 48119 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1388 | the given syntax constant, having checked that it has been declared | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1389 |   via some @{command syntax} commands within the theory context.  Note
 | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1390 | that the usual naming convention makes syntax constants start with | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1391 | underscore, to reduce the chance of accidental clashes with other | 
| 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1392 | names occurring in parse trees (unqualified constants etc.). | 
| 58618 | 1393 | \<close> | 
| 48118 | 1394 | |
| 48119 
55c305e29f4b
cover @{class_syntax}, @{type_syntax}, @{const_syntax}, @{syntax_const} in isar-ref, in contrast to other ML antiquotations in implementation manual;
 wenzelm parents: 
48118diff
changeset | 1395 | |
| 58618 | 1396 | subsubsection \<open>The translation strategy\<close> | 
| 28762 | 1397 | |
| 58618 | 1398 | text \<open>The different kinds of translation functions are invoked during | 
| 48118 | 1399 | the transformations between parse trees, ASTs and syntactic terms | 
| 1400 |   (cf.\ \figref{fig:parse-print}).  Whenever a combination of the form
 | |
| 61493 | 1401 | \<open>c x\<^sub>1 \<dots> x\<^sub>n\<close> is encountered, and a translation function | 
| 1402 | \<open>f\<close> of appropriate kind is declared for \<open>c\<close>, the | |
| 1403 | result is produced by evaluation of \<open>f [x\<^sub>1, \<dots>, x\<^sub>n]\<close> in ML. | |
| 48118 | 1404 | |
| 61493 | 1405 | For AST translations, the arguments \<open>x\<^sub>1, \<dots>, x\<^sub>n\<close> are ASTs. A | 
| 1406 |   combination has the form @{ML "Ast.Constant"}~\<open>c\<close> or @{ML
 | |
| 1407 |   "Ast.Appl"}~\<open>[\<close>@{ML Ast.Constant}~\<open>c, x\<^sub>1, \<dots>, x\<^sub>n]\<close>.
 | |
| 48118 | 1408 | For term translations, the arguments are terms and a combination has | 
| 61493 | 1409 |   the form @{ML Const}~\<open>(c, \<tau>)\<close> or @{ML Const}~\<open>(c, \<tau>)
 | 
| 1410 | $ x\<^sub>1 $ \<dots> $ x\<^sub>n\<close>. Terms allow more sophisticated transformations | |
| 48118 | 1411 | than ASTs do, typically involving abstractions and bound | 
| 61477 | 1412 | variables. \<^emph>\<open>Typed\<close> print translations may even peek at the type | 
| 61493 | 1413 | \<open>\<tau>\<close> of the constant they are invoked on, although some | 
| 52413 | 1414 | information might have been suppressed for term output already. | 
| 48118 | 1415 | |
| 1416 | Regardless of whether they act on ASTs or terms, translation | |
| 1417 | functions called during the parsing process differ from those for | |
| 1418 | printing in their overall behaviour: | |
| 1419 | ||
| 61439 | 1420 | \<^descr>[Parse translations] are applied bottom-up. The arguments are | 
| 48118 | 1421 | already in translated form. The translations must not fail; | 
| 1422 | exceptions trigger an error message. There may be at most one | |
| 1423 | function associated with any syntactic name. | |
| 46294 | 1424 | |
| 61439 | 1425 | \<^descr>[Print translations] are applied top-down. They are supplied | 
| 48118 | 1426 | with arguments that are partly still in internal form. The result | 
| 1427 | again undergoes translation; therefore a print translation should | |
| 1428 | not introduce as head the very constant that invoked it. The | |
| 1429 |   function may raise exception @{ML Match} to indicate failure; in
 | |
| 1430 | this event it has no effect. Multiple functions associated with | |
| 1431 | some syntactic name are tried in the order of declaration in the | |
| 1432 | theory. | |
| 1433 | ||
| 1434 | ||
| 1435 |   Only constant atoms --- constructor @{ML Ast.Constant} for ASTs and
 | |
| 1436 |   @{ML Const} for terms --- can invoke translation functions.  This
 | |
| 1437 | means that parse translations can only be associated with parse tree | |
| 1438 | heads of concrete syntax, or syntactic constants introduced via | |
| 1439 | other translations. For plain identifiers within the term language, | |
| 1440 | the status of constant versus variable is not yet know during | |
| 1441 | parsing. This is in contrast to print translations, where constants | |
| 1442 | are explicitly known from the given term in its fully internal form. | |
| 58618 | 1443 | \<close> | 
| 28762 | 1444 | |
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1445 | |
| 58618 | 1446 | subsection \<open>Built-in syntax transformations\<close> | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1447 | |
| 58618 | 1448 | text \<open> | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1449 | Here are some further details of the main syntax transformation | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1450 |   phases of \figref{fig:parse-print}.
 | 
| 58618 | 1451 | \<close> | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1452 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1453 | |
| 58618 | 1454 | subsubsection \<open>Transforming parse trees to ASTs\<close> | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1455 | |
| 58618 | 1456 | text \<open>The parse tree is the raw output of the parser. It is | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1457 | transformed into an AST according to some basic scheme that may be | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1458 | augmented by AST translation functions as explained in | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1459 |   \secref{sec:tr-funs}.
 | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1460 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1461 | The parse tree is constructed by nesting the right-hand sides of the | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1462 | productions used to recognize the input. Such parse trees are | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1463 | simply lists of tokens and constituent parse trees, the latter | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1464 | representing the nonterminals of the productions. Ignoring AST | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1465 | translation functions, parse trees are transformed to ASTs by | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1466 | stripping out delimiters and copy productions, while retaining some | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1467 | source position information from input tokens. | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1468 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1469 | The Pure syntax provides predefined AST translations to make the | 
| 61493 | 1470 | basic \<open>\<lambda>\<close>-term structure more apparent within the | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1471 | (first-order) AST representation, and thus facilitate the use of | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1472 |   @{command translations} (see also \secref{sec:syn-trans}).  This
 | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1473 | covers ordinary term application, type application, nested | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1474 | abstraction, iterated meta implications and function types. The | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1475 | effect is illustrated on some representative input strings is as | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1476 | follows: | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1477 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1478 |   \begin{center}
 | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1479 |   \begin{tabular}{ll}
 | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1480 | input source & AST \\ | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1481 | \hline | 
| 61503 | 1482 | \<open>f x y z\<close> & \<^verbatim>\<open>(f x y z)\<close> \\ | 
| 1483 | \<open>'a ty\<close> & \<^verbatim>\<open>(ty 'a)\<close> \\ | |
| 1484 |   \<open>('a, 'b)ty\<close> & \<^verbatim>\<open>(ty 'a 'b)\<close> \\
 | |
| 1485 |   \<open>\<lambda>x y z. t\<close> & \<^verbatim>\<open>("_abs" x ("_abs" y ("_abs" z t)))\<close> \\
 | |
| 1486 |   \<open>\<lambda>x :: 'a. t\<close> & \<^verbatim>\<open>("_abs" ("_constrain" x 'a) t)\<close> \\
 | |
| 1487 |   \<open>\<lbrakk>P; Q; R\<rbrakk> \<Longrightarrow> S\<close> & \<^verbatim>\<open>("Pure.imp" P ("Pure.imp" Q ("Pure.imp" R S)))\<close> \\
 | |
| 1488 |    \<open>['a, 'b, 'c] \<Rightarrow> 'd\<close> & \<^verbatim>\<open>("fun" 'a ("fun" 'b ("fun" 'c 'd)))\<close> \\
 | |
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1489 |   \end{tabular}
 | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1490 |   \end{center}
 | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1491 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1492 | Note that type and sort constraints may occur in further places --- | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1493 | translations need to be ready to cope with them. The built-in | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1494 | syntax transformation from parse trees to ASTs insert additional | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1495 | constraints that represent source positions. | 
| 58618 | 1496 | \<close> | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1497 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1498 | |
| 58618 | 1499 | subsubsection \<open>Transforming ASTs to terms\<close> | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1500 | |
| 58618 | 1501 | text \<open>After application of macros (\secref{sec:syn-trans}), the AST
 | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1502 | is transformed into a term. This term still lacks proper type | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1503 | information, but it might contain some constraints consisting of | 
| 61503 | 1504 | applications with head \<^verbatim>\<open>_constrain\<close>, where the second | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1505 | argument is a type encoded as a pre-term within the syntax. Type | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1506 | inference later introduces correct types, or indicates type errors | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1507 | in the input. | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1508 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1509 | Ignoring parse translations, ASTs are transformed to terms by | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1510 | mapping AST constants to term constants, AST variables to term | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1511 | variables or constants (according to the name space), and AST | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1512 | applications to iterated term applications. | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1513 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1514 | The outcome is still a first-order term. Proper abstractions and | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1515 | bound variables are introduced by parse translations associated with | 
| 61503 | 1516 |   certain syntax constants.  Thus \<^verbatim>\<open>("_abs" x x)\<close> eventually
 | 
| 1517 |   becomes a de-Bruijn term \<^verbatim>\<open>Abs ("x", _, Bound 0)\<close>.
 | |
| 58618 | 1518 | \<close> | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1519 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1520 | |
| 58618 | 1521 | subsubsection \<open>Printing of terms\<close> | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1522 | |
| 58618 | 1523 | text \<open>The output phase is essentially the inverse of the input | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1524 | phase. Terms are translated via abstract syntax trees into | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1525 | pretty-printed text. | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1526 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1527 | Ignoring print translations, the transformation maps term constants, | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1528 | variables and applications to the corresponding constructs on ASTs. | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1529 | Abstractions are mapped to applications of the special constant | 
| 61503 | 1530 | \<^verbatim>\<open>_abs\<close> as seen before. Type constraints are represented | 
| 1531 | via special \<^verbatim>\<open>_constrain\<close> forms, according to various | |
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1532 | policies of type annotation determined elsewhere. Sort constraints | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1533 | of type variables are handled in a similar fashion. | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1534 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1535 |   After application of macros (\secref{sec:syn-trans}), the AST is
 | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1536 | finally pretty-printed. The built-in print AST translations reverse | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1537 | the corresponding parse AST translations. | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1538 | |
| 61421 | 1539 | \<^medskip> | 
| 1540 | For the actual printing process, the priority grammar | |
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1541 |   (\secref{sec:priority-grammar}) plays a vital role: productions are
 | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1542 | used as templates for pretty printing, with argument slots stemming | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1543 | from nonterminals, and syntactic sugar stemming from literal tokens. | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1544 | |
| 61493 | 1545 | Each AST application with constant head \<open>c\<close> and arguments | 
| 1546 | \<open>t\<^sub>1\<close>, \dots, \<open>t\<^sub>n\<close> (for \<open>n = 0\<close> the AST is | |
| 1547 | just the constant \<open>c\<close> itself) is printed according to the | |
| 1548 | first grammar production of result name \<open>c\<close>. The required | |
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1549 | syntax priority of the argument slot is given by its nonterminal | 
| 61493 | 1550 | \<open>A\<^sup>(\<^sup>p\<^sup>)\<close>. The argument \<open>t\<^sub>i\<close> that corresponds to the | 
| 1551 | position of \<open>A\<^sup>(\<^sup>p\<^sup>)\<close> is printed recursively, and then put in | |
| 1552 | parentheses \<^emph>\<open>if\<close> its priority \<open>p\<close> requires this. The | |
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1553 | resulting output is concatenated with the syntactic sugar according | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1554 | to the grammar production. | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1555 | |
| 61493 | 1556 | If an AST application \<open>(c x\<^sub>1 \<dots> x\<^sub>m)\<close> has more arguments than | 
| 1557 | the corresponding production, it is first split into \<open>((c x\<^sub>1 | |
| 1558 | \<dots> x\<^sub>n) x\<^sub>n\<^sub>+\<^sub>1 \<dots> x\<^sub>m)\<close> and then printed recursively as above. | |
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1559 | |
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1560 | Applications with too few arguments or with non-constant head or | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1561 | without a corresponding production are printed in prefix-form like | 
| 61493 | 1562 | \<open>f t\<^sub>1 \<dots> t\<^sub>n\<close> for terms. | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1563 | |
| 61493 | 1564 | Multiple productions associated with some name \<open>c\<close> are tried | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1565 | in order of appearance within the grammar. An occurrence of some | 
| 61493 | 1566 | AST variable \<open>x\<close> is printed as \<open>x\<close> outright. | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1567 | |
| 61421 | 1568 | \<^medskip> | 
| 61477 | 1569 | White space is \<^emph>\<open>not\<close> inserted automatically. If | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1570 | blanks (or breaks) are required to separate tokens, they need to be | 
| 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1571 |   specified in the mixfix declaration (\secref{sec:mixfix}).
 | 
| 58618 | 1572 | \<close> | 
| 52414 
8429123bc58a
more on built-in syntax transformations, based on reduced version of old material;
 wenzelm parents: 
52413diff
changeset | 1573 | |
| 28762 | 1574 | end |