misc tuning and modernization;
authorwenzelm
Mon, 13 Jul 2020 21:20:36 +0200
changeset 72030 eece87547736
parent 72029 83456d9f0ed5
child 72031 b7cec26e41d1
misc tuning and modernization;
src/HOL/Examples/Records.thy
--- a/src/HOL/Examples/Records.thy	Mon Jul 13 17:08:45 2020 +0200
+++ b/src/HOL/Examples/Records.thy	Mon Jul 13 21:20:36 2020 +0200
@@ -7,7 +7,7 @@
 section \<open>Using extensible records in HOL -- points and coloured points\<close>
 
 theory Records
-imports Main
+  imports Main
 begin
 
 subsection \<open>Points\<close>
@@ -21,7 +21,6 @@
   following theorems:
 \<close>
 
-
 thm point.simps
 thm point.iffs
 thm point.defs
@@ -31,138 +30,133 @@
   automatically to the standard simpset, @{thm [source] point.iffs} is
   added to the Classical Reasoner and Simplifier context.
 
-  \medskip Record declarations define new types and type abbreviations:
+  \<^medskip> Record declarations define new types and type abbreviations:
   @{text [display]
 \<open>point = \<lparr>xpos :: nat, ypos :: nat\<rparr> = () point_ext_type
 'a point_scheme = \<lparr>xpos :: nat, ypos :: nat, ... :: 'a\<rparr>  = 'a point_ext_type\<close>}
 \<close>
 
-consts foo2 :: "(| xpos :: nat, ypos :: nat |)"
-consts foo4 :: "'a => (| xpos :: nat, ypos :: nat, ... :: 'a |)"
+consts foo2 :: "\<lparr>xpos :: nat, ypos :: nat\<rparr>"
+consts foo4 :: "'a \<Rightarrow> \<lparr>xpos :: nat, ypos :: nat, \<dots> :: 'a\<rparr>"
 
 
 subsubsection \<open>Introducing concrete records and record schemes\<close>
 
 definition foo1 :: point
-  where "foo1 = (| xpos = 1, ypos = 0 |)"
+  where "foo1 = \<lparr>xpos = 1, ypos = 0\<rparr>"
 
-definition foo3 :: "'a => 'a point_scheme"
-  where "foo3 ext = (| xpos = 1, ypos = 0, ... = ext |)"
+definition foo3 :: "'a \<Rightarrow> 'a point_scheme"
+  where "foo3 ext = \<lparr>xpos = 1, ypos = 0, \<dots> = ext\<rparr>"
 
 
 subsubsection \<open>Record selection and record update\<close>
 
-definition getX :: "'a point_scheme => nat"
+definition getX :: "'a point_scheme \<Rightarrow> nat"
   where "getX r = xpos r"
 
-definition setX :: "'a point_scheme => nat => 'a point_scheme"
-  where "setX r n = r (| xpos := n |)"
+definition setX :: "'a point_scheme \<Rightarrow> nat \<Rightarrow> 'a point_scheme"
+  where "setX r n = r \<lparr>xpos := n\<rparr>"
 
 
 subsubsection \<open>Some lemmas about records\<close>
 
 text \<open>Basic simplifications.\<close>
 
-lemma "point.make n p = (| xpos = n, ypos = p |)"
+lemma "point.make n p = \<lparr>xpos = n, ypos = p\<rparr>"
   by (simp only: point.make_def)
 
-lemma "xpos (| xpos = m, ypos = n, ... = p |) = m"
+lemma "xpos \<lparr>xpos = m, ypos = n, \<dots> = p\<rparr> = m"
   by simp
 
-lemma "(| xpos = m, ypos = n, ... = p |) (| xpos:= 0 |) = (| xpos = 0, ypos = n, ... = p |)"
+lemma "\<lparr>xpos = m, ypos = n, \<dots> = p\<rparr>\<lparr>xpos:= 0\<rparr> = \<lparr>xpos = 0, ypos = n, \<dots> = p\<rparr>"
   by simp
 
 
-text \<open>\medskip Equality of records.\<close>
+text \<open>\<^medskip> Equality of records.\<close>
 
-lemma "n = n' ==> p = p' ==> (| xpos = n, ypos = p |) = (| xpos = n', ypos = p' |)"
+lemma "n = n' \<Longrightarrow> p = p' \<Longrightarrow> \<lparr>xpos = n, ypos = p\<rparr> = \<lparr>xpos = n', ypos = p'\<rparr>"
   \<comment> \<open>introduction of concrete record equality\<close>
   by simp
 
-lemma "(| xpos = n, ypos = p |) = (| xpos = n', ypos = p' |) ==> n = n'"
+lemma "\<lparr>xpos = n, ypos = p\<rparr> = \<lparr>xpos = n', ypos = p'\<rparr> \<Longrightarrow> n = n'"
   \<comment> \<open>elimination of concrete record equality\<close>
   by simp
 
-lemma "r (| xpos := n |) (| ypos := m |) = r (| ypos := m |) (| xpos := n |)"
+lemma "r\<lparr>xpos := n\<rparr>\<lparr>ypos := m\<rparr> = r\<lparr>ypos := m\<rparr>\<lparr>xpos := n\<rparr>"
   \<comment> \<open>introduction of abstract record equality\<close>
   by simp
 
-lemma "r (| xpos := n |) = r (| xpos := n' |) ==> n = n'"
+lemma "r\<lparr>xpos := n\<rparr> = r\<lparr>xpos := n'\<rparr>" if "n = n'"
   \<comment> \<open>elimination of abstract record equality (manual proof)\<close>
 proof -
-  assume "r (| xpos := n |) = r (| xpos := n' |)" (is "?lhs = ?rhs")
-  then have "xpos ?lhs = xpos ?rhs" by simp
+  let "?lhs = ?rhs" = ?thesis
+  from that have "xpos ?lhs = xpos ?rhs" by simp
   then show ?thesis by simp
 qed
 
 
-text \<open>\medskip Surjective pairing\<close>
+text \<open>\<^medskip> Surjective pairing\<close>
 
-lemma "r = (| xpos = xpos r, ypos = ypos r |)"
+lemma "r = \<lparr>xpos = xpos r, ypos = ypos r\<rparr>"
   by simp
 
-lemma "r = (| xpos = xpos r, ypos = ypos r, ... = point.more r |)"
+lemma "r = \<lparr>xpos = xpos r, ypos = ypos r, \<dots> = point.more r\<rparr>"
   by simp
 
 
-text \<open>
-  \medskip Representation of records by cases or (degenerate)
-  induction.
-\<close>
+text \<open>\<^medskip> Representation of records by cases or (degenerate) induction.\<close>
 
-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 |)"
+lemma "r\<lparr>xpos := n\<rparr>\<lparr>ypos := m\<rparr> = r\<lparr>ypos := m\<rparr>\<lparr>xpos := n\<rparr>"
 proof (cases r)
   fix xpos ypos more
   assume "r = \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>"
   then show ?thesis by simp
 qed
 
