| author | nipkow | 
| Wed, 10 Jan 2018 15:25:09 +0100 | |
| changeset 67399 | eab6ce8368fa | 
| parent 66251 | cd935b7cb3fb | 
| child 68782 | 8ff34c1ad580 | 
| permissions | -rw-r--r-- | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
1  | 
(* Title: HOL/Library/Mapping.thy  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
2  | 
Author: Florian Haftmann and Ondrej Kuncar  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
3  | 
*)  | 
| 29708 | 4  | 
|
| 60500 | 5  | 
section \<open>An abstract view on maps for code generation.\<close>  | 
| 29708 | 6  | 
|
7  | 
theory Mapping  | 
|
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
8  | 
imports Main  | 
| 29708 | 9  | 
begin  | 
10  | 
||
| 60500 | 11  | 
subsection \<open>Parametricity transfer rules\<close>  | 
| 51379 | 12  | 
|
| 63462 | 13  | 
lemma map_of_foldr: "map_of xs = foldr (\<lambda>(k, v) m. m(k \<mapsto> v)) xs Map.empty" (* FIXME move *)  | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
14  | 
using map_add_map_of_foldr [of Map.empty] by auto  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
15  | 
|
| 63343 | 16  | 
context includes lifting_syntax  | 
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
17  | 
begin  | 
| 56528 | 18  | 
|
| 63462 | 19  | 
lemma empty_parametric: "(A ===> rel_option B) Map.empty Map.empty"  | 
| 56528 | 20  | 
by transfer_prover  | 
| 51379 | 21  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
22  | 
lemma lookup_parametric: "((A ===> B) ===> A ===> B) (\<lambda>m k. m k) (\<lambda>m k. m k)"  | 
| 56528 | 23  | 
by transfer_prover  | 
| 51379 | 24  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
25  | 
lemma update_parametric:  | 
| 51379 | 26  | 
assumes [transfer_rule]: "bi_unique A"  | 
| 56528 | 27  | 
shows "(A ===> B ===> (A ===> rel_option B) ===> A ===> rel_option B)  | 
28  | 
(\<lambda>k v m. m(k \<mapsto> v)) (\<lambda>k v m. m(k \<mapsto> v))"  | 
|
29  | 
by transfer_prover  | 
|
| 51379 | 30  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
31  | 
lemma delete_parametric:  | 
| 51379 | 32  | 
assumes [transfer_rule]: "bi_unique A"  | 
| 63462 | 33  | 
shows "(A ===> (A ===> rel_option B) ===> A ===> rel_option B)  | 
| 56528 | 34  | 
(\<lambda>k m. m(k := None)) (\<lambda>k m. m(k := None))"  | 
35  | 
by transfer_prover  | 
|
| 51379 | 36  | 
|
| 56528 | 37  | 
lemma is_none_parametric [transfer_rule]:  | 
38  | 
"(rel_option A ===> HOL.eq) Option.is_none Option.is_none"  | 
|
| 61068 | 39  | 
by (auto simp add: Option.is_none_def rel_fun_def rel_option_iff split: option.split)  | 
| 51379 | 40  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
41  | 
lemma dom_parametric:  | 
| 51379 | 42  | 
assumes [transfer_rule]: "bi_total A"  | 
| 63462 | 43  | 
shows "((A ===> rel_option B) ===> rel_set A) dom dom"  | 
| 61068 | 44  | 
unfolding dom_def [abs_def] Option.is_none_def [symmetric] by transfer_prover  | 
| 51379 | 45  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
46  | 
lemma map_of_parametric [transfer_rule]:  | 
| 51379 | 47  | 
assumes [transfer_rule]: "bi_unique R1"  | 
| 55944 | 48  | 
shows "(list_all2 (rel_prod R1 R2) ===> R1 ===> rel_option R2) map_of map_of"  | 
| 56528 | 49  | 
unfolding map_of_def by transfer_prover  | 
| 51379 | 50  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
51  | 
lemma map_entry_parametric [transfer_rule]:  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
52  | 
assumes [transfer_rule]: "bi_unique A"  | 
| 63462 | 53  | 
shows "(A ===> (B ===> B) ===> (A ===> rel_option B) ===> A ===> rel_option B)  | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
54  | 
(\<lambda>k f m. (case m k of None \<Rightarrow> m  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
55  | 
| Some v \<Rightarrow> m (k \<mapsto> (f v)))) (\<lambda>k f m. (case m k of None \<Rightarrow> m  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
56  | 
| Some v \<Rightarrow> m (k \<mapsto> (f v))))"  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
57  | 
by transfer_prover  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
58  | 
|
| 63462 | 59  | 
lemma tabulate_parametric:  | 
| 51379 | 60  | 
assumes [transfer_rule]: "bi_unique A"  | 
| 63462 | 61  | 
shows "(list_all2 A ===> (A ===> B) ===> A ===> rel_option B)  | 
| 56528 | 62  | 
(\<lambda>ks f. (map_of (map (\<lambda>k. (k, f k)) ks))) (\<lambda>ks f. (map_of (map (\<lambda>k. (k, f k)) ks)))"  | 
63  | 
by transfer_prover  | 
|
| 51379 | 64  | 
|
| 63462 | 65  | 
lemma bulkload_parametric:  | 
66  | 
"(list_all2 A ===> HOL.eq ===> rel_option A)  | 
|
67  | 
(\<lambda>xs k. if k < length xs then Some (xs ! k) else None)  | 
|
68  | 
(\<lambda>xs k. if k < length xs then Some (xs ! k) else None)"  | 
|
| 56528 | 69  | 
proof  | 
70  | 
fix xs ys  | 
|
71  | 
assume "list_all2 A xs ys"  | 
|
| 63462 | 72  | 
then show  | 
73  | 
"(HOL.eq ===> rel_option A)  | 
|
74  | 
(\<lambda>k. if k < length xs then Some (xs ! k) else None)  | 
|
75  | 
(\<lambda>k. if k < length ys then Some (ys ! k) else None)"  | 
|
| 56528 | 76  | 
apply induct  | 
| 63476 | 77  | 
apply auto  | 
| 56528 | 78  | 
unfolding rel_fun_def  | 
| 63462 | 79  | 
apply clarsimp  | 
80  | 
apply (case_tac xa)  | 
|
| 63476 | 81  | 
apply (auto dest: list_all2_lengthD list_all2_nthD)  | 
| 56528 | 82  | 
done  | 
83  | 
qed  | 
|
| 51379 | 84  | 
|
| 63462 | 85  | 
lemma map_parametric:  | 
86  | 
"((A ===> B) ===> (C ===> D) ===> (B ===> rel_option C) ===> A ===> rel_option D)  | 
|
| 56528 | 87  | 
(\<lambda>f g m. (map_option g \<circ> m \<circ> f)) (\<lambda>f g m. (map_option g \<circ> m \<circ> f))"  | 
88  | 
by transfer_prover  | 
|
| 63462 | 89  | 
|
90  | 
lemma combine_with_key_parametric:  | 
|
91  | 
"((A ===> B ===> B ===> B) ===> (A ===> rel_option B) ===> (A ===> rel_option B) ===>  | 
|
92  | 
(A ===> rel_option B)) (\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x))  | 
|
93  | 
(\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x))"  | 
|
| 63194 | 94  | 
unfolding combine_options_def by transfer_prover  | 
| 63462 | 95  | 
|
96  | 
lemma combine_parametric:  | 
|
97  | 
"((B ===> B ===> B) ===> (A ===> rel_option B) ===> (A ===> rel_option B) ===>  | 
|
98  | 
(A ===> rel_option B)) (\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x))  | 
|
99  | 
(\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x))"  | 
|
| 63194 | 100  | 
unfolding combine_options_def by transfer_prover  | 
| 51379 | 101  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
102  | 
end  | 
| 51379 | 103  | 
|
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
104  | 
|
| 60500 | 105  | 
subsection \<open>Type definition and primitive operations\<close>  | 
| 29708 | 106  | 
|
| 49834 | 107  | 
typedef ('a, 'b) mapping = "UNIV :: ('a \<rightharpoonup> 'b) set"
 | 
