src/HOL/ex/Records.thy
author wenzelm
Tue, 05 Nov 2019 14:28:00 +0100
changeset 71047 87c132cf5860
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
more options;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/ex/Records.thy
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
     2
    Author:     Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel,
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 12591
diff changeset
     3
                TU Muenchen
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
     4
*)
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
     5
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     6
section \<open>Using extensible records in HOL -- points and coloured points\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
     7
33596
27c5023ee818 tuned header
haftmann
parents: 31723
diff changeset
     8
theory Records
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
     9
imports Main
33596
27c5023ee818 tuned header
haftmann
parents: 31723
diff changeset
    10
begin
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    11
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    12
subsection \<open>Points\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    13
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    14
record point =
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    15
  xpos :: nat
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    16
  ypos :: nat
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    17
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    18
text \<open>
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    19
  Apart many other things, above record declaration produces the
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    20
  following theorems:
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    21
\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    22
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 12591
diff changeset
    23
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    24
thm point.simps
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    25
thm point.iffs
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    26
thm point.defs
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    27
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    28
text \<open>
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    29
  The set of theorems @{thm [source] point.simps} is added
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    30
  automatically to the standard simpset, @{thm [source] point.iffs} is
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    31
  added to the Classical Reasoner and Simplifier context.
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    32
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 12591
diff changeset
    33
  \medskip Record declarations define new types and type abbreviations:
