src/HOL/Examples/Records.thy
author paulson <lp15@cam.ac.uk>
Mon, 03 Jul 2023 11:45:59 +0100
changeset 78248 740b23f1138a
parent 76042 e076b1b42c44
child 78624 8d7394e533f8
permissions -rw-r--r--
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Examples/Records.thy
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
     2
    Author:     Wolfgang Naraschewski, TU Muenchen
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
     3
    Author:     Norbert Schirmer, TU Muenchen
76039
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
     4
    Author:     Norbert Schirmer, Apple, 2022
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
     5
    Author:     Markus Wenzel, TU Muenchen
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
     6
*)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
     7
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
     8
section \<open>Using extensible records in HOL -- points and coloured points\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
     9
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    10
theory Records
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    11
  imports Main
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    12
begin
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    13
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    14
subsection \<open>Points\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    15
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    16
record point =
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    17
  xpos :: nat
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    18
  ypos :: nat
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    19
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    20
text \<open>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    21
  Apart many other things, above record declaration produces the
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    22
  following theorems:
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    23
\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    24
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    25
thm point.simps
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    26
thm point.iffs
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    27
thm point.defs
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    28
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    29
text \<open>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    30
  The set of theorems @{thm [source] point.simps} is added
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    31
  automatically to the standard simpset, @{thm [source] point.iffs} is
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    32
  added to the Classical Reasoner and Simplifier context.
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    33
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    34
  \<^medskip> Record declarations define new types and type abbreviations:
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    35
  @{text [display]
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    36
\<open>point = \<lparr>xpos :: nat, ypos :: nat\<rparr> = () point_ext_type
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    37
'a point_scheme = \<lparr>xpos :: nat, ypos :: nat, ... :: 'a\<rparr>  = 'a point_ext_type\<close>}
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    38
\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    39
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    40
consts foo2 :: "\<lparr>xpos :: nat, ypos :: nat\<rparr>"
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    41
consts foo4 :: "'a \<Rightarrow> \<lparr>xpos :: nat, ypos :: nat, \<dots> :: 'a\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    42
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    43
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    44
subsubsection \<open>Introducing concrete records and record schemes\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    45
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    46
definition foo1 :: point
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    47
  where "foo1 = \<lparr>xpos = 1, ypos = 0\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    48
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    49
definition foo3 :: "'a \<Rightarrow> 'a point_scheme"
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    50
  where "foo3 ext = \<lparr>xpos = 1, ypos = 0, \<dots> = ext\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    51
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    52
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    53
subsubsection \<open>Record selection and record update\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    54
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    55
definition getX :: "'a point_scheme \<Rightarrow> nat"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    56
  where "getX r = xpos r"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    57
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    58
definition setX :: "'a point_scheme \<Rightarrow> nat \<Rightarrow> 'a point_scheme"
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    59
  where "setX r n = r \<lparr>xpos := n\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    60
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    61
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    62
subsubsection \<open>Some lemmas about records\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    63
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    64
text \<open>Basic simplifications.\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    65
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    66
lemma "point.make n p = \<lparr>xpos = n, ypos = p\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    67
  by (simp only: point.make_def)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    68
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    69
lemma "xpos \<lparr>xpos = m, ypos = n, \<dots> = p\<rparr> = m"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    70
  by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    71
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    72
lemma "\<lparr>xpos = m, ypos = n, \<dots> = p\<rparr>\<lparr>xpos:= 0\<rparr> = \<lparr>xpos = 0, ypos = n, \<dots> = p\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    73
  by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    74
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    75
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    76
text \<open>\<^medskip> Equality of records.\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    77
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    78
lemma "n = n' \<Longrightarrow> p = p' \<Longrightarrow> \<lparr>xpos = n, ypos = p\<rparr> = \<lparr>xpos = n', ypos = p'\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    79
  \<comment> \<open>introduction of concrete record equality\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    80
  by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    81
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    82
lemma "\<lparr>xpos = n, ypos = p\<rparr> = \<lparr>xpos = n', ypos = p'\<rparr> \<Longrightarrow> n = n'"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    83
  \<comment> \<open>elimination of concrete record equality\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    84
  by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    85
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    86
lemma "r\<lparr>xpos := n\<rparr>\<lparr>ypos := m\<rparr> = r\<lparr>ypos := m\<rparr>\<lparr>xpos := n\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    87
  \<comment> \<open>introduction of abstract record equality\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    88
  by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    89
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    90
lemma "r\<lparr>xpos := n\<rparr> = r\<lparr>xpos := n'\<rparr>" if "n = n'"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    91
  \<comment> \<open>elimination of abstract record equality (manual proof)\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    92
proof -
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    93
  let "?lhs = ?rhs" = ?thesis
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    94
  from that have "xpos ?lhs = xpos ?rhs" by simp
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    95
  then show ?thesis by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    96
qed
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    97
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
    98
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
    99
text \<open>\<^medskip> Surjective pairing\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   100
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   101
lemma "r = \<lparr>xpos = xpos r, ypos = ypos r\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   102
  by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   103
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   104
lemma "r = \<lparr>xpos = xpos r, ypos = ypos r, \<dots> = point.more r\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   105
  by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   106
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   107
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   108
text \<open>\<^medskip> Representation of records by cases or (degenerate) induction.\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   109
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   110
lemma "r\<lparr>xpos := n\<rparr>\<lparr>ypos := m\<rparr> = r\<lparr>ypos := m\<rparr>\<lparr>xpos := n\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   111
proof (cases r)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   112
  fix xpos ypos more
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   113
  assume "r = \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   114
  then show ?thesis by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   115
qed
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   116
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   117
lemma "r\<lparr>xpos := n\<rparr>\<lparr>ypos := m\<rparr> = r\<lparr>ypos := m\<rparr>\<lparr>xpos := n\<rparr>"
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   118
proof (induct r)
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   119
  fix xpos ypos more
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   120
  show "\<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>\<lparr>xpos := n, ypos := m\<rparr> =
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   121
      \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>\<lparr>ypos := m, xpos := n\<rparr>"
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   122
    by simp
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   123
qed
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   124
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   125
lemma "r\<lparr>xpos := n\<rparr>\<lparr>xpos := m\<rparr> = r\<lparr>xpos := m\<rparr>"
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   126
proof (cases r)
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   127
  fix xpos ypos more
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   128
  assume "r = \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>"
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   129
  then show ?thesis by simp
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   130
qed
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   131
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   132
lemma "r\<lparr>xpos := n\<rparr>\<lparr>xpos := m\<rparr> = r\<lparr>xpos := m\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   133
proof (cases r)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   134
  case fields
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   135
  then show ?thesis by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   136
qed
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   137
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   138
lemma "r\<lparr>xpos := n\<rparr>\<lparr>xpos := m\<rparr> = r\<lparr>xpos := m\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   139
  by (cases r) simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   140
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   141
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   142
text \<open>\<^medskip> Concrete records are type instances of record schemes.\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   143
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   144
definition foo5 :: nat
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   145
  where "foo5 = getX \<lparr>xpos = 1, ypos = 0\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   146
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   147
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   148
text \<open>\<^medskip> Manipulating the ``\<open>...\<close>'' (more) part.\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   149
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   150
definition incX :: "'a point_scheme \<Rightarrow> 'a point_scheme"
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   151
  where "incX r = \<lparr>xpos = xpos r + 1, ypos = ypos r, \<dots> = point.more r\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   152
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   153
lemma "incX r = setX r (Suc (getX r))"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   154
  by (simp add: getX_def setX_def incX_def)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   155
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   156
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   157
text \<open>\<^medskip> An alternative definition.\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   158
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   159
definition incX' :: "'a point_scheme \<Rightarrow> 'a point_scheme"
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   160
  where "incX' r = r\<lparr>xpos := xpos r + 1\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   161
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   162
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   163
subsection \<open>Coloured points: record extension\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   164
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   165
datatype colour = Red | Green | Blue
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   166
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   167
record cpoint = point +
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   168
  colour :: colour
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   169
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   170
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   171
text \<open>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   172
  The record declaration defines a new type constructor and abbreviations:
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   173
  @{text [display]
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   174
\<open>cpoint = \<lparr>xpos :: nat, ypos :: nat, colour :: colour\<rparr> =
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   175
  () cpoint_ext_type point_ext_type
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   176
'a cpoint_scheme = \<lparr>xpos :: nat, ypos :: nat, colour :: colour, \<dots> :: 'a\<rparr> =
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   177
  'a cpoint_ext_type point_ext_type\<close>}
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   178
\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   179
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   180
consts foo6 :: cpoint
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   181
consts foo7 :: "\<lparr>xpos :: nat, ypos :: nat, colour :: colour\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   182
consts foo8 :: "'a cpoint_scheme"
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   183
consts foo9 :: "\<lparr>xpos :: nat, ypos :: nat, colour :: colour, \<dots> :: 'a\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   184
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   185
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   186
text \<open>Functions on \<open>point\<close> schemes work for \<open>cpoints\<close> as well.\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   187
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   188
definition foo10 :: nat
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   189
  where "foo10 = getX \<lparr>xpos = 2, ypos = 0, colour = Blue\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   190
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   191
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   192
subsubsection \<open>Non-coercive structural subtyping\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   193
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   194
text \<open>Term \<^term>\<open>foo11\<close> has type \<^typ>\<open>cpoint\<close>, not type \<^typ>\<open>point\<close> --- Great!\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   195
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   196
definition foo11 :: cpoint
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   197
  where "foo11 = setX \<lparr>xpos = 2, ypos = 0, colour = Blue\<rparr> 0"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   198
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   199
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   200
subsection \<open>Other features\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   201
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   202
text \<open>Field names contribute to record identity.\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   203
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   204
record point' =
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   205
  xpos' :: nat
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   206
  ypos' :: nat
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   207
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   208
text \<open>
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   209
  \<^noindent> May not apply \<^term>\<open>getX\<close> to @{term [source] "\<lparr>xpos' = 2, ypos' = 0\<rparr>"}
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   210
  --- type error.
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   211
\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   212
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   213
text \<open>\<^medskip> Polymorphic records.\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   214
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   215
record 'a point'' = point +
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   216
  content :: 'a
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   217
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   218
type_synonym cpoint'' = "colour point''"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   219
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   220
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   221
text \<open>Updating a record field with an identical value is simplified.\<close>
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   222
lemma "r\<lparr>xpos := xpos r\<rparr> = r"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   223
  by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   224
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   225
text \<open>Only the most recent update to a component survives simplification.\<close>
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   226
lemma "r\<lparr>xpos := x, ypos := y, xpos := x'\<rparr> = r\<lparr>ypos := y, xpos := x'\<rparr>"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   227
  by simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   228
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   229
text \<open>
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   230
  In some cases its convenient to automatically split (quantified) records.
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   231
  For this purpose there is the simproc @{ML [source] "Record.split_simproc"}
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   232
  and the tactic @{ML [source] "Record.split_simp_tac"}. The simplification
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   233
  procedure only splits the records, whereas the tactic also simplifies the
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   234
  resulting goal with the standard record simplification rules. A
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   235
  (generalized) predicate on the record is passed as parameter that decides
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   236
  whether or how `deep' to split the record. It can peek on the subterm
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   237
  starting at the quantified occurrence of the record (including the
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   238
  quantifier). The value \<^ML>\<open>0\<close> indicates no split, a value greater
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   239
  \<^ML>\<open>0\<close> splits up to the given bound of record extension and finally the
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   240
  value \<^ML>\<open>~1\<close> completely splits the record. @{ML [source]
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   241
  "Record.split_simp_tac"} additionally takes a list of equations for
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   242
  simplification and can also split fixed record variables.
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   243
\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   244
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   245
lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   246
  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   247
    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   248
  apply simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   249
  done
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   250
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   251
lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   252
  apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   253
  apply simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   254
  done
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   255
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   256
lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   257
  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   258
    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   259
  apply simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   260
  done
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   261
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   262
lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   263
  apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   264
  apply simp
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   265
  done
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   266
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   267
lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   268
  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   269
    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   270
  apply auto
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   271
  done
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   272
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   273
lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   274
  apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   275
  apply auto
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   276
  done
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   277
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   278
lemma "P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   279
  apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   280
  apply auto
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   281
  done
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   282
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   283
notepad
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   284
begin
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   285
  have "\<exists>x. P x"
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   286
    if "P (xpos r)" for P r
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   287
    apply (insert that)
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   288
    apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   289
    apply auto
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   290
    done
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   291
end
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   292
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   293
text \<open>
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   294
  The effect of simproc @{ML [source] Record.ex_sel_eq_simproc} is illustrated
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   295
  by the following lemma.\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   296
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   297
lemma "\<exists>r. xpos r = x"
76039
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   298
  supply [[simproc add: Record.ex_sel_eq]]
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   299
  apply (simp)
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   300
  done
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   301
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   302
subsection \<open>Simprocs for update and equality\<close>
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   303
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   304
record alph1 =
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   305
a::nat
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   306
b::nat
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   307
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   308
record alph2 = alph1 +
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   309
c::nat
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   310
d::nat
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   311
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   312
record alph3 = alph2 + 
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   313
  e::nat
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   314
  f::nat
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   315
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   316
text \<open>The simprocs that are activated by default are:
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   317
\<^item> @{ML [source] Record.simproc}: field selection of (nested) record updates.
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   318
\<^item> @{ML [source] Record.upd_simproc}: nested record updates.
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   319
\<^item> @{ML [source] Record.eq_simproc}: (componentwise) equality of records.
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   320
\<close>
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   321
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   322
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   323
text \<open>By default record updates are not ordered by simplification.\<close>
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   324
schematic_goal "r\<lparr>b := x, a:= y\<rparr> = ?X"
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   325
  by simp
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   326
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   327
text \<open>Normalisation towards an update ordering (string ordering of update function names) can
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   328
be configured as follows.\<close>
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   329
schematic_goal "r\<lparr>b := y, a := x\<rparr> = ?X"
76042
wenzelm
parents: 76039
diff changeset
   330
  supply [[record_sort_updates]]
76039
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   331
  by simp
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   332
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   333
text \<open>Note the interplay between update ordering and record equality. Without update ordering
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   334
the following equality is handled by @{ML [source] Record.eq_simproc}. Record equality is thus
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   335
solved by componentwise comparison of all the fields of the records which can be expensive 
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   336
in the presence of many fields.\<close>
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   337
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   338
lemma "r\<lparr>f := x1, a:= x2\<rparr> = r\<lparr>a := x2, f:= x1\<rparr>"
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   339
  by simp
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   340
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   341
lemma "r\<lparr>f := x1, a:= x2\<rparr> = r\<lparr>a := x2, f:= x1\<rparr>"
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   342
  supply [[simproc del: Record.eq]]
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   343
  apply (simp?)
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   344
  oops
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   345
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   346
text \<open>With update ordering the equality is already established after update normalisation. There
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   347
is no need for componentwise comparison.\<close>
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   348
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   349
lemma "r\<lparr>f := x1, a:= x2\<rparr> = r\<lparr>a := x2, f:= x1\<rparr>"
76042
wenzelm
parents: 76039
diff changeset
   350
  supply [[record_sort_updates, simproc del: Record.eq]]
76039
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   351
  apply simp
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   352
  done
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   353
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   354
schematic_goal "r\<lparr>f := x1, e := x2, d:= x3, c:= x4, b:=x5, a:= x6\<rparr> = ?X"
76042
wenzelm
parents: 76039
diff changeset
   355
  supply [[record_sort_updates]]
76039
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   356
  by simp
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   357
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   358
schematic_goal "r\<lparr>f := x1, e := x2, d:= x3, c:= x4, e:=x5, a:= x6\<rparr> = ?X"
76042
wenzelm
parents: 76039
diff changeset
   359
  supply [[record_sort_updates]]
76039
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   360
  by simp
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   361
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   362
schematic_goal "r\<lparr>f := x1, e := x2, d:= x3, c:= x4, e:=x5, a:= x6\<rparr> = ?X"
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   363
  by simp
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   364
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   365
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   366
subsection \<open>A more complex record expression\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   367
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   368
record ('a, 'b, 'c) bar = bar1 :: 'a
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   369
  bar2 :: 'b
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   370
  bar3 :: 'c
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   371
  bar21 :: "'b \<times> 'a"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   372
  bar32 :: "'c \<times> 'b"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   373
  bar31 :: "'c \<times> 'a"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   374
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   375
print_record "('a,'b,'c) bar"
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   376
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   377
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   378
subsection \<open>Some code generation\<close>
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   379
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   380
export_code foo1 foo3 foo5 foo10 checking SML
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   381
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   382
text \<open>
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   383
  Code generation can also be switched off, for instance for very large
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   384
  records:\<close>
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   385
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   386
declare [[record_codegen = false]]
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   387
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   388
record not_so_large_record =
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   389
  bar520 :: nat
72030
eece87547736 misc tuning and modernization;
wenzelm
parents: 72029
diff changeset
   390
  bar521 :: "nat \<times> nat"
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   391
76039
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   392
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   393
setup \<open>
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   394
let
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   395
 val N = 300
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   396
in
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   397
  Record.add_record {overloaded=false} ([], \<^binding>\<open>large_record\<close>) NONE 
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   398
    (map (fn i => (Binding.make ("fld_" ^ string_of_int i, \<^here>), @{typ nat}, Mixfix.NoSyn)) 
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   399
      (1 upto N))
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   400
end
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   401
\<close>
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   402
76042
wenzelm
parents: 76039
diff changeset
   403
declare [[record_codegen]]
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   404
76039
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   405
schematic_goal \<open>fld_1 (r\<lparr>fld_300 := x300, fld_20 := x20, fld_200 := x200\<rparr>) = ?X\<close>
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   406
  by simp
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   407
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   408
schematic_goal \<open>r\<lparr>fld_300 := x300, fld_20 := x20, fld_200 := x200\<rparr> = ?X\<close>
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   409
  supply [[record_sort_updates]]
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   410
  by simp
ca7737249aa4 option "sort_updates" for record update simproc. Make proper record simproc definitions.
Norbert Schirmer <nschirmer@apple.com>
parents: 72030
diff changeset
   411
72029
83456d9f0ed5 clarified examples;
wenzelm
parents:
diff changeset
   412
end