| 63462 | 108  | 
morphisms rep Mapping ..  | 
| 
37700
 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 
haftmann 
parents: 
37299 
diff
changeset
 | 
109  | 
|
| 59485 | 110  | 
setup_lifting type_definition_mapping  | 
| 
37700
 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 
haftmann 
parents: 
37299 
diff
changeset
 | 
111  | 
|
| 56528 | 112  | 
lift_definition empty :: "('a, 'b) mapping"
 | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
113  | 
is Map.empty parametric empty_parametric .  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
114  | 
|
| 56528 | 115  | 
lift_definition lookup :: "('a, 'b) mapping \<Rightarrow> 'a \<Rightarrow> 'b option"
 | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
116  | 
is "\<lambda>m k. m k" parametric lookup_parametric .  | 
| 56528 | 117  | 
|
| 63194 | 118  | 
definition "lookup_default d m k = (case Mapping.lookup m k of None \<Rightarrow> d | Some v \<Rightarrow> v)"  | 
119  | 
||
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
63476 
diff
changeset
 | 
120  | 
lemma [code abstract]:  | 
| 
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
63476 
diff
changeset
 | 
121  | 
"lookup (Mapping f) = f"  | 
| 
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
63476 
diff
changeset
 | 
122  | 
by (fact Mapping.lookup.abs_eq) (* FIXME lifting *)  | 
| 59485 | 123  | 
|
| 56528 | 124  | 
lift_definition update :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
 | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
125  | 
is "\<lambda>k v m. m(k \<mapsto> v)" parametric update_parametric .  | 
| 
37700
 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 
haftmann 
parents: 
37299 
diff
changeset
 | 
126  | 
|
| 56528 | 127  | 
lift_definition delete :: "'a \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
 | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
128  | 
is "\<lambda>k m. m(k := None)" parametric delete_parametric .  | 
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
129  | 
|
| 63194 | 130  | 
lift_definition filter :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
 | 
| 63462 | 131  | 
is "\<lambda>P m k. case m k of None \<Rightarrow> None | Some v \<Rightarrow> if P k v then Some v else None" .  | 
| 63194 | 132  | 
|
| 56528 | 133  | 
lift_definition keys :: "('a, 'b) mapping \<Rightarrow> 'a set"
 | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
134  | 
is dom parametric dom_parametric .  | 
| 29708 | 135  | 
|
| 56528 | 136  | 
lift_definition tabulate :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping"
 | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
137  | 
is "\<lambda>ks f. (map_of (List.map (\<lambda>k. (k, f k)) ks))" parametric tabulate_parametric .  | 
| 29708 | 138  | 
|
| 56528 | 139  | 
lift_definition bulkload :: "'a list \<Rightarrow> (nat, 'a) mapping"  | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
140  | 
is "\<lambda>xs k. if k < length xs then Some (xs ! k) else None" parametric bulkload_parametric .  | 
| 29708 | 141  | 
|
| 56528 | 142  | 
lift_definition map :: "('c \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('c, 'd) mapping"
 | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
143  | 
is "\<lambda>f g m. (map_option g \<circ> m \<circ> f)" parametric map_parametric .  | 
| 63462 | 144  | 
|
| 63194 | 145  | 
lift_definition map_values :: "('c \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> ('c, 'a) mapping \<Rightarrow> ('c, 'b) mapping"
 | 
| 63462 | 146  | 
is "\<lambda>f m x. map_option (f x) (m x)" .  | 
| 63194 | 147  | 
|
| 63462 | 148  | 
lift_definition combine_with_key ::  | 
| 63194 | 149  | 
  "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping"
 | 
150  | 
is "\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x)" parametric combine_with_key_parametric .  | 
|
151  | 
||
| 63462 | 152  | 
lift_definition combine ::  | 
| 63194 | 153  | 
  "('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping"
 | 
154  | 
is "\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x)" parametric combine_parametric .  | 
|
155  | 
||
| 63462 | 156  | 
definition "All_mapping m P \<longleftrightarrow>  | 
157  | 
(\<forall>x. case Mapping.lookup m x of None \<Rightarrow> True | Some y \<Rightarrow> P x y)"  | 
|
| 29708 | 158  | 
|
| 59485 | 159  | 
declare [[code drop: map]]  | 
160  | 
||
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
161  | 
|
| 60500 | 162  | 
subsection \<open>Functorial structure\<close>  | 
| 40605 | 163  | 
|
| 
55467
 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 
