| author | wenzelm | 
| Thu, 14 Mar 2019 16:35:58 +0100 | |
| changeset 69912 | dd55d2c926d9 | 
| parent 69605 | a96320074298 | 
| child 69913 | ca515cf61651 | 
| permissions | -rw-r--r-- | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 1 | (* Title: HOL/Record.thy | 
| 32763 | 2 | Author: Wolfgang Naraschewski, TU Muenchen | 
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | Author: Norbert Schirmer, TU Muenchen | |
| 5 | Author: Thomas Sewell, NICTA | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 6 | Author: Florian Haftmann, TU Muenchen | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 7 | *) | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 8 | |
| 60758 | 9 | section \<open>Extensible records with structural subtyping\<close> | 
| 22817 | 10 | |
| 15131 | 11 | theory Record | 
| 56048 
d311c6377e08
moved 'Quickcheck_Narrowing' further down the theory graph
 blanchet parents: 
56047diff
changeset | 12 | imports Quickcheck_Exhaustive | 
| 62117 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61955diff
changeset | 13 | keywords | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61955diff
changeset | 14 | "record" :: thy_decl and | 
| 
86a31308a8e1
print_record: diagnostic printing of record definitions
 kleing parents: 
61955diff
changeset | 15 | "print_record" :: diag | 
| 15131 | 16 | begin | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 17 | |
| 60758 | 18 | subsection \<open>Introduction\<close> | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 19 | |
| 60758 | 20 | text \<open> | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 21 | Records are isomorphic to compound tuple types. To implement | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 22 | efficient records, we make this isomorphism explicit. Consider the | 
| 61799 | 23 | record access/update simplification \<open>alpha (beta_update f | 
| 24 | rec) = alpha rec\<close> for distinct fields alpha and beta of some record | |
| 25 | rec with n fields. There are \<open>n ^ 2\<close> such theorems, which | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 26 | prohibits storage of all of them for large n. The rules can be | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 27 | proved on the fly by case decomposition and simplification in O(n) | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 28 | time. By creating O(n) isomorphic-tuple types while defining the | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 29 | record, however, we can prove the access/update simplification in | 
| 61799 | 30 | \<open>O(log(n)^2)\<close> time. | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 31 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 32 | The O(n) cost of case decomposition is not because O(n) steps are | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 33 | taken, but rather because the resulting rule must contain O(n) new | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 34 | variables and an O(n) size concrete record construction. To sidestep | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 35 | this cost, we would like to avoid case decomposition in proving | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 36 | access/update theorems. | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 37 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 38 | Record types are defined as isomorphic to tuple types. For instance, | 
| 61799 | 39 | a record type with fields \<open>'a\<close>, \<open>'b\<close>, \<open>'c\<close> | 
| 40 | and \<open>'d\<close> might be introduced as isomorphic to \<open>'a \<times> | |
| 41 |   ('b \<times> ('c \<times> 'd))\<close>. If we balance the tuple tree to \<open>('a \<times>
 | |
| 42 |   'b) \<times> ('c \<times> 'd)\<close> then accessors can be defined by converting to the
 | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 43 | underlying type then using O(log(n)) fst or snd operations. | 
| 61799 | 44 | Updators can be defined similarly, if we introduce a \<open>fst_update\<close> and \<open>snd_update\<close> function. Furthermore, we can | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 45 | prove the access/update theorem in O(log(n)) steps by using simple | 
| 61799 | 46 | rewrites on fst, snd, \<open>fst_update\<close> and \<open>snd_update\<close>. | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 47 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 48 | The catch is that, although O(log(n)) steps were taken, the | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 49 | underlying type we converted to is a tuple tree of size | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 50 | O(n). Processing this term type wastes performance. We avoid this | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 51 | for large n by taking each subtree of size K and defining a new type | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 52 | isomorphic to that tuple subtree. A record can now be defined as | 
| 61799 | 53 | isomorphic to a tuple tree of these O(n/K) new types, or, if \<open>n > K*K\<close>, we can repeat the process, until the record can be | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 54 | defined in terms of a tuple tree of complexity less than the | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 55 | constant K. | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 56 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 57 | If we prove the access/update theorem on this type with the | 
| 61799 | 58 | analogous steps to the tuple tree, we consume \<open>O(log(n)^2)\<close> | 
| 59 | time as the intermediate terms are \<open>O(log(n))\<close> in size and | |
| 44922 | 60 | the types needed have size bounded by K. To enable this analogous | 
| 61799 | 61 | traversal, we define the functions seen below: \<open>iso_tuple_fst\<close>, \<open>iso_tuple_snd\<close>, \<open>iso_tuple_fst_update\<close> | 
| 62 | and \<open>iso_tuple_snd_update\<close>. These functions generalise tuple | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 63 | operations by taking a parameter that encapsulates a tuple | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 64 | isomorphism. The rewrites needed on these functions now need an | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 65 | additional assumption which is that the isomorphism works. | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 66 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 67 | These rewrites are typically used in a structured way. They are here | 
| 61799 | 68 | presented as the introduction rule \<open>isomorphic_tuple.intros\<close> | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 69 | rather than as a rewrite rule set. The introduction form is an | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 70 | optimisation, as net matching can be performed at one term location | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 71 | for each step rather than the simplifier searching the term for | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 72 | possible pattern matches. The rule set is used as it is viewed | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 73 | outside the locale, with the locale assumption (that the isomorphism | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 74 | is valid) left as a rule assumption. All rules are structured to aid | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 75 | net matching, using either a point-free form or an encapsulating | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 76 | predicate. | 
| 60758 | 77 | \<close> | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 78 | |
| 60758 | 79 | subsection \<open>Operators and lemmas for types isomorphic to tuples\<close> | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 80 | |
| 58310 | 81 | datatype (dead 'a, dead 'b, dead 'c) tuple_isomorphism = | 
| 35132 | 82 | Tuple_Isomorphism "'a \<Rightarrow> 'b \<times> 'c" "'b \<times> 'c \<Rightarrow> 'a" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 83 | |
| 35132 | 84 | primrec | 
| 85 |   repr :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'c" where
 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 86 | "repr (Tuple_Isomorphism r a) = r" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 87 | |
| 35132 | 88 | primrec | 
| 89 |   abst :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'b \<times> 'c \<Rightarrow> 'a" where
 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 90 | "abst (Tuple_Isomorphism r a) = a" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 91 | |
| 35132 | 92 | definition | 
| 93 |   iso_tuple_fst :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'a \<Rightarrow> 'b" where
 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 94 | "iso_tuple_fst isom = fst \<circ> repr isom" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 95 | |
| 35132 | 96 | definition | 
| 97 |   iso_tuple_snd :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'a \<Rightarrow> 'c" where
 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 98 | "iso_tuple_snd isom = snd \<circ> repr isom" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 99 | |
| 35132 | 100 | definition | 
| 101 | iso_tuple_fst_update :: | |
| 102 |     "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'a)" where
 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 103 | "iso_tuple_fst_update isom f = abst isom \<circ> apfst f \<circ> repr isom" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 104 | |
| 35132 | 105 | definition | 
| 106 | iso_tuple_snd_update :: | |
| 107 |     "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> ('c \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'a)" where
 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 108 | "iso_tuple_snd_update isom f = abst isom \<circ> apsnd f \<circ> repr isom" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 109 | |
| 35132 | 110 | definition | 
| 111 | iso_tuple_cons :: | |
| 112 |     "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'a" where
 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 113 | "iso_tuple_cons isom = curry (abst isom)" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 114 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 115 | |
| 60758 | 116 | subsection \<open>Logical infrastructure for records\<close> | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 117 | |
| 35132 | 118 | definition | 
| 119 |   iso_tuple_surjective_proof_assist :: "'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 120 | "iso_tuple_surjective_proof_assist x y f \<longleftrightarrow> f x = y" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 121 | |
| 35132 | 122 | definition | 
| 123 | iso_tuple_update_accessor_cong_assist :: | |
| 124 |     "(('b \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'a)) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
 | |
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 125 | "iso_tuple_update_accessor_cong_assist upd ac \<longleftrightarrow> | 
| 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 126 | (\<forall>f v. upd (\<lambda>x. f (ac v)) v = upd f v) \<and> (\<forall>v. upd id v = v)" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 127 | |
| 35132 | 128 | definition | 
| 129 | iso_tuple_update_accessor_eq_assist :: | |
| 130 |     "(('b \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'a)) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" where
 | |
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 131 | "iso_tuple_update_accessor_eq_assist upd ac v f v' x \<longleftrightarrow> | 
| 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 132 | upd f v = v' \<and> ac v = x \<and> iso_tuple_update_accessor_cong_assist upd ac" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 133 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 134 | lemma update_accessor_congruence_foldE: | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 135 | assumes uac: "iso_tuple_update_accessor_cong_assist upd ac" | 
| 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 136 | and r: "r = r'" and v: "ac r' = v'" | 
| 35132 | 137 | and f: "\<And>v. v' = v \<Longrightarrow> f v = f' v" | 
| 138 | shows "upd f r = upd f' r'" | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 139 | using uac r v [symmetric] | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 140 | apply (subgoal_tac "upd (\<lambda>x. f (ac r')) r' = upd (\<lambda>x. f' (ac r')) r'") | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 141 | apply (simp add: iso_tuple_update_accessor_cong_assist_def) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 142 | apply (simp add: f) | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 143 | done | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 144 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 145 | lemma update_accessor_congruence_unfoldE: | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 146 | "iso_tuple_update_accessor_cong_assist upd ac \<Longrightarrow> | 
| 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 147 | r = r' \<Longrightarrow> ac r' = v' \<Longrightarrow> (\<And>v. v = v' \<Longrightarrow> f v = f' v) \<Longrightarrow> | 
| 35132 | 148 | upd f r = upd f' r'" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 149 | apply (erule(2) update_accessor_congruence_foldE) | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 150 | apply simp | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 151 | done | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 152 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 153 | lemma iso_tuple_update_accessor_cong_assist_id: | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 154 | "iso_tuple_update_accessor_cong_assist upd ac \<Longrightarrow> upd id = id" | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 155 | by rule (simp add: iso_tuple_update_accessor_cong_assist_def) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 156 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 157 | lemma update_accessor_noopE: | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 158 | assumes uac: "iso_tuple_update_accessor_cong_assist upd ac" | 
| 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 159 | and ac: "f (ac x) = ac x" | 
| 35132 | 160 | shows "upd f x = x" | 
| 161 | using uac | |
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 162 | by (simp add: ac iso_tuple_update_accessor_cong_assist_id [OF uac, unfolded id_def] | 
| 35132 | 163 | cong: update_accessor_congruence_unfoldE [OF uac]) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 164 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 165 | lemma update_accessor_noop_compE: | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 166 | assumes uac: "iso_tuple_update_accessor_cong_assist upd ac" | 
| 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 167 | and ac: "f (ac x) = ac x" | 
| 35132 | 168 | shows "upd (g \<circ> f) x = upd g x" | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 169 | by (simp add: ac cong: update_accessor_congruence_unfoldE[OF uac]) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 170 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 171 | lemma update_accessor_cong_assist_idI: | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 172 | "iso_tuple_update_accessor_cong_assist id id" | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 173 | by (simp add: iso_tuple_update_accessor_cong_assist_def) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 174 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 175 | lemma update_accessor_cong_assist_triv: | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 176 | "iso_tuple_update_accessor_cong_assist upd ac \<Longrightarrow> | 
| 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 177 | iso_tuple_update_accessor_cong_assist upd ac" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 178 | by assumption | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 179 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 180 | lemma update_accessor_accessor_eqE: | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 181 | "iso_tuple_update_accessor_eq_assist upd ac v f v' x \<Longrightarrow> ac v = x" | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 182 | by (simp add: iso_tuple_update_accessor_eq_assist_def) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 183 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 184 | lemma update_accessor_updator_eqE: | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 185 | "iso_tuple_update_accessor_eq_assist upd ac v f v' x \<Longrightarrow> upd f v = v'" | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 186 | by (simp add: iso_tuple_update_accessor_eq_assist_def) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 187 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 188 | lemma iso_tuple_update_accessor_eq_assist_idI: | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 189 | "v' = f v \<Longrightarrow> iso_tuple_update_accessor_eq_assist id id v f v' v" | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 190 | by (simp add: iso_tuple_update_accessor_eq_assist_def update_accessor_cong_assist_idI) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 191 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 192 | lemma iso_tuple_update_accessor_eq_assist_triv: | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 193 | "iso_tuple_update_accessor_eq_assist upd ac v f v' x \<Longrightarrow> | 
| 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 194 | iso_tuple_update_accessor_eq_assist upd ac v f v' x" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 195 | by assumption | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 196 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 197 | lemma iso_tuple_update_accessor_cong_from_eq: | 
| 38394 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 198 | "iso_tuple_update_accessor_eq_assist upd ac v f v' x \<Longrightarrow> | 
| 
3142c1e21a0e
moved Record.thy from session Plain to Main; avoid variable name acc
 haftmann parents: 
36176diff
changeset | 199 | iso_tuple_update_accessor_cong_assist upd ac" | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 200 | by (simp add: iso_tuple_update_accessor_eq_assist_def) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 201 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 202 | lemma iso_tuple_surjective_proof_assistI: | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 203 | "f x = y \<Longrightarrow> iso_tuple_surjective_proof_assist x y f" | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 204 | by (simp add: iso_tuple_surjective_proof_assist_def) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 205 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 206 | lemma iso_tuple_surjective_proof_assist_idE: | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 207 | "iso_tuple_surjective_proof_assist x y id \<Longrightarrow> x = y" | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 208 | by (simp add: iso_tuple_surjective_proof_assist_def) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 209 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 210 | locale isomorphic_tuple = | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 211 |   fixes isom :: "('a, 'b, 'c) tuple_isomorphism"
 | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 212 | assumes repr_inv: "\<And>x. abst isom (repr isom x) = x" | 
| 35132 | 213 | and abst_inv: "\<And>y. repr isom (abst isom y) = y" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 214 | begin | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 215 | |
| 35132 | 216 | lemma repr_inj: "repr isom x = repr isom y \<longleftrightarrow> x = y" | 
| 217 | by (auto dest: arg_cong [of "repr isom x" "repr isom y" "abst isom"] | |
| 218 | simp add: repr_inv) | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 219 | |
| 35132 | 220 | lemma abst_inj: "abst isom x = abst isom y \<longleftrightarrow> x = y" | 
| 221 | by (auto dest: arg_cong [of "abst isom x" "abst isom y" "repr isom"] | |
| 222 | simp add: abst_inv) | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 223 | |
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 224 | lemmas simps = Let_def repr_inv abst_inv repr_inj abst_inj | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 225 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 226 | lemma iso_tuple_access_update_fst_fst: | 
| 67091 | 227 | "f \<circ> h g = j \<circ> f \<Longrightarrow> | 
| 228 | (f \<circ> iso_tuple_fst isom) \<circ> (iso_tuple_fst_update isom \<circ> h) g = | |
| 229 | j \<circ> (f \<circ> iso_tuple_fst isom)" | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 230 | by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_fst_def simps | 
| 44922 | 231 | fun_eq_iff) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 232 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 233 | lemma iso_tuple_access_update_snd_snd: | 
| 67091 | 234 | "f \<circ> h g = j \<circ> f \<Longrightarrow> | 
| 235 | (f \<circ> iso_tuple_snd isom) \<circ> (iso_tuple_snd_update isom \<circ> h) g = | |
| 236 | j \<circ> (f \<circ> iso_tuple_snd isom)" | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 237 | by (clarsimp simp: iso_tuple_snd_update_def iso_tuple_snd_def simps | 
| 44922 | 238 | fun_eq_iff) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 239 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 240 | lemma iso_tuple_access_update_fst_snd: | 
| 67091 | 241 | "(f \<circ> iso_tuple_fst isom) \<circ> (iso_tuple_snd_update isom \<circ> h) g = | 
| 242 | id \<circ> (f \<circ> iso_tuple_fst isom)" | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 243 | by (clarsimp simp: iso_tuple_snd_update_def iso_tuple_fst_def simps | 
| 44922 | 244 | fun_eq_iff) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 245 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 246 | lemma iso_tuple_access_update_snd_fst: | 
| 67091 | 247 | "(f \<circ> iso_tuple_snd isom) \<circ> (iso_tuple_fst_update isom \<circ> h) g = | 
| 248 | id \<circ> (f \<circ> iso_tuple_snd isom)" | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 249 | by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_snd_def simps | 
| 44922 | 250 | fun_eq_iff) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 251 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 252 | lemma iso_tuple_update_swap_fst_fst: | 
| 67091 | 253 | "h f \<circ> j g = j g \<circ> h f \<Longrightarrow> | 
| 254 | (iso_tuple_fst_update isom \<circ> h) f \<circ> (iso_tuple_fst_update isom \<circ> j) g = | |
| 255 | (iso_tuple_fst_update isom \<circ> j) g \<circ> (iso_tuple_fst_update isom \<circ> h) f" | |
| 44922 | 256 | by (clarsimp simp: iso_tuple_fst_update_def simps apfst_compose fun_eq_iff) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 257 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 258 | lemma iso_tuple_update_swap_snd_snd: | 
| 67091 | 259 | "h f \<circ> j g = j g \<circ> h f \<Longrightarrow> | 
| 260 | (iso_tuple_snd_update isom \<circ> h) f \<circ> (iso_tuple_snd_update isom \<circ> j) g = | |
| 261 | (iso_tuple_snd_update isom \<circ> j) g \<circ> (iso_tuple_snd_update isom \<circ> h) f" | |
| 44922 | 262 | by (clarsimp simp: iso_tuple_snd_update_def simps apsnd_compose fun_eq_iff) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 263 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 264 | lemma iso_tuple_update_swap_fst_snd: | 
| 67091 | 265 | "(iso_tuple_snd_update isom \<circ> h) f \<circ> (iso_tuple_fst_update isom \<circ> j) g = | 
| 266 | (iso_tuple_fst_update isom \<circ> j) g \<circ> (iso_tuple_snd_update isom \<circ> h) f" | |
| 35132 | 267 | by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_snd_update_def | 
| 44922 | 268 | simps fun_eq_iff) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 269 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 270 | lemma iso_tuple_update_swap_snd_fst: | 
| 67091 | 271 | "(iso_tuple_fst_update isom \<circ> h) f \<circ> (iso_tuple_snd_update isom \<circ> j) g = | 
| 272 | (iso_tuple_snd_update isom \<circ> j) g \<circ> (iso_tuple_fst_update isom \<circ> h) f" | |
| 44922 | 273 | by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_snd_update_def simps | 
| 274 | fun_eq_iff) | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 275 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 276 | lemma iso_tuple_update_compose_fst_fst: | 
| 67091 | 277 | "h f \<circ> j g = k (f \<circ> g) \<Longrightarrow> | 
| 278 | (iso_tuple_fst_update isom \<circ> h) f \<circ> (iso_tuple_fst_update isom \<circ> j) g = | |
| 279 | (iso_tuple_fst_update isom \<circ> k) (f \<circ> g)" | |
| 44922 | 280 | by (clarsimp simp: iso_tuple_fst_update_def simps apfst_compose fun_eq_iff) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 281 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 282 | lemma iso_tuple_update_compose_snd_snd: | 
| 67091 | 283 | "h f \<circ> j g = k (f \<circ> g) \<Longrightarrow> | 
| 284 | (iso_tuple_snd_update isom \<circ> h) f \<circ> (iso_tuple_snd_update isom \<circ> j) g = | |
| 285 | (iso_tuple_snd_update isom \<circ> k) (f \<circ> g)" | |
| 44922 | 286 | by (clarsimp simp: iso_tuple_snd_update_def simps apsnd_compose fun_eq_iff) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 287 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 288 | lemma iso_tuple_surjective_proof_assist_step: | 
| 67091 | 289 | "iso_tuple_surjective_proof_assist v a (iso_tuple_fst isom \<circ> f) \<Longrightarrow> | 
| 290 | iso_tuple_surjective_proof_assist v b (iso_tuple_snd isom \<circ> f) \<Longrightarrow> | |
| 35132 | 291 | iso_tuple_surjective_proof_assist v (iso_tuple_cons isom a b) f" | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 292 | by (clarsimp simp: iso_tuple_surjective_proof_assist_def simps | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 293 | iso_tuple_fst_def iso_tuple_snd_def iso_tuple_cons_def) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 294 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 295 | lemma iso_tuple_fst_update_accessor_cong_assist: | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 296 | assumes "iso_tuple_update_accessor_cong_assist f g" | 
| 35132 | 297 | shows "iso_tuple_update_accessor_cong_assist | 
| 67091 | 298 | (iso_tuple_fst_update isom \<circ> f) (g \<circ> iso_tuple_fst isom)" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 299 | proof - | 
| 35132 | 300 | from assms have "f id = id" | 
| 301 | by (rule iso_tuple_update_accessor_cong_assist_id) | |
| 302 | with assms show ?thesis | |
| 303 | by (clarsimp simp: iso_tuple_update_accessor_cong_assist_def simps | |
| 304 | iso_tuple_fst_update_def iso_tuple_fst_def) | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 305 | qed | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 306 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 307 | lemma iso_tuple_snd_update_accessor_cong_assist: | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 308 | assumes "iso_tuple_update_accessor_cong_assist f g" | 
| 35132 | 309 | shows "iso_tuple_update_accessor_cong_assist | 
| 67091 | 310 | (iso_tuple_snd_update isom \<circ> f) (g \<circ> iso_tuple_snd isom)" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 311 | proof - | 
| 35132 | 312 | from assms have "f id = id" | 
| 313 | by (rule iso_tuple_update_accessor_cong_assist_id) | |
| 314 | with assms show ?thesis | |
| 315 | by (clarsimp simp: iso_tuple_update_accessor_cong_assist_def simps | |
| 316 | iso_tuple_snd_update_def iso_tuple_snd_def) | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 317 | qed | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 318 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 319 | lemma iso_tuple_fst_update_accessor_eq_assist: | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 320 | assumes "iso_tuple_update_accessor_eq_assist f g a u a' v" | 
| 35132 | 321 | shows "iso_tuple_update_accessor_eq_assist | 
| 67091 | 322 | (iso_tuple_fst_update isom \<circ> f) (g \<circ> iso_tuple_fst isom) | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 323 | (iso_tuple_cons isom a b) u (iso_tuple_cons isom a' b) v" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 324 | proof - | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 325 | from assms have "f id = id" | 
| 35132 | 326 | by (auto simp add: iso_tuple_update_accessor_eq_assist_def | 
| 327 | intro: iso_tuple_update_accessor_cong_assist_id) | |
| 328 | with assms show ?thesis | |
| 329 | by (clarsimp simp: iso_tuple_update_accessor_eq_assist_def | |
| 330 | iso_tuple_fst_update_def iso_tuple_fst_def | |
| 331 | iso_tuple_update_accessor_cong_assist_def iso_tuple_cons_def simps) | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 332 | qed | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 333 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 334 | lemma iso_tuple_snd_update_accessor_eq_assist: | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 335 | assumes "iso_tuple_update_accessor_eq_assist f g b u b' v" | 
| 35132 | 336 | shows "iso_tuple_update_accessor_eq_assist | 
| 67091 | 337 | (iso_tuple_snd_update isom \<circ> f) (g \<circ> iso_tuple_snd isom) | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 338 | (iso_tuple_cons isom a b) u (iso_tuple_cons isom a b') v" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 339 | proof - | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 340 | from assms have "f id = id" | 
| 35132 | 341 | by (auto simp add: iso_tuple_update_accessor_eq_assist_def | 
| 342 | intro: iso_tuple_update_accessor_cong_assist_id) | |
| 343 | with assms show ?thesis | |
| 344 | by (clarsimp simp: iso_tuple_update_accessor_eq_assist_def | |
| 345 | iso_tuple_snd_update_def iso_tuple_snd_def | |
| 346 | iso_tuple_update_accessor_cong_assist_def iso_tuple_cons_def simps) | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 347 | qed | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 348 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 349 | lemma iso_tuple_cons_conj_eqI: | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 350 | "a = c \<and> b = d \<and> P \<longleftrightarrow> Q \<Longrightarrow> | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 351 | iso_tuple_cons isom a b = iso_tuple_cons isom c d \<and> P \<longleftrightarrow> Q" | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 352 | by (clarsimp simp: iso_tuple_cons_def simps) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 353 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 354 | lemmas intros = | 
| 35132 | 355 | iso_tuple_access_update_fst_fst | 
| 356 | iso_tuple_access_update_snd_snd | |
| 357 | iso_tuple_access_update_fst_snd | |
| 358 | iso_tuple_access_update_snd_fst | |
| 359 | iso_tuple_update_swap_fst_fst | |
| 360 | iso_tuple_update_swap_snd_snd | |
| 361 | iso_tuple_update_swap_fst_snd | |
| 362 | iso_tuple_update_swap_snd_fst | |
| 363 | iso_tuple_update_compose_fst_fst | |
| 364 | iso_tuple_update_compose_snd_snd | |
| 365 | iso_tuple_surjective_proof_assist_step | |
| 366 | iso_tuple_fst_update_accessor_eq_assist | |
| 367 | iso_tuple_snd_update_accessor_eq_assist | |
| 368 | iso_tuple_fst_update_accessor_cong_assist | |
| 369 | iso_tuple_snd_update_accessor_cong_assist | |
| 370 | iso_tuple_cons_conj_eqI | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 371 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 372 | end | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 373 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 374 | lemma isomorphic_tuple_intro: | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 375 | fixes repr abst | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 376 | assumes repr_inj: "\<And>x y. repr x = repr y \<longleftrightarrow> x = y" | 
| 35132 | 377 | and abst_inv: "\<And>z. repr (abst z) = z" | 
| 378 | and v: "v \<equiv> Tuple_Isomorphism repr abst" | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 379 | shows "isomorphic_tuple v" | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 380 | proof | 
| 35132 | 381 | fix x have "repr (abst (repr x)) = repr x" | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 382 | by (simp add: abst_inv) | 
| 35132 | 383 | then show "Record.abst v (Record.repr v x) = x" | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 384 | by (simp add: v repr_inj) | 
| 35132 | 385 | next | 
| 386 | fix y | |
| 387 | show "Record.repr v (Record.abst v y) = y" | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 388 | by (simp add: v) (fact abst_inv) | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 389 | qed | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 390 | |
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 391 | definition | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 392 | "tuple_iso_tuple \<equiv> Tuple_Isomorphism id id" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 393 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 394 | lemma tuple_iso_tuple: | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 395 | "isomorphic_tuple tuple_iso_tuple" | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 396 | by (simp add: isomorphic_tuple_intro [OF _ _ reflexive] tuple_iso_tuple_def) | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 397 | |
| 35132 | 398 | lemma refl_conj_eq: "Q = R \<Longrightarrow> P \<and> Q \<longleftrightarrow> P \<and> R" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 399 | by simp | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 400 | |
| 54147 
97a8ff4e4ac9
killed most "no_atp", to make Sledgehammer more complete
 blanchet parents: 
51112diff
changeset | 401 | lemma iso_tuple_UNIV_I: "x \<in> UNIV \<equiv> True" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 402 | by simp | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 403 | |
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 404 | lemma iso_tuple_True_simp: "(True \<Longrightarrow> PROP P) \<equiv> PROP P" | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 405 | by simp | 
| 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 406 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14080diff
changeset | 407 | lemma prop_subst: "s = t \<Longrightarrow> PROP P t \<Longrightarrow> PROP P s" | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14080diff
changeset | 408 | by simp | 
| 11826 | 409 | |
| 35132 | 410 | lemma K_record_comp: "(\<lambda>x. c) \<circ> f = (\<lambda>x. c)" | 
| 25705 | 411 | by (simp add: comp_def) | 
| 11821 | 412 | |
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 413 | |
| 60758 | 414 | subsection \<open>Concrete record syntax\<close> | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 415 | |
| 41229 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
38539diff
changeset | 416 | nonterminal | 
| 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
38539diff
changeset | 417 | ident and | 
| 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
38539diff
changeset | 418 | field_type and | 
| 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
38539diff
changeset | 419 | field_types and | 
| 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
38539diff
changeset | 420 | field and | 
| 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
38539diff
changeset | 421 | fields and | 
| 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
38539diff
changeset | 422 | field_update and | 
| 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
38539diff
changeset | 423 | field_updates | 
| 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
38539diff
changeset | 424 | |
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 425 | syntax | 
| 11821 | 426 |   "_constify"           :: "id => ident"                        ("_")
 | 
| 427 |   "_constify"           :: "longid => ident"                    ("_")
 | |
| 5198 | 428 | |
| 35144 | 429 |   "_field_type"         :: "ident => type => field_type"        ("(2_ ::/ _)")
 | 
| 11821 | 430 |   ""                    :: "field_type => field_types"          ("_")
 | 
| 35144 | 431 |   "_field_types"        :: "field_type => field_types => field_types"    ("_,/ _")
 | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 432 |   "_record_type"        :: "field_types => type"                ("(3\<lparr>_\<rparr>)")
 | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 433 |   "_record_type_scheme" :: "field_types => type => type"        ("(3\<lparr>_,/ (2\<dots> ::/ _)\<rparr>)")
 | 
| 5198 | 434 | |
| 35144 | 435 |   "_field"              :: "ident => 'a => field"               ("(2_ =/ _)")
 | 
| 11821 | 436 |   ""                    :: "field => fields"                    ("_")
 | 
| 35144 | 437 |   "_fields"             :: "field => fields => fields"          ("_,/ _")
 | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 438 |   "_record"             :: "fields => 'a"                       ("(3\<lparr>_\<rparr>)")
 | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 439 |   "_record_scheme"      :: "fields => 'a => 'a"                 ("(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)")
 | 
| 5198 | 440 | |
| 35146 | 441 |   "_field_update"       :: "ident => 'a => field_update"        ("(2_ :=/ _)")
 | 
| 442 |   ""                    :: "field_update => field_updates"      ("_")
 | |
| 443 |   "_field_updates"      :: "field_update => field_updates => field_updates"  ("_,/ _")
 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 444 |   "_record_update"      :: "'a => field_updates => 'b"          ("_/(3\<lparr>_\<rparr>)" [900, 0] 900)
 | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 445 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 446 | syntax (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 447 |   "_record_type"        :: "field_types => type"                ("(3'(| _ |'))")
 | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 448 |   "_record_type_scheme" :: "field_types => type => type"        ("(3'(| _,/ (2... ::/ _) |'))")
 | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 449 |   "_record"             :: "fields => 'a"                       ("(3'(| _ |'))")
 | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 450 |   "_record_scheme"      :: "fields => 'a => 'a"                 ("(3'(| _,/ (2... =/ _) |'))")
 | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 451 |   "_record_update"      :: "'a => field_updates => 'b"          ("_/(3'(| _ |'))" [900, 0] 900)
 | 
| 9729 | 452 | |
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32745diff
changeset | 453 | |
| 60758 | 454 | subsection \<open>Record package\<close> | 
| 32752 
f65d74a264dd
Initial response to feedback from Norbert, Makarius on record patch
 tsewell@rubicon.NSW.bigpond.net.au parents: 
32745diff
changeset | 455 | |
| 69605 | 456 | ML_file \<open>Tools/record.ML\<close> | 
| 10641 | 457 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35146diff
changeset | 458 | hide_const (open) Tuple_Isomorphism repr abst iso_tuple_fst iso_tuple_snd | 
| 34151 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 459 | iso_tuple_fst_update iso_tuple_snd_update iso_tuple_cons | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 460 | iso_tuple_surjective_proof_assist iso_tuple_update_accessor_cong_assist | 
| 
8d57ce46b3f7
prefer prefix "iso" over potentially misleading "is"; tuned
 haftmann parents: 
33595diff
changeset | 461 | iso_tuple_update_accessor_eq_assist tuple_iso_tuple | 
| 33595 
7264824baf66
substantial simplification restores code generation
 haftmann parents: 
32799diff
changeset | 462 | |
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 463 | end |