src/HOL/ex/Records.thy
changeset 61343 5b5656a63bd6
parent 58889 5b7a9633cfa8
child 61499 4efe9a6dd212
--- a/src/HOL/ex/Records.thy	Tue Oct 06 17:46:07 2015 +0200
+++ b/src/HOL/ex/Records.thy	Tue Oct 06 17:47:28 2015 +0200
@@ -3,29 +3,29 @@
                 TU Muenchen
 *)
 
-section {* Using extensible records in HOL -- points and coloured points *}
+section \<open>Using extensible records in HOL -- points and coloured points\<close>
 
 theory Records
 imports Main
 begin
 
-subsection {* Points *}
+subsection \<open>Points\<close>
 
 record point =
   xpos :: nat
   ypos :: nat
 
-text {*
+text \<open>
   Apart many other things, above record declaration produces the
   following theorems:
-*}
+\<close>
 
 
 thm point.simps
 thm point.iffs
 thm point.defs
 
-text {*
+text \<open>
   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.
@@ -34,13 +34,13 @@
   @{text [display]
 "  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>
 
 consts foo2 :: "(| xpos :: nat, ypos :: nat |)"
 consts foo4 :: "'a => (| xpos :: nat, ypos :: nat, ... :: 'a |)"
 
 
-subsubsection {* Introducing concrete records and record schemes *}
+subsubsection \<open>Introducing concrete records and record schemes\<close>
 
 definition foo1 :: point
   where "foo1 = (| xpos = 1, ypos = 0 |)"
@@ -49,7 +49,7 @@
   where "foo3 ext = (| xpos = 1, ypos = 0, ... = ext |)"
 
 
-subsubsection {* Record selection and record update *}
+subsubsection \<open>Record selection and record update\<close>
 
 definition getX :: "'a point_scheme => nat"
   where "getX r = xpos r"
@@ -58,9 +58,9 @@
   where "setX r n = r (| xpos := n |)"
 
 
-subsubsection {* Some lemmas about records *}
+subsubsection \<open>Some lemmas about records\<close>
 
-text {* Basic simplifications. *}
+text \<open>Basic simplifications.\<close>
 
 lemma "point.make n p = (| xpos = n, ypos = p |)"
   by (simp only: point.make_def)
@@ -72,7 +72,7 @@
   by simp
 
 
-text {* \medskip Equality of records. *}
+text \<open>\medskip Equality of records.\<close>
 
 lemma "n = n' ==> p = p' ==> (| xpos = n, ypos = p |) = (| xpos = n', ypos = p' |)"
   -- "introduction of concrete record equality"
@@ -95,7 +95,7 @@
 qed
 
 
-text {* \medskip Surjective pairing *}
+text \<open>\medskip Surjective pairing\<close>
 
 lemma "r = (| xpos = xpos r, ypos = ypos r |)"
   by simp
@@ -104,10 +104,10 @@
   by simp
 
 
-text {*
+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)
@@ -141,15 +141,15 @@
   by (cases r) simp
 
 
-text {*
+text \<open>
  \medskip Concrete records are type instances of record schemes.
-*}
+\<close>
 
 definition foo5 :: nat
   where "foo5 = getX (| xpos = 1, ypos = 0 |)"
 
 
-text {* \medskip Manipulating the ``@{text "..."}'' (more) part. *}
+text \<open>\medskip Manipulating the ``@{text "..."}'' (more) part.\<close>
 
 definition incX :: "'a point_scheme => 'a point_scheme"
   where "incX r = (| xpos = xpos r + 1, ypos = ypos r, ... = point.more r |)"
@@ -158,13 +158,13 @@
   by (simp add: getX_def setX_def incX_def)
 
 
-text {* An alternative definition. *}
+text \<open>An alternative definition.\<close>
 
 definition incX' :: "'a point_scheme => 'a point_scheme"
   where "incX' r = r (| xpos := xpos r + 1 |)"
 
 
-subsection {* Coloured points: record extension *}
+subsection \<open>Coloured points: record extension\<close>
 
 datatype colour = Red | Green | Blue
 
@@ -172,14 +172,14 @@
   colour :: colour
 
 
-text {*
+text \<open>
   The record declaration defines a new type constructure and abbreviations:
   @{text [display]
 "  cpoint = (| xpos :: nat, ypos :: nat, colour :: colour |) =
      () cpoint_ext_type point_ext_type
    'a cpoint_scheme = (| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |) =
      'a cpoint_ext_type point_ext_type"}
-*}
+\<close>
 
 consts foo6 :: cpoint
 consts foo7 :: "(| xpos :: nat, ypos :: nat, colour :: colour |)"
@@ -187,39 +187,39 @@
 consts foo9 :: "(| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |)"
 
 
-text {*
+text \<open>
  Functions on @{text point} schemes work for @{text cpoints} as well.
-*}
+\<close>
 
 definition foo10 :: nat
   where "foo10 = getX (| xpos = 2, ypos = 0, colour = Blue |)"
 
 
-subsubsection {* Non-coercive structural subtyping *}
+subsubsection \<open>Non-coercive structural subtyping\<close>
 
-text {*
+text \<open>
  Term @{term foo11} has type @{typ cpoint}, not type @{typ point} ---
  Great!
-*}
+\<close>
 
 definition foo11 :: cpoint
   where "foo11 = setX (| xpos = 2, ypos = 0, colour = Blue |) 0"
 
 
-subsection {* Other features *}
+subsection \<open>Other features\<close>
 
-text {* Field names contribute to record identity. *}
+text \<open>Field names contribute to record identity.\<close>
 
 record point' =
   xpos' :: nat
   ypos' :: nat
 
-text {*
+text \<open>
   \noindent May not apply @{term getX} to @{term [source] "(| xpos' =
   2, ypos' = 0 |)"} -- type error.
-*}
+\<close>
 
-text {* \medskip Polymorphic records. *}
+text \<open>\medskip Polymorphic records.\<close>
 
 record 'a point'' = point +
   content :: 'a
@@ -228,15 +228,15 @@
 
 
 
-text {* Updating a record field with an identical value is simplified.*}
+text \<open>Updating a record field with an identical value is simplified.\<close>
 lemma "r (| xpos := xpos r |) = r"
   by simp
 
-text {* Only the most recent update to a component survives simplification. *}
+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' |)"
   by simp
 
-text {* In some cases its convenient to automatically split
+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
@@ -251,43 +251,43 @@
 @{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)"
-  apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context}
-    addsimprocs [Record.split_simproc (K ~1)]) 1 *})
+  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss @{context}
+    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
   apply simp
   done
 
 lemma "(\<forall>r. P (xpos r)) \<longrightarrow> (\<forall>x. P x)"
-  apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*})
+  apply (tactic \<open>Record.split_simp_tac @{context} [] (K ~1) 1\<close>)
   apply simp
   done
 
 lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
