| author | wenzelm |
| Thu, 19 Jun 2014 22:07:44 +0200 | |
| changeset 57345 | 8a9639888639 |
| parent 57344 | 3355a0657f10 |
| child 57346 | 1d6d44a0583f |
| permissions | -rw-r--r-- |
| 57344 | 1 |
(*:wrap=hard:maxLineLen=78:*) |
2 |
||
| 30124 | 3 |
theory Syntax |
4 |
imports Base |
|
5 |
begin |
|
6 |
||
| 34924 | 7 |
chapter {* Concrete syntax and type-checking *}
|
| 30124 | 8 |
|
| 45258 | 9 |
text {* Pure @{text "\<lambda>"}-calculus as introduced in \chref{ch:logic} is
|
10 |
an adequate foundation for logical languages --- in the tradition of |
|
11 |
\emph{higher-order abstract syntax} --- but end-users require
|
|
12 |
additional means for reading and printing of terms and types. This |
|
13 |
important add-on outside the logical core is called \emph{inner
|
|
14 |
syntax} in Isabelle jargon, as opposed to the \emph{outer syntax} of
|
|
| 46484 | 15 |
the theory and proof language (cf.\ \cite{isabelle-isar-ref}).
|
| 45258 | 16 |
|
17 |
For example, according to \cite{church40} quantifiers are
|
|
18 |
represented as higher-order constants @{text "All :: ('a \<Rightarrow> bool) \<Rightarrow>
|
|
19 |
bool"} such that @{text "All (\<lambda>x::'a. B x)"} faithfully represents
|
|
20 |
the idea that is displayed as @{text "\<forall>x::'a. B x"} via @{keyword
|
|
21 |
"binder"} notation. Moreover, type-inference in the style of |
|
22 |
Hindley-Milner \cite{hindleymilner} (and extensions) enables users
|
|
23 |
to write @{text "\<forall>x. B x"} concisely, when the type @{text "'a"} is
|
|
24 |
already clear from the context.\footnote{Type-inference taken to the
|
|
25 |
extreme can easily confuse users, though. Beginners often stumble |
|
26 |
over unexpectedly general types inferred by the system.} |
|
27 |
||
28 |
\medskip The main inner syntax operations are \emph{read} for
|
|
29 |
parsing together with type-checking, and \emph{pretty} for formatted
|
|
30 |
output. See also \secref{sec:read-print}.
|
|
31 |
||
32 |
Furthermore, the input and output syntax layers are sub-divided into |
|
33 |
separate phases for \emph{concrete syntax} versus \emph{abstract
|
|
34 |
syntax}, see also \secref{sec:parse-unparse} and
|
|
35 |
\secref{sec:term-check}, respectively. This results in the
|
|
36 |
following decomposition of the main operations: |
|
37 |
||
38 |
\begin{itemize}
|
|
39 |
||
40 |
\item @{text "read = parse; check"}
|
|
41 |
||
42 |
\item @{text "pretty = uncheck; unparse"}
|
|
43 |
||
44 |
\end{itemize}
|
|
45 |
||
46 |
Some specification package might thus intercept syntax processing at |
|
47 |
a well-defined stage after @{text "parse"}, to a augment the
|
|
48 |
resulting pre-term before full type-reconstruction is performed by |
|
49 |
@{text "check"}, for example. Note that the formal status of bound
|
|
50 |
variables, versus free variables, versus constants must not be |
|
| 57345 | 51 |
changed between these phases! |
52 |
||
53 |
\medskip In general, @{text check} and @{text uncheck} operate
|
|
54 |
simultaneously on a list of terms. This is particular important for |
|
55 |
type-checking, to reconstruct types for several terms of the same context |
|
56 |
and scope. In contrast, @{text parse} and @{text unparse} operate separately
|
|
57 |
in single terms. |
|
58 |
||
59 |
There are analogous operations to read and print types, with the same |
|
60 |
sub-division into phases. |
|
61 |
*} |
|
| 45258 | 62 |
|
| 30272 | 63 |
|
|
39876
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
64 |
section {* Reading and pretty printing \label{sec:read-print} *}
|
| 34924 | 65 |
|
| 45258 | 66 |
text {* Read and print operations are roughly dual to each other, such
|
67 |
that for the user @{text "s' = pretty (read s)"} looks similar to
|
|
68 |
the original source text @{text "s"}, but the details depend on many
|
|
69 |
side-conditions. There are also explicit options to control |
|
70 |
suppressing of type information in the output. The default |
|
71 |
configuration routinely looses information, so @{text "t' = read
|
|
| 57344 | 72 |
(pretty t)"} might fail, or produce a differently typed term, or a |
| 45258 | 73 |
completely different term in the face of syntactic overloading! *} |
| 34924 | 74 |
|
|
39876
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
75 |
text %mlref {*
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
76 |
\begin{mldecls}
|
| 57345 | 77 |
@{index_ML Syntax.read_typs: "Proof.context -> string list -> typ list"} \\
|
78 |
@{index_ML Syntax.read_terms: "Proof.context -> string list -> term list"} \\
|
|
79 |
@{index_ML Syntax.read_props: "Proof.context -> string list -> term list"} \\[0.5ex]
|
|
|
39876
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
80 |
@{index_ML Syntax.read_typ: "Proof.context -> string -> typ"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
81 |
@{index_ML Syntax.read_term: "Proof.context -> string -> term"} \\
|
| 57345 | 82 |
@{index_ML Syntax.read_prop: "Proof.context -> string -> term"} \\[0.5ex]
|
|
39876
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
83 |
@{index_ML Syntax.pretty_typ: "Proof.context -> typ -> Pretty.T"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
84 |
@{index_ML Syntax.pretty_term: "Proof.context -> term -> Pretty.T"} \\
|
| 57345 | 85 |
@{index_ML Syntax.string_of_typ: "Proof.context -> typ -> string"} \\
|
86 |
@{index_ML Syntax.string_of_term: "Proof.context -> term -> string"} \\
|
|
|
39876
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
87 |
\end{mldecls}
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
88 |
|
| 57345 | 89 |
\begin{description}
|
90 |
||
91 |
\item @{ML Syntax.read_typs}~@{text "ctxt strs"} reads and checks a
|
|
92 |
simultaneous list of source strings as types of the logic. |
|
93 |
||
94 |
\item @{ML Syntax.read_terms}~@{text "ctxt strs"} reads and checks a
|
|
95 |
simultaneous list of source strings as terms of the logic. |
|
96 |
Type-reconstruction puts all parsed terms into the same scope. |
|
97 |
||
98 |
If particular type-constraints are required for some of the arguments, the |
|
99 |
read operations needs to be split into its parse and check phases, using |
|
100 |
@{ML Type.constraint} on the intermediate pre-terms.
|
|
101 |
||
102 |
\item @{ML Syntax.read_props} ~@{text "ctxt strs"} reads and checks a
|
|
103 |
simultaneous list of source strings as propositions of the logic, with an |
|
104 |
implicit type-constraint for each argument to make it of type @{typ prop};
|
|
105 |
this also affects the inner syntax for parsing. The remaining |
|
106 |
type-reconstructions works as for @{ML Syntax.read_terms} above.
|
|
107 |
||
108 |
\item @{ML Syntax.read_typ}, @{ML Syntax.read_term}, @{ML Syntax.read_prop}
|
|
109 |
are like the simultaneous versions above, but operate on a single argument |
|
110 |
only. This convenient shorthand is adequate in situations where a single |
|
111 |
item in its own scope is processed. Never use @{ML "map o Syntax.read_term"}
|
|
112 |
where @{ML Syntax.read_terms} is actually intended!
|
|
113 |
||
114 |
\item @{ML Syntax.pretty_typ}~@{text "ctxt T"} and @{ML
|
|
115 |
Syntax.pretty_term}~@{text "ctxt t"} uncheck and pretty-print the given type
|
|
116 |
or term, respectively. Although the uncheck phase acts on a simultaneous |
|
117 |
list as well, this is rarely relevant in practice, so only the singleton |
|
118 |
case is provided as combined pretty operation. Moreover, the distinction of |
|
119 |
term vs.\ proposition is ignored here. |
|
120 |
||
121 |
\item @{ML Syntax.string_of_typ} and @{ML Syntax.string_of_term} are
|
|
122 |
convenient compositions of @{ML Syntax.pretty_typ} and @{ML
|
|
123 |
Syntax.pretty_term} with @{ML Pretty.string_of} for output. The result may
|
|
124 |
be concatenated with other strings, as long as there is no further |
|
125 |
formatting and line-breaking involved. |
|
126 |
||
127 |
\end{description}
|
|
128 |
||
129 |
The most important operations in practice are @{ML Syntax.read_term}, @{ML
|
|
130 |
Syntax.read_prop}, and @{ML Syntax.string_of_term}.
|
|
131 |
||
132 |
\medskip Note that the string values that are passed in and out here are |
|
133 |
actually annotated by the system, to carry further markup that is relevant |
|
134 |
for the Prover IDE \cite{isabelle-jedit}. User code should neither compose
|
|
135 |
its own strings for input, nor try to analyze the string for output. |
|
136 |
||
137 |
The standard way to provide the required position markup for input works via |
|
138 |
the outer syntax parser wrapper @{ML Parse.inner_syntax}, which is already
|
|
139 |
part of @{ML Parse.typ}, @{ML Parse.term}, @{ML Parse.prop}. So a string
|
|
140 |
obtained from one of the latter may be directly passed to the corresponding |
|
141 |
read operation, in order to get PIDE markup of the input and precise |
|
142 |
positions for warnings and errors. |
|
|
39876
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
143 |
*} |
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
144 |
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
145 |
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
146 |
section {* Parsing and unparsing \label{sec:parse-unparse} *}
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
147 |
|
| 45258 | 148 |
text {* Parsing and unparsing converts between actual source text and
|
149 |
a certain \emph{pre-term} format, where all bindings and scopes are
|
|
150 |
resolved faithfully. Thus the names of free variables or constants |
|
151 |
are already determined in the sense of the logical context, but type |
|
| 57344 | 152 |
information might be still missing. Pre-terms support an explicit |
| 45258 | 153 |
language of \emph{type constraints} that may be augmented by user
|
| 57344 | 154 |
code to guide the later \emph{check} phase.
|
| 45258 | 155 |
|
156 |
Actual parsing is based on traditional lexical analysis and Earley |
|
157 |
parsing for arbitrary context-free grammars. The user can specify |
|
| 57344 | 158 |
the grammar via mixfix annotations. Moreover, there are \emph{syntax
|
| 45258 | 159 |
translations} that can be augmented by the user, either |
160 |
declaratively via @{command translations} or programmatically via
|
|
161 |
@{command parse_translation}, @{command print_translation} etc. The
|
|
162 |
final scope resolution is performed by the system, according to name |
|
| 57344 | 163 |
spaces for types, term variables and constants etc.\ determined by |
164 |
the context. |
|
| 45258 | 165 |
*} |
|
39876
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
166 |
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
167 |
text %mlref {*
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
168 |
\begin{mldecls}
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
169 |
@{index_ML Syntax.parse_typ: "Proof.context -> string -> typ"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
170 |
@{index_ML Syntax.parse_term: "Proof.context -> string -> term"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
171 |
@{index_ML Syntax.parse_prop: "Proof.context -> string -> term"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
172 |
@{index_ML Syntax.unparse_typ: "Proof.context -> typ -> Pretty.T"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
173 |
@{index_ML Syntax.unparse_term: "Proof.context -> term -> Pretty.T"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
174 |
\end{mldecls}
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
175 |
|
| 52422 | 176 |
%FIXME description |
|
39876
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
177 |
*} |
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
178 |
|
| 39852 | 179 |
|
| 34924 | 180 |
section {* Checking and unchecking \label{sec:term-check} *}
|
181 |
||
| 45258 | 182 |
text {* These operations define the transition from pre-terms and
|
183 |
fully-annotated terms in the sense of the logical core |
|
184 |
(\chref{ch:logic}).
|
|
185 |
||
186 |
The \emph{check} phase is meant to subsume a variety of mechanisms
|
|
187 |
in the manner of ``type-inference'' or ``type-reconstruction'' or |
|
188 |
``type-improvement'', not just type-checking in the narrow sense. |
|
189 |
The \emph{uncheck} phase is roughly dual, it prunes type-information
|
|
190 |
before pretty printing. |
|
191 |
||
192 |
A typical add-on for the check/uncheck syntax layer is the @{command
|
|
193 |
abbreviation} mechanism. Here the user specifies syntactic |
|
194 |
definitions that are managed by the system as polymorphic @{text
|
|
195 |
"let"} bindings. These are expanded during the @{text "check"}
|
|
196 |
phase, and contracted during the @{text "uncheck"} phase, without
|
|
197 |
affecting the type-assignment of the given terms. |
|
198 |
||
199 |
\medskip The precise meaning of type checking depends on the context |
|
| 57344 | 200 |
--- additional check/uncheck plugins might be defined in user space. |
| 45258 | 201 |
|
202 |
For example, the @{command class} command defines a context where
|
|
203 |
@{text "check"} treats certain type instances of overloaded
|
|
204 |
constants according to the ``dictionary construction'' of its |
|
205 |
logical foundation. This involves ``type improvement'' |
|
206 |
(specialization of slightly too general types) and replacement by |
|
207 |
certain locale parameters. See also \cite{Haftmann-Wenzel:2009}.
|
|
208 |
*} |
|
| 39852 | 209 |
|
|
39876
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
210 |
text %mlref {*
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
211 |
\begin{mldecls}
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
212 |
@{index_ML Syntax.check_typs: "Proof.context -> typ list -> typ list"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
213 |
@{index_ML Syntax.check_terms: "Proof.context -> term list -> term list"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
214 |
@{index_ML Syntax.check_props: "Proof.context -> term list -> term list"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
215 |
@{index_ML Syntax.uncheck_typs: "Proof.context -> typ list -> typ list"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
216 |
@{index_ML Syntax.uncheck_terms: "Proof.context -> term list -> term list"} \\
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
217 |
\end{mldecls}
|
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
218 |
|
| 52422 | 219 |
%FIXME description |
|
39876
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
220 |
*} |
|
1ff9bce085bd
preliminary material on "Concrete syntax and type-checking";
wenzelm
parents:
39865
diff
changeset
|
221 |
|
| 30124 | 222 |
end |