author | wenzelm |
Thu, 29 Apr 2010 17:47:53 +0200 | |
changeset 36508 | 03d2a2d0ee4a |
parent 35413 | 4c7cba1f7ce9 |
child 36745 | 403585a89772 |
permissions | -rw-r--r-- |
28762 | 1 |
theory Inner_Syntax |
2 |
imports Main |
|
3 |
begin |
|
4 |
||
28778 | 5 |
chapter {* Inner syntax --- the term language \label{ch:inner-syntax} *} |
28762 | 6 |
|
7 |
section {* Printing logical entities *} |
|
8 |
||
9 |
subsection {* Diagnostic commands *} |
|
10 |
||
11 |
text {* |
|
12 |
\begin{matharray}{rcl} |
|
28766
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
13 |
@{command_def "typ"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
28762 | 14 |
@{command_def "term"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
15 |
@{command_def "prop"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
|
28766
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
16 |
@{command_def "thm"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
28762 | 17 |
@{command_def "prf"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
18 |
@{command_def "full_prf"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
|
28766
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
19 |
@{command_def "pr"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\ |
28762 | 20 |
\end{matharray} |
21 |
||
22 |
These diagnostic commands assist interactive development by printing |
|
23 |
internal logical entities in a human-readable fashion. |
|
24 |
||
25 |
\begin{rail} |
|
28766
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
26 |
'typ' modes? type |
28762 | 27 |
; |
28 |
'term' modes? term |
|
29 |
; |
|
30 |
'prop' modes? prop |
|
31 |
; |
|
28766
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
32 |
'thm' modes? thmrefs |
28762 | 33 |
; |
28766
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
34 |
( 'prf' | 'full\_prf' ) modes? thmrefs? |
28762 | 35 |
; |
28766
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
36 |
'pr' modes? nat? (',' nat)? |
28762 | 37 |
; |
38 |
||
39 |
modes: '(' (name + ) ')' |
|
40 |
; |
|
41 |
\end{rail} |
|
42 |
||
43 |
\begin{description} |
|
44 |
||
28766
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
45 |
\item @{command "typ"}~@{text \<tau>} reads and prints types of the |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
46 |
meta-logic according to the current theory or proof context. |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
47 |
|
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
48 |
\item @{command "term"}~@{text t} and @{command "prop"}~@{text \<phi>} |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
49 |
read, type-check and print terms or propositions according to the |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
50 |
current theory or proof context; the inferred type of @{text t} is |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
51 |
output as well. Note that these commands are also useful in |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
52 |
inspecting the current environment of term abbreviations. |
28762 | 53 |
|
54 |
\item @{command "thm"}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} retrieves |
|
55 |
theorems from the current theory or proof context. Note that any |
|
56 |
attributes included in the theorem specifications are applied to a |
|
57 |
temporary context derived from the current theory or proof; the |
|
58 |
result is discarded, i.e.\ attributes involved in @{text "a\<^sub>1, |
|
59 |
\<dots>, a\<^sub>n"} do not have any permanent effect. |
|
60 |
||
61 |
\item @{command "prf"} displays the (compact) proof term of the |
|
62 |
current proof state (if present), or of the given theorems. Note |
|
63 |
that this requires proof terms to be switched on for the current |
|
64 |
object logic (see the ``Proof terms'' section of the Isabelle |
|
65 |
reference manual for information on how to do this). |
|
66 |
||
67 |
\item @{command "full_prf"} is like @{command "prf"}, but displays |
|
68 |
the full proof term, i.e.\ also displays information omitted in the |
|
69 |
compact proof term, which is denoted by ``@{text _}'' placeholders |
|
70 |
there. |
|
71 |
||
28766
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
72 |
\item @{command "pr"}~@{text "goals, prems"} prints the current |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
73 |
proof state (if present), including the proof context, current facts |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
74 |
and goals. The optional limit arguments affect the number of goals |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
75 |
and premises to be displayed, which is initially 10 for both. |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
76 |
Omitting limit values leaves the current setting unchanged. |
accab7594b8e
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28765
diff
changeset
|
77 |
|
28762 | 78 |
\end{description} |
79 |
||
80 |
All of the diagnostic commands above admit a list of @{text modes} |
|
81 |
to be specified, which is appended to the current print mode (see |
|
82 |
also \cite{isabelle-ref}). Thus the output behavior may be modified |
|
83 |
according particular print mode features. For example, @{command |
|
84 |
"pr"}~@{text "(latex xsymbols)"} would print the current proof state |
|
85 |
with mathematical symbols and special characters represented in |
|
86 |
{\LaTeX} source, according to the Isabelle style |
|
87 |
\cite{isabelle-sys}. |
|
88 |
||
89 |
Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more |
|
90 |
systematic way to include formal items into the printed text |
|
91 |
document. |
|
92 |
*} |
|
93 |
||
94 |
||
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
95 |
subsection {* Details of printed content *} |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
96 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
97 |
text {* |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
98 |
\begin{mldecls} |
32833 | 99 |
@{index_ML show_types: "bool Unsynchronized.ref"} & default @{ML false} \\ |
100 |
@{index_ML show_sorts: "bool Unsynchronized.ref"} & default @{ML false} \\ |
|
101 |
@{index_ML show_consts: "bool Unsynchronized.ref"} & default @{ML false} \\ |
|
102 |
@{index_ML long_names: "bool Unsynchronized.ref"} & default @{ML false} \\ |
|
103 |
@{index_ML short_names: "bool Unsynchronized.ref"} & default @{ML false} \\ |
|
104 |
@{index_ML unique_names: "bool Unsynchronized.ref"} & default @{ML true} \\ |
|
105 |
@{index_ML show_brackets: "bool Unsynchronized.ref"} & default @{ML false} \\ |
|
106 |
@{index_ML eta_contract: "bool Unsynchronized.ref"} & default @{ML true} \\ |
|
107 |
@{index_ML goals_limit: "int Unsynchronized.ref"} & default @{ML 10} \\ |
|
108 |
@{index_ML Proof.show_main_goal: "bool Unsynchronized.ref"} & default @{ML false} \\ |
|
109 |
@{index_ML show_hyps: "bool Unsynchronized.ref"} & default @{ML false} \\ |
|
110 |
@{index_ML show_tags: "bool Unsynchronized.ref"} & default @{ML false} \\ |
|
111 |
@{index_ML show_question_marks: "bool Unsynchronized.ref"} & default @{ML true} \\ |
|
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
112 |
\end{mldecls} |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
113 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
114 |
These global ML variables control the detail of information that is |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
115 |
displayed for types, terms, theorems, goals etc. |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
116 |
|
28765
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
117 |
In interactive sessions, the user interface usually manages these |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
118 |
global parameters of the Isabelle process, even with some concept of |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
119 |
persistence. Nonetheless it is occasionally useful to manipulate ML |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
120 |
variables directly, e.g.\ using @{command "ML_val"} or @{command |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
121 |
"ML_command"}. |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
122 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
123 |
Batch-mode logic sessions may be configured by putting appropriate |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
124 |
ML text directly into the @{verbatim ROOT.ML} file. |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
125 |
|
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
126 |
\begin{description} |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
127 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
128 |
\item @{ML show_types} and @{ML show_sorts} control printing of type |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
129 |
constraints for term variables, and sort constraints for type |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
130 |
variables. By default, neither of these are shown in output. If |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
131 |
@{ML show_sorts} is set to @{ML true}, types are always shown as |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
132 |
well. |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
133 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
134 |
Note that displaying types and sorts may explain why a polymorphic |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
135 |
inference rule fails to resolve with some goal, or why a rewrite |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
136 |
rule does not apply as expected. |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
137 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
138 |
\item @{ML show_consts} controls printing of types of constants when |
28765
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
139 |
displaying a goal state. |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
140 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
141 |
Note that the output can be enormous, because polymorphic constants |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
142 |
often occur at several different type instances. |
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
143 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
144 |
\item @{ML long_names}, @{ML short_names}, and @{ML unique_names} |
28765
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
145 |
control the way of printing fully qualified internal names in |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
146 |
external form. See also \secref{sec:antiq} for the document |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
147 |
antiquotation options of the same names. |
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
148 |
|
28765
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
149 |
\item @{ML show_brackets} controls bracketing in pretty printed |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
150 |
output. If set to @{ML true}, all sub-expressions of the pretty |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
151 |
printing tree will be parenthesized, even if this produces malformed |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
152 |
term syntax! This crude way of showing the internal structure of |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
153 |
pretty printed entities may occasionally help to diagnose problems |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
154 |
with operator priorities, for example. |
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
155 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
156 |
\item @{ML eta_contract} controls @{text "\<eta>"}-contracted printing of |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
157 |
terms. |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
158 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
159 |
The @{text \<eta>}-contraction law asserts @{prop "(\<lambda>x. f x) \<equiv> f"}, |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
160 |
provided @{text x} is not free in @{text f}. It asserts |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
161 |
\emph{extensionality} of functions: @{prop "f \<equiv> g"} if @{prop "f x \<equiv> |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
162 |
g x"} for all @{text x}. Higher-order unification frequently puts |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
163 |
terms into a fully @{text \<eta>}-expanded form. For example, if @{text |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
164 |
F} has type @{text "(\<tau> \<Rightarrow> \<tau>) \<Rightarrow> \<tau>"} then its expanded form is @{term |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
165 |
"\<lambda>h. F (\<lambda>x. h x)"}. |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
166 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
167 |
Setting @{ML eta_contract} makes Isabelle perform @{text |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
168 |
\<eta>}-contractions before printing, so that @{term "\<lambda>h. F (\<lambda>x. h x)"} |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
169 |
appears simply as @{text F}. |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
170 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
171 |
Note that the distinction between a term and its @{text \<eta>}-expanded |
28765
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
172 |
form occasionally matters. While higher-order resolution and |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
173 |
rewriting operate modulo @{text "\<alpha>\<beta>\<eta>"}-conversion, some other tools |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
174 |
might look at terms more discretely. |
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
175 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
176 |
\item @{ML goals_limit} controls the maximum number of subgoals to |
28765
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
177 |
be shown in goal output. |
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
178 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
179 |
\item @{ML Proof.show_main_goal} controls whether the main result to |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
180 |
be proven should be displayed. This information might be relevant |
28765
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
181 |
for schematic goals, to inspect the current claim that has been |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
182 |
synthesized so far. |
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
183 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
184 |
\item @{ML show_hyps} controls printing of implicit hypotheses of |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
185 |
local facts. Normally, only those hypotheses are displayed that are |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
186 |
\emph{not} covered by the assumptions of the current context: this |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
187 |
situation indicates a fault in some tool being used. |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
188 |
|
28765
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
189 |
By setting @{ML show_hyps} to @{ML true}, output of \emph{all} |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
190 |
hypotheses can be enforced, which is occasionally useful for |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
191 |
diagnostic purposes. |
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
192 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
193 |
\item @{ML show_tags} controls printing of extra annotations within |
28765
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
194 |
theorems, such as internal position information, or the case names |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
195 |
being attached by the attribute @{attribute case_names}. |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
196 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
197 |
Note that the @{attribute tagged} and @{attribute untagged} |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
198 |
attributes provide low-level access to the collection of tags |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
199 |
associated with a theorem. |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
200 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
201 |
\item @{ML show_question_marks} controls printing of question marks |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
202 |
for schematic variables, such as @{text ?x}. Only the leading |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
203 |
question mark is affected, the remaining text is unchanged |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
204 |
(including proper markup for schematic variables that might be |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
205 |
relevant for user interfaces). |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
206 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
207 |
\end{description} |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
208 |
*} |
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 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
211 |
subsection {* Printing limits *} |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
212 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
213 |
text {* |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
214 |
\begin{mldecls} |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
215 |
@{index_ML Pretty.setdepth: "int -> unit"} \\ |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
216 |
@{index_ML Pretty.setmargin: "int -> unit"} \\ |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
217 |
@{index_ML print_depth: "int -> unit"} \\ |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
218 |
\end{mldecls} |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
219 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
220 |
These ML functions set limits for pretty printed text. |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
221 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
222 |
\begin{description} |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
223 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
224 |
\item @{ML Pretty.setdepth}~@{text d} tells the pretty printer to |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
225 |
limit the printing depth to @{text d}. This affects the display of |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
226 |
types, terms, theorems etc. The default value is 0, which permits |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
227 |
printing to an arbitrary depth. Other useful values for @{text d} |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
228 |
are 10 and 20. |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
229 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
230 |
\item @{ML Pretty.setmargin}~@{text m} tells the pretty printer to |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
231 |
assume a right margin (page width) of @{text m}. The initial margin |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
232 |
is 76, but user interfaces might adapt the margin automatically when |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
233 |
resizing windows. |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
234 |
|
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
235 |
\item @{ML print_depth}~@{text n} limits the printing depth of the |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
236 |
ML toplevel pretty printer; the precise effect depends on the ML |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
237 |
compiler and run-time system. Typically @{text n} should be less |
da8f6f4a74be
misc tuning and rearrangement of section "Printing logical entities";
wenzelm
parents:
28763
diff
changeset
|
238 |
than 10. Bigger values such as 100--1000 are useful for debugging. |
28763
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
239 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
240 |
\end{description} |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
241 |
*} |
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
242 |
|
b5e6122ff575
added pretty printing options (from old ref manual);
wenzelm
parents:
28762
diff
changeset
|
243 |
|
28762 | 244 |
section {* Mixfix annotations *} |
245 |
||
246 |
text {* Mixfix annotations specify concrete \emph{inner syntax} of |
|
35351
7425aece4ee3
allow general mixfix syntax for type constructors;
wenzelm
parents:
32833
diff
changeset
|
247 |
Isabelle types and terms. Locally fixed parameters in toplevel |
7425aece4ee3
allow general mixfix syntax for type constructors;
wenzelm
parents:
32833
diff
changeset
|
248 |
theorem statements, locale specifications etc.\ also admit mixfix |
7425aece4ee3
allow general mixfix syntax for type constructors;
wenzelm
parents:
32833
diff
changeset
|
249 |
annotations. |
28762 | 250 |
|
251 |
\indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix} |
|
252 |
\begin{rail} |
|
253 |
infix: '(' ('infix' | 'infixl' | 'infixr') string nat ')' |
|
254 |
; |
|
255 |
mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')' |
|
256 |
; |
|
257 |
structmixfix: mixfix | '(' 'structure' ')' |
|
258 |
; |
|
259 |
||
260 |
prios: '[' (nat + ',') ']' |
|
261 |
; |
|
262 |
\end{rail} |
|
263 |
||
264 |
Here the \railtok{string} specifications refer to the actual mixfix |
|
265 |
template, which may include literal text, spacing, blocks, and |
|
266 |
arguments (denoted by ``@{text _}''); the special symbol |
|
267 |
``@{verbatim "\<index>"}'' (printed as ``@{text "\<index>"}'') represents an index |
|
268 |
argument that specifies an implicit structure reference (see also |
|
269 |
\secref{sec:locale}). Infix and binder declarations provide common |
|
270 |
abbreviations for particular mixfix declarations. So in practice, |
|
271 |
mixfix templates mostly degenerate to literal text for concrete |
|
272 |
syntax, such as ``@{verbatim "++"}'' for an infix symbol. |
|
273 |
||
274 |
\medskip In full generality, mixfix declarations work as follows. |
|
275 |
Suppose a constant @{text "c :: \<tau>\<^sub>1 \<Rightarrow> \<dots> \<tau>\<^sub>n \<Rightarrow> \<tau>"} is |
|
276 |
annotated by @{text "(mixfix [p\<^sub>1, \<dots>, p\<^sub>n] p)"}, where @{text |
|
277 |
"mixfix"} is a string @{text "d\<^sub>0 _ d\<^sub>1 _ \<dots> _ d\<^sub>n"} consisting of |
|
278 |
delimiters that surround argument positions as indicated by |
|
279 |
underscores. |
|
280 |
||
281 |
Altogether this determines a production for a context-free priority |
|
282 |
grammar, where for each argument @{text "i"} the syntactic category |
|
283 |
is determined by @{text "\<tau>\<^sub>i"} (with priority @{text "p\<^sub>i"}), and |
|
284 |
the result category is determined from @{text "\<tau>"} (with |
|
285 |
priority @{text "p"}). Priority specifications are optional, with |
|
286 |
default 0 for arguments and 1000 for the result. |
|
287 |
||
288 |
Since @{text "\<tau>"} may be again a function type, the constant |
|
289 |
type scheme may have more argument positions than the mixfix |
|
290 |
pattern. Printing a nested application @{text "c t\<^sub>1 \<dots> t\<^sub>m"} for |
|
291 |
@{text "m > n"} works by attaching concrete notation only to the |
|
292 |
innermost part, essentially by printing @{text "(c t\<^sub>1 \<dots> t\<^sub>n) \<dots> t\<^sub>m"} |
|
293 |
instead. If a term has fewer arguments than specified in the mixfix |
|
294 |
template, the concrete syntax is ignored. |
|
295 |
||
296 |
\medskip A mixfix template may also contain additional directives |
|
297 |
for pretty printing, notably spaces, blocks, and breaks. The |
|
298 |
general template format is a sequence over any of the following |
|
299 |
entities. |
|
300 |
||
28778 | 301 |
\begin{description} |
28762 | 302 |
|
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
303 |
\item @{text "d"} is a delimiter, namely a non-empty sequence of |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
304 |
characters other than the following special characters: |
28762 | 305 |
|
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
306 |
\smallskip |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
307 |
\begin{tabular}{ll} |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
308 |
@{verbatim "'"} & single quote \\ |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
309 |
@{verbatim "_"} & underscore \\ |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
310 |
@{text "\<index>"} & index symbol \\ |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
311 |
@{verbatim "("} & open parenthesis \\ |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
312 |
@{verbatim ")"} & close parenthesis \\ |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
313 |
@{verbatim "/"} & slash \\ |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
314 |
\end{tabular} |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
315 |
\medskip |
28762 | 316 |
|
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
317 |
\item @{verbatim "'"} escapes the special meaning of these |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
318 |
meta-characters, producing a literal version of the following |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
319 |
character, unless that is a blank. |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
320 |
|
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
321 |
A single quote followed by a blank separates delimiters, without |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
322 |
affecting printing, but input tokens may have additional white space |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
323 |
here. |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
324 |
|
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
325 |
\item @{verbatim "_"} is an argument position, which stands for a |
28762 | 326 |
certain syntactic category in the underlying grammar. |
327 |
||
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
328 |
\item @{text "\<index>"} is an indexed argument position; this is the place |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
329 |
where implicit structure arguments can be attached. |
28762 | 330 |
|
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
331 |
\item @{text "s"} is a non-empty sequence of spaces for printing. |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
332 |
This and the following specifications do not affect parsing at all. |
28762 | 333 |
|
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
334 |
\item @{verbatim "("}@{text n} opens a pretty printing block. The |
28762 | 335 |
optional number specifies how much indentation to add when a line |
336 |
break occurs within the block. If the parenthesis is not followed |
|
337 |
by digits, the indentation defaults to 0. A block specified via |
|
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
338 |
@{verbatim "(00"} is unbreakable. |
28762 | 339 |
|
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
340 |
\item @{verbatim ")"} closes a pretty printing block. |
28762 | 341 |
|
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
342 |
\item @{verbatim "//"} forces a line break. |
28762 | 343 |
|
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
344 |
\item @{verbatim "/"}@{text s} allows a line break. Here @{text s} |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
345 |
stands for the string of spaces (zero or more) right after the |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
346 |
slash. These spaces are printed if the break is \emph{not} taken. |
28762 | 347 |
|
28778 | 348 |
\end{description} |
28762 | 349 |
|
28771
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
350 |
For example, the template @{verbatim "(_ +/ _)"} specifies an infix |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
351 |
operator. There are two argument positions; the delimiter |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
352 |
@{verbatim "+"} is preceded by a space and followed by a space or |
4510201c6aaf
mixfix annotations: verbatim for special symbols;
wenzelm
parents:
28770
diff
changeset
|
353 |
line break; the entire phrase is a pretty printing block. |
28762 | 354 |
|
355 |
The general idea of pretty printing with blocks and breaks is also |
|
356 |
described in \cite{paulson-ml2}. |
|
357 |
*} |
|
358 |
||
359 |
||
35413 | 360 |
section {* Explicit notation *} |
28762 | 361 |
|
362 |
text {* |
|
363 |
\begin{matharray}{rcll} |
|
35413 | 364 |
@{command_def "type_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
365 |
@{command_def "no_type_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
|
28762 | 366 |
@{command_def "notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
367 |
@{command_def "no_notation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
|
36508
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
wenzelm
parents:
35413
diff
changeset
|
368 |
@{command_def "write"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\ |
28762 | 369 |
\end{matharray} |
370 |
||
371 |
\begin{rail} |
|
35413 | 372 |
('type\_notation' | 'no\_type\_notation') target? mode? \\ (nameref mixfix + 'and') |
373 |
; |
|
29741 | 374 |
('notation' | 'no\_notation') target? mode? \\ (nameref structmixfix + 'and') |
28762 | 375 |
; |
36508
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
wenzelm
parents:
35413
diff
changeset
|
376 |
'write' mode? (nameref structmixfix + 'and') |
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
wenzelm
parents:
35413
diff
changeset
|
377 |
; |
28762 | 378 |
\end{rail} |
379 |
||
380 |
\begin{description} |
|
381 |
||
35413 | 382 |
\item @{command "type_notation"}~@{text "c (mx)"} associates mixfix |
383 |
syntax with an existing type constructor. The arity of the |
|
384 |
constructor is retrieved from the context. |
|
385 |
||
386 |
\item @{command "no_type_notation"} is similar to @{command |
|
387 |
"type_notation"}, but removes the specified syntax annotation from |
|
388 |
the present context. |
|
389 |
||
28762 | 390 |
\item @{command "notation"}~@{text "c (mx)"} associates mixfix |
35413 | 391 |
syntax with an existing constant or fixed variable. The type |
392 |
declaration of the given entity is retrieved from the context. |
|
28762 | 393 |
|
394 |
\item @{command "no_notation"} is similar to @{command "notation"}, |
|
395 |
but removes the specified syntax annotation from the present |
|
396 |
context. |
|
397 |
||
36508
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
wenzelm
parents:
35413
diff
changeset
|
398 |
\item @{command "write"} is similar to @{command "notation"}, but |
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
wenzelm
parents:
35413
diff
changeset
|
399 |
works within an Isar proof body. |
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
wenzelm
parents:
35413
diff
changeset
|
400 |
|
28762 | 401 |
\end{description} |
35413 | 402 |
|
36508
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
wenzelm
parents:
35413
diff
changeset
|
403 |
Note that the more primitive commands @{command "syntax"} and |
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
wenzelm
parents:
35413
diff
changeset
|
404 |
@{command "no_syntax"} (\secref{sec:syn-trans}) provide raw access |
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
wenzelm
parents:
35413
diff
changeset
|
405 |
to the syntax tables of a global theory. |
28762 | 406 |
*} |
407 |
||
28778 | 408 |
|
409 |
section {* The Pure syntax \label{sec:pure-syntax} *} |
|
28769
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
410 |
|
28777 | 411 |
subsection {* Priority grammars \label{sec:priority-grammar} *} |
28769
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
412 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
413 |
text {* A context-free grammar consists of a set of \emph{terminal |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
414 |
symbols}, a set of \emph{nonterminal symbols} and a set of |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
415 |
\emph{productions}. Productions have the form @{text "A = \<gamma>"}, |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
416 |
where @{text A} is a nonterminal and @{text \<gamma>} is a string of |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
417 |
terminals and nonterminals. One designated nonterminal is called |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
418 |
the \emph{root symbol}. The language defined by the grammar |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
419 |
consists of all strings of terminals that can be derived from the |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
420 |
root symbol by applying productions as rewrite rules. |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
421 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
422 |
The standard Isabelle parser for inner syntax uses a \emph{priority |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
423 |
grammar}. Each nonterminal is decorated by an integer priority: |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
424 |
@{text "A\<^sup>(\<^sup>p\<^sup>)"}. In a derivation, @{text "A\<^sup>(\<^sup>p\<^sup>)"} may be rewritten |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
425 |
using a production @{text "A\<^sup>(\<^sup>q\<^sup>) = \<gamma>"} only if @{text "p \<le> q"}. Any |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
426 |
priority grammar can be translated into a normal context-free |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
427 |
grammar by introducing new nonterminals and productions. |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
428 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
429 |
\medskip Formally, a set of context free productions @{text G} |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
430 |
induces a derivation relation @{text "\<longrightarrow>\<^sub>G"} as follows. Let @{text |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
431 |
\<alpha>} and @{text \<beta>} denote strings of terminal or nonterminal symbols. |
28774 | 432 |
Then @{text "\<alpha> A\<^sup>(\<^sup>p\<^sup>) \<beta> \<longrightarrow>\<^sub>G \<alpha> \<gamma> \<beta>"} holds if and only if @{text G} |
433 |
contains some production @{text "A\<^sup>(\<^sup>q\<^sup>) = \<gamma>"} for @{text "p \<le> q"}. |
|
28769
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
434 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
435 |
\medskip The following grammar for arithmetic expressions |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
436 |
demonstrates how binding power and associativity of operators can be |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
437 |
enforced by priorities. |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
438 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
439 |
\begin{center} |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
440 |
\begin{tabular}{rclr} |
28774 | 441 |
@{text "A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "="} & @{verbatim "("} @{text "A\<^sup>(\<^sup>0\<^sup>)"} @{verbatim ")"} \\ |
28769
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
442 |
@{text "A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "="} & @{verbatim 0} \\ |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
443 |
@{text "A\<^sup>(\<^sup>0\<^sup>)"} & @{text "="} & @{text "A\<^sup>(\<^sup>0\<^sup>)"} @{verbatim "+"} @{text "A\<^sup>(\<^sup>1\<^sup>)"} \\ |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
444 |
@{text "A\<^sup>(\<^sup>2\<^sup>)"} & @{text "="} & @{text "A\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "*"} @{text "A\<^sup>(\<^sup>2\<^sup>)"} \\ |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
445 |
@{text "A\<^sup>(\<^sup>3\<^sup>)"} & @{text "="} & @{verbatim "-"} @{text "A\<^sup>(\<^sup>3\<^sup>)"} \\ |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
446 |
\end{tabular} |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
447 |
\end{center} |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
448 |
The choice of priorities determines that @{verbatim "-"} binds |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
449 |
tighter than @{verbatim "*"}, which binds tighter than @{verbatim |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
450 |
"+"}. Furthermore @{verbatim "+"} associates to the left and |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
451 |
@{verbatim "*"} to the right. |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
452 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
453 |
\medskip For clarity, grammars obey these conventions: |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
454 |
\begin{itemize} |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
455 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
456 |
\item All priorities must lie between 0 and 1000. |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
457 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
458 |
\item Priority 0 on the right-hand side and priority 1000 on the |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
459 |
left-hand side may be omitted. |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
460 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
461 |
\item The production @{text "A\<^sup>(\<^sup>p\<^sup>) = \<alpha>"} is written as @{text "A = \<alpha> |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
462 |
(p)"}, i.e.\ the priority of the left-hand side actually appears in |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
463 |
a column on the far right. |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
464 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
465 |
\item Alternatives are separated by @{text "|"}. |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
466 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
467 |
\item Repetition is indicated by dots @{text "(\<dots>)"} in an informal |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
468 |
but obvious way. |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
469 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
470 |
\end{itemize} |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
471 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
472 |
Using these conventions, the example grammar specification above |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
473 |
takes the form: |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
474 |
\begin{center} |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
475 |
\begin{tabular}{rclc} |
28774 | 476 |
@{text A} & @{text "="} & @{verbatim "("} @{text A} @{verbatim ")"} \\ |
477 |
& @{text "|"} & @{verbatim 0} & \qquad\qquad \\ |
|
28769
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
478 |
& @{text "|"} & @{text A} @{verbatim "+"} @{text "A\<^sup>(\<^sup>1\<^sup>)"} & @{text "(0)"} \\ |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
479 |
& @{text "|"} & @{text "A\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "*"} @{text "A\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
480 |
& @{text "|"} & @{verbatim "-"} @{text "A\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\ |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
481 |
\end{tabular} |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
482 |
\end{center} |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
483 |
*} |
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
484 |
|
8fc228f21861
added section "Priority grammars" (variant from old ref manual);
wenzelm
parents:
28767
diff
changeset
|
485 |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
486 |
subsection {* The Pure grammar *} |
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
487 |
|
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
488 |
text {* |
28773 | 489 |
The priority grammar of the @{text "Pure"} theory is defined as follows: |
490 |
||
28774 | 491 |
%FIXME syntax for "index" (?) |
492 |
%FIXME "op" versions of ==> etc. (?) |
|
493 |
||
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
494 |
\begin{center} |
28773 | 495 |
\begin{supertabular}{rclr} |
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
496 |
|
28778 | 497 |
@{syntax_def (inner) any} & = & @{text "prop | logic"} \\\\ |
28772 | 498 |
|
28778 | 499 |
@{syntax_def (inner) prop} & = & @{verbatim "("} @{text prop} @{verbatim ")"} \\ |
28772 | 500 |
& @{text "|"} & @{text "prop\<^sup>(\<^sup>4\<^sup>)"} @{verbatim "::"} @{text type} & @{text "(3)"} \\ |
28773 | 501 |
& @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "=?="} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ |
28772 | 502 |
& @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "=="} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ |
28773 | 503 |
& @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{text "\<equiv>"} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ |
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28779
diff
changeset
|
504 |
& @{text "|"} & @{text "prop\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "&&&"} @{text "prop\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ |
28772 | 505 |
& @{text "|"} & @{text "prop\<^sup>(\<^sup>2\<^sup>)"} @{verbatim "==>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ |
28773 | 506 |
& @{text "|"} & @{text "prop\<^sup>(\<^sup>2\<^sup>)"} @{text "\<Longrightarrow>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ |
28772 | 507 |
& @{text "|"} & @{verbatim "[|"} @{text prop} @{verbatim ";"} @{text "\<dots>"} @{verbatim ";"} @{text prop} @{verbatim "|]"} @{verbatim "==>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ |
28773 | 508 |
& @{text "|"} & @{text "\<lbrakk>"} @{text prop} @{verbatim ";"} @{text "\<dots>"} @{verbatim ";"} @{text prop} @{text "\<rbrakk>"} @{text "\<Longrightarrow>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ |
28772 | 509 |
& @{text "|"} & @{verbatim "!!"} @{text idts} @{verbatim "."} @{text prop} & @{text "(0)"} \\ |
28773 | 510 |
& @{text "|"} & @{text "\<And>"} @{text idts} @{verbatim "."} @{text prop} & @{text "(0)"} \\ |
511 |
& @{text "|"} & @{verbatim OFCLASS} @{verbatim "("} @{text type} @{verbatim ","} @{text logic} @{verbatim ")"} \\ |
|
512 |
& @{text "|"} & @{verbatim SORT_CONSTRAINT} @{verbatim "("} @{text type} @{verbatim ")"} \\ |
|
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28779
diff
changeset
|
513 |
& @{text "|"} & @{verbatim TERM} @{text logic} \\ |
28773 | 514 |
& @{text "|"} & @{verbatim PROP} @{text aprop} \\\\ |
28772 | 515 |
|
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28779
diff
changeset
|
516 |
@{syntax_def (inner) aprop} & = & @{verbatim "("} @{text aprop} @{verbatim ")"} \\ |
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28779
diff
changeset
|
517 |
& @{text "|"} & @{text "id | longid | var | "}@{verbatim "_"}@{text " | "}@{verbatim "..."} \\ |
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28779
diff
changeset
|
518 |
& @{text "|"} & @{verbatim CONST} @{text "id | "}@{verbatim CONST} @{text "longid"} \\ |
28773 | 519 |
& @{text "|"} & @{text "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>)"} & @{text "(999)"} \\\\ |
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
520 |
|
28778 | 521 |
@{syntax_def (inner) logic} & = & @{verbatim "("} @{text logic} @{verbatim ")"} \\ |
28772 | 522 |
& @{text "|"} & @{text "logic\<^sup>(\<^sup>4\<^sup>)"} @{verbatim "::"} @{text type} & @{text "(3)"} \\ |
28773 | 523 |
& @{text "|"} & @{text "id | longid | var | "}@{verbatim "_"}@{text " | "}@{verbatim "..."} \\ |
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28779
diff
changeset
|
524 |
& @{text "|"} & @{verbatim CONST} @{text "id | "}@{verbatim CONST} @{text "longid"} \\ |
28773 | 525 |
& @{text "|"} & @{text "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>)"} & @{text "(999)"} \\ |
28772 | 526 |
& @{text "|"} & @{verbatim "%"} @{text pttrns} @{verbatim "."} @{text "any\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\ |
28773 | 527 |
& @{text "|"} & @{text \<lambda>} @{text pttrns} @{verbatim "."} @{text "any\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\ |
28772 | 528 |
& @{text "|"} & @{verbatim TYPE} @{verbatim "("} @{text type} @{verbatim ")"} \\\\ |
529 |
||
28778 | 530 |
@{syntax_def (inner) idt} & = & @{verbatim "("} @{text idt} @{verbatim ")"}@{text " | id | "}@{verbatim "_"} \\ |
28773 | 531 |
& @{text "|"} & @{text id} @{verbatim "::"} @{text type} & @{text "(0)"} \\ |
532 |
& @{text "|"} & @{verbatim "_"} @{verbatim "::"} @{text type} & @{text "(0)"} \\\\ |
|
28772 | 533 |
|
28778 | 534 |
@{syntax_def (inner) idts} & = & @{text "idt | idt\<^sup>(\<^sup>1\<^sup>) idts"} & @{text "(0)"} \\\\ |
28772 | 535 |
|
28778 | 536 |
@{syntax_def (inner) pttrn} & = & @{text idt} \\\\ |
28772 | 537 |
|
28778 | 538 |
@{syntax_def (inner) pttrns} & = & @{text "pttrn | pttrn\<^sup>(\<^sup>1\<^sup>) pttrns"} & @{text "(0)"} \\\\ |
28774 | 539 |
|
28778 | 540 |
@{syntax_def (inner) type} & = & @{verbatim "("} @{text type} @{verbatim ")"} \\ |
28773 | 541 |
& @{text "|"} & @{text "tid | tvar | "}@{verbatim "_"} \\ |
542 |
& @{text "|"} & @{text "tid"} @{verbatim "::"} @{text "sort | tvar "}@{verbatim "::"} @{text "sort | "}@{verbatim "_"} @{verbatim "::"} @{text "sort"} \\ |
|
28772 | 543 |
& @{text "|"} & @{text "id | type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) id | "}@{verbatim "("} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim ")"} @{text id} \\ |
29741 | 544 |
& @{text "|"} & @{text "longid | type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) longid"} \\ |
545 |
& @{text "|"} & @{verbatim "("} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim ")"} @{text longid} \\ |
|
28772 | 546 |
& @{text "|"} & @{text "type\<^sup>(\<^sup>1\<^sup>)"} @{verbatim "=>"} @{text type} & @{text "(0)"} \\ |
28773 | 547 |
& @{text "|"} & @{text "type\<^sup>(\<^sup>1\<^sup>)"} @{text "\<Rightarrow>"} @{text type} & @{text "(0)"} \\ |
548 |
& @{text "|"} & @{verbatim "["} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim "]"} @{verbatim "=>"} @{text type} & @{text "(0)"} \\ |
|
549 |
& @{text "|"} & @{verbatim "["} @{text type} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text type} @{verbatim "]"} @{text "\<Rightarrow>"} @{text type} & @{text "(0)"} \\\\ |
|
28772 | 550 |
|
29741 | 551 |
@{syntax_def (inner) sort} & = & @{text "id | longid | "}@{verbatim "{}"} \\ |
552 |
& @{text "|"} & @{verbatim "{"} @{text "(id | longid)"} @{verbatim ","} @{text "\<dots>"} @{verbatim ","} @{text "(id | longid)"} @{verbatim "}"} \\ |
|
28773 | 553 |
\end{supertabular} |
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
554 |
\end{center} |
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
555 |
|
28774 | 556 |
\medskip Here literal terminals are printed @{verbatim "verbatim"}; |
557 |
see also \secref{sec:inner-lex} for further token categories of the |
|
558 |
inner syntax. The meaning of the nonterminals defined by the above |
|
559 |
grammar is as follows: |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
560 |
|
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
561 |
\begin{description} |
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
562 |
|
28778 | 563 |
\item @{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
|
564 |
|
28778 | 565 |
\item @{syntax_ref (inner) prop} denotes meta-level propositions, |
566 |
which are terms of type @{typ prop}. The syntax of such formulae of |
|
567 |
the meta-logic is carefully distinguished from usual conventions for |
|
568 |
object-logics. In particular, plain @{text "\<lambda>"}-term notation is |
|
569 |
\emph{not} recognized as @{syntax (inner) prop}. |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
570 |
|
28778 | 571 |
\item @{syntax_ref (inner) aprop} denotes atomic propositions, which |
572 |
are embedded into regular @{syntax (inner) prop} by means of an |
|
573 |
explicit @{verbatim PROP} token. |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
574 |
|
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
575 |
Terms of type @{typ prop} with non-constant head, e.g.\ a plain |
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
576 |
variable, are printed in this form. Constants that yield type @{typ |
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
577 |
prop} are expected to provide their own concrete syntax; otherwise |
28778 | 578 |
the printed version will appear like @{syntax (inner) logic} and |
579 |
cannot be parsed again as @{syntax (inner) prop}. |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
580 |
|
28778 | 581 |
\item @{syntax_ref (inner) logic} denotes arbitrary terms of a |
582 |
logical type, excluding type @{typ prop}. This is the main |
|
583 |
syntactic category of object-logic entities, covering plain @{text |
|
584 |
\<lambda>}-term notation (variables, abstraction, application), plus |
|
585 |
anything defined by the user. |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
586 |
|
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
587 |
When specifying notation for logical entities, all logical types |
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
588 |
(excluding @{typ prop}) are \emph{collapsed} to this single category |
28778 | 589 |
of @{syntax (inner) logic}. |
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
590 |
|
28778 | 591 |
\item @{syntax_ref (inner) idt} denotes identifiers, possibly |
592 |
constrained by types. |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
593 |
|
28778 | 594 |
\item @{syntax_ref (inner) idts} denotes a sequence of @{syntax_ref |
595 |
(inner) idt}. This is the most basic category for variables in |
|
596 |
iterated binders, such as @{text "\<lambda>"} or @{text "\<And>"}. |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
597 |
|
28778 | 598 |
\item @{syntax_ref (inner) pttrn} and @{syntax_ref (inner) pttrns} |
599 |
denote patterns for abstraction, cases bindings etc. In Pure, these |
|
600 |
categories start as a merely copy of @{syntax (inner) idt} and |
|
601 |
@{syntax (inner) idts}, respectively. Object-logics may add |
|
602 |
additional productions for binding forms. |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
603 |
|
28778 | 604 |
\item @{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
|
605 |
|
28778 | 606 |
\item @{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
|
607 |
|
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
608 |
\end{description} |
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
609 |
|
28774 | 610 |
Here are some further explanations of certain syntax features. |
28773 | 611 |
|
612 |
\begin{itemize} |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
613 |
|
28778 | 614 |
\item In @{syntax (inner) idts}, note that @{text "x :: nat y"} is |
615 |
parsed as @{text "x :: (nat y)"}, treating @{text y} like a type |
|
616 |
constructor applied to @{text nat}. To avoid this interpretation, |
|
617 |
write @{text "(x :: nat) y"} with explicit parentheses. |
|
28773 | 618 |
|
619 |
\item Similarly, @{text "x :: nat y :: nat"} is parsed as @{text "x :: |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
620 |
(nat y :: nat)"}. The correct form is @{text "(x :: nat) (y :: |
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
621 |
nat)"}, or @{text "(x :: nat) y :: nat"} if @{text y} is last in the |
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
622 |
sequence of identifiers. |
28773 | 623 |
|
624 |
\item Type constraints for terms bind very weakly. For example, |
|
625 |
@{text "x < y :: nat"} is normally parsed as @{text "(x < y) :: |
|
626 |
nat"}, unless @{text "<"} has a very low priority, in which case the |
|
627 |
input is likely to be ambiguous. The correct form is @{text "x < (y |
|
628 |
:: nat)"}. |
|
629 |
||
630 |
\item Constraints may be either written with two literal colons |
|
631 |
``@{verbatim "::"}'' or the double-colon symbol @{verbatim "\<Colon>"}, |
|
28774 | 632 |
which actually looks exactly the same in some {\LaTeX} styles. |
28773 | 633 |
|
28774 | 634 |
\item Dummy variables (written as underscore) may occur in different |
635 |
roles. |
|
28773 | 636 |
|
637 |
\begin{description} |
|
638 |
||
28774 | 639 |
\item A type ``@{text "_"}'' or ``@{text "_ :: sort"}'' acts like an |
640 |
anonymous inference parameter, which is filled-in according to the |
|
641 |
most general type produced by the type-checking phase. |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
642 |
|
28774 | 643 |
\item A bound ``@{text "_"}'' refers to a vacuous abstraction, where |
644 |
the body does not refer to the binding introduced here. As in the |
|
645 |
term @{term "\<lambda>x _. x"}, which is @{text "\<alpha>"}-equivalent to @{text |
|
646 |
"\<lambda>x y. x"}. |
|
28773 | 647 |
|
28774 | 648 |
\item A free ``@{text "_"}'' refers to an implicit outer binding. |
649 |
Higher definitional packages usually allow forms like @{text "f x _ |
|
650 |
= x"}. |
|
28773 | 651 |
|
28774 | 652 |
\item A schematic ``@{text "_"}'' (within a term pattern, see |
653 |
\secref{sec:term-decls}) refers to an anonymous variable that is |
|
654 |
implicitly abstracted over its context of locally bound variables. |
|
655 |
For example, this allows pattern matching of @{text "{x. f x = g |
|
656 |
x}"} against @{text "{x. _ = _}"}, or even @{text "{_. _ = _}"} by |
|
657 |
using both bound and schematic dummies. |
|
28773 | 658 |
|
659 |
\end{description} |
|
660 |
||
28774 | 661 |
\item The three literal dots ``@{verbatim "..."}'' may be also |
662 |
written as ellipsis symbol @{verbatim "\<dots>"}. In both cases this |
|
663 |
refers to a special schematic variable, which is bound in the |
|
664 |
context. This special term abbreviation works nicely with |
|
665 |
calculational reasoning (\secref{sec:calculation}). |
|
666 |
||
28773 | 667 |
\end{itemize} |
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
668 |
*} |
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
669 |
|
28777 | 670 |
|
28774 | 671 |
section {* Lexical matters \label{sec:inner-lex} *} |
672 |
||
28777 | 673 |
text {* The inner lexical syntax vaguely resembles the outer one |
674 |
(\secref{sec:outer-lex}), but some details are different. There are |
|
675 |
two main categories of inner syntax tokens: |
|
676 |
||
677 |
\begin{enumerate} |
|
678 |
||
679 |
\item \emph{delimiters} --- the literal tokens occurring in |
|
680 |
productions of the given priority grammar (cf.\ |
|
681 |
\secref{sec:priority-grammar}); |
|
682 |
||
683 |
\item \emph{named tokens} --- various categories of identifiers etc. |
|
684 |
||
685 |
\end{enumerate} |
|
686 |
||
687 |
Delimiters override named tokens and may thus render certain |
|
688 |
identifiers inaccessible. Sometimes the logical context admits |
|
689 |
alternative ways to refer to the same entity, potentially via |
|
690 |
qualified names. |
|
691 |
||
692 |
\medskip The categories for named tokens are defined once and for |
|
693 |
all as follows, reusing some categories of the outer token syntax |
|
694 |
(\secref{sec:outer-lex}). |
|
695 |
||
696 |
\begin{center} |
|
697 |
\begin{supertabular}{rcl} |
|
698 |
@{syntax_def (inner) id} & = & @{syntax_ref ident} \\ |
|
699 |
@{syntax_def (inner) longid} & = & @{syntax_ref longident} \\ |
|
700 |
@{syntax_def (inner) var} & = & @{syntax_ref var} \\ |
|
701 |
@{syntax_def (inner) tid} & = & @{syntax_ref typefree} \\ |
|
702 |
@{syntax_def (inner) tvar} & = & @{syntax_ref typevar} \\ |
|
703 |
@{syntax_def (inner) num} & = & @{syntax_ref nat}@{text " | "}@{verbatim "-"}@{syntax_ref nat} \\ |
|
29157 | 704 |
@{syntax_def (inner) float_token} & = & @{syntax_ref nat}@{verbatim "."}@{syntax_ref nat}@{text " | "}@{verbatim "-"}@{syntax_ref nat}@{verbatim "."}@{syntax_ref nat} \\ |
28777 | 705 |
@{syntax_def (inner) xnum} & = & @{verbatim "#"}@{syntax_ref nat}@{text " | "}@{verbatim "#-"}@{syntax_ref nat} \\ |
706 |
||
707 |
@{syntax_def (inner) xstr} & = & @{verbatim "''"} @{text "\<dots>"} @{verbatim "''"} \\ |
|
708 |
\end{supertabular} |
|
709 |
\end{center} |
|
710 |
||
29157 | 711 |
The token categories @{syntax (inner) num}, @{syntax (inner) |
712 |
float_token}, @{syntax (inner) xnum}, and @{syntax (inner) xstr} are |
|
713 |
not used in Pure. Object-logics may implement numerals and string |
|
714 |
constants by adding appropriate syntax declarations, together with |
|
715 |
some translation functions (e.g.\ see Isabelle/HOL). |
|
716 |
||
717 |
The derived categories @{syntax_def (inner) num_const} and |
|
718 |
@{syntax_def (inner) float_const} provide robust access to @{syntax |
|
719 |
(inner) num}, and @{syntax (inner) float_token}, respectively: the |
|
720 |
syntax tree holds a syntactic constant instead of a free variable. |
|
28777 | 721 |
*} |
28774 | 722 |
|
28770
93a372e2dc7a
added section "The Pure grammar" (incomplete version, based on old ref manual);
wenzelm
parents:
28769
diff
changeset
|
723 |
|
28762 | 724 |
section {* Syntax and translations \label{sec:syn-trans} *} |
725 |
||
726 |
text {* |
|
727 |
\begin{matharray}{rcl} |
|
728 |
@{command_def "nonterminals"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
729 |
@{command_def "syntax"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
730 |
@{command_def "no_syntax"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
731 |
@{command_def "translations"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
732 |
@{command_def "no_translations"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
733 |
\end{matharray} |
|
734 |
||
735 |
\begin{rail} |
|
736 |
'nonterminals' (name +) |
|
737 |
; |
|
738 |
('syntax' | 'no\_syntax') mode? (constdecl +) |
|
739 |
; |
|
740 |
('translations' | 'no\_translations') (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +) |
|
741 |
; |
|
742 |
||
743 |
mode: ('(' ( name | 'output' | name 'output' ) ')') |
|
744 |
; |
|
745 |
transpat: ('(' nameref ')')? string |
|
746 |
; |
|
747 |
\end{rail} |
|
748 |
||
749 |
\begin{description} |
|
750 |
||
751 |
\item @{command "nonterminals"}~@{text c} declares a type |
|
752 |
constructor @{text c} (without arguments) to act as purely syntactic |
|
753 |
type: a nonterminal symbol of the inner syntax. |
|
754 |
||
755 |
\item @{command "syntax"}~@{text "(mode) decls"} is similar to |
|
756 |
@{command "consts"}~@{text decls}, except that the actual logical |
|
757 |
signature extension is omitted. Thus the context free grammar of |
|
758 |
Isabelle's inner syntax may be augmented in arbitrary ways, |
|
759 |
independently of the logic. The @{text mode} argument refers to the |
|
760 |
print mode that the grammar rules belong; unless the @{keyword_ref |
|
761 |
"output"} indicator is given, all productions are added both to the |
|
762 |
input and output grammar. |
|
763 |
||
764 |
\item @{command "no_syntax"}~@{text "(mode) decls"} removes grammar |
|
765 |
declarations (and translations) resulting from @{text decls}, which |
|
766 |
are interpreted in the same manner as for @{command "syntax"} above. |
|
767 |
||
768 |
\item @{command "translations"}~@{text rules} specifies syntactic |
|
769 |
translation rules (i.e.\ macros): parse~/ print rules (@{text "\<rightleftharpoons>"}), |
|
770 |
parse rules (@{text "\<rightharpoonup>"}), or print rules (@{text "\<leftharpoondown>"}). |
|
771 |
Translation patterns may be prefixed by the syntactic category to be |
|
772 |
used for parsing; the default is @{text logic}. |
|
773 |
||
774 |
\item @{command "no_translations"}~@{text rules} removes syntactic |
|
775 |
translation rules, which are interpreted in the same manner as for |
|
776 |
@{command "translations"} above. |
|
777 |
||
778 |
\end{description} |
|
779 |
*} |
|
780 |
||
781 |
||
28779
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
782 |
section {* Syntax translation functions \label{sec:tr-funs} *} |
28762 | 783 |
|
784 |
text {* |
|
785 |
\begin{matharray}{rcl} |
|
786 |
@{command_def "parse_ast_translation"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
787 |
@{command_def "parse_translation"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
788 |
@{command_def "print_translation"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
789 |
@{command_def "typed_print_translation"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
790 |
@{command_def "print_ast_translation"} & : & @{text "theory \<rightarrow> theory"} \\ |
|
791 |
\end{matharray} |
|
792 |
||
793 |
\begin{rail} |
|
794 |
( 'parse\_ast\_translation' | 'parse\_translation' | 'print\_translation' | |
|
795 |
'typed\_print\_translation' | 'print\_ast\_translation' ) ('(advanced)')? text |
|
796 |
; |
|
797 |
\end{rail} |
|
798 |
||
799 |
Syntax translation functions written in ML admit almost arbitrary |
|
800 |
manipulations of Isabelle's inner syntax. Any of the above commands |
|
801 |
have a single \railqtok{text} argument that refers to an ML |
|
802 |
expression of appropriate type, which are as follows by default: |
|
803 |
||
804 |
%FIXME proper antiquotations |
|
805 |
\begin{ttbox} |
|
806 |
val parse_ast_translation : (string * (ast list -> ast)) list |
|
807 |
val parse_translation : (string * (term list -> term)) list |
|
808 |
val print_translation : (string * (term list -> term)) list |
|
809 |
val typed_print_translation : |
|
810 |
(string * (bool -> typ -> term list -> term)) list |
|
811 |
val print_ast_translation : (string * (ast list -> ast)) list |
|
812 |
\end{ttbox} |
|
813 |
||
814 |
If the @{text "(advanced)"} option is given, the corresponding |
|
815 |
translation functions may depend on the current theory or proof |
|
816 |
context. This allows to implement advanced syntax mechanisms, as |
|
817 |
translations functions may refer to specific theory declarations or |
|
818 |
auxiliary proof data. |
|
819 |
||
30397 | 820 |
See also \cite{isabelle-ref} for more information on the general |
821 |
concept of syntax transformations in Isabelle. |
|
28762 | 822 |
|
823 |
%FIXME proper antiquotations |
|
824 |
\begin{ttbox} |
|
825 |
val parse_ast_translation: |
|
826 |
(string * (Proof.context -> ast list -> ast)) list |
|
827 |
val parse_translation: |
|
828 |
(string * (Proof.context -> term list -> term)) list |
|
829 |
val print_translation: |
|
830 |
(string * (Proof.context -> term list -> term)) list |
|
831 |
val typed_print_translation: |
|
832 |
(string * (Proof.context -> bool -> typ -> term list -> term)) list |
|
833 |
val print_ast_translation: |
|
834 |
(string * (Proof.context -> ast list -> ast)) list |
|
835 |
\end{ttbox} |
|
836 |
*} |
|
837 |
||
28779
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
838 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
839 |
section {* Inspecting the syntax *} |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
840 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
841 |
text {* |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
842 |
\begin{matharray}{rcl} |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
843 |
@{command_def "print_syntax"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
844 |
\end{matharray} |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
845 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
846 |
\begin{description} |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
847 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
848 |
\item @{command "print_syntax"} prints the inner syntax of the |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
849 |
current context. The output can be quite large; the most important |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
850 |
sections are explained below. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
851 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
852 |
\begin{description} |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
853 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
854 |
\item @{text "lexicon"} lists the delimiters of the inner token |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
855 |
language; see \secref{sec:inner-lex}. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
856 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
857 |
\item @{text "prods"} lists the productions of the underlying |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
858 |
priority grammar; see \secref{sec:priority-grammar}. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
859 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
860 |
The nonterminal @{text "A\<^sup>(\<^sup>p\<^sup>)"} is rendered in plain text as @{text |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
861 |
"A[p]"}; delimiters are quoted. Many productions have an extra |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
862 |
@{text "\<dots> => name"}. These names later become the heads of parse |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
863 |
trees; they also guide the pretty printer. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
864 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
865 |
Productions without such parse tree names are called \emph{copy |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
866 |
productions}. Their right-hand side must have exactly one |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
867 |
nonterminal symbol (or named token). The parser does not create a |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
868 |
new parse tree node for copy productions, but simply returns the |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
869 |
parse tree of the right-hand symbol. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
870 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
871 |
If the right-hand side of a copy production consists of a single |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
872 |
nonterminal without any delimiters, then it is called a \emph{chain |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
873 |
production}. Chain productions act as abbreviations: conceptually, |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
874 |
they are removed from the grammar by adding new productions. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
875 |
Priority information attached to chain productions is ignored; only |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
876 |
the dummy value @{text "-1"} is displayed. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
877 |
|
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28779
diff
changeset
|
878 |
\item @{text "print modes"} lists the alternative print modes |
28779
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
879 |
provided by this grammar; see \secref{sec:print-modes}. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
880 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
881 |
\item @{text "parse_rules"} and @{text "print_rules"} relate to |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
882 |
syntax translations (macros); see \secref{sec:syn-trans}. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
883 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
884 |
\item @{text "parse_ast_translation"} and @{text |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
885 |
"print_ast_translation"} list sets of constants that invoke |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
886 |
translation functions for abstract syntax trees, which are only |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
887 |
required in very special situations; see \secref{sec:tr-funs}. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
888 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
889 |
\item @{text "parse_translation"} and @{text "print_translation"} |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
890 |
list the sets of constants that invoke regular translation |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
891 |
functions; see \secref{sec:tr-funs}. |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
892 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
893 |
\end{description} |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
894 |
|
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
895 |
\end{description} |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
896 |
*} |
698960f08652
separate section "Inspecting the syntax" for print_syntax;
wenzelm
parents:
28778
diff
changeset
|
897 |
|
28762 | 898 |
end |