blanchet 
parents: 
55466 
diff
changeset
 | 
164  | 
functor map: map  | 
| 55466 | 165  | 
by (transfer, auto simp add: fun_eq_iff option.map_comp option.map_id)+  | 
| 40605 | 166  | 
|
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
167  | 
|
| 60500 | 168  | 
subsection \<open>Derived operations\<close>  | 
| 29708 | 169  | 
|
| 61076 | 170  | 
definition ordered_keys :: "('a::linorder, 'b) mapping \<Rightarrow> 'a list"
 | 
| 63462 | 171  | 
where "ordered_keys m = (if finite (keys m) then sorted_list_of_set (keys m) else [])"  | 
| 35194 | 172  | 
|
| 56528 | 173  | 
definition is_empty :: "('a, 'b) mapping \<Rightarrow> bool"
 | 
| 63462 | 174  | 
  where "is_empty m \<longleftrightarrow> keys m = {}"
 | 
| 35157 | 175  | 
|
| 56528 | 176  | 
definition size :: "('a, 'b) mapping \<Rightarrow> nat"
 | 
| 63462 | 177  | 
where "size m = (if finite (keys m) then card (keys m) else 0)"  | 
| 35157 | 178  | 
|
| 56528 | 179  | 
definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
 | 
| 63462 | 180  | 
where "replace k v m = (if k \<in> keys m then update k v m else m)"  | 
| 29814 | 181  | 
|
| 56528 | 182  | 
definition default :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
 | 
| 63462 | 183  | 
where "default k v m = (if k \<in> keys m then m else update k v m)"  | 
| 
37026
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
184  | 
|
| 60500 | 185  | 
text \<open>Manual derivation of transfer rule is non-trivial\<close>  | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
186  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
187  | 
lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is
 | 
| 63462 | 188  | 
"\<lambda>k f m.  | 
189  | 
(case m k of  | 
|
190  | 
None \<Rightarrow> m  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
191  | 
| Some v \<Rightarrow> m (k \<mapsto> (f v)))" parametric map_entry_parametric .  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
192  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
193  | 
lemma map_entry_code [code]:  | 
| 63462 | 194  | 
"map_entry k f m =  | 
195  | 
(case lookup m k of  | 
|
196  | 
None \<Rightarrow> m  | 
|
| 
49975
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
49973 
diff
changeset
 | 
197  | 
| Some v \<Rightarrow> update k (f v) m)"  | 
| 
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
49973 
diff
changeset
 | 
198  | 
by transfer rule  | 
| 
37026
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
199  | 
|
| 56528 | 200  | 
definition map_default :: "'a \<Rightarrow> 'b \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
 | 
| 63462 | 201  | 
where "map_default k v f m = map_entry k f (default k v m)"  | 
| 
37026
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
202  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
203  | 
definition of_alist :: "('k \<times> 'v) list \<Rightarrow> ('k, 'v) mapping"
 | 
