doc-src/TutorialI/Types/Records.thy
author wenzelm
Thu, 20 Dec 2001 21:11:04 +0100
changeset 12567 614ef5ca41ed
parent 12409 25bf458af885
child 12586 6bf380202adb
permissions -rw-r--r--
turned into proper Isabelle document;
paulson@11387
     1
wenzelm@12567
     2
header {* Records \label{sec:records} *}
paulson@11387
     3
wenzelm@12567
     4
(*<*)
paulson@11387
     5
theory Records = Main:
wenzelm@12567
     6
(*>*)
paulson@11387
     7
wenzelm@12567
     8
text {*
wenzelm@12567
     9
  \index{records|(}%
wenzelm@12567
    10
  Records are familiar from programming languages.  A record of $n$
wenzelm@12567
    11
  fields is essentially an $n$-tuple, but the record's components have
wenzelm@12567
    12
  names, which can make expressions easier to read and reduces the
wenzelm@12567
    13
  risk of confusing one field for another.
wenzelm@12567
    14
wenzelm@12567
    15
  A basic Isabelle record covers a certain set of fields, with select
wenzelm@12567
    16
  and update operations.  Each field has a specified type, which may
wenzelm@12567
    17
  be polymorphic.  The field names are part of the record type, and
wenzelm@12567
    18
  the order of the fields is significant --- as it is in Pascal but
wenzelm@12567
    19
  not in Standard ML.  If two different record types have field names
wenzelm@12567
    20
  in common, then the ambiguity is resolved in the usual way, by
wenzelm@12567
    21
  qualified names.
wenzelm@12567
    22
wenzelm@12567
    23
  Record types can also be defined by extending other record types.
wenzelm@12567
    24
  Extensible records make use of the reserved pseudo-field \cdx{more},
wenzelm@12567
    25
  which is present in every record type.  Generic record operations
wenzelm@12567
    26
  work on all possible extensions of a given type scheme; naive
wenzelm@12567
    27
  polymorphism takes care of structural sub-typing behind the scenes.
wenzelm@12567
    28
  There are also explicit coercion functions between fixed record
wenzelm@12567
    29
  types.
wenzelm@12567
    30
*}
wenzelm@12567
    31
wenzelm@12567
    32
wenzelm@12567
    33
subsection {* Record Basics *}
wenzelm@12567
    34
wenzelm@12567
    35
text {*
wenzelm@12567
    36
  Record types are not primitive in Isabelle and have a subtle
wenzelm@12567
    37
  internal representation based on nested copies of the primitive
wenzelm@12567
    38
  product type.  A \commdx{record} declaration introduces a new record
wenzelm@12567
    39
  type scheme by specifying its fields, which are packaged internally
wenzelm@12567
    40
  to hold up the perception of records as a separate concept.
wenzelm@12567
    41
*}
paulson@11387
    42
paulson@11387
    43
record point =
paulson@11387
    44
  Xcoord :: int
paulson@11387
    45
  Ycoord :: int
paulson@11387
    46
paulson@11387
    47
text {*
wenzelm@12567
    48
  Records of type @{typ point} have two fields named @{text Xcoord}
wenzelm@12567
    49
  and @{text Ycoord}, both of type~@{typ int}.  We now define a
wenzelm@12567
    50
  constant of type @{typ point}:
paulson@11387
    51
*}
paulson@11387
    52
wenzelm@12567
    53
constdefs
wenzelm@12567
    54
  pt1 :: point
wenzelm@12567
    55
  "pt1 \<equiv> (| Xcoord = 999, Ycoord = 23 |)"
paulson@11428
    56
paulson@11387
    57
text {*
wenzelm@12567
    58
  We see above the ASCII notation for record brackets.  You can also
wenzelm@12567
    59
  use the symbolic brackets @{text \<lparr>} and @{text \<rparr>}.  Record type
wenzelm@12567
    60
  expressions can be written directly as well, without referring to
wenzelm@12567
    61
  previously declared names (which happen to be mere type
wenzelm@12567
    62
  abbreviations):
paulson@11387
    63
*}
paulson@11387
    64
wenzelm@12567
    65
constdefs
wenzelm@12567
    66
  pt2 :: "\<lparr>Xcoord :: int, Ycoord :: int\<rparr>"
wenzelm@12567
    67
  "pt2 \<equiv> \<lparr>Xcoord = -45, Ycoord = 97\<rparr>"
wenzelm@12567
    68
paulson@11387
    69
text {*
wenzelm@12567
    70
  For each field, there is a \emph{selector} function of the same
wenzelm@12567
    71
  name.  For example, if @{text p} has type @{typ point} then @{text
wenzelm@12567
    72
  "Xcoord p"} denotes the value of the @{text Xcoord} field of~@{text
wenzelm@12567
    73
  p}.  Expressions involving field selection of explicit records are
wenzelm@12567
    74
  simplified automatically:
paulson@11387
    75
*}
paulson@11387
    76
wenzelm@12567
    77
lemma "Xcoord \<lparr>Xcoord = a, Ycoord = b\<rparr> = a"
wenzelm@12567
    78
  by simp
wenzelm@12567
    79
paulson@11387
    80
text {*
wenzelm@12567
    81
  The \emph{update} operation is functional.  For example, @{term
wenzelm@12567
    82
  "p\<lparr>Xcoord := 0\<rparr>"} is a record whose @{text Xcoord} value is zero
wenzelm@12567
    83
  and whose @{text Ycoord} value is copied from~@{text p}.  Updates
wenzelm@12567
    84
  are also simplified automatically:
paulson@11387
    85
*}
paulson@11387
    86
wenzelm@12567
    87
lemma "\<lparr>Xcoord = a, Ycoord = b\<rparr>\<lparr>Xcoord := 0\<rparr> =
wenzelm@12567
    88
    \<lparr>Xcoord = 0, Ycoord = b\<rparr>"
wenzelm@12567
    89
  by simp
wenzelm@12567
    90
paulson@11387
    91
text {*
wenzelm@12567
    92
  \begin{warn}
wenzelm@12567
    93
  Field names are declared as constants and can no longer be used as
wenzelm@12567
    94
  variables.  It would be unwise, for example, to call the fields of
wenzelm@12567
    95
  type @{typ point} simply @{text x} and~@{text y}.  Each record
wenzelm@12567
    96
  declaration introduces a constant \cdx{more}.
wenzelm@12567
    97
  \end{warn}
paulson@11387
    98
*}
paulson@11387
    99
paulson@11387
   100
wenzelm@12567
   101
subsection {* Extensible Records and Generic Operations *}
paulson@11387
   102
paulson@11387
   103
text {*
wenzelm@12567
   104
  \index{records!extensible|(}%
wenzelm@12567
   105
wenzelm@12567
   106
  Now, let us define coloured points (type @{text cpoint}) to be
wenzelm@12567
   107
  points extended with a field @{text col} of type @{text colour}:
paulson@11387
   108
*}
paulson@11387
   109
paulson@11387
   110
datatype colour = Red | Green | Blue
paulson@11387
   111
paulson@11387
   112
record cpoint = point +
paulson@11387
   113
  col :: colour
paulson@11387
   114
wenzelm@12567
   115
text {*
wenzelm@12567
   116
  The fields of this new type are @{text Xcoord}, @{text Ycoord} and
wenzelm@12567
   117
  @{text col}, in that order:
wenzelm@12567
   118
*}
paulson@11387
   119
wenzelm@12567
   120
constdefs
paulson@11387
   121
  cpt1 :: cpoint
wenzelm@12567
   122
  "cpt1 \<equiv> \<lparr>Xcoord = 999, Ycoord = 23, col = Green\<rparr>"
paulson@11387
   123
wenzelm@12567
   124
text {*
wenzelm@12567
   125
  We can define generic operations that work on arbitrary instances of
wenzelm@12567
   126
  a record scheme, e.g.\ covering @{typ point}, @{typ cpoint} and any
wenzelm@12567
   127
  further extensions.  Every record structure has an implicit
wenzelm@12567
   128
  pseudo-field, \cdx{more}, that keeps the extension as an explicit
wenzelm@12567
   129
  value.  Its type is declared as completely polymorphic:~@{typ 'a}.
wenzelm@12567
   130
  When a fixed record value is expressed using just its standard
wenzelm@12567
   131
  fields, the value of @{text more} is implicitly set to @{text "()"},
wenzelm@12567
   132
  the empty tuple, which has type @{typ unit}.  Within the record
wenzelm@12567
   133
  brackets, you can refer to the @{text more} field by writing @{text
wenzelm@12567
   134
  "\<dots>"} (three dots):
wenzelm@12567
   135
*}
paulson@11387
   136
wenzelm@12567
   137
lemma "Xcoord \<lparr>Xcoord = a, Ycoord = b, \<dots> = p\<rparr> = a"
wenzelm@12567
   138
  by simp
paulson@11387
   139
wenzelm@12567
   140
text {*
wenzelm@12567
   141
  This lemma applies to any record whose first two fields are @{text
wenzelm@12567
   142
  Xcoord} and~@{text Ycoord}.  Note that @{text "\<lparr>Xcoord = a, Ycoord
wenzelm@12567
   143
  = b, \<dots> = ()\<rparr>"} is actually the same as @{text "\<lparr>Xcoord = a,
wenzelm@12567
   144
  Ycoord = b\<rparr>"}.
paulson@11387
   145
wenzelm@12567
   146
  The pseudo-field @{text more} can be selected in the usual way, but
wenzelm@12567
   147
  the identifier must be qualified:
wenzelm@12567
   148
*}
paulson@11387
   149
paulson@11387
   150
lemma "point.more cpt1 = \<lparr>col = Green\<rparr>"
wenzelm@12567
   151
  by (simp add: cpt1_def)
paulson@11387
   152
paulson@11387
   153
text {*
wenzelm@12567
   154
  We see that the colour attached to this @{typ point} is a
wenzelm@12567
   155
  (rudimentary) record in its own right, namely @{text "\<lparr>col =
wenzelm@12567
   156
  Green\<rparr>"}.  In order to select or update @{text col} in the above
wenzelm@12567
   157
  fragment, @{text "\<lparr>col = Green\<rparr>"} needs to be put back into the
wenzelm@12567
   158
  context of its parent type scheme, say as @{text more} part of a
wenzelm@12567
   159
  @{typ point}.
wenzelm@12567
   160
wenzelm@12567
   161
  To define generic operations, we need to know a bit more about
wenzelm@12567
   162
  records.  Our declaration of @{typ point} above generated two type
wenzelm@12567
   163
  abbreviations:
wenzelm@12567
   164
wenzelm@12567
   165
  \smallskip
wenzelm@12567
   166
  \begin{tabular}{l}
wenzelm@12567
   167
  @{typ point}~@{text "="}~@{typ "\<lparr>Xcoord :: int, Ycoord :: int\<rparr>"} \\
wenzelm@12567
   168
  @{typ "'a point_scheme"}~@{text "="}~@{typ "\<lparr>Xcoord :: int, Ycoord :: int, \<dots> :: 'a\<rparr>"} \\
wenzelm@12567
   169
  \end{tabular}
wenzelm@12567
   170
  \smallskip
wenzelm@12567
   171
wenzelm@12567
   172
  Type @{typ point} is for rigid records having exactly the two fields
wenzelm@12567
   173
  @{text Xcoord} and~@{text Ycoord}, while the polymorphic type @{typ
wenzelm@12567
   174
  "'a point_scheme"} comprises all possible extensions to those two
wenzelm@12567
   175
  fields, recall that @{typ "unit point_scheme"} coincides with @{typ
wenzelm@12567
   176
  point}.  For example, let us define two operations --- methods, if
wenzelm@12567
   177
  we regard records as objects --- to get and set any point's @{text
wenzelm@12567
   178
  Xcoord} field.
paulson@11387
   179
*}
paulson@11387
   180
paulson@11387
   181
constdefs
wenzelm@11942
   182
  getX :: "'a point_scheme \<Rightarrow> int"
wenzelm@12567
   183
  "getX r \<equiv> Xcoord r"
wenzelm@12567
   184
  setX :: "'a point_scheme \<Rightarrow> int \<Rightarrow> 'a point_scheme"
wenzelm@12567
   185
  "setX r a \<equiv> r\<lparr>Xcoord := a\<rparr>"
paulson@11387
   186
paulson@11387
   187
text {*
wenzelm@12567
   188
  Here is a generic method that modifies a point, incrementing its
wenzelm@12567
   189
  @{text Xcoord} field.  The @{text Ycoord} and @{text more} fields
wenzelm@12567
   190
  are copied across.  It works for any record type scheme derived from
wenzelm@12567
   191
  @{typ point}, such as @{typ cpoint}:
paulson@11387
   192
*}
paulson@11387
   193
paulson@11387
   194
constdefs
wenzelm@11942
   195
  incX :: "'a point_scheme \<Rightarrow> 'a point_scheme"
wenzelm@12567
   196
  "incX r \<equiv> \<lparr>Xcoord = Xcoord r + 1,
wenzelm@12567
   197
    Ycoord = Ycoord r, \<dots> = point.more r\<rparr>"
paulson@11387
   198
paulson@11387
   199
text {*
wenzelm@12567
   200
  Generic theorems can be proved about generic methods.  This trivial
wenzelm@12567
   201
  lemma relates @{text incX} to @{text getX} and @{text setX}:
wenzelm@12567
   202
*}
paulson@11387
   203
wenzelm@12567
   204
lemma "incX r = setX r (getX r + 1)"
wenzelm@12567
   205
  by (simp add: getX_def setX_def incX_def)
paulson@11387
   206
paulson@11387
   207
text {*
wenzelm@12567
   208
  \begin{warn}
wenzelm@12567
   209
  If you use the symbolic record brackets @{text \<lparr>} and @{text \<rparr>},
wenzelm@12567
   210
  then you must also use the symbolic ellipsis, ``@{text \<dots>}'', rather
wenzelm@12567
   211
  than three consecutive periods, ``@{text "..."}''.  Mixing the ASCII
wenzelm@12567
   212
  and symbolic versions causes a syntax error.  (The two versions are
wenzelm@12567
   213
  more distinct on screen than they are on paper.)
wenzelm@12567
   214
  \end{warn}%\index{records!extensible|)}
wenzelm@12567
   215
*}
paulson@11387
   216
paulson@11387
   217
wenzelm@12567
   218
subsection {* Record Equality *}
paulson@11387
   219
wenzelm@12567
   220
text {*
wenzelm@12567
   221
  Two records are equal\index{equality!of records} if all pairs of
wenzelm@12567
   222
  corresponding fields are equal.  Record equalities are simplified
wenzelm@12567
   223
  automatically:
wenzelm@12567
   224
*}
paulson@11387
   225
wenzelm@12567
   226
lemma "(\<lparr>Xcoord = a, Ycoord = b\<rparr> = \<lparr>Xcoord = a', Ycoord = b'\<rparr>) =
wenzelm@12567
   227
    (a = a' \<and> b = b')"
wenzelm@12567
   228
  by simp
paulson@11387
   229
wenzelm@12567
   230
text {*
wenzelm@12567
   231
  The following equality is similar, but generic, in that @{text r}
wenzelm@12567
   232
  can be any instance of @{text point_scheme}:
wenzelm@12567
   233
*}
paulson@11387
   234
wenzelm@12567
   235
lemma "r\<lparr>Xcoord := a, Ycoord := b\<rparr> = r\<lparr>Ycoord := b, Xcoord := a\<rparr>"
wenzelm@12567
   236
  by simp
paulson@11387
   237
wenzelm@12567
   238
text {*
wenzelm@12567
   239
  We see above the syntax for iterated updates.  We could equivalently
wenzelm@12567
   240
  have written the left-hand side as @{term "r\<lparr>Xcoord := a\<rparr>\<lparr>Ycoord
wenzelm@12567
   241
  := b\<rparr>"}.
paulson@11387
   242
wenzelm@12567
   243
  Record equality is \emph{extensional}: \index{extensionality!for
wenzelm@12567
   244
  records} a record is determined entirely by the values of its
wenzelm@12567
   245
  fields.
wenzelm@12567
   246
*}
paulson@11387
   247
paulson@11387
   248
lemma "r = \<lparr>Xcoord = Xcoord r, Ycoord = Ycoord r\<rparr>"
wenzelm@12567
   249
  by simp
paulson@11387
   250
paulson@11387
   251
text {*
wenzelm@12567
   252
  The generic version of this equality includes the field @{text
wenzelm@12567
   253
  more}:
paulson@11387
   254
*}
paulson@11387
   255
paulson@11387
   256
lemma "r = \<lparr>Xcoord = Xcoord r, Ycoord = Ycoord r, \<dots> = point.more r\<rparr>"
wenzelm@12567
   257
  by simp
paulson@11387
   258
paulson@11387
   259
text {*
wenzelm@12567
   260
  Note that the @{text r} is of a different (more general) type than
wenzelm@12567
   261
  the previous one.
wenzelm@12567
   262
wenzelm@12567
   263
  \medskip The simplifier can prove many record equalities
wenzelm@12567
   264
  automatically, but general equality reasoning can be tricky.
wenzelm@12567
   265
  Consider proving this obvious fact:
paulson@11387
   266
*}
paulson@11387
   267
wenzelm@12567
   268
lemma "r\<lparr>Xcoord := a\<rparr> = r\<lparr>Xcoord := a'\<rparr> \<Longrightarrow> a = a'"
wenzelm@12567
   269
  apply simp?
wenzelm@12567
   270
  oops
paulson@11387
   271
wenzelm@12567
   272
text {*
wenzelm@12567
   273
  The simplifier can do nothing, since general record equality is not
wenzelm@12567
   274
  eliminated automatically.  One way to proceed is by an explicit
wenzelm@12567
   275
  forward step that applies the selector @{text Xcoord} to both sides
wenzelm@12567
   276
  of the assumed record equality:
wenzelm@12567
   277
*}
wenzelm@12567
   278
wenzelm@12567
   279
lemma "r\<lparr>Xcoord := a\<rparr> = r\<lparr>Xcoord := a'\<rparr> \<Longrightarrow> a = a'"
wenzelm@12567
   280
  apply (drule_tac f = Xcoord in arg_cong)
wenzelm@12567
   281
  txt {* @{subgoals [display, indent = 0, margin = 65]}
wenzelm@12567
   282
    Now, @{text simp} will reduce the assumption to the desired
wenzelm@12567
   283
    conclusion. *}
wenzelm@12567
   284
  apply simp
wenzelm@12567
   285
  done
wenzelm@12567
   286
wenzelm@12567
   287
text {*
wenzelm@12567
   288
  The @{text cases} method is preferable to such a forward proof.
wenzelm@12567
   289
  State the desired lemma again:
wenzelm@12567
   290
*}
wenzelm@12567
   291
wenzelm@12567
   292
lemma "r\<lparr>Xcoord := a\<rparr> = r\<lparr>Xcoord := a'\<rparr> \<Longrightarrow> a = a'"
wenzelm@12567
   293
  txt {*
wenzelm@12567
   294
    The \methdx{cases} method adds an equality to replace the named
wenzelm@12567
   295
    record variable by an explicit record, listing all fields.  It
wenzelm@12567
   296
    even includes the pseudo-field @{text more}, since the record
wenzelm@12567
   297
    equality stated above is generic. *}
wenzelm@12567
   298
  apply (cases r)
wenzelm@12567
   299
wenzelm@12567
   300
  txt {* @{subgoals [display, indent = 0, margin = 65]}
wenzelm@12567
   301
    Again, @{text simp} finishes the proof.  Because @{text r} has
wenzelm@12567
   302
    become an explicit record expression, the updates can be applied
wenzelm@12567
   303
    and the record equality can be replaced by equality of the
wenzelm@12567
   304
    corresponding fields (due to injectivity). *}
wenzelm@12567
   305
  apply simp
wenzelm@12567
   306
  done
wenzelm@12567
   307
wenzelm@12567
   308
wenzelm@12567
   309
subsection {* Extending and Truncating Records *}
wenzelm@12567
   310
wenzelm@12567
   311
text {*
wenzelm@12567
   312
  Each record declaration introduces functions to refer collectively
wenzelm@12567
   313
  to a record's fields and to convert between related record types.
wenzelm@12567
   314
  They can, for instance, convert between types @{typ point} and @{typ
wenzelm@12567
   315
  cpoint}.  We can add a colour to a point or to convert a @{typ
wenzelm@12567
   316
  cpoint} to a @{typ point} by forgetting its colour.
wenzelm@12567
   317
wenzelm@12567
   318
  \begin{itemize}
wenzelm@12567
   319
wenzelm@12567
   320
  \item Function \cdx{make} takes as arguments all of the record's
wenzelm@12567
   321
  fields.  It returns the corresponding record.
wenzelm@12567
   322
wenzelm@12567
   323
  \item Function \cdx{fields} takes the record's new fields and
wenzelm@12567
   324
  returns a record fragment consisting of just those fields.  This may
wenzelm@12567
   325
  be filled into the @{text more} part of the parent record scheme.
wenzelm@12567
   326
wenzelm@12567
   327
  \item Function \cdx{extend} takes two arguments: a record to be
wenzelm@12567
   328
  extended and a record containing the new fields.
wenzelm@12567
   329
wenzelm@12567
   330
  \item Function \cdx{truncate} takes a record (possibly an extension
wenzelm@12567
   331
  of the original record type) and returns a fixed record, removing
wenzelm@12567
   332
  any additional fields.
wenzelm@12567
   333
wenzelm@12567
   334
  \end{itemize}
wenzelm@12567
   335
wenzelm@12567
   336
  These functions merely provide handsome abbreviations for standard
wenzelm@12567
   337
  record expressions involving constructors and selectors.  The
wenzelm@12567
   338
  definitions, which are \emph{not} unfolded by default, are made
wenzelm@12567
   339
  available by the collective name of @{text defs} (e.g.\ @{text
wenzelm@12567
   340
  point.defs} or @{text cpoint.defs}).
wenzelm@12567
   341
wenzelm@12567
   342
  For example, here are the versions of those functions generated for
wenzelm@12567
   343
  record @{typ point}.  We omit @{text point.fields}, which happens to
wenzelm@12567
   344
  be the same as @{text point.make}.
wenzelm@12567
   345
wenzelm@12567
   346
  @{thm [display, indent = 0, margin = 65] point.make_def
wenzelm@12567
   347
  point.extend_def point.truncate_def}
wenzelm@12567
   348
wenzelm@12567
   349
  Contrast those with the corresponding functions for record @{typ
wenzelm@12567
   350
  cpoint}.  Observe @{text cpoint.fields} in particular.
wenzelm@12567
   351
wenzelm@12567
   352
  @{thm [display, indent = 0, margin = 65] cpoint.make_def
wenzelm@12567
   353
  cpoint.extend_def cpoint.truncate_def}
wenzelm@12567
   354
wenzelm@12567
   355
  To demonstrate these functions, we declare a new coloured point by
wenzelm@12567
   356
  extending an ordinary point.  Function @{text point.extend} augments
wenzelm@12567
   357
  @{text pt1} with a colour, which is converted into an appropriate
wenzelm@12567
   358
  record fragment by @{text cpoint.fields}.
wenzelm@12567
   359
*}
paulson@11387
   360
paulson@12407
   361
constdefs
paulson@12407
   362
  cpt2 :: cpoint
wenzelm@12567
   363
  "cpt2 \<equiv> point.extend pt1 (cpoint.fields Green)"
paulson@12407
   364
paulson@12407
   365
text {*
wenzelm@12567
   366
  The coloured points @{text cpt1} and @{text cpt2} are equal.  The
wenzelm@12567
   367
  proof is trivial, by unfolding all the definitions.  We deliberately
wenzelm@12567
   368
  omit the definition of~@{text pt1} in order to reveal the underlying
wenzelm@12567
   369
  comparison on type @{typ point}.
wenzelm@12567
   370
*}
wenzelm@12567
   371
wenzelm@12567
   372
lemma "cpt1 = cpt2"
wenzelm@12567
   373
  apply (simp add: cpt1_def cpt2_def point.defs cpoint.defs)
wenzelm@12567
   374
  txt {* @{subgoals [display, indent = 0, margin = 65]} *}
wenzelm@12567
   375
  apply (simp add: pt1_def)
wenzelm@12567
   376
  done
paulson@12407
   377
paulson@12407
   378
text {*
wenzelm@12567
   379
  In the example below, a coloured point is truncated to leave a
wenzelm@12567
   380
  point.  We must use the @{text truncate} function of the shorter
wenzelm@12567
   381
  record.
wenzelm@12567
   382
*}
paulson@12407
   383
paulson@12407
   384
lemma "point.truncate cpt2 = pt1"
wenzelm@12567
   385
  by (simp add: pt1_def cpt2_def point.defs)
paulson@12407
   386
wenzelm@12567
   387
text {*
wenzelm@12567
   388
  \begin{exercise}
wenzelm@12567
   389
  Extend record @{typ cpoint} to have a further field, @{text
wenzelm@12567
   390
  intensity}, of type~@{typ nat}.  Experiment with coercions among the
wenzelm@12567
   391
  three record types.
wenzelm@12567
   392
  \end{exercise}
wenzelm@12567
   393
wenzelm@12567
   394
  \begin{exercise}
wenzelm@12567
   395
  (For Java programmers.)
wenzelm@12567
   396
  Model a small class hierarchy using records.
wenzelm@12567
   397
  \end{exercise}
wenzelm@12567
   398
  \index{records|)}
wenzelm@12567
   399
*}
wenzelm@12567
   400
wenzelm@12567
   401
(*<*)
paulson@11387
   402
end
wenzelm@12567
   403
(*>*)