-lemma "r (| xpos := n |) (| xpos := m |) = r (| xpos := m |)"
+lemma "r\<lparr>xpos := n\<rparr>\<lparr>ypos := m\<rparr> = r\<lparr>ypos := m\<rparr>\<lparr>xpos := n\<rparr>"
+proof (induct r)
+  fix xpos ypos more
+  show "\<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>\<lparr>xpos := n, ypos := m\<rparr> =
+      \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>\<lparr>ypos := m, xpos := n\<rparr>"
+    by simp
+qed
+
+lemma "r\<lparr>xpos := n\<rparr>\<lparr>xpos := m\<rparr> = r\<lparr>xpos := m\<rparr>"
+proof (cases r)
+  fix xpos ypos more
+  assume "r = \<lparr>xpos = xpos, ypos = ypos, \<dots> = more\<rparr>"
+  then show ?thesis by simp
+qed
+
+lemma "r\<lparr>xpos := n\<rparr>\<lparr>xpos := m\<rparr> = r\<lparr>xpos := m\<rparr>"
 proof (cases r)
   case fields
   then show ?thesis by simp
 qed
 
-lemma "r (| xpos := n |) (| xpos := m |) = r (| xpos := m |)"
+lemma "r\<lparr>xpos := n\<rparr>\<lparr>xpos := m\<rparr> = r\<lparr>xpos := m\<rparr>"
   by (cases r) simp
 
 
-text \<open>
- \medskip Concrete records are type instances of record schemes.
-\<close>
+text \<open>\<^medskip> Concrete records are type instances of record schemes.\<close>
 
 definition foo5 :: nat