| 63462 | 204  | 
where "of_alist xs = foldr (\<lambda>(k, v) m. update k v m) xs empty"  | 
| 51379 | 205  | 
|
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
206  | 
instantiation mapping :: (type, type) equal  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
207  | 
begin  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
208  | 
|
| 63462 | 209  | 
definition "HOL.equal m1 m2 \<longleftrightarrow> (\<forall>k. lookup m1 k = lookup m2 k)"  | 
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
210  | 
|
| 60679 | 211  | 
instance  | 
| 63462 | 212  | 
apply standard  | 
213  | 
unfolding equal_mapping_def  | 
|
214  | 
apply transfer  | 
|
215  | 
apply auto  | 
|
216  | 
done  | 
|
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
217  | 
|
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
218  | 
end  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
219  | 
|
| 63343 | 220  | 
context includes lifting_syntax  | 
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
221  | 
begin  | 
| 56528 | 222  | 
|
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
223  | 
lemma [transfer_rule]:  | 
| 51379 | 224  | 
assumes [transfer_rule]: "bi_total A"  | 
| 63462 | 225  | 
and [transfer_rule]: "bi_unique B"  | 
| 67399 | 226  | 
shows "(pcr_mapping A B ===> pcr_mapping A B ===> (=)) HOL.eq HOL.equal"  | 
| 63462 | 227  | 
unfolding equal by transfer_prover  | 
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
228  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
229  | 
lemma of_alist_transfer [transfer_rule]:  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
230  | 
assumes [transfer_rule]: "bi_unique R1"  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
231  | 
shows "(list_all2 (rel_prod R1 R2) ===> pcr_mapping R1 R2) map_of of_alist"  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
232  | 
unfolding of_alist_def [abs_def] map_of_foldr [abs_def] by transfer_prover  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
233  | 
|
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
234  | 
end  | 
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
235  | 
|
| 56528 | 236  | 
|
| 60500 | 237  | 
subsection \<open>Properties\<close>  | 
| 29708 | 238  | 
|
| 63462 | 239  | 
lemma mapping_eqI: "(\<And>x. lookup m x = lookup m' x) \<Longrightarrow> m = m'"  | 
| 63195 | 240  | 
by transfer (simp add: fun_eq_iff)  | 
241  | 
||
| 63462 | 242  | 
lemma mapping_eqI':  | 
243  | 
assumes "\<And>x. x \<in> Mapping.keys m \<Longrightarrow> Mapping.lookup_default d m x = Mapping.lookup_default d m' x"  | 
|
244  | 
and "Mapping.keys m = Mapping.keys m'"  | 
|
245  | 
shows "m = m'"  | 
|
| 63195 | 246  | 
proof (intro mapping_eqI)  | 
| 63462 | 247  | 
show "Mapping.lookup m x = Mapping.lookup m' x" for x  | 
| 63195 | 248  | 
proof (cases "Mapping.lookup m x")  | 
249  | 
case None  | 
|
| 63462 | 250  | 
then have "x \<notin> Mapping.keys m"  | 
251  | 
by transfer (simp add: dom_def)  | 
|
252  | 
then have "x \<notin> Mapping.keys m'"  | 
|
253  | 
by (simp add: assms)  | 
|
254  | 
then have "Mapping.lookup m' x = None"  | 
|
255  | 
by transfer (simp add: dom_def)  | 
|
256  | 
with None show ?thesis  | 
|
257  | 
by simp  | 
|
| 63195 | 258  | 
next  | 
259  | 
case (Some y)  | 
|
| 63462 | 260  | 
then have A: "x \<in> Mapping.keys m"  | 
261  | 
by transfer (simp add: dom_def)  | 
|
262  | 
then have "x \<in> Mapping.keys m'"  | 
|
263  | 
by (simp add: assms)  | 
|
264  | 
then have "\<exists>y'. Mapping.lookup m' x = Some y'"  | 
|
265  | 
by transfer (simp add: dom_def)  | 
|
266  | 
with Some assms(1)[OF A] show ?thesis  | 
|
267  | 
by (auto simp add: lookup_default_def)  | 
|
| 63195 | 268  | 
qed  | 
269  | 
qed  | 
|
270  | 
||
| 63462 | 271  | 
lemma lookup_update: "lookup (update k v m) k = Some v"  | 
| 49973 | 272  | 
by transfer simp  | 
273  | 
||
| 63462 | 274  | 
lemma lookup_update_neq: "k \<noteq> k' \<Longrightarrow> lookup (update k v m) k' = lookup m k'"  | 
| 49973 | 275  | 
by transfer simp  | 
276  | 
||
| 63462 | 277  | 
lemma lookup_update': "Mapping.lookup (update k v m) k' = (if k = k' then Some v else lookup m k')"  | 
| 63194 | 278  | 
by (auto simp: lookup_update lookup_update_neq)  | 
279  | 
||
| 63462 | 280  | 
lemma lookup_empty: "lookup empty k = None"  | 
| 49973 | 281  | 
by transfer simp  | 
282  | 
||
| 63194 | 283  | 
lemma lookup_filter:  | 
| 63462 | 284  | 
"lookup (filter P m) k =  | 
285  | 
(case lookup m k of  | 
|
286  | 
None \<Rightarrow> None  | 
|
287  | 
| Some v \<Rightarrow> if P k v then Some v else None)"  | 
|
| 63194 | 288  | 
by transfer simp_all  | 
289  | 
||
| 63462 | 290  | 
lemma lookup_map_values: "lookup (map_values f m) k = map_option (f k) (lookup m k)"  | 
| 63194 | 291  | 
by transfer simp_all  | 
292  | 
||
293  | 
lemma lookup_default_empty: "lookup_default d empty k = d"  | 
|
294  | 
by (simp add: lookup_default_def lookup_empty)  | 
|
295  | 
||
| 63462 | 296  | 
lemma lookup_default_update: "lookup_default d (update k v m) k = v"  | 
| 63194 | 297  | 
by (simp add: lookup_default_def lookup_update)  | 
298  | 
||
299  | 
lemma lookup_default_update_neq:  | 
|
| 63462 | 300  | 
"k \<noteq> k' \<Longrightarrow> lookup_default d (update k v m) k' = lookup_default d m k'"  | 
| 63194 | 301  | 
by (simp add: lookup_default_def lookup_update_neq)  | 
302  | 
||
| 63462 | 303  | 
lemma lookup_default_update':  | 
| 63194 | 304  | 
"lookup_default d (update k v m) k' = (if k = k' then v else lookup_default d m k')"  | 
305  | 
by (auto simp: lookup_default_update lookup_default_update_neq)  | 
|
306  | 
||
307  | 
lemma lookup_default_filter:  | 
|
| 63462 | 308  | 
"lookup_default d (filter P m) k =  | 
| 63194 | 309  | 
(if P k (lookup_default d m k) then lookup_default d m k else d)"  | 
310  | 
by (simp add: lookup_default_def lookup_filter split: option.splits)  | 
|
311  | 
||
312  | 
lemma lookup_default_map_values:  | 
|
313  | 
"lookup_default (f k d) (map_values f m) k = f k (lookup_default d m k)"  | 
|
| 63462 | 314  | 
by (simp add: lookup_default_def lookup_map_values split: option.splits)  | 
| 63194 | 315  | 
|
316  | 
lemma lookup_combine_with_key:  | 
|
| 63462 | 317  | 
"Mapping.lookup (combine_with_key f m1 m2) x =  | 
318  | 
combine_options (f x) (Mapping.lookup m1 x) (Mapping.lookup m2 x)"  | 
|
| 63194 | 319  | 
by transfer (auto split: option.splits)  | 
| 63462 | 320  | 
|
| 63194 | 321  | 
lemma combine_altdef: "combine f m1 m2 = combine_with_key (\<lambda>_. f) m1 m2"  | 
322  | 
by transfer' (rule refl)  | 
|
323  | 
||
324  | 
lemma lookup_combine:  | 
|
| 63462 | 325  | 
"Mapping.lookup (combine f m1 m2) x =  | 
| 63194 | 326  | 
combine_options f (Mapping.lookup m1 x) (Mapping.lookup m2 x)"  | 
327  | 
by transfer (auto split: option.splits)  | 
|
| 63462 | 328  | 
|
329  | 
lemma lookup_default_neutral_combine_with_key:  | 
|
| 63194 | 330  | 
assumes "\<And>x. f k d x = x" "\<And>x. f k x d = x"  | 
| 63462 | 331  | 
shows "Mapping.lookup_default d (combine_with_key f m1 m2) k =  | 
332  | 
f k (Mapping.lookup_default d m1 k) (Mapping.lookup_default d m2 k)"  | 
|
| 63194 | 333  | 
by (auto simp: lookup_default_def lookup_combine_with_key assms split: option.splits)  | 
| 63462 | 334  | 
|
335  | 
lemma lookup_default_neutral_combine:  | 
|
| 63194 | 336  | 
assumes "\<And>x. f d x = x" "\<And>x. f x d = x"  | 
| 63462 | 337  | 
shows "Mapping.lookup_default d (combine f m1 m2) x =  | 
338  | 
f (Mapping.lookup_default d m1 x) (Mapping.lookup_default d m2 x)"  | 
|
| 63194 | 339  | 
by (auto simp: lookup_default_def lookup_combine assms split: option.splits)  | 
340  | 
||
| 63462 | 341  | 
lemma lookup_map_entry: "lookup (map_entry x f m) x = map_option f (lookup m x)"  | 
| 63195 | 342  | 
by transfer (auto split: option.splits)  | 
343  | 
||
| 63462 | 344  | 
lemma lookup_map_entry_neq: "x \<noteq> y \<Longrightarrow> lookup (map_entry x f m) y = lookup m y"  | 
| 63195 | 345  | 
by transfer (auto split: option.splits)  | 
346  | 
||
347  | 
lemma lookup_map_entry':  | 
|
| 63462 | 348  | 
"lookup (map_entry x f m) y =  | 
| 63195 | 349  | 
(if x = y then map_option f (lookup m y) else lookup m y)"  | 
350  | 
by transfer (auto split: option.splits)  | 
|
351  | 
||
| 63462 | 352  | 
lemma lookup_default: "lookup (default x d m) x = Some (lookup_default d m x)"  | 
353  | 
unfolding lookup_default_def default_def  | 
|
354  | 
by transfer (auto split: option.splits)  | 
|
355  | 
||
356  | 
lemma lookup_default_neq: "x \<noteq> y \<Longrightarrow> lookup (default x d m) y = lookup m y"  | 
|
357  | 
unfolding lookup_default_def default_def  | 
|
358  | 
by transfer (auto split: option.splits)  | 
|
| 63195 | 359  | 
|
360  | 
lemma lookup_default':  | 
|
| 63462 | 361  | 
"lookup (default x d m) y =  | 
362  | 
(if x = y then Some (lookup_default d m x) else lookup m y)"  | 
|
| 63195 | 363  | 
unfolding lookup_default_def default_def  | 
364  | 
by transfer (auto split: option.splits)  | 
|
365  | 
||
| 63462 | 366  | 
lemma lookup_map_default: "lookup (map_default x d f m) x = Some (f (lookup_default d m x))"  | 
367  | 
unfolding lookup_default_def default_def  | 
|
368  | 
by (simp add: map_default_def lookup_map_entry lookup_default lookup_default_def)  | 
|
369  | 
||
370  | 
lemma lookup_map_default_neq: "x \<noteq> y \<Longrightarrow> lookup (map_default x d f m) y = lookup m y"  | 
|
371  | 
unfolding lookup_default_def default_def  | 
|
372  | 
by (simp add: map_default_def lookup_map_entry_neq lookup_default_neq)  | 
|
| 63195 | 373  | 
|
374  | 
lemma lookup_map_default':  | 
|
| 63462 | 375  | 
"lookup (map_default x d f m) y =  | 
376  | 
(if x = y then Some (f (lookup_default d m x)) else lookup m y)"  | 
|
377  | 
unfolding lookup_default_def default_def  | 
|
378  | 
by (simp add: map_default_def lookup_map_entry' lookup_default' lookup_default_def)  | 
|
| 63195 | 379  | 
|
| 63462 | 380  | 
lemma lookup_tabulate:  | 
| 63194 | 381  | 
assumes "distinct xs"  | 
| 63462 | 382  | 
shows "Mapping.lookup (Mapping.tabulate xs f) x = (if x \<in> set xs then Some (f x) else None)"  | 
| 63194 | 383  | 
using assms by transfer (auto simp: map_of_eq_None_iff o_def dest!: map_of_SomeD)  | 
384  | 
||
385  | 
lemma lookup_of_alist: "Mapping.lookup (Mapping.of_alist xs) k = map_of xs k"  | 
|
386  | 
by transfer simp_all  | 
|
387  | 
||
| 63462 | 388  | 
lemma keys_is_none_rep [code_unfold]: "k \<in> keys m \<longleftrightarrow> \<not> (Option.is_none (lookup m k))"  | 
| 61068 | 389  | 
by transfer (auto simp add: Option.is_none_def)  | 
| 29708 | 390  | 
|
391  | 
lemma update_update:  | 
|
392  | 
"update k v (update k w m) = update k v m"  | 
|
393  | 
"k \<noteq> l \<Longrightarrow> update k v (update l w m) = update l w (update k v m)"  | 
|
| 63462 | 394  | 
by (transfer; simp add: fun_upd_twist)+  | 
| 29708 | 395  | 
|
| 63462 | 396  | 
lemma update_delete [simp]: "update k v (delete k m) = update k v m"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
397  | 
by transfer simp  | 
| 29708 | 398  | 
|
399  | 
lemma delete_update:  | 
|
400  | 
"delete k (update k v m) = delete k m"  | 
|
401  | 
"k \<noteq> l \<Longrightarrow> delete k (update l v m) = update l v (delete k m)"  | 
|
| 63462 | 402  | 
by (transfer; simp add: fun_upd_twist)+  | 
| 29708 | 403  | 
|
| 63462 | 404  | 
lemma delete_empty [simp]: "delete k empty = empty"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
405  | 
by transfer simp  | 
| 29708 | 406  | 
|
| 35157 | 407  | 
lemma replace_update:  | 
| 37052 | 408  | 
"k \<notin> keys m \<Longrightarrow> replace k v m = m"  | 
409  | 
"k \<in> keys m \<Longrightarrow> replace k v m = update k v m"  | 
|
| 63462 | 410  | 
by (transfer; auto simp add: replace_def fun_upd_twist)+  | 
411  | 
||
| 63194 | 412  | 
lemma map_values_update: "map_values f (update k v m) = update k (f k v) (map_values f m)"  | 
413  | 
by transfer (simp_all add: fun_eq_iff)  | 
|
| 63462 | 414  | 
|
415  | 
lemma size_mono: "finite (keys m') \<Longrightarrow> keys m \<subseteq> keys m' \<Longrightarrow> size m \<le> size m'"  | 
|
| 63194 | 416  | 
unfolding size_def by (auto intro: card_mono)  | 
| 29708 | 417  | 
|
| 63462 | 418  | 
lemma size_empty [simp]: "size empty = 0"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
419  | 
unfolding size_def by transfer simp  | 
| 29708 | 420  | 
|
421  | 
lemma size_update:  | 
|
| 37052 | 422  | 
"finite (keys m) \<Longrightarrow> size (update k v m) =  | 
423  | 
(if k \<in> keys m then size m else Suc (size m))"  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
424  | 
unfolding size_def by transfer (auto simp add: insert_dom)  | 
| 29708 | 425  | 
|
| 63462 | 426  | 
lemma size_delete: "size (delete k m) = (if k \<in> keys m then size m - 1 else size m)"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
427  | 
unfolding size_def by transfer simp  | 
| 29708 | 428  | 
|
| 63462 | 429  | 
lemma size_tabulate [simp]: "size (tabulate ks f) = length (remdups ks)"  | 
430  | 
unfolding size_def by transfer (auto simp add: map_of_map_restrict card_set comp_def)  | 
|
| 29708 | 431  | 
|
| 63194 | 432  | 
lemma keys_filter: "keys (filter P m) \<subseteq> keys m"  | 
433  | 
by transfer (auto split: option.splits)  | 
|
434  | 
||
435  | 
lemma size_filter: "finite (keys m) \<Longrightarrow> size (filter P m) \<le> size m"  | 
|
436  | 
by (intro size_mono keys_filter)  | 
|
437  | 
||
| 63462 | 438  | 
lemma bulkload_tabulate: "bulkload xs = tabulate [0..<length xs] (nth xs)"  | 
| 56528 | 439  | 
by transfer (auto simp add: map_of_map_restrict)  | 
| 29826 | 440  | 
|
| 63462 | 441  | 
lemma is_empty_empty [simp]: "is_empty empty"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
442  | 
unfolding is_empty_def by transfer simp  | 
| 37052 | 443  | 
|
| 63462 | 444  | 
lemma is_empty_update [simp]: "\<not> is_empty (update k v m)"  | 
445  | 
unfolding is_empty_def by transfer simp  | 
|
446  | 
||
447  | 
lemma is_empty_delete: "is_empty (delete k m) \<longleftrightarrow> is_empty m \<or> keys m = {k}"
 | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
