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