Theory Records
section ‹Records \label{sec:records}›
theory Records imports Main begin
text ‹
\index{records|(}%
Records are familiar from programming languages. A record of $n$
fields is essentially an $n$-tuple, but the record's components have
names, which can make expressions easier to read and reduces the
risk of confusing one field for another.
A record of Isabelle/HOL covers a collection of fields, with select
and update operations. Each field has a specified type, which may
be polymorphic. The field names are part of the record type, and
the order of the fields is significant --- as it is in Pascal but
not in Standard ML. If two different record types have field names
in common, then the ambiguity is resolved in the usual way, by
qualified names.
Record types can also be defined by extending other record types.
Extensible records make use of the reserved pseudo-field \cdx{more},
which is present in every record type. Generic record operations
work on all possible extensions of a given type scheme; polymorphism
takes care of structural sub-typing behind the scenes. There are
also explicit coercion functions between fixed record types.
›
subsection ‹Record Basics›
text ‹
Record types are not primitive in Isabelle and have a delicate
internal representation \<^cite>‹"NaraschewskiW-TPHOLs98"›, based on
nested copies of the primitive product type. A \commdx{record}
declaration introduces a new record type scheme by specifying its
fields, which are packaged internally to hold up the perception of
the record as a distinguished entity. Here is a simple example:
›
record point =
Xcoord :: int
Ycoord :: int
text ‹\noindent
Records of type \<^typ>‹point› have two fields named \<^const>‹Xcoord›
and \<^const>‹Ycoord›, both of type~\<^typ>‹int›. We now define a
constant of type \<^typ>‹point›:
›
definition pt1 :: point where
"pt1 ≡ (| Xcoord = 999, Ycoord = 23 |)"
text ‹\noindent
We see above the ASCII notation for record brackets. You can also
use the symbolic brackets ‹⦇› and ‹⦈›. Record type
expressions can be also written directly with individual fields.
The type name above is merely an abbreviation.
›
definition pt2 :: "⦇Xcoord :: int, Ycoord :: int⦈" where
"pt2 ≡ ⦇Xcoord = -45, Ycoord = 97⦈"
text ‹
For each field, there is a \emph{selector}\index{selector!record}
function of the same name. For example, if ‹p› has type \<^typ>‹point› then ‹Xcoord p› denotes the value of the ‹Xcoord› field of~‹p›. Expressions involving field selection
of explicit records are simplified automatically:
›
lemma "Xcoord ⦇Xcoord = a, Ycoord = b⦈ = a"
by simp
text ‹
The \emph{update}\index{update!record} operation is functional. For
example, \<^term>‹p⦇Xcoord := 0⦈› is a record whose \<^const>‹Xcoord›
value is zero and whose \<^const>‹Ycoord› value is copied from~‹p›. Updates of explicit records are also simplified automatically:
›
lemma "⦇Xcoord = a, Ycoord = b⦈⦇Xcoord := 0⦈ =
⦇Xcoord = 0, Ycoord = b⦈"
by simp
text ‹
\begin{warn}
Field names are declared as constants and can no longer be used as
variables. It would be unwise, for example, to call the fields of
type \<^typ>‹point› simply ‹x› and~‹y›.
\end{warn}
›
subsection ‹Extensible Records and Generic Operations›
text ‹
\index{records!extensible|(}%
Now, let us define coloured points (type ‹cpoint›) to be
points extended with a field ‹col› of type ‹colour›:
›