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