10357
wenzelm
parents: 10052
diff changeset
    34
  @{text [display]
61499
4efe9a6dd212 tuned document;
wenzelm
parents: 61343
diff changeset
    35
\<open>point = \<lparr>xpos :: nat, ypos :: nat\<rparr> = () point_ext_type
4efe9a6dd212 tuned document;
wenzelm
parents: 61343
diff changeset
    36
'a point_scheme = \<lparr>xpos :: nat, ypos :: nat, ... :: 'a\<rparr>  = 'a point_ext_type\<close>}
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    37
\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    38
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    39
consts foo2 :: "(| xpos :: nat, ypos :: nat |)"
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    40
consts foo4 :: "'a => (| xpos :: nat, ypos :: nat, ... :: 'a |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    41
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    42
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    43
subsubsection \<open>Introducing concrete records and record schemes\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    44
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    45
definition foo1 :: point
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    46
  where "foo1 = (| xpos = 1, ypos = 0 |)"
22737
haftmann
parents: 21404
diff changeset
    47
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    48
definition foo3 :: "'a => 'a point_scheme"
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    49
  where "foo3 ext = (| xpos = 1, ypos = 0, ... = ext |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    50
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    51
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    52
subsubsection \<open>Record selection and record update\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    53
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    54
definition getX :: "'a point_scheme => nat"
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    55
  where "getX r = xpos r"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    56
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    57
definition setX :: "'a point_scheme => nat => 'a point_scheme"
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    58
  where "setX r n = r (| xpos := n |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    59
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    60
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    61
subsubsection \<open>Some lemmas about records\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    62
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    63
text \<open>Basic simplifications.\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    64
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    65
lemma "point.make n p = (| xpos = n, ypos = p |)"
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    66
  by (simp only: point.make_def)
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    67
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    68
lemma "xpos (| xpos = m, ypos = n, ... = p |) = m"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    69
  by simp
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    70
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    71
lemma "(| xpos = m, ypos = n, ... = p |) (| xpos:= 0 |) = (| xpos = 0, ypos = n, ... = p |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    72
  by simp
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    73
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    74
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    75
text \<open>\medskip Equality of records.\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    76
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    77
lemma "n = n' ==> p = p' ==> (| xpos = n, ypos = p |) = (| xpos = n', ypos = p' |)"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62119
diff changeset
    78
  \<comment> \<open>introduction of concrete record equality\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    79
  by simp
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    80
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    81
lemma "(| xpos = n, ypos = p |) = (| xpos = n', ypos = p' |) ==> n = n'"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62119
diff changeset
    82
  \<comment> \<open>elimination of concrete record equality\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    83
  by simp
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    84
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    85
lemma "r (| xpos := n |) (| ypos := m |) = r (| ypos := m |) (| xpos := n |)"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62119
diff changeset
    86
  \<comment> \<open>introduction of abstract record equality\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    87
  by simp
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    88
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    89
lemma "r (| xpos := n |) = r (| xpos := n' |) ==> n = n'"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62119
diff changeset
    90
  \<comment> \<open>elimination of abstract record equality (manual proof)\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    91
proof -
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
    92
  assume "r (| xpos := n |) = r (| xpos := n' |)" (is "?lhs = ?rhs")
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    93
  then have "xpos ?lhs = xpos ?rhs" by simp
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
    94
  then show ?thesis by simp
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    95
qed
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    96
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    97
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    98
text \<open>\medskip Surjective pairing\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
    99
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   100
lemma "r = (| xpos = xpos r, ypos = ypos r |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   101
  by simp
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   102
12591
5a46569d2b05 qualified point.more;
wenzelm
parents: 12266
diff changeset
   103
lemma "r = (| xpos = xpos r, ypos = ypos r, ... = point.more r |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   104
  by simp
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   105
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   106
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   107
text \<open>
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   108
  \medskip Representation of records by cases or (degenerate)
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   109
  induction.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   110
\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   111
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 12591
diff changeset
   112
lemma "r(| xpos := n |) (| ypos := m |) = r (| ypos := m |) (| xpos := n |)"
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   113
proof (cases r)
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   114
  fix xpos ypos more
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   115
  assume "r = (| xpos = xpos, ypos = ypos, ... = more |)"
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   116
  then show ?thesis by simp
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   117
qed
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   118
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   119
lemma "r (| xpos := n |) (| ypos := m |) = r (| ypos := m |) (| xpos := n |)"
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   120
proof (induct r)
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   121
  fix xpos ypos more
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   122
  show "(| xpos = xpos, ypos = ypos, ... = more |) (| xpos := n, ypos := m |) =
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   123
      (| xpos = xpos, ypos = ypos, ... = more |) (| ypos := m, xpos := n |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   124
    by simp
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   125
qed
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   126
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   127
lemma "r (| xpos := n |) (| xpos := m |) = r (| xpos := m |)"
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   128
proof (cases r)
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   129
  fix xpos ypos more
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   130
  assume "r = \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>"
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   131
  then show ?thesis by simp
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   132
qed
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   133
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   134
lemma "r (| xpos := n |) (| xpos := m |) = r (| xpos := m |)"
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   135
proof (cases r)
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   136
  case fields
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   137
  then show ?thesis by simp
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   138
qed
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   139
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   140
lemma "r (| xpos := n |) (| xpos := m |) = r (| xpos := m |)"
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   141
  by (cases r) simp
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   142
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   143
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   144
text \<open>
10357
wenzelm
parents: 10052
diff changeset
   145
 \medskip Concrete records are type instances of record schemes.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   146
\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   147
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   148
definition foo5 :: nat
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   149
  where "foo5 = getX (| xpos = 1, ypos = 0 |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   150
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   151
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61499
diff changeset
   152
text \<open>\medskip Manipulating the ``\<open>...\<close>'' (more) part.\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   153
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   154
definition incX :: "'a point_scheme => 'a point_scheme"
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   155
  where "incX r = (| xpos = xpos r + 1, ypos = ypos r, ... = point.more r |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   156
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   157
lemma "incX r = setX r (Suc (getX r))"
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   158
  by (simp add: getX_def setX_def incX_def)
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   159
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   160
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   161
text \<open>An alternative definition.\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   162
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   163
definition incX' :: "'a point_scheme => 'a point_scheme"
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   164
  where "incX' r = r (| xpos := xpos r + 1 |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   165
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   166
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   167
subsection \<open>Coloured points: record extension\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   168
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
   169
datatype colour = Red | Green | Blue
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   170
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   171
record cpoint = point +
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   172
  colour :: colour
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   173
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   174
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   175
text \<open>
61499
4efe9a6dd212 tuned document;
wenzelm
parents: 61343
diff changeset
   176
  The record declaration defines a new type constructor and abbreviations:
10357
wenzelm
parents: 10052
diff changeset
   177
  @{text [display]
61499
4efe9a6dd212 tuned document;
wenzelm
parents: 61343
diff changeset
   178
\<open>cpoint = (| xpos :: nat, ypos :: nat, colour :: colour |) =
4efe9a6dd212 tuned document;
wenzelm
parents: 61343
diff changeset
   179
  () cpoint_ext_type point_ext_type
4efe9a6dd212 tuned document;
wenzelm
parents: 61343
diff changeset
   180
'a cpoint_scheme = (| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |) =
4efe9a6dd212 tuned document;
wenzelm
parents: 61343
diff changeset
   181
  'a cpoint_ext_type point_ext_type\<close>}
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   182
\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   183
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   184
consts foo6 :: cpoint
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   185
consts foo7 :: "(| xpos :: nat, ypos :: nat, colour :: colour |)"
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   186
consts foo8 :: "'a cpoint_scheme"
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   187
consts foo9 :: "(| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   188
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   189
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   190
text \<open>
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61499
diff changeset
   191
 Functions on \<open>point\<close> schemes work for \<open>cpoints\<close> as well.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   192
\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   193
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   194
definition foo10 :: nat
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   195
  where "foo10 = getX (| xpos = 2, ypos = 0, colour = Blue |)"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   196
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   197
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   198
subsubsection \<open>Non-coercive structural subtyping\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   199
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   200
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   201
 Term \<^term>\<open>foo11\<close> has type \<^typ>\<open>cpoint\<close>, not type \<^typ>\<open>point\<close> ---
10357
wenzelm
parents: 10052
diff changeset
   202
 Great!
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   203
\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   204
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   205
definition foo11 :: cpoint
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   206
  where "foo11 = setX (| xpos = 2, ypos = 0, colour = Blue |) 0"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   207
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   208
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   209
subsection \<open>Other features\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   210
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   211
text \<open>Field names contribute to record identity.\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   212
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   213
record point' =
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   214
  xpos' :: nat
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   215
  ypos' :: nat
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   216
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   217
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   218
  \noindent May not apply \<^term>\<open>getX\<close> to @{term [source] "(| xpos' =
11939
c5e69470f03b updated records;
wenzelm
parents: 11928
diff changeset
   219
  2, ypos' = 0 |)"} -- type error.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   220
\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   221
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   222
text \<open>\medskip Polymorphic records.\<close>
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   223
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   224
record 'a point'' = point +
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   225
  content :: 'a
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   226
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 38012
diff changeset
   227
type_synonym cpoint'' = "colour point''"
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   228
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   229
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   230
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   231
text \<open>Updating a record field with an identical value is simplified.\<close>
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   232
lemma "r (| xpos := xpos r |) = r"
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   233
  by simp
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   234
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   235
text \<open>Only the most recent update to a component survives simplification.\<close>
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   236
lemma "r (| xpos := x, ypos := y, xpos := x' |) = r (| ypos := y, xpos := x' |)"
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   237
  by simp
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   238
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   239
text \<open>In some cases its convenient to automatically split
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   240
(quantified) records. For this purpose there is the simproc @{ML [source]
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37826
diff changeset
   241
"Record.split_simproc"} and the tactic @{ML [source]
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37826
diff changeset
   242
"Record.split_simp_tac"}.  The simplification procedure
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   243
only splits the records, whereas the tactic also simplifies the
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   244
resulting goal with the standard record simplification rules. A
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   245
(generalized) predicate on the record is passed as parameter that
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   246
decides whether or how `deep' to split the record. It can peek on the
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   247
subterm starting at the quantified occurrence of the record (including
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   248
the quantifier). The value \<^ML>\<open>0\<close> indicates no split, a value
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   249
greater \<^ML>\<open>0\<close> splits up to the given bound of record extension and
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   250
finally the value \<^ML>\<open>~1\<close> completely splits the record.
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37826
diff changeset
   251
@{ML [source] "Record.split_simp_tac"} additionally takes a list of
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   252
equations for simplification and can also split fixed record variables.
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   253
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   254
\<close>
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   255
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   256
lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   257
  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   258
    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   259
  apply simp
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   260
  done
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   261
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   262
lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   263
  apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   264
  apply simp
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   265
  done
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   266
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   267
lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   268
  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   269
    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   270
  apply simp
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   271
  done
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   272
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   273
lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   274
  apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   275
  apply simp
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   276
  done
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   277
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   278
lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   279
  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   280
    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   281
  apply auto
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   282
  done
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   283
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   284
lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   285
  apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   286
  apply auto
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   287
  done
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   288
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   289
lemma "P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   290
  apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   291
  apply auto
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   292
  done
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   293
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   294
lemma True
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   295
proof -
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   296
  {
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25707
diff changeset
   297
    fix P r
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   298
    assume pre: "P (xpos r)"
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   299
    then have "\<exists>x. P x"
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   300
      apply -
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   301
      apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
46231
76e32c39dd43 tuned example;
wenzelm
parents: 42463
diff changeset
   302
      apply auto
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   303
      done
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   304
  }
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   305
  show ?thesis ..
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   306
qed
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   307
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   308
text \<open>The effect of simproc @{ML [source] Record.ex_sel_eq_simproc} is
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   309
  illustrated by the following lemma.\<close>
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   310
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   311
lemma "\<exists>r. xpos r = x"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   312
  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   313
    addsimprocs [Record.ex_sel_eq_simproc]) 1\<close>)
25707
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   314
  done
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   315
0a599404f5a1 more examples
schirmer
parents: 22737
diff changeset
   316
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   317
subsection \<open>A more complex record expression\<close>
34971
5c290f56ebf7 a more complex record expression -- cf. src/HOL/Tools/quickcheck_generators.ML
haftmann
parents: 33613
diff changeset
   318
5c290f56ebf7 a more complex record expression -- cf. src/HOL/Tools/quickcheck_generators.ML
haftmann
parents: 33613
diff changeset
   319
record ('a, 'b, 'c) bar = bar1 :: 'a
5c290f56ebf7 a more complex record expression -- cf. src/HOL/Tools/quickcheck_generators.ML
haftmann
parents: 33613
diff changeset
   320
  bar2 :: 'b
5c290f56ebf7 a more complex record expression -- cf. src/HOL/Tools/quickcheck_generators.ML
haftmann
parents: 33613
diff changeset
   321
  bar3 :: 'c
5c290f56ebf7 a more complex record expression -- cf. src/HOL/Tools/quickcheck_generators.ML
haftmann
parents: 33613
diff changeset
   322
  bar21 :: "'b \<times> 'a"
5c290f56ebf7 a more complex record expression -- cf. src/HOL/Tools/quickcheck_generators.ML
haftmann
parents: 33613
diff changeset
   323
  bar32 :: "'c \<times> 'b"
5c290f56ebf7 a more complex record expression -- cf. src/HOL/Tools/quickcheck_generators.ML
haftmann
parents: 33613
diff changeset
   324
  bar31 :: "'c \<times> 'a"
5c290f56ebf7 a more complex record expression -- cf. src/HOL/Tools/quickcheck_generators.ML
haftmann
parents: 33613
diff changeset
   325
62119
b8c973d90ae7 add more frequently-run test for print_record
kleing
parents: 61933
diff changeset
   326
print_record "('a,'b,'c) bar"
34971
5c290f56ebf7 a more complex record expression -- cf. src/HOL/Tools/quickcheck_generators.ML
haftmann
parents: 33613
diff changeset
   327
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   328
subsection \<open>Some code generation\<close>
33613
ad27f52ee759 explicit invocation of code generation
haftmann
parents: 33596
diff changeset
   329
37826
4c0a5e35931a avoid export_code ... file -
haftmann
parents: 34971
diff changeset
   330
export_code foo1 foo3 foo5 foo10 checking SML
33613
ad27f52ee759 explicit invocation of code generation
haftmann
parents: 33596
diff changeset
   331
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   332
text \<open>
47842
bfc08ce7b7b9 provide [[record_codegen]] option for skipping codegen setup for records
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 46231
diff changeset
   333
  Code generation can also be switched off, for instance for very large records
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   334
\<close>
47842
bfc08ce7b7b9 provide [[record_codegen]] option for skipping codegen setup for records
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 46231
diff changeset
   335
bfc08ce7b7b9 provide [[record_codegen]] option for skipping codegen setup for records
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 46231
diff changeset
   336
declare [[record_codegen = false]]
bfc08ce7b7b9 provide [[record_codegen]] option for skipping codegen setup for records
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 46231
diff changeset
   337
bfc08ce7b7b9 provide [[record_codegen]] option for skipping codegen setup for records
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 46231
diff changeset
   338
record not_so_large_record =
bfc08ce7b7b9 provide [[record_codegen]] option for skipping codegen setup for records
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 46231
diff changeset
   339
  bar520 :: nat
bfc08ce7b7b9 provide [[record_codegen]] option for skipping codegen setup for records
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 46231
diff changeset
   340
  bar521 :: "nat * nat"
bfc08ce7b7b9 provide [[record_codegen]] option for skipping codegen setup for records
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 46231
diff changeset
   341
bfc08ce7b7b9 provide [[record_codegen]] option for skipping codegen setup for records
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 46231
diff changeset
   342
declare [[record_codegen = true]]
bfc08ce7b7b9 provide [[record_codegen]] option for skipping codegen setup for records
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 46231
diff changeset
   343
10052
5fa8d8d5c852 renamed HOL/ex/Points to HOL/ex/Records;
wenzelm
parents:
diff changeset
   344
end