-  where "foo5 = getX (| xpos = 1, ypos = 0 |)"
+  where "foo5 = getX \<lparr>xpos = 1, ypos = 0\<rparr>"
 
 
-text \<open>\medskip Manipulating the ``\<open>...\<close>'' (more) part.\<close>
+text \<open>\<^medskip> Manipulating the ``\<open>...\<close>'' (more) part.\<close>
 
-definition incX :: "'a point_scheme => 'a point_scheme"
-  where "incX r = (| xpos = xpos r + 1, ypos = ypos r, ... = point.more r |)"
+definition incX :: "'a point_scheme \<Rightarrow> 'a point_scheme"
+  where "incX r = \<lparr>xpos = xpos r + 1, ypos = ypos r, \<dots> = point.more r\<rparr>"
 
 lemma "incX r = setX r (Suc (getX r))"
   by (simp add: getX_def setX_def incX_def)
 
 
-text \<open>An alternative definition.\<close>
+text \<open>\<^medskip> An alternative definition.\<close>
 
-definition incX' :: "'a point_scheme => 'a point_scheme"
-  where "incX' r = r (| xpos := xpos r + 1 |)"
+definition incX' :: "'a point_scheme \<Rightarrow> 'a point_scheme"
+  where "incX' r = r\<lparr>xpos := xpos r + 1\<rparr>"
 
 
 subsection \<open>Coloured points: record extension\<close>
