updated records;
authorwenzelm
Thu, 25 Oct 2001 22:43:05 +0200
changeset 11939 c5e69470f03b
parent 11938 7b594aba1300
child 11940 80365073b8b3
updated records;
src/HOL/ex/MonoidGroup.thy
src/HOL/ex/Records.thy
--- a/src/HOL/ex/MonoidGroup.thy	Thu Oct 25 22:42:50 2001 +0200
+++ b/src/HOL/ex/MonoidGroup.thy	Thu Oct 25 22:43:05 2001 +0200
@@ -18,15 +18,15 @@
   inv :: "'a => 'a"
 
 constdefs
-  monoid :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b::more |) => bool"
+  monoid :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |) => bool"
   "monoid M == \<forall>x y z.
     times M (times M x y) z = times M x (times M y z) \<and>
     times M (one M) x = x \<and> times M x (one M) = x"
 
-  group :: "(| times :: 'a => 'a => 'a, one :: 'a, inv :: 'a => 'a, ... :: 'b::more |) => bool"
+  group :: "(| times :: 'a => 'a => 'a, one :: 'a, inv :: 'a => 'a, ... :: 'b |) => bool"
   "group G == monoid G \<and> (\<forall>x. times G (inv G x) x = one G)"
 
-  reverse :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b::more |) =>
+  reverse :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |) =>
     (| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |)"
   "reverse M == M (| times := \<lambda>x y. times M y x |)"
 
--- a/src/HOL/ex/Records.thy	Thu Oct 25 22:42:50 2001 +0200
+++ b/src/HOL/ex/Records.thy	Thu Oct 25 22:43:05 2001 +0200
@@ -11,121 +11,133 @@
 subsection {* Points *}
 
 record point =
-  x :: nat
-  y :: nat
+  xpos :: nat
+  ypos :: nat
 
 text {*
- Apart many other things, above record declaration produces the
- following theorems:
+  Apart many other things, above record declaration produces the
+  following theorems:
 *}
 
 thm "point.simps"
 thm "point.iffs"
-thm "point.update_defs"
+thm "point.derived_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.
-*}
+  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.
 
