5753
|
1 |
(* Title: HOL/ex/Points.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Wolfgang Naraschewski and Markus Wenzel, TU Muenchen
|
|
4 |
*)
|
|
5 |
|
6737
|
6 |
theory Points = Main:;
|
|
7 |
|
|
8 |
title "Basic use of extensible records in HOL, including the famous points
|
|
9 |
and coloured points.";
|
5741
|
10 |
|
|
11 |
|
6737
|
12 |
section "Points";
|
5741
|
13 |
|
|
14 |
record point =
|
|
15 |
x :: nat
|
6737
|
16 |
y :: nat;
|
5741
|
17 |
|
6737
|
18 |
text {|
|
|
19 |
Apart many other things, above record declaration produces the
|
|
20 |
following theorems:
|
5753
|
21 |
|
6737
|
22 |
thms "point.simps"
|
|
23 |
thms "point.iffs"
|
|
24 |
thms "point.update_defs"
|
5753
|
25 |
|
|
26 |
The set of theorems "point.simps" is added automatically to the
|
|
27 |
standard simpset, "point.iffs" is added to the claset and simpset.
|
6737
|
28 |
|};
|
5741
|
29 |
|
6737
|
30 |
text {|
|
5753
|
31 |
Record declarations define new type abbreviations:
|
|
32 |
|
6737
|
33 |
point = "(| x :: nat, y :: nat |)"
|
|
34 |
'a point_scheme = "(| x :: nat, y :: nat, ... :: 'a |)"
|
5753
|
35 |
|
|
36 |
Extensions `...' must be in type class `more'!
|
6737
|
37 |
|};
|
|
38 |
|
|
39 |
consts foo1 :: point;
|
|
40 |
consts foo2 :: "(| x :: nat, y :: nat |)";
|
|
41 |
consts foo3 :: "'a => ('a::more) point_scheme";
|
|
42 |
consts foo4 :: "'a => (| x :: nat, y :: nat, ... :: 'a |)";
|
|
43 |
|
|
44 |
|
|
45 |
subsection "Introducing concrete records and record schemes";
|
|
46 |
|
|
47 |
defs
|
|
48 |
foo1_def: "foo1 == (| x = 1, y = 0 |)"
|
|
49 |
foo3_def: "foo3 ext == (| x = 1, y = 0, ... = ext |)";
|
5741
|
50 |
|
6737
|
51 |
|
|
52 |
subsection "Record selection and record update";
|
|
53 |
|
|
54 |
constdefs
|
|
55 |
getX :: "('a::more) point_scheme => nat"
|
|
56 |
"getX r == x r"
|
|
57 |
setX :: "('a::more) point_scheme => nat => 'a point_scheme"
|
|
58 |
"setX r n == r (| x := n |)";
|
|
59 |
|
|
60 |
|
|
61 |
subsection "Some lemmas about records";
|
|
62 |
|
|
63 |
text "Note: any of these goals may be solved with a single
|
|
64 |
stroke of auto or force.";
|
5741
|
65 |
|
|
66 |
|
6737
|
67 |
text "basic simplifications";
|
5741
|
68 |
|
6737
|
69 |
lemma "x (| x = m, y = n, ... = p |) = m";
|
|
70 |
by simp;
|
|
71 |
|
|
72 |
lemma "(| x = m, y = n, ... = p |) (| x:= 0 |) = (| x = 0, y = n, ... = p |)";
|
|
73 |
by simp;
|
5741
|
74 |
|
|
75 |
|
6737
|
76 |
text "splitting quantifiers: the !!r is NECESSARY";
|
5741
|
77 |
|
6737
|
78 |
lemma "!!r. r (| x := n |) (| y := m |) = r (| y := m |) (| x := n |)";
|
|
79 |
proof record_split;
|
|
80 |
fix a b rest;
|
|
81 |
show "(| x = a, y = b, ... = rest |)(| x := n, y := m |) =
|
|
82 |
(| x = a, y = b, ... = rest |)(| y := m, x := n |)";
|
|
83 |
by simp;
|
|
84 |
qed;
|
|
85 |
|
|
86 |
lemma "!!r. r (| x := n |) (| x := m |) = r (| x := m |)";
|
|
87 |
proof record_split;
|
|
88 |
fix a b rest;
|
|
89 |
show "(| x = a, y = b, ... = rest |)(| x := n, x := m |) =
|
|
90 |
(| x = a, y = b, ... = rest |)(| x := m |)";
|
|
91 |
by simp;
|
|
92 |
qed;
|
5741
|
93 |
|
|
94 |
|
6737
|
95 |
text "equality of records";
|
|
96 |
|
|
97 |
lemma "(| x = n, y = p |) = (| x = n', y = p' |) --> n = n'" (is "??EQ --> _");
|
|
98 |
proof;
|
|
99 |
assume ??EQ;
|
|
100 |
thus "n = n'"; by simp;
|
|
101 |
qed;
|
|
102 |
|
|
103 |
|
|
104 |
text "concrete records are type instances of record schemes";
|
5741
|
105 |
|
|
106 |
constdefs
|
|
107 |
foo5 :: nat
|
6737
|
108 |
"foo5 == getX (| x = 1, y = 0 |)";
|
5741
|
109 |
|
|
110 |
|
6737
|
111 |
text "manipulating the `...' (more) part";
|
5741
|
112 |
|
|
113 |
constdefs
|
6737
|
114 |
incX :: "('a::more) point_scheme => 'a point_scheme"
|
|
115 |
"incX r == (| x = Suc (x r), y = y r, ... = point.more r |)";
|
5741
|
116 |
|
6737
|
117 |
lemma "!!r n. incX r = setX r (Suc (getX r))";
|
|
118 |
proof (unfold getX_def setX_def incX_def);
|
|
119 |
show "!! r n. (| x = Suc (x r), y = y r, ... = more r |) = r(| x := Suc (x r) |)";
|
|
120 |
by (record_split, simp);
|
|
121 |
qed;
|
|
122 |
|
|
123 |
|
|
124 |
text "alternative definition";
|
|
125 |
|
5741
|
126 |
constdefs
|
6737
|
127 |
incX' :: "('a::more) point_scheme => 'a point_scheme"
|
|
128 |
"incX' r == r (| x := Suc (x r) |)";
|
5741
|
129 |
|
|
130 |
|
|
131 |
|
6737
|
132 |
section "coloured points: record extension";
|
5741
|
133 |
|
6737
|
134 |
datatype colour = Red | Green | Blue;
|
5741
|
135 |
|
|
136 |
record cpoint = point +
|
6737
|
137 |
colour :: colour;
|
5741
|
138 |
|
|
139 |
|
6737
|
140 |
text {|
|
5753
|
141 |
The record declaration defines new type constructors:
|
5741
|
142 |
|
5753
|
143 |
cpoint = (| x :: nat, y :: nat, colour :: colour |)
|
|
144 |
'a cpoint_scheme = (| x :: nat, y :: nat, colour :: colour, ... :: 'a |)
|
6737
|
145 |
|};
|
5741
|
146 |
|
6737
|
147 |
consts foo6 :: cpoint;
|
|
148 |
consts foo7 :: "(| x :: nat, y :: nat, colour :: colour |)";
|
|
149 |
consts foo8 :: "('a::more) cpoint_scheme";
|
|
150 |
consts foo9 :: "(| x :: nat, y :: nat, colour :: colour, ... :: 'a |)";
|
5741
|
151 |
|
|
152 |
|
6737
|
153 |
text "functions on point schemes work for cpoints as well";
|
5741
|
154 |
|
5753
|
155 |
constdefs
|
5741
|
156 |
foo10 :: nat
|
6737
|
157 |
"foo10 == getX (| x = 2, y = 0, colour = Blue |)";
|
5741
|
158 |
|
|
159 |
|
6737
|
160 |
subsection "Non-coercive structural subtyping";
|
5741
|
161 |
|
6737
|
162 |
text "foo11 has type cpoint, not type point --- Great!";
|
|
163 |
|
5753
|
164 |
constdefs
|
5741
|
165 |
foo11 :: cpoint
|
6737
|
166 |
"foo11 == setX (| x = 2, y = 0, colour = Blue |) 0";
|
5741
|
167 |
|
|
168 |
|
|
169 |
|
6737
|
170 |
section "Other features";
|
5753
|
171 |
|
6737
|
172 |
text "field names contribute to record identity";
|
5741
|
173 |
|
|
174 |
record point' =
|
|
175 |
x' :: nat
|
6737
|
176 |
y' :: nat;
|
5741
|
177 |
|
|
178 |
consts
|
6737
|
179 |
foo12 :: nat;
|
|
180 |
|
|
181 |
text {| Definition @{term "foo12 == getX (| x' = 2, y' = 0 |)"} is invalid |};
|
5741
|
182 |
|
|
183 |
|
6737
|
184 |
text "polymorphic records";
|
5741
|
185 |
|
|
186 |
record 'a point'' = point +
|
6737
|
187 |
content :: 'a;
|
5741
|
188 |
|
6737
|
189 |
types cpoint'' = "colour point''";
|
5741
|
190 |
|
|
191 |
|
6737
|
192 |
text "Have a lot of fun!";
|
5741
|
193 |
|
6737
|
194 |
end;
|