448  | 
unfolding is_empty_def by transfer (auto simp del: dom_eq_empty_conv)  | 
| 37052 | 449  | 
|
| 63462 | 450  | 
lemma is_empty_replace [simp]: "is_empty (replace k v m) \<longleftrightarrow> is_empty m"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
451  | 
unfolding is_empty_def replace_def by transfer auto  | 
| 37052 | 452  | 
|
| 63462 | 453  | 
lemma is_empty_default [simp]: "\<not> is_empty (default k v m)"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
454  | 
unfolding is_empty_def default_def by transfer auto  | 
| 37052 | 455  | 
|
| 63462 | 456  | 
lemma is_empty_map_entry [simp]: "is_empty (map_entry k f m) \<longleftrightarrow> is_empty m"  | 
| 56528 | 457  | 
unfolding is_empty_def by transfer (auto split: option.split)  | 
| 37052 | 458  | 
|
| 63462 | 459  | 
lemma is_empty_map_values [simp]: "is_empty (map_values f m) \<longleftrightarrow> is_empty m"  | 
| 63194 | 460  | 
unfolding is_empty_def by transfer (auto simp: fun_eq_iff)  | 
461  | 
||
| 63462 | 462  | 
lemma is_empty_map_default [simp]: "\<not> is_empty (map_default k v f m)"  | 
| 37052 | 463  | 
by (simp add: map_default_def)  | 
464  | 
||
| 63462 | 465  | 
lemma keys_dom_lookup: "keys m = dom (Mapping.lookup m)"  | 
| 56545 | 466  | 
by transfer rule  | 
467  | 
||
| 63462 | 468  | 
lemma keys_empty [simp]: "keys empty = {}"
 | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
