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