-text {*
-  Record declarations define new type abbreviations:
+  \medskip Record declarations define new type abbreviations:
   @{text [display]
-"    point = (| x :: nat, y :: nat |)
-    'a point_scheme = (| x :: nat, y :: nat, ... :: 'a |)"}
-  Extensions `...' must be in type class @{text more}.
+"    point = (| xpos :: nat, ypos :: nat |)
+    'a point_scheme = (| xpos :: nat, ypos :: nat, ... :: 'a |)"}
 *}
 
 consts foo1 :: point
-consts foo2 :: "(| x :: nat, y :: nat |)"
-consts foo3 :: "'a => ('a::more) point_scheme"
-consts foo4 :: "'a => (| x :: nat, y :: nat, ... :: 'a |)"
+consts foo2 :: "(| xpos :: nat, ypos :: nat |)"
+consts foo3 :: "'a => 'a point_scheme"
+consts foo4 :: "'a => (| xpos :: nat, ypos :: nat, ... :: 'a |)"
 
 
 subsubsection {* Introducing concrete records and record schemes *}
 
 defs
-  foo1_def: "foo1 == (| x = 1, y = 0 |)"
-  foo3_def: "foo3 ext == (| x = 1, y = 0, ... = ext |)"
+  foo1_def: "foo1 == (| xpos = 1, ypos = 0 |)"
+  foo3_def: "foo3 ext == (| xpos = 1, ypos = 0, ... = ext |)"
 
 
 subsubsection {* Record selection and record update *}
 
 constdefs
-  getX :: "('a::more) point_scheme => nat"
-  "getX r == x r"
-  setX :: "('a::more) point_scheme => nat => 'a point_scheme"
-  "setX r n == r (| x := n |)"
+  getX :: "'a point_scheme => nat"
+  "getX r == xpos r"
+  setX :: "'a point_scheme => nat => 'a point_scheme"
+  "setX r n == r (| xpos := n |)"
 
 
 subsubsection {* Some lemmas about records *}
 
 text {* Basic simplifications. *}
 
-lemma "point.make n p = (| x = n, y = p |)"
-  by (simp add: point.make_def)
+lemma "point.make n p = (| xpos = n, ypos = p |)"
+  by (simp only: point.make_def)
 
-lemma "x (| x = m, y = n, ... = p |) = m"
+lemma "xpos (| xpos = m, ypos = n, ... = p |) = m"
   by simp
 
-lemma "(| x = m, y = n, ... = p |) (| x:= 0 |) = (| x = 0, y = n, ... = p |)"
+lemma "(| xpos = m, ypos = n, ... = p |) (| xpos:= 0 |) = (| xpos = 0, ypos = n, ... = p |)"
   by simp
 
 
 text {* \medskip Equality of records. *}
 
-lemma "n = n' ==> p = p' ==> (| x = n, y = p |) = (| x = n', y = p' |)"
+lemma "n = n' ==> p = p' ==> (| xpos = n, ypos = p |) = (| xpos = n', ypos = p' |)"
   -- "introduction of concrete record equality"
   by simp
 
-lemma "(| x = n, y = p |) = (| x = n', y = p' |) ==> n = n'"
+lemma "(| xpos = n, ypos = p |) = (| xpos = n', ypos = p' |) ==> n = n'"
   -- "elimination of concrete record equality"
   by simp
 
-lemma "r (| x := n |) (| y := m |) = r (| y := m |) (| x := n |)"
+lemma "r (| xpos := n |) (| ypos := m |) = r (| ypos := m |) (| xpos := n |)"
   -- "introduction of abstract record equality"
   by simp
 
-lemma "r (| x := n |) = r (| x := n' |) ==> n = n'"
+lemma "r (| xpos := n |) = r (| xpos := n' |) ==> n = n'"
   -- "elimination of abstract record equality (manual proof)"
 proof -
-  assume "r (| x := n |) = r (| x := n' |)" (is "?lhs = ?rhs")
-  hence "x ?lhs = x ?rhs" by simp
+  assume "r (| xpos := n |) = r (| xpos := n' |)" (is "?lhs = ?rhs")
+  hence "xpos ?lhs = xpos ?rhs" by simp
   thus ?thesis by simp
 qed
 
 
 text {* \medskip Surjective pairing *}
 
-lemma "r = (| x = x r, y = y r |)"
+lemma "r = (| xpos = xpos r, ypos = ypos r |)"
   by simp
 
-lemma "r = (| x = x r, y = y r, ... = more r |)"
+lemma "r = (| xpos = xpos r, ypos = ypos r, ... = more r |)"
   by simp
 
 
 text {*
- \medskip Splitting quantifiers: the @{text "!!r"} is \emph{necessary}
- here!
+  \medskip Representation of records by cases or (degenerate)
+  induction.
 *}
 
-lemma "!!r. r (| x := n |) (| y := m |) = r (| y := m |) (| x := n |)"
-proof record_split
-  fix x y more
-  show "(| x = x, y = y, ... = more |)(| x := n, y := m |) =
-        (| x = x, y = y, ... = more |)(| y := m, x := n |)"
+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 |)"
+  thus ?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. r (| x := n |) (| x := m |) = r (| x := m |)"
-proof record_split
-  fix x y more
-  show "(| x = x, y = y, ... = more |)(| x := n, x := m |) =
-        (| x = x, y = y, ... = more |)(| x := m |)"
-    by simp
+lemma "r (| xpos := n |) (| xpos := m |) = r (| xpos := m |)"
+proof (cases r)
+  fix xpos ypos more
+  assume "r = \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>"
+  thus ?thesis by simp
 qed
 
+lemma "r (| xpos := n |) (| xpos := m |) = r (| xpos := m |)"
+proof (cases r)
+  case fields
+  thus ?thesis by simp
+qed
+
+lemma "r (| xpos := n |) (| xpos := m |) = r (| xpos := m |)"
+  by (cases r) simp
+
 
 text {*
  \medskip Concrete records are type instances of record schemes.
@@ -133,26 +145,24 @@
 
 constdefs
   foo5 :: nat
-  "foo5 == getX (| x = 1, y = 0 |)"
+  "foo5 == getX (| xpos = 1, ypos = 0 |)"
 
 
-text {* \medskip Manipulating the `...' (more) part. *}
+text {* \medskip Manipulating the ``@{text "..."}'' (more) part. *}
 
 constdefs
-  incX :: "('a::more) point_scheme => 'a point_scheme"
-  "incX r == (| x = Suc (x r), y = y r, ... = point.more r |)"
+  incX :: "'a point_scheme => 'a point_scheme"
+  "incX r == (| xpos = xpos r + 1, ypos = ypos r, ... = point.more r |)"
 
-lemma "!!r n. incX r = setX r (Suc (getX r))"
-proof (unfold getX_def setX_def incX_def)
-  show "!!r n. (| x = Suc (x r), y = y r, ... = more r |) = r(| x := Suc (x r) |)"
-    by record_split simp
-qed
+lemma "incX r = setX r (Suc (getX r))"
+  by (simp add: getX_def setX_def incX_def)
+
 
 text {* An alternative definition. *}
 
 constdefs
-  incX' :: "('a::more) point_scheme => 'a point_scheme"
-  "incX' r == r (| x := Suc (x r) |)"
+  incX' :: "'a point_scheme => 'a point_scheme"
+  "incX' r == r (| xpos := xpos r + 1 |)"
 
 
 subsection {* Coloured points: record extension *}
