Theory HOL-Examples.Records

(*  Title:      HOL/Examples/Records.thy
    Author:     Wolfgang Naraschewski, TU Muenchen
    Author:     Norbert Schirmer, TU Muenchen
    Author:     Norbert Schirmer, Apple, 2022
    Author:     Markus Wenzel, TU Muenchen
*)

section ‹Using extensible records in HOL -- points and coloured points›

theory Records
  imports Main
begin

subsection ‹Points›

record point =
  xpos :: nat
  ypos :: nat

text ‹
  Apart many other things, above record declaration produces the
  following theorems:
›

thm point.simps
thm point.iffs
thm point.defs

text ‹
  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.

   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"


subsubsection ‹Introducing concrete records and record schemes›

definition foo1 :: point
  where "foo1 = xpos = 1, ypos = 0"

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"
  where "getX r = xpos r"

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"
  by (simp only: point.make_def)

lemma "xpos xpos = m, ypos = n,  = p = m"
  by simp

lemma "xpos = m, ypos = n,  = pxpos:= 0 = xpos = 0, ypos = n,  = p"
  by simp


text  Equality of records.›

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'"
  ― ‹elimination of concrete record equality›
  by simp

lemma "rxpos := nypos := m = rypos := mxpos := n"
  ― ‹introduction of abstract record equality›
  by simp

lemma "rxpos := n = rxpos := n'" if "n = n'"
  ― ‹elimination of abstract record equality (manual proof)›
proof -
  let "?lhs = ?rhs" = ?thesis
  from that have "xpos ?lhs = xpos ?rhs" by simp
  then show ?thesis by simp
qed


text  Surjective pairing›

lemma "r = xpos = xpos r, ypos = ypos r"
  by simp

lemma "r = xpos = xpos r, ypos = ypos r,  = point.more r"
  by simp


text  Representation of records by cases or (degenerate) induction.›

lemma "rxpos := nypos := m = rypos := mxpos := n"
proof (cases r)
  fix xpos ypos more
  assume "r = xpos = xpos, ypos = ypos,  = more"
  then show ?thesis by simp
qed

lemma "rxpos := nypos := m = rypos := mxpos := n"
proof (induct r)
  fix xpos ypos more
  show "xpos = xpos, ypos = ypos,  = morexpos := n, ypos := m =
      xpos = xpos, ypos = ypos,  = moreypos := m, xpos := n"
    by simp
qed

lemma "rxpos := nxpos := m = rxpos := m"
proof (cases r)
  fix xpos ypos more
  assume "r = xpos = xpos, ypos = ypos,  = more"
  then show ?thesis by simp
qed

lemma "rxpos := nxpos := m = rxpos := m"
proof (cases r)
  case fields
  then show ?thesis by simp
qed

lemma "rxpos := nxpos := m = rxpos := m"
  by (cases r) simp


text  Concrete records are type instances of record schemes.›

definition foo5 :: nat
  where "foo5 = getX xpos = 1, ypos = 0"


text  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"

lemma "incX r = setX r (Suc (getX r))"
  by (simp add: getX_def setX_def incX_def)


text  An alternative definition.›

definition incX' :: "'a point_scheme  'a point_scheme"
  where "incX' r = rxpos := xpos r + 1"


subsection ‹Coloured points: record extension›

datatype colour = Red | Green | Blue

record cpoint = point +
  colour :: colour


text ‹
  The record declaration defines a new type constructor 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›}

consts foo6 :: cpoint
consts foo7 :: "xpos :: nat, ypos :: nat, colour :: colour"
consts foo8 :: "'a cpoint_scheme"
consts foo9 :: "xpos :: nat, ypos :: nat, colour :: colour,  :: 'a"


text ‹Functions on point› schemes work for cpoints› as well.›

definition foo10 :: nat
  where "foo10 = getX xpos = 2, ypos = 0, colour = Blue"


subsubsection ‹Non-coercive structural subtyping›

text ‹Term termfoo11 has type typcpoint, not type typpoint --- Great!›

definition foo11 :: cpoint
  where "foo11 = setX xpos = 2, ypos = 0, colour = Blue 0"


subsection ‹Other features›

text ‹Field names contribute to record identity.›

record point' =
  xpos' :: nat
  ypos' :: nat

text  May not apply termgetX to @{term [source] "xpos' = 2, ypos' = 0"}
  --- type error.
›

text  Polymorphic records.›

record 'a point'' = point +
  content :: 'a

type_synonym cpoint'' = "colour point''"


text ‹Updating a record field with an identical value is simplified.›
lemma "rxpos := xpos r = r"
  by simp