-  apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context}
-    addsimprocs [Record.split_simproc (K ~1)]) 1 *})
+  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss @{context}
+    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
   apply simp
   done
 
 lemma "(\<exists>r. P (xpos r)) \<longrightarrow> (\<exists>x. P x)"
-  apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1 *})
+  apply (tactic \<open>Record.split_simp_tac @{context} [] (K ~1) 1\<close>)
   apply simp
   done
 
 lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
-  apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context}
-    addsimprocs [Record.split_simproc (K ~1)]) 1 *})
+  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss @{context}
+    addsimprocs [Record.split_simproc (K ~1)]) 1\<close>)
   apply auto
   done
 
 lemma "\<And>r. P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
-  apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*})
+  apply (tactic \<open>Record.split_simp_tac @{context} [] (K ~1) 1\<close>)
   apply auto
   done
 
 lemma "P (xpos r) \<Longrightarrow> (\<exists>x. P x)"
-  apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*})
+  apply (tactic \<open>Record.split_simp_tac @{context} [] (K ~1) 1\<close>)
   apply auto
   done
 
@@ -298,23 +298,23 @@
     assume pre: "P (xpos r)"
     then have "\<exists>x. P x"
       apply -
-      apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1 *})
+      apply (tactic \<open>Record.split_simp_tac @{context} [] (K ~1) 1\<close>)
       apply auto
       done
   }
   show ?thesis ..
 qed
 
-text {* The effect of simproc @{ML [source] Record.ex_sel_eq_simproc} is
-  illustrated by the following lemma. *}
+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 {* simp_tac (put_simpset HOL_basic_ss @{context}
-    addsimprocs [Record.ex_sel_eq_simproc]) 1 *})
+  apply (tactic \<open>simp_tac (put_simpset HOL_basic_ss @{context}
+    addsimprocs [Record.ex_sel_eq_simproc]) 1\<close>)
   done
 
 
-subsection {* A more complex record expression *}
+subsection \<open>A more complex record expression\<close>
 
 record ('a, 'b, 'c) bar = bar1 :: 'a
   bar2 :: 'b
@@ -324,13 +324,13 @@
   bar31 :: "'c \<times> 'a"
 
 
-subsection {* Some code generation *}
+subsection \<open>Some code generation\<close>
 
 export_code foo1 foo3 foo5 foo10 checking SML
 
-text {*
+text \<open>
   Code generation can also be switched off, for instance for very large records
-*}
+\<close>
 
 declare [[record_codegen = false]]