Theory HOL-Examples.Records
section ‹Using extensible records in HOL -- points and coloured points›
theory Records
imports Main
begin
subsection ‹Points›
record point =
xpos :: nat
ypos :: nat
text ‹
Apart many other things, above record declaration produces the
following theorems:
›
thm point.simps
thm point.iffs
thm point.defs
text ‹
The set of theorems @{thm [source] point.simps} is added
automatically to the standard simpset, @{thm [source] point.iffs} is
added to the Classical Reasoner and Simplifier context.
┉ Record declarations define new types and type abbreviations:
@{text [display]
‹point = ⦇xpos :: nat, ypos :: nat⦈ = () point_ext_type
'a point_scheme = ⦇xpos :: nat, ypos :: nat, ... :: 'a⦈ = 'a point_ext_type›}
›
consts foo2 :: "⦇xpos :: nat, ypos :: nat⦈"
consts foo4 :: "'a ⇒ ⦇xpos :: nat, ypos :: nat, … :: 'a⦈"
subsubsection ‹Introducing concrete records and record schemes›
definition foo1 :: point
where "foo1 = ⦇xpos = 1, ypos = 0⦈"
definition foo3 :: "'a ⇒ 'a point_scheme"
where "foo3 ext = ⦇xpos = 1, ypos = 0, … = ext⦈"
subsubsection ‹Record selection and record update›
definition getX :: "'a point_scheme ⇒ nat"
where "getX r = xpos r"
definition setX :: "'a point_scheme ⇒ nat ⇒ 'a point_scheme"
where "setX r n = r ⦇xpos := n⦈"
subsubsection ‹Some lemmas about records›
text ‹Basic simplifications.›
lemma "point.make n p = ⦇xpos = n, ypos = p⦈"
by (simp only: point.make_def)
lemma "xpos ⦇xpos = m, ypos = n, … = p⦈ = m"
by simp
lemma "⦇xpos = m, ypos = n, … = p⦈⦇xpos:= 0⦈ = ⦇xpos = 0, ypos = n, … = p⦈"
by simp
text ‹┉ Equality of records.›
lemma "n = n' ⟹ p = p' ⟹ ⦇xpos = n, ypos = p⦈ = ⦇xpos = n', ypos = p'⦈"
by simp
lemma "⦇xpos = n, ypos = p⦈ = ⦇xpos = n', ypos = p'⦈ ⟹ n = n'"
by simp
lemma "r⦇xpos := n⦈⦇ypos := m⦈ = r⦇ypos := m⦈⦇xpos := n⦈"
by simp
lemma "r⦇xpos := n⦈ = r⦇xpos := n'⦈" if "n = n'"
proof -
let "?lhs = ?rhs" = ?thesis
from that have "xpos ?lhs = xpos ?rhs" by simp
then show ?thesis by simp
qed
text ‹┉ Surjective pairing›
lemma "r = ⦇xpos = xpos r, ypos = ypos r⦈"
by simp
lemma "r = ⦇xpos = xpos r, ypos = ypos r, … = point.more r⦈"
by simp
text ‹┉ Representation of records by cases or (degenerate) induction.›
lemma "r⦇xpos := n⦈⦇ypos := m⦈ = r⦇ypos := m⦈⦇xpos := n⦈"
proof (cases r)
fix xpos ypos more
assume "r = ⦇xpos = xpos, ypos = ypos, … = more⦈"
then show ?thesis by simp
qed
lemma "r⦇xpos := n⦈⦇ypos := m⦈ = r⦇ypos := m⦈⦇xpos := n⦈"
proof (induct r)
fix xpos ypos more
show "⦇xpos = xpos, ypos = ypos, … = more⦈⦇xpos := n, ypos := m⦈ =
⦇xpos = xpos, ypos = ypos, … = more⦈⦇ypos := m, xpos := n⦈"
by simp
qed
lemma "r⦇xpos := n⦈⦇xpos := m⦈ = r⦇xpos := m⦈"
proof (cases r)
fix xpos ypos more
assume "r = ⦇xpos = xpos, ypos = ypos, … = more⦈"
then show ?thesis by simp
qed
lemma "r⦇xpos := n⦈⦇xpos := m⦈ = r⦇xpos := m⦈"
proof (cases r)
case fields
then show ?thesis by simp
qed
lemma "r⦇xpos := n⦈⦇xpos := m⦈ = r⦇xpos := m⦈"
by (cases r) simp
text ‹┉ Concrete records are type instances of record schemes.›
definition foo5 :: nat
where "foo5 = getX ⦇xpos = 1, ypos = 0⦈"
text ‹┉ Manipulating the ``‹...›'' (more) part.›
definition incX :: "'a point_scheme ⇒ 'a point_scheme"
where "incX r = ⦇xpos = xpos r + 1, ypos = ypos r, … = point.more r⦈"
lemma "incX r = setX r (Suc (getX r))"
by (simp add: getX_def setX_def incX_def)
text ‹┉ An alternative definition.›
definition incX' :: "'a point_scheme ⇒ 'a point_scheme"
where "incX' r = r⦇xpos := xpos r + 1⦈"
subsection ‹Coloured points: record extension›