@@ -166,14 +176,14 @@
 text {*
   The record declaration defines new type constructors:
   @{text [display]
-"    cpoint = (| x :: nat, y :: nat, colour :: colour |)
-    'a cpoint_scheme = (| x :: nat, y :: nat, colour :: colour, ... :: 'a |)"}
+"    cpoint = (| xpos :: nat, ypos :: nat, colour :: colour |)
+    'a cpoint_scheme = (| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |)"}
 *}
 
 consts foo6 :: cpoint
-consts foo7 :: "(| x :: nat, y :: nat, colour :: colour |)"
-consts foo8 :: "('a::more) cpoint_scheme"
-consts foo9 :: "(| x :: nat, y :: nat, colour :: colour, ... :: 'a |)"
+consts foo7 :: "(| xpos :: nat, ypos :: nat, colour :: colour |)"
+consts foo8 :: "'a cpoint_scheme"
+consts foo9 :: "(| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |)"
 
 
 text {*
@@ -182,7 +192,7 @@
 
 constdefs
   foo10 :: nat
-  "foo10 == getX (| x = 2, y = 0, colour = Blue |)"
+  "foo10 == getX (| xpos = 2, ypos = 0, colour = Blue |)"
 
 
 subsubsection {* Non-coercive structural subtyping *}
@@ -194,7 +204,7 @@
 
 constdefs
   foo11 :: cpoint
-  "foo11 == setX (| x = 2, y = 0, colour = Blue |) 0"
+  "foo11 == setX (| xpos = 2, ypos = 0, colour = Blue |) 0"
 
 
 subsection {* Other features *}
@@ -202,12 +212,12 @@
 text {* Field names contribute to record identity. *}
 
 record point' =
-  x' :: nat
-  y' :: nat
+  xpos' :: nat
+  ypos' :: nat
 
 text {*
- \noindent May not apply @{term getX} to 
- @{term [source] "(| x' = 2, y' = 0 |)"}.
+  \noindent May not apply @{term getX} to @{term [source] "(| xpos' =
+  2, ypos' = 0 |)"} -- type error.
 *}
 
 text {* \medskip Polymorphic records. *}