text ‹Only the most recent update to a component survives simplification.›
lemma "rxpos := x, ypos := y, xpos := x' = rypos := 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 ML0 indicates no split, a value greater
  ML0 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)"
  apply (tactic simp_tac (put_simpset HOL_basic_ss context
    addsimprocs [Record.split_simproc (K ~1)]) 1)
  apply simp
  done

lemma "(r. P (xpos r))  (x. P x)"
  apply (tactic Record.split_simp_tac context [] (K ~1) 1)
  apply simp
  done

lemma "(r. P (xpos r))  (x. P x)"
  apply (tactic simp_tac (put_simpset HOL_basic_ss context
    addsimprocs [Record.split_simproc (K ~1)]) 1)
  apply simp
  done

lemma "(r. P (xpos r))  (x. P x)"
  apply (tactic Record.split_simp_tac context [] (K ~1) 1)
  apply simp
  done

lemma "r. P (xpos r)  (x. P x)"
  apply (tactic simp_tac (put_simpset HOL_basic_ss context
    addsimprocs [Record.split_simproc (K ~1)]) 1)
  apply auto
  done

lemma "r. P (xpos r)  (x. P x)"
  apply (tactic Record.split_simp_tac context [] (K ~1) 1)
  apply auto
  done

lemma "P (xpos r)  (x. P x)"
  apply (tactic Record.split_simp_tac context [] (K ~1) 1)
  apply auto
  done

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.›

lemma "r. xpos r = x"
  supply [[simproc add: Record.ex_sel_eq]]
  apply (simp)
  done


subsection ‹Simprocs for update and equality›

record alph1 =
  a :: nat
  b :: nat

record alph2 = alph1 +
  c :: nat
  d :: nat

record alph3 = alph2 +
  e :: nat
  f :: nat

text ‹
  The simprocs that are activated by default are:
     @{ML [source] Record.simproc}: field selection of (nested) record updates.
     @{ML [source] Record.upd_simproc}: nested record updates.
     @{ML [source] Record.eq_simproc}: (componentwise) equality of records.
›


text ‹By default record updates are not ordered by simplification.›
schematic_goal "rb := x, a:= y = ?X"
  by simp

text ‹Normalisation towards an update ordering (string ordering of update function names) can
  be configured as follows.›
schematic_goal "rb := y, a := x = ?X"
  supply [[record_sort_updates]]
  by simp

text ‹Note the interplay between update ordering and record equality. Without update ordering
  the following equality is handled by @{ML [source] Record.eq_simproc}. Record equality is thus
  solved by componentwise comparison of all the fields of the records which can be expensive
  in the presence of many fields.›

lemma "rf := x1, a:= x2 = ra := x2, f:= x1"
  by simp

lemma "rf := x1, a:= x2 = ra := x2, f:= x1"
  supply [[simproc del: Record.eq]]
  apply (simp?)
  oops

text ‹With update ordering the equality is already established after update normalisation. There
  is no need for componentwise comparison.›

lemma "rf := x1, a:= x2 = ra := x2, f:= x1"
  supply [[record_sort_updates, simproc del: Record.eq]]
  apply simp
  done

schematic_goal "rf := x1, e := x2, d:= x3, c:= x4, b:=x5, a:= x6 = ?X"
  supply [[record_sort_updates]]
  by simp

schematic_goal "rf := x1, e := x2, d:= x3, c:= x4, e:=x5, a:= x6 = ?X"
  supply [[record_sort_updates]]
  by simp

schematic_goal "rf := x1, e := x2, d:= x3, c:= x4, e:=x5, a:= x6 = ?X"
  by simp


subsection ‹A more complex record expression›

record ('a, 'b, 'c) bar = bar1 :: 'a
  bar2 :: 'b
  bar3 :: 'c
  bar21 :: "'b × 'a"
  bar32 :: "'c × 'b"
  bar31 :: "'c × 'a"

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:›

declare [[record_codegen = false]]

record not_so_large_record =
  bar520 :: nat
  bar521 :: "nat × nat"


setup let
   val N = 300
  in
    Record.add_record {overloaded = false} ([], bindinglarge_record) NONE
      (map (fn i => (Binding.make ("fld_" ^ string_of_int i, ), @{typ nat}, Mixfix.NoSyn))
        (1 upto N))
  end

declare [[record_codegen]]

schematic_goal fld_1 (rfld_300 := x300, fld_20 := x20, fld_200 := x200) = ?X
  by simp

schematic_goal rfld_300 := x300, fld_20 := x20, fld_200 := x200 = ?X
  supply [[record_sort_updates]]
  by simp

end