469  | 
by transfer simp  | 
| 37052 | 470  | 
|
| 63462 | 471  | 
lemma keys_update [simp]: "keys (update k v m) = insert k (keys m)"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
472  | 
by transfer simp  | 
| 37052 | 473  | 
|
| 63462 | 474  | 
lemma keys_delete [simp]: "keys (delete k m) = keys m - {k}"
 | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
475  | 
by transfer simp  | 
| 37052 | 476  | 
|
| 63462 | 477  | 
lemma keys_replace [simp]: "keys (replace k v m) = keys m"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
478  | 
unfolding replace_def by transfer (simp add: insert_absorb)  | 
| 37052 | 479  | 
|
| 63462 | 480  | 
lemma keys_default [simp]: "keys (default k v m) = insert k (keys m)"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
481  | 
unfolding default_def by transfer (simp add: insert_absorb)  | 
| 37052 | 482  | 
|
| 63462 | 483  | 
lemma keys_map_entry [simp]: "keys (map_entry k f m) = keys m"  | 
| 56528 | 484  | 
by transfer (auto split: option.split)  | 
| 37052 | 485  | 
|
| 63462 | 486  | 
lemma keys_map_default [simp]: "keys (map_default k v f m) = insert k (keys m)"  | 
| 37052 | 487  | 
by (simp add: map_default_def)  | 
488  | 
||
| 63462 | 489  | 
lemma keys_map_values [simp]: "keys (map_values f m) = keys m"  | 
| 63194 | 490  | 
by transfer (simp_all add: dom_def)  | 
491  | 
||
| 63462 | 492  | 
lemma keys_combine_with_key [simp]:  | 
| 63194 | 493  | 
"Mapping.keys (combine_with_key f m1 m2) = Mapping.keys m1 \<union> Mapping.keys m2"  | 
| 63462 | 494  | 
by transfer (auto simp: dom_def combine_options_def split: option.splits)  | 
| 63194 | 495  | 
|
496  | 
lemma keys_combine [simp]: "Mapping.keys (combine f m1 m2) = Mapping.keys m1 \<union> Mapping.keys m2"  | 
|
497  | 
by (simp add: combine_altdef)  | 
|
498  | 
||
| 63462 | 499  | 
lemma keys_tabulate [simp]: "keys (tabulate ks f) = set ks"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
500  | 
by transfer (simp add: map_of_map_restrict o_def)  | 
| 
37026
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
501  | 
|
| 63194 | 502  | 
lemma keys_of_alist [simp]: "keys (of_alist xs) = set (List.map fst xs)"  | 
503  | 
by transfer (simp_all add: dom_map_of_conv_image_fst)  | 
|
504  | 
||
| 63462 | 505  | 
lemma keys_bulkload [simp]: "keys (bulkload xs) = {0..<length xs}"
 | 
