| 72029 |      1 | (*  Title:      HOL/Examples/Records.thy
 | 
|  |      2 |     Author:     Wolfgang Naraschewski, TU Muenchen
 | 
|  |      3 |     Author:     Norbert Schirmer, TU Muenchen
 | 
|  |      4 |     Author:     Markus Wenzel, TU Muenchen
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | section \<open>Using extensible records in HOL -- points and coloured points\<close>
 | 
|  |      8 | 
 | 
|  |      9 | theory Records
 | 
| 72030 |     10 |   imports Main
 | 
| 72029 |     11 | begin
 | 
|  |     12 | 
 | 
|  |     13 | subsection \<open>Points\<close>
 | 
|  |     14 | 
 | 
|  |     15 | record point =
 | 
|  |     16 |   xpos :: nat
 | 
|  |     17 |   ypos :: nat
 | 
|  |     18 | 
 | 
|  |     19 | text \<open>
 | 
|  |     20 |   Apart many other things, above record declaration produces the
 | 
|  |     21 |   following theorems:
 | 
|  |     22 | \<close>
 | 
|  |     23 | 
 | 
|  |     24 | thm point.simps
 | 
|  |     25 | thm point.iffs
 | 
|  |     26 | thm point.defs
 | 
|  |     27 | 
 | 
|  |     28 | text \<open>
 | 
|  |     29 |   The set of theorems @{thm [source] point.simps} is added
 | 
|  |     30 |   automatically to the standard simpset, @{thm [source] point.iffs} is
 | 
|  |     31 |   added to the Classical Reasoner and Simplifier context.
 | 
|  |     32 | 
 | 
| 72030 |     33 |   \<^medskip> Record declarations define new types and type abbreviations:
 | 
| 72029 |     34 |   @{text [display]
 | 
|  |     35 | \<open>point = \<lparr>xpos :: nat, ypos :: nat\<rparr> = () point_ext_type
 | 
|  |     36 | 'a point_scheme = \<lparr>xpos :: nat, ypos :: nat, ... :: 'a\<rparr>  = 'a point_ext_type\<close>}
 | 
|  |     37 | \<close>
 | 
|  |     38 | 
 | 
| 72030 |     39 | consts foo2 :: "\<lparr>xpos :: nat, ypos :: nat\<rparr>"
 | 
|  |     40 | consts foo4 :: "'a \<Rightarrow> \<lparr>xpos :: nat, ypos :: nat, \<dots> :: 'a\<rparr>"
 | 
| 72029 |     41 | 
 | 
|  |     42 | 
 | 
|  |     43 | subsubsection \<open>Introducing concrete records and record schemes\<close>
 | 
|  |     44 | 
 | 
|  |     45 | definition foo1 :: point
 | 
| 72030 |     46 |   where "foo1 = \<lparr>xpos = 1, ypos = 0\<rparr>"
 | 
| 72029 |     47 | 
 | 
| 72030 |     48 | definition foo3 :: "'a \<Rightarrow> 'a point_scheme"
 | 
|  |     49 |   where "foo3 ext = \<lparr>xpos = 1, ypos = 0, \<dots> = ext\<rparr>"
 | 
| 72029 |     50 | 
 | 
|  |     51 | 
 | 
|  |     52 | subsubsection \<open>Record selection and record update\<close>
 | 
|  |     53 | 
 | 
| 72030 |     54 | definition getX :: "'a point_scheme \<Rightarrow> nat"
 | 
| 72029 |     55 |   where "getX r = xpos r"
 | 
|  |     56 | 
 | 
| 72030 |     57 | definition setX :: "'a point_scheme \<Rightarrow> nat \<Rightarrow> 'a point_scheme"
 | 
|  |     58 |   where "setX r n = r \<lparr>xpos := n\<rparr>"
 | 
| 72029 |     59 | 
 | 
|  |     60 | 
 | 
|  |     61 | subsubsection \<open>Some lemmas about records\<close>
 | 
|  |     62 | 
 | 
|  |     63 | text \<open>Basic simplifications.\<close>
 | 
|  |     64 | 
 | 
| 72030 |     65 | lemma "point.make n p = \<lparr>xpos = n, ypos = p\<rparr>"
 | 
| 72029 |     66 |   by (simp only: point.make_def)
 | 
|  |     67 | 
 | 
| 72030 |     68 | lemma "xpos \<lparr>xpos = m, ypos = n, \<dots> = p\<rparr> = m"
 | 
| 72029 |     69 |   by simp
 | 
|  |     70 | 
 | 
| 72030 |     71 | lemma "\<lparr>xpos = m, ypos = n, \<dots> = p\<rparr>\<lparr>xpos:= 0\<rparr> = \<lparr>xpos = 0, ypos = n, \<dots> = p\<rparr>"
 | 
| 72029 |     72 |   by simp
 | 
|  |     73 | 
 | 
|  |     74 | 
 | 
| 72030 |     75 | text \<open>\<^medskip> Equality of records.\<close>
 | 
| 72029 |     76 | 
 | 
| 72030 |     77 | lemma "n = n' \<Longrightarrow> p = p' \<Longrightarrow> \<lparr>xpos = n, ypos = p\<rparr> = \<lparr>xpos = n', ypos = p'\<rparr>"
 | 
| 72029 |     78 |   \<comment> \<open>introduction of concrete record equality\<close>
 | 
|  |     79 |   by simp
 | 
|  |     80 | 
 | 
| 72030 |     81 | lemma "\<lparr>xpos = n, ypos = p\<rparr> = \<lparr>xpos = n', ypos = p'\<rparr> \<Longrightarrow> n = n'"
 | 
| 72029 |     82 |   \<comment> \<open>elimination of concrete record equality\<close>
 | 
|  |     83 |   by simp
 | 
|  |     84 | 
 | 
| 72030 |     85 | lemma "r\<lparr>xpos := n\<rparr>\<lparr>ypos := m\<rparr> = r\<lparr>ypos := m\<rparr>\<lparr>xpos := n\<rparr>"
 | 
| 72029 |     86 |   \<comment> \<open>introduction of abstract record equality\<close>
 | 
|  |     87 |   by simp
 | 
|  |     88 | 
 | 
| 72030 |     89 | lemma "r\<lparr>xpos := n\<rparr> = r\<lparr>xpos := n'\<rparr>" if "n = n'"
 | 
| 72029 |     90 |   \<comment> \<open>elimination of abstract record equality (manual proof)\<close>
 | 
|  |     91 | proof -
 | 
| 72030 |     92 |   let "?lhs = ?rhs" = ?thesis
 | 
|  |     93 |   from that have "xpos ?lhs = xpos ?rhs" by simp
 | 
| 72029 |     94 |   then show ?thesis by simp
 | 
|  |     95 | qed
 | 
|  |     96 | 
 | 
|  |     97 | 
 | 
| 72030 |     98 | text \<open>\<^medskip> Surjective pairing\<close>
 | 
| 72029 |     99 | 
 | 
| 72030 |    100 | lemma "r = \<lparr>xpos = xpos r, ypos = ypos r\<rparr>"
 | 
| 72029 |    101 |   by simp
 | 
|  |    102 | 
 | 
| 72030 |    103 | lemma "r = \<lparr>xpos = xpos r, ypos = ypos r, \<dots> = point.more r\<rparr>"
 | 
| 72029 |    104 |   by simp
 | 
|  |    105 | 
 | 
|  |    106 | 
 | 
| 72030 |    107 | text \<open>\<^medskip> Representation of records by cases or (degenerate) induction.\<close>
 | 
| 72029 |    108 | 
 | 
| 72030 |    109 | lemma "r\<lparr>xpos := n\<rparr>\<lparr>ypos := m\<rparr> = r\<lparr>ypos := m\<rparr>\<lparr>xpos := n\<rparr>"
 | 
| 72029 |    110 | proof (cases r)
 | 
|  |    111 |   fix xpos ypos more
 | 
|  |    112 |   assume "r = \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>"
 | 
|  |    113 |   then show ?thesis by simp
 | 
|  |    114 | qed
 | 
|  |    115 | 
 | 
| 72030 |    116 | lemma "r\<lparr>xpos := n\<rparr>\<lparr>ypos := m\<rparr> = r\<lparr>ypos := m\<rparr>\<lparr>xpos := n\<rparr>"
 | 
|  |    117 | proof (induct r)
 | 
|  |    118 |   fix xpos ypos more
 | 
|  |    119 |   show "\<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>\<lparr>xpos := n, ypos := m\<rparr> =
 | 
|  |    120 |       \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>\<lparr>ypos := m, xpos := n\<rparr>"
 | 
|  |    121 |     by simp
 | 
|  |    122 | qed
 | 
|  |    123 | 
 | 
|  |    124 | lemma "r\<lparr>xpos := n\<rparr>\<lparr>xpos := m\<rparr> = r\<lparr>xpos := m\<rparr>"
 | 
|  |    125 | proof (cases r)
 | 
|  |    126 |   fix xpos ypos more
 | 
|  |    127 |   assume "r = \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>"
 | 
|  |    128 |   then show ?thesis by simp
 | 
|  |    129 | qed
 | 
|  |    130 | 
 | 
|  |    131 | lemma "r\<lparr>xpos := n\<rparr>\<lparr>xpos := m\<rparr> = r\<lparr>xpos := m\<rparr>"
 | 
| 72029 |    132 | proof (cases r)
 | 
|  |    133 |   case fields
 | 
|  |    134 |   then show ?thesis by simp
 | 
|  |    135 | qed
 | 
|  |    136 | 
 | 
| 72030 |    137 | lemma "r\<lparr>xpos := n\<rparr>\<lparr>xpos := m\<rparr> = r\<lparr>xpos := m\<rparr>"
 | 
| 72029 |    138 |   by (cases r) simp
 | 
|  |    139 | 
 | 
|  |    140 | 
 | 
| 72030 |    141 | text \<open>\<^medskip> Concrete records are type instances of record schemes.\<close>
 | 
| 72029 |    142 | 
 | 
|  |    143 | definition foo5 :: nat
 | 
| 72030 |    144 |   where "foo5 = getX \<lparr>xpos = 1, ypos = 0\<rparr>"
 | 
| 72029 |    145 | 
 | 
|  |    146 | 
 | 
| 72030 |    147 | text \<open>\<^medskip> Manipulating the ``\<open>...\<close>'' (more) part.\<close>
 | 
| 72029 |    148 | 
 | 
| 72030 |    149 | definition incX :: "'a point_scheme \<Rightarrow> 'a point_scheme"
 | 
|  |    150 |   where "incX r = \<lparr>xpos = xpos r + 1, ypos = ypos r, \<dots> = point.more r\<rparr>"
 | 
| 72029 |    151 | 
 | 
|  |    152 | lemma "incX r = setX r (Suc (getX r))"
 | 
|  |    153 |   by (simp add: getX_def setX_def incX_def)
 | 
|  |    154 | 
 | 
|  |    155 | 
 | 
| 72030 |    156 | text \<open>\<^medskip> An alternative definition.\<close>
 | 
| 72029 |    157 | 
 | 
| 72030 |    158 | definition incX' :: "'a point_scheme \<Rightarrow> 'a point_scheme"
 | 
|  |    159 |   where "incX' r = r\<lparr>xpos := xpos r + 1\<rparr>"
 | 
| 72029 |    160 | 
 | 
|  |    161 | 
 | 
|  |    162 | subsection \<open>Coloured points: record extension\<close>
 | 
|  |    163 | 
 | 
|  |    164 | datatype colour = Red | Green | Blue
 | 
|  |    165 | 
 | 
|  |    166 | record cpoint = point +
 | 
|  |    167 |   colour :: colour
 | 
|  |    168 | 
 | 
|  |    169 | 
 | 
|  |    170 | text \<open>
 | 
|  |    171 |   The record declaration defines a new type constructor and abbreviations:
 | 
|  |    172 |   @{text [display]
 | 
| 72030 |    173 | \<open>cpoint = \<lparr>xpos :: nat, ypos :: nat, colour :: colour\<rparr> =
 | 
| 72029 |    174 |   () cpoint_ext_type point_ext_type
 | 
| 72030 |    175 | 'a cpoint_scheme = \<lparr>xpos :: nat, ypos :: nat, colour :: colour, \<dots> :: 'a\<rparr> =
 | 
| 72029 |    176 |   'a cpoint_ext_type point_ext_type\<close>}
 | 
|  |    177 | \<close>
 | 
|  |    178 | 
 | 
|  |    179 | consts foo6 :: cpoint
 | 
| 72030 |    180 | consts foo7 :: "\<lparr>xpos :: nat, ypos :: nat, colour :: colour\<rparr>"
 | 
| 72029 |    181 | consts foo8 :: "'a cpoint_scheme"
 | 
| 72030 |    182 | consts foo9 :: "\<lparr>xpos :: nat, ypos :: nat, colour :: colour, \<dots> :: 'a\<rparr>"
 | 
| 72029 |    183 | 
 | 
|  |    184 | 
 | 
| 72030 |    185 | text \<open>Functions on \<open>point\<close> schemes work for \<open>cpoints\<close> as well.\<close>
 | 
| 72029 |    186 | 
 | 
|  |    187 | definition foo10 :: nat
 | 
| 72030 |    188 |   where "foo10 = getX \<lparr>xpos = 2, ypos = 0, colour = Blue\<rparr>"
 | 
| 72029 |    189 | 
 | 
|  |    190 | 
 | 
|  |    191 | subsubsection \<open>Non-coercive structural subtyping\<close>
 | 
|  |    192 | 
 | 
| 72030 |    193 | text \<open>Term \<^term>\<open>foo11\<close> has type \<^typ>\<open>cpoint\<close>, not type \<^typ>\<open>point\<close> --- Great!\<close>
 | 
| 72029 |    194 | 
 | 
|  |    195 | definition foo11 :: cpoint
 | 
| 72030 |    196 |   where "foo11 = setX \<lparr>xpos = 2, ypos = 0, colour = Blue\<rparr> 0"
 | 
| 72029 |    197 | 
 | 
|  |    198 | 
 | 
|  |    199 | subsection \<open>Other features\<close>
 | 
|  |    200 | 
 | 
|  |    201 | text \<open>Field names contribute to record identity.\<close>
 | 
|  |    202 | 
 | 
|  |    203 | record point' =
 | 
|  |    204 |   xpos' :: nat
 | 
|  |    205 |   ypos' :: nat
 | 
|  |    206 | 
 | 
|  |    207 | text \<open>
 | 
| 72030 |    208 |   \<^noindent> May not apply \<^term>\<open>getX\<close> to @{term [source] "\<lparr>xpos' = 2, ypos' = 0\<rparr>"}
 | 
|  |    209 |   --- type error.
 | 
| 72029 |    210 | \<close>
 | 
|  |    211 | 
 | 
| 72030 |    212 | text \<open>\<^medskip> Polymorphic records.\<close>
 | 
| 72029 |    213 | 
 | 
|  |    214 | record 'a point'' = point +
 | 
|  |    215 |   content :: 'a
 | 
|  |    216 | 
 | 
|  |    217 | type_synonym cpoint'' = "colour point''"
 | 
|  |    218 | 
 | 
|  |    219 | 
 | 
|  |    220 | text \<open>Updating a record field with an identical value is simplified.\<close>
 | 
| 72030 |    221 | lemma "r\<lparr>xpos := xpos r\<rparr> = r"
 | 
| 72029 |    222 |   by simp
 | 
|  |    223 | 
 | 
|  |    224 | text \<open>Only the most recent update to a component survives simplification.\<close>
 | 
| 72030 |    225 | lemma "r\<lparr>xpos := x, ypos := y, xpos := x'\<rparr> = r\<lparr>ypos := y, xpos := x'\<rparr>"
 | 
| 72029 |    226 |   by simp
 | 
|  |    227 | 
 | 
| 72030 |    228 | text \<open>
 | 
|  |    229 |   In some cases its convenient to automatically split (quantified) records.
 | 
|  |    230 |   For this purpose there is the simproc @{ML [source] "Record.split_simproc"}
 | 
|  |    231 |   and the tactic @{ML [source] "Record.split_simp_tac"}. The simplification
 | 
|  |    232 |   procedure only splits the records, whereas the tactic also simplifies the
 | 
|  |    233 |   resulting goal with the standard record simplification rules. A
 | 
|  |    234 |   (generalized) predicate on the record is passed as parameter that decides
 | 
|  |    235 |   whether or how `deep' to split the record. It can peek on the subterm
 | 
|  |    236 |   starting at the quantified occurrence of the record (including the
 | 
|  |    237 |   quantifier). The value \<^ML>\<open>0\<close> indicates no split, a value greater
 | 
|  |    238 |   \<^ML>\<open>0\<close> splits up to the given bound of record extension and finally the
 | 
|  |    239 |   value \<^ML>\<open>~1\<close> completely splits the record. @{ML [source]
 | 
|  |    240 |   "Record.split_simp_tac"} additionally takes a list of equations for
 | 
|  |    241 |   simplification and can also split fixed record variables.
 | 
| 72029 |    242 | \<close>
 | 
|  |    243 | 
 | 
|  |    244 | lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
 | 
|  |    245 |   apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
 | 
|  |    246 |     addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
 | 
|  |    247 |   apply simp
 | 
|  |    248 |   done
 | 
|  |    249 | 
 | 
|  |    250 | lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
 | 
|  |    251 |   apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
 | 
|  |    252 |   apply simp
 | 
|  |    253 |   done
 | 
|  |    254 | 
 | 
|  |    255 | lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
 | 
|  |    256 |   apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
 | 
|  |    257 |     addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
 | 
|  |    258 |   apply simp
 | 
|  |    259 |   done
 | 
|  |    260 | 
 | 
|  |    261 | lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
 | 
|  |    262 |   apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
 | 
|  |    263 |   apply simp
 | 
|  |    264 |   done
 | 
|  |    265 | 
 | 
|  |    266 | lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
 | 
|  |    267 |   apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
 | 
|  |    268 |     addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
 | 
|  |    269 |   apply auto
 | 
|  |    270 |   done
 | 
|  |    271 | 
 | 
|  |    272 | lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
 | 
|  |    273 |   apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
 | 
|  |    274 |   apply auto
 | 
|  |    275 |   done
 | 
|  |    276 | 
 | 
|  |    277 | lemma "P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
 | 
|  |    278 |   apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
 | 
|  |    279 |   apply auto
 | 
|  |    280 |   done
 | 
|  |    281 | 
 | 
| 72030 |    282 | notepad
 | 
|  |    283 | begin
 | 
|  |    284 |   have "\<exists>x. P x"
 | 
|  |    285 |     if "P (xpos r)" for P r
 | 
|  |    286 |     apply (insert that)
 | 
|  |    287 |     apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
 | 
|  |    288 |     apply auto
 | 
|  |    289 |     done
 | 
|  |    290 | end
 | 
| 72029 |    291 | 
 | 
| 72030 |    292 | text \<open>
 | 
|  |    293 |   The effect of simproc @{ML [source] Record.ex_sel_eq_simproc} is illustrated
 | 
|  |    294 |   by the following lemma.\<close>
 | 
| 72029 |    295 | 
 | 
|  |    296 | lemma "\<exists>r. xpos r = x"
 | 
| 72030 |    297 |   by (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
 | 
| 72029 |    298 |     addsimprocs [Record.ex_sel_eq_simproc]) 1\<close>)
 | 
|  |    299 | 
 | 
|  |    300 | 
 | 
|  |    301 | subsection \<open>A more complex record expression\<close>
 | 
|  |    302 | 
 | 
|  |    303 | record ('a, 'b, 'c) bar = bar1 :: 'a
 | 
|  |    304 |   bar2 :: 'b
 | 
|  |    305 |   bar3 :: 'c
 | 
|  |    306 |   bar21 :: "'b \<times> 'a"
 | 
|  |    307 |   bar32 :: "'c \<times> 'b"
 | 
|  |    308 |   bar31 :: "'c \<times> 'a"
 | 
|  |    309 | 
 | 
|  |    310 | print_record "('a,'b,'c) bar"
 | 
|  |    311 | 
 | 
| 72030 |    312 | 
 | 
| 72029 |    313 | subsection \<open>Some code generation\<close>
 | 
|  |    314 | 
 | 
|  |    315 | export_code foo1 foo3 foo5 foo10 checking SML
 | 
|  |    316 | 
 | 
|  |    317 | text \<open>
 | 
| 72030 |    318 |   Code generation can also be switched off, for instance for very large
 | 
|  |    319 |   records:\<close>
 | 
| 72029 |    320 | 
 | 
|  |    321 | declare [[record_codegen = false]]
 | 
|  |    322 | 
 | 
|  |    323 | record not_so_large_record =
 | 
|  |    324 |   bar520 :: nat
 | 
| 72030 |    325 |   bar521 :: "nat \<times> nat"
 | 
| 72029 |    326 | 
 | 
|  |    327 | declare [[record_codegen = true]]
 | 
|  |    328 | 
 | 
|  |    329 | end
 |