@@ -176,35 +170,30 @@
 text \<open>
   The record declaration defines a new type constructor and abbreviations:
   @{text [display]
-\<open>cpoint = (| xpos :: nat, ypos :: nat, colour :: colour |) =
+\<open>cpoint = \<lparr>xpos :: nat, ypos :: nat, colour :: colour\<rparr> =
   () cpoint_ext_type point_ext_type
-'a cpoint_scheme = (| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |) =
+'a cpoint_scheme = \<lparr>xpos :: nat, ypos :: nat, colour :: colour, \<dots> :: 'a\<rparr> =
   'a cpoint_ext_type point_ext_type\<close>}
 \<close>
 
 consts foo6 :: cpoint
-consts foo7 :: "(| xpos :: nat, ypos :: nat, colour :: colour |)"
+consts foo7 :: "\<lparr>xpos :: nat, ypos :: nat, colour :: colour\<rparr>"
 consts foo8 :: "'a cpoint_scheme"
-consts foo9 :: "(| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |)"
+consts foo9 :: "\<lparr>xpos :: nat, ypos :: nat, colour :: colour, \<dots> :: 'a\<rparr>"
 
 
-text \<open>
- Functions on \<open>point\<close> schemes work for \<open>cpoints\<close> as well.
-\<close>
+text \<open>Functions on \<open>point\<close> schemes work for \<open>cpoints\<close> as well.\<close>
 
 definition foo10 :: nat
-  where "foo10 = getX (| xpos = 2, ypos = 0, colour = Blue |)"
+  where "foo10 = getX \<lparr>xpos = 2, ypos = 0, colour = Blue\<rparr>"
 
 
 subsubsection \<open>Non-coercive structural subtyping\<close>
 
-text \<open>
- Term \<^term>\<open>foo11\<close> has type \<^typ>\<open>cpoint\<close>, not type \<^typ>\<open>point\<close> ---
- Great!
-\<close>
+text \<open>Term \<^term>\<open>foo11\<close> has type \<^typ>\<open>cpoint\<close>, not type \<^typ>\<open>point\<close> --- Great!\<close>
 
 definition foo11 :: cpoint
-  where "foo11 = setX (| xpos = 2, ypos = 0, colour = Blue |) 0"
+  where "foo11 = setX \<lparr>xpos = 2, ypos = 0, colour = Blue\<rparr> 0"
 
 
 subsection \<open>Other features\<close>
@@ -216,11 +205,11 @@
   ypos' :: nat
 
 text \<open>
-  \noindent May not apply \<^term>\<open>getX\<close> to @{term [source] "(| xpos' =
-  2, ypos' = 0 |)"} -- type error.
+  \<^noindent> May not apply \<^term>\<open>getX\<close> to @{term [source] "\<lparr>xpos' = 2, ypos' = 0\<rparr>"}
+  --- type error.
 \<close>
 
-text \<open>\medskip Polymorphic records.\<close>
+text \<open>\<^medskip> Polymorphic records.\<close>
 
 record 'a point'' = point +
   content :: 'a
@@ -228,30 +217,28 @@
 type_synonym cpoint'' = "colour point''"
 
 
-
 text \<open>Updating a record field with an identical value is simplified.\<close>
-lemma "r (| xpos := xpos r |) = r"
+lemma "r\<lparr>xpos := xpos r\<rparr> = r"
   by simp
 
 text \<open>Only the most recent update to a component survives simplification.\<close>
-lemma "r (| xpos := x, ypos := y, xpos := x' |) = r (| ypos := y, xpos := x' |)"
+lemma "r\<lparr>xpos := x, ypos := y, xpos := x'\<rparr> = r\<lparr>ypos := y, xpos := x'\<rparr>"
   by simp
 
-text \<open>In some cases its convenient to automatically split
-(quantified) records. For this purpose there is the simproc @{ML [source]
-"Record.split_simproc"} and the tactic @{ML [source]
-"Record.split_simp_tac"}.  The simplification procedure
-only splits the records, whereas the tactic also simplifies the
-resulting goal with the standard record simplification rules. A
-(generalized) predicate on the record is passed as parameter that
-decides whether or how `deep' to split the record. It can peek on the
-subterm starting at the quantified occurrence of the record (including
-the quantifier). The value \<^ML>\<open>0\<close> indicates no split, a value
-greater \<^ML>\<open>0\<close> splits up to the given bound of record extension and
-finally the value \<^ML>\<open>~1\<close> completely splits the record.
-@{ML [source] "Record.split_simp_tac"} additionally takes a list of
-equations for simplification and can also split fixed record variables.
-
+text \<open>
+  In some cases its convenient to automatically split (quantified) records.
+  For this purpose there is the simproc @{ML [source] "Record.split_simproc"}
+  and the tactic @{ML [source] "Record.split_simp_tac"}. The simplification
+  procedure only splits the records, whereas the tactic also simplifies the
+  resulting goal with the standard record simplification rules. A
+  (generalized) predicate on the record is passed as parameter that decides
+  whether or how `deep' to split the record. It can peek on the subterm
+  starting at the quantified occurrence of the record (including the
+  quantifier). The value \<^ML>\<open>0\<close> indicates no split, a value greater
+  \<^ML>\<open>0\<close> splits up to the given bound of record extension and finally the
+  value \<^ML>\<open>~1\<close> completely splits the record. @{ML [source]
+  "Record.split_simp_tac"} additionally takes a list of equations for
+  simplification and can also split fixed record variables.
 \<close>
 
 lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
@@ -292,27 +279,23 @@
   apply auto
   done
 
-lemma True
-proof -
-  {
-    fix P r
-    assume pre: "P (xpos r)"
-    then have "\<exists>x. P x"
-      apply -
-      apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
-      apply auto
-      done
-  }
-  show ?thesis ..
-qed
+notepad
+begin
+  have "\<exists>x. P x"
+    if "P (xpos r)" for P r
+    apply (insert that)
+    apply (tactic \<open>Record.split_simp_tac \<^context> [] (K ~1) 1\<close>)
+    apply auto
+    done
+end
 
-text \<open>The effect of simproc @{ML [source] Record.ex_sel_eq_simproc} is
-  illustrated by the following lemma.\<close>
+text \<open>
+  The effect of simproc @{ML [source] Record.ex_sel_eq_simproc} is illustrated
+  by the following lemma.\<close>
 
 lemma "\<exists>r. xpos r = x"
-  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
+  by (tactic \<open>simp_tac (put_simpset HOL_basic_ss \<^context>
     addsimprocs [Record.ex_sel_eq_simproc]) 1\<close>)
-  done
 
 
 subsection \<open>A more complex record expression\<close>
@@ -326,19 +309,20 @@
 
 print_record "('a,'b,'c) bar"
 
+
 subsection \<open>Some code generation\<close>
 
 export_code foo1 foo3 foo5 foo10 checking SML
 
 text \<open>
-  Code generation can also be switched off, for instance for very large records
-\<close>
+  Code generation can also be switched off, for instance for very large
+  records:\<close>
 
 declare [[record_codegen = false]]
 
 record not_so_large_record =
   bar520 :: nat
-  bar521 :: "nat * nat"
+  bar521 :: "nat \<times> nat"
 
 declare [[record_codegen = true]]