| 56528 | 506  | 
by (simp add: bulkload_tabulate)  | 
| 
37026
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
507  | 
|
| 63462 | 508  | 
lemma distinct_ordered_keys [simp]: "distinct (ordered_keys m)"  | 
| 37052 | 509  | 
by (simp add: ordered_keys_def)  | 
510  | 
||
| 63462 | 511  | 
lemma ordered_keys_infinite [simp]: "\<not> finite (keys m) \<Longrightarrow> ordered_keys m = []"  | 
| 37052 | 512  | 
by (simp add: ordered_keys_def)  | 
513  | 
||
| 63462 | 514  | 
lemma ordered_keys_empty [simp]: "ordered_keys empty = []"  | 
| 37052 | 515  | 
by (simp add: ordered_keys_def)  | 
516  | 
||
517  | 
lemma ordered_keys_update [simp]:  | 
|
518  | 
"k \<in> keys m \<Longrightarrow> ordered_keys (update k v m) = ordered_keys m"  | 
|
| 63462 | 519  | 
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow>  | 
520  | 
ordered_keys (update k v m) = insort k (ordered_keys m)"  | 
|
521  | 
by (simp_all add: ordered_keys_def)  | 
|
522  | 
(auto simp only: sorted_list_of_set_insert [symmetric] insert_absorb)  | 
|
| 37052 | 523  | 
|
| 63462 | 524  | 
lemma ordered_keys_delete [simp]: "ordered_keys (delete k m) = remove1 k (ordered_keys m)"  | 
| 37052 | 525  | 
proof (cases "finite (keys m)")  | 
| 63462 | 526  | 
case False  | 
527  | 
then show ?thesis by simp  | 
|
| 37052 | 528  | 
next  | 
| 63462 | 529  | 
case fin: True  | 
| 37052 | 530  | 
show ?thesis  | 
531  | 
proof (cases "k \<in> keys m")  | 
|
| 63462 | 532  | 
case False  | 
533  | 
with fin have "k \<notin> set (sorted_list_of_set (keys m))"  | 
|
534  | 
by simp  | 
|
535  | 
with False show ?thesis  | 
|
536  | 
by (simp add: ordered_keys_def remove1_idem)  | 
|
| 37052 | 537  | 
next  | 
| 63462 | 538  | 
case True  | 
539  | 
with fin show ?thesis  | 
|
540  | 
by (simp add: ordered_keys_def sorted_list_of_set_remove)  | 
|
| 37052 | 541  | 
qed  | 
542  | 
qed  | 
|
543  | 
||
| 63462 | 544  | 
lemma ordered_keys_replace [simp]: "ordered_keys (replace k v m) = ordered_keys m"  | 
| 37052 | 545  | 
by (simp add: replace_def)  | 
546  | 
||
547  | 
lemma ordered_keys_default [simp]:  | 
|
548  | 
"k \<in> keys m \<Longrightarrow> ordered_keys (default k v m) = ordered_keys m"  | 
|
549  | 
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (default k v m) = insort k (ordered_keys m)"  | 
|
550  | 
by (simp_all add: default_def)  | 
|
551  | 
||
| 63462 | 552  | 
lemma ordered_keys_map_entry [simp]: "ordered_keys (map_entry k f m) = ordered_keys m"  | 
| 37052 | 553  | 
by (simp add: ordered_keys_def)  | 
554  | 
||
555  | 
lemma ordered_keys_map_default [simp]:  | 
|
556  | 
"k \<in> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = ordered_keys m"  | 
|
557  | 
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = insort k (ordered_keys m)"  | 
|
558  | 
by (simp_all add: map_default_def)  | 
|
559  | 
||
| 63462 | 560  | 
lemma ordered_keys_tabulate [simp]: "ordered_keys (tabulate ks f) = sort (remdups ks)"  | 
| 37052 | 561  | 
by (simp add: ordered_keys_def sorted_list_of_set_sort_remdups)  | 
562  | 
||
| 63462 | 563  | 
lemma ordered_keys_bulkload [simp]: "ordered_keys (bulkload ks) = [0..<length ks]"  | 
| 37052 | 564  | 
by (simp add: ordered_keys_def)  | 
| 36110 | 565  | 
|
| 63462 | 566  | 
lemma tabulate_fold: "tabulate xs f = fold (\<lambda>k m. update k (f k) m) xs empty"  | 
| 56528 | 567  | 
proof transfer  | 
568  | 
fix f :: "'a \<Rightarrow> 'b" and xs  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
569  | 
have "map_of (List.map (\<lambda>k. (k, f k)) xs) = foldr (\<lambda>k m. m(k \<mapsto> f k)) xs Map.empty"  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
570  | 
by (simp add: foldr_map comp_def map_of_foldr)  | 
| 56528 | 571  | 
also have "foldr (\<lambda>k m. m(k \<mapsto> f k)) xs = fold (\<lambda>k m. m(k \<mapsto> f k)) xs"  | 
572  | 
by (rule foldr_fold) (simp add: fun_eq_iff)  | 
|
573  | 
ultimately show "map_of (List.map (\<lambda>k. (k, f k)) xs) = fold (\<lambda>k m. m(k \<mapsto> f k)) xs Map.empty"  | 
|
574  | 
by simp  | 
|
575  | 
qed  | 
|
576  | 
||
| 63194 | 577  | 
lemma All_mapping_mono:  | 
578  | 
"(\<And>k v. k \<in> keys m \<Longrightarrow> P k v \<Longrightarrow> Q k v) \<Longrightarrow> All_mapping m P \<Longrightarrow> All_mapping m Q"  | 
|
579  | 
unfolding All_mapping_def by transfer (auto simp: All_mapping_def dom_def split: option.splits)  | 
|
| 31459 | 580  | 
|
| 63194 | 581  | 
lemma All_mapping_empty [simp]: "All_mapping Mapping.empty P"  | 
582  | 
by (auto simp: All_mapping_def lookup_empty)  | 
|
| 63462 | 583  | 
|
584  | 
lemma All_mapping_update_iff:  | 
|
| 63194 | 585  | 
"All_mapping (Mapping.update k v m) P \<longleftrightarrow> P k v \<and> All_mapping m (\<lambda>k' v'. k = k' \<or> P k' v')"  | 
| 63462 | 586  | 
unfolding All_mapping_def  | 
| 63194 | 587  | 
proof safe  | 
588  | 
assume "\<forall>x. case Mapping.lookup (Mapping.update k v m) x of None \<Rightarrow> True | Some y \<Rightarrow> P x y"  | 
|
| 63462 | 589  | 
then have *: "case Mapping.lookup (Mapping.update k v m) x of None \<Rightarrow> True | Some y \<Rightarrow> P x y" for x  | 
| 63194 | 590  | 
by blast  | 
| 63462 | 591  | 
from *[of k] show "P k v"  | 
592  | 
by (simp add: lookup_update)  | 
|
| 63194 | 593  | 
show "case Mapping.lookup m x of None \<Rightarrow> True | Some v' \<Rightarrow> k = x \<or> P x v'" for x  | 
| 63462 | 594  | 
using *[of x] by (auto simp add: lookup_update' split: if_splits option.splits)  | 
| 63194 | 595  | 
next  | 
596  | 
assume "P k v"  | 
|
597  | 
assume "\<forall>x. case Mapping.lookup m x of None \<Rightarrow> True | Some v' \<Rightarrow> k = x \<or> P x v'"  | 
|
| 63462 | 598  | 
then have A: "case Mapping.lookup m x of None \<Rightarrow> True | Some v' \<Rightarrow> k = x \<or> P x v'" for x  | 
599  | 
by blast  | 
|
| 63194 | 600  | 
show "case Mapping.lookup (Mapping.update k v m) x of None \<Rightarrow> True | Some xa \<Rightarrow> P x xa" for x  | 
601  | 
using \<open>P k v\<close> A[of x] by (auto simp: lookup_update' split: option.splits)  | 
|
602  | 
qed  | 
|
603  | 
||
604  | 
lemma All_mapping_update:  | 
|
605  | 
"P k v \<Longrightarrow> All_mapping m (\<lambda>k' v'. k = k' \<or> P k' v') \<Longrightarrow> All_mapping (Mapping.update k v m) P"  | 
|
606  | 
by (simp add: All_mapping_update_iff)  | 
|
607  | 
||
| 63462 | 608  | 
lemma All_mapping_filter_iff: "All_mapping (filter P m) Q \<longleftrightarrow> All_mapping m (\<lambda>k v. P k v \<longrightarrow> Q k v)"  | 
| 63194 | 609  | 
by (auto simp: All_mapping_def lookup_filter split: option.splits)  | 
610  | 
||
| 63462 | 611  | 
lemma All_mapping_filter: "All_mapping m Q \<Longrightarrow> All_mapping (filter P m) Q"  | 
| 63194 | 612  | 
by (auto simp: All_mapping_filter_iff intro: All_mapping_mono)  | 
| 31459 | 613  | 
|
| 63462 | 614  | 
lemma All_mapping_map_values: "All_mapping (map_values f m) P \<longleftrightarrow> All_mapping m (\<lambda>k v. P k (f k v))"  | 
| 63194 | 615  | 
by (auto simp: All_mapping_def lookup_map_values split: option.splits)  | 
616  | 
||
| 63462 | 617  | 
lemma All_mapping_tabulate: "(\<forall>x\<in>set xs. P x (f x)) \<Longrightarrow> All_mapping (Mapping.tabulate xs f) P"  | 
618  | 
unfolding All_mapping_def  | 
|
619  | 
apply (intro allI)  | 
|
620  | 
apply transfer  | 
|
621  | 
apply (auto split: option.split dest!: map_of_SomeD)  | 
|
622  | 
done  | 
|
| 63194 | 623  | 
|
624  | 
lemma All_mapping_alist:  | 
|
625  | 
"(\<And>k v. (k, v) \<in> set xs \<Longrightarrow> P k v) \<Longrightarrow> All_mapping (Mapping.of_alist xs) P"  | 
|
626  | 
by (auto simp: All_mapping_def lookup_of_alist dest!: map_of_SomeD split: option.splits)  | 
|
627  | 
||
| 63462 | 628  | 
lemma combine_empty [simp]: "combine f Mapping.empty y = y" "combine f y Mapping.empty = y"  | 
629  | 
by (transfer; force)+  | 
|
| 63194 | 630  | 
|
631  | 
lemma (in abel_semigroup) comm_monoid_set_combine: "comm_monoid_set (combine f) Mapping.empty"  | 
|
632  | 
by standard (transfer fixing: f, simp add: combine_options_ac[of f] ac_simps)+  | 
|
633  | 
||
634  | 
locale combine_mapping_abel_semigroup = abel_semigroup  | 
|
635  | 
begin  | 
|
636  | 
||
637  | 
sublocale combine: comm_monoid_set "combine f" Mapping.empty  | 
|
638  | 
by (rule comm_monoid_set_combine)  | 
|
639  | 
||
640  | 
lemma fold_combine_code:  | 
|
641  | 
"combine.F g (set xs) = foldr (\<lambda>x. combine f (g x)) (remdups xs) Mapping.empty"  | 
|
642  | 
proof -  | 
|
643  | 
have "combine.F g (set xs) = foldr (\<lambda>x. combine f (g x)) xs Mapping.empty"  | 
|
644  | 
if "distinct xs" for xs  | 
|
645  | 
using that by (induction xs) simp_all  | 
|
646  | 
from this[of "remdups xs"] show ?thesis by simp  | 
|
647  | 
qed  | 
|
| 63462 | 648  | 
|
649  | 
lemma keys_fold_combine: "finite A \<Longrightarrow> Mapping.keys (combine.F g A) = (\<Union>x\<in>A. Mapping.keys (g x))"  | 
|
650  | 
by (induct A rule: finite_induct) simp_all  | 
|
| 35157 | 651  | 
|
| 
49975
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
49973 
diff
changeset
 | 
652  | 
end  | 
| 59485 | 653  | 
|
| 63462 | 654  | 
|
| 63194 | 655  | 
subsection \<open>Code generator setup\<close>  | 
656  | 
||
657  | 
hide_const (open) empty is_empty rep lookup lookup_default filter update delete ordered_keys  | 
|
658  | 
keys size replace default map_entry map_default tabulate bulkload map map_values combine of_alist  | 
|
659  | 
||
660  | 
end  |