# HG changeset patch # User wenzelm # Date 1594668036 -7200 # Node ID eece87547736b4a7a4648620d26c22a31df9cc2f # Parent 83456d9f0ed5a16a9fd3d37917ce8763b48fe519 misc tuning and modernization; diff -r 83456d9f0ed5 -r eece87547736 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 \Using extensible records in HOL -- points and coloured points\ theory Records -imports Main + imports Main begin subsection \Points\ @@ -21,7 +21,6 @@ following theorems: \ - 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] \point = \xpos :: nat, ypos :: nat\ = () point_ext_type 'a point_scheme = \xpos :: nat, ypos :: nat, ... :: 'a\ = 'a point_ext_type\} \ -consts foo2 :: "(| xpos :: nat, ypos :: nat |)" -consts foo4 :: "'a => (| xpos :: nat, ypos :: nat, ... :: 'a |)" +consts foo2 :: "\xpos :: nat, ypos :: nat\" +consts foo4 :: "'a \ \xpos :: nat, ypos :: nat, \ :: 'a\" subsubsection \Introducing concrete records and record schemes\ definition foo1 :: point - where "foo1 = (| xpos = 1, ypos = 0 |)" + where "foo1 = \xpos = 1, ypos = 0\" -definition foo3 :: "'a => 'a point_scheme" - where "foo3 ext = (| xpos = 1, ypos = 0, ... = ext |)" +definition foo3 :: "'a \ 'a point_scheme" + where "foo3 ext = \xpos = 1, ypos = 0, \ = ext\" subsubsection \Record selection and record update\ -definition getX :: "'a point_scheme => nat" +definition getX :: "'a point_scheme \ 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 \ nat \ 'a point_scheme" + where "setX r n = r \xpos := n\" subsubsection \Some lemmas about records\ text \Basic simplifications.\ -lemma "point.make n p = (| xpos = n, ypos = p |)" +lemma "point.make n p = \xpos = n, ypos = p\" by (simp only: point.make_def) -lemma "xpos (| xpos = m, ypos = n, ... = p |) = m" +lemma "xpos \xpos = m, ypos = n, \ = p\ = m" by simp -lemma "(| xpos = m, ypos = n, ... = p |) (| xpos:= 0 |) = (| xpos = 0, ypos = n, ... = p |)" +lemma "\xpos = m, ypos = n, \ = p\\xpos:= 0\ = \xpos = 0, ypos = n, \ = p\" by simp -text \\medskip Equality of records.\ +text \\<^medskip> Equality of records.\ -lemma "n = n' ==> p = p' ==> (| xpos = n, ypos = p |) = (| xpos = n', ypos = p' |)" +lemma "n = n' \ p = p' \ \xpos = n, ypos = p\ = \xpos = n', ypos = p'\" \ \introduction of concrete record equality\ by simp -lemma "(| xpos = n, ypos = p |) = (| xpos = n', ypos = p' |) ==> n = n'" +lemma "\xpos = n, ypos = p\ = \xpos = n', ypos = p'\ \ n = n'" \ \elimination of concrete record equality\ by simp -lemma "r (| xpos := n |) (| ypos := m |) = r (| ypos := m |) (| xpos := n |)" +lemma "r\xpos := n\\ypos := m\ = r\ypos := m\\xpos := n\" \ \introduction of abstract record equality\ by simp -lemma "r (| xpos := n |) = r (| xpos := n' |) ==> n = n'" +lemma "r\xpos := n\ = r\xpos := n'\" if "n = n'" \ \elimination of abstract record equality (manual proof)\ 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 \\medskip Surjective pairing\ +text \\<^medskip> Surjective pairing\ -lemma "r = (| xpos = xpos r, ypos = ypos r |)" +lemma "r = \xpos = xpos r, ypos = ypos r\" by simp -lemma "r = (| xpos = xpos r, ypos = ypos r, ... = point.more r |)" +lemma "r = \xpos = xpos r, ypos = ypos r, \ = point.more r\" by simp -text \ - \medskip Representation of records by cases or (degenerate) - induction. -\ +text \\<^medskip> Representation of records by cases or (degenerate) induction.\ -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\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 |) (| xpos := m |) = r (| xpos := m |)" +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\" +proof (cases r) + fix xpos ypos more + assume "r = \xpos = xpos, ypos = ypos, \ = more\" + then show ?thesis by simp +qed + +lemma "r\xpos := n\\xpos := m\ = r\xpos := m\" proof (cases r) case fields then show ?thesis by simp qed -lemma "r (| xpos := n |) (| xpos := m |) = r (| xpos := m |)" +lemma "r\xpos := n\\xpos := m\ = r\xpos := m\" by (cases r) simp -text \ - \medskip Concrete records are type instances of record schemes. -\ +text \\<^medskip> Concrete records are type instances of record schemes.\ definition foo5 :: nat - where "foo5 = getX (| xpos = 1, ypos = 0 |)" + where "foo5 = getX \xpos = 1, ypos = 0\" -text \\medskip Manipulating the ``\...\'' (more) part.\ +text \\<^medskip> Manipulating the ``\...\'' (more) part.\ -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 \ 'a point_scheme" + where "incX r = \xpos = xpos r + 1, ypos = ypos r, \ = point.more r\" lemma "incX r = setX r (Suc (getX r))" by (simp add: getX_def setX_def incX_def) -text \An alternative definition.\ +text \\<^medskip> An alternative definition.\ -definition incX' :: "'a point_scheme => 'a point_scheme" - where "incX' r = r (| xpos := xpos r + 1 |)" +definition incX' :: "'a point_scheme \ 'a point_scheme" + where "incX' r = r\xpos := xpos r + 1\" subsection \Coloured points: record extension\ @@ -176,35 +170,30 @@ text \ The record declaration defines a new type constructor and abbreviations: @{text [display] -\cpoint = (| xpos :: nat, ypos :: nat, colour :: colour |) = +\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_scheme = \xpos :: nat, ypos :: nat, colour :: colour, \ :: 'a\ = 'a cpoint_ext_type point_ext_type\} \ consts foo6 :: cpoint -consts foo7 :: "(| xpos :: nat, ypos :: nat, colour :: colour |)" +consts foo7 :: "\xpos :: nat, ypos :: nat, colour :: colour\" consts foo8 :: "'a cpoint_scheme" -consts foo9 :: "(| xpos :: nat, ypos :: nat, colour :: colour, ... :: 'a |)" +consts foo9 :: "\xpos :: nat, ypos :: nat, colour :: colour, \ :: 'a\" -text \ - Functions on \point\ schemes work for \cpoints\ as well. -\ +text \Functions on \point\ schemes work for \cpoints\ as well.\ definition foo10 :: nat - where "foo10 = getX (| xpos = 2, ypos = 0, colour = Blue |)" + where "foo10 = getX \xpos = 2, ypos = 0, colour = Blue\" subsubsection \Non-coercive structural subtyping\ -text \ - Term \<^term>\foo11\ has type \<^typ>\cpoint\, not type \<^typ>\point\ --- - Great! -\ +text \Term \<^term>\foo11\ has type \<^typ>\cpoint\, not type \<^typ>\point\ --- Great!\ definition foo11 :: cpoint - where "foo11 = setX (| xpos = 2, ypos = 0, colour = Blue |) 0" + where "foo11 = setX \xpos = 2, ypos = 0, colour = Blue\ 0" subsection \Other features\ @@ -216,11 +205,11 @@ ypos' :: nat text \ - \noindent May not apply \<^term>\getX\ to @{term [source] "(| xpos' = - 2, ypos' = 0 |)"} -- type error. + \<^noindent> May not apply \<^term>\getX\ to @{term [source] "\xpos' = 2, ypos' = 0\"} + --- type error. \ -text \\medskip Polymorphic records.\ +text \\<^medskip> Polymorphic records.\ record 'a point'' = point + content :: 'a @@ -228,30 +217,28 @@ type_synonym cpoint'' = "colour point''" - text \Updating a record field with an identical value is simplified.\ -lemma "r (| xpos := xpos r |) = r" +lemma "r\xpos := xpos r\ = r" by simp text \Only the most recent update to a component survives simplification.\ -lemma "r (| xpos := x, ypos := y, xpos := x' |) = r (| ypos := y, xpos := x' |)" +lemma "r\xpos := x, ypos := y, xpos := x'\ = r\ypos := y, xpos := x'\" by simp -text \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>\0\ indicates no split, a value -greater \<^ML>\0\ splits up to the given bound of record extension and -finally the value \<^ML>\~1\ 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 \ + 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>\0\ indicates no split, a value greater + \<^ML>\0\ splits up to the given bound of record extension and finally the + value \<^ML>\~1\ 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. \ lemma "(\r. P (xpos r)) \ (\x. P x)" @@ -292,27 +279,23 @@ apply auto done -lemma True -proof - - { - fix P r - assume pre: "P (xpos r)" - then have "\x. P x" - apply - - apply (tactic \Record.split_simp_tac \<^context> [] (K ~1) 1\) - apply auto - done - } - show ?thesis .. -qed +notepad +begin + have "\x. P x" + if "P (xpos r)" for P r + apply (insert that) + apply (tactic \Record.split_simp_tac \<^context> [] (K ~1) 1\) + apply auto + done +end -text \The effect of simproc @{ML [source] Record.ex_sel_eq_simproc} is - illustrated by the following lemma.\ +text \ + The effect of simproc @{ML [source] Record.ex_sel_eq_simproc} is illustrated + by the following lemma.\ lemma "\r. xpos r = x" - apply (tactic \simp_tac (put_simpset HOL_basic_ss \<^context> + by (tactic \simp_tac (put_simpset HOL_basic_ss \<^context> addsimprocs [Record.ex_sel_eq_simproc]) 1\) - done subsection \A more complex record expression\ @@ -326,19 +309,20 @@ print_record "('a,'b,'c) bar" + subsection \Some code generation\ export_code foo1 foo3 foo5 foo10 checking SML text \ - Code generation can also be switched off, for instance for very large records -\ + Code generation can also be switched off, for instance for very large + records:\ declare [[record_codegen = false]] record not_so_large_record = bar520 :: nat - bar521 :: "nat * nat" + bar521 :: "nat \ nat" declare [[record_codegen = true]]