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