| author | wenzelm | 
| Thu, 16 Apr 2015 15:22:44 +0200 | |
| changeset 60097 | d20ca79d50e4 | 
| parent 59485 | 792272e6ee6b | 
| child 60500 | 903bb1495239 | 
| 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  | 
|
| 58881 | 5  | 
section {* An abstract view on maps for code generation. *}
 | 
| 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  | 
||
| 51379 | 11  | 
subsection {* Parametricity transfer rules *}
 | 
12  | 
||
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
13  | 
lemma map_of_foldr: -- {* FIXME move *}
 | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
14  | 
"map_of xs = foldr (\<lambda>(k, v) m. m(k \<mapsto> v)) xs Map.empty"  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
15  | 
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
 | 
16  | 
|
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
17  | 
context  | 
| 
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
18  | 
begin  | 
| 56528 | 19  | 
|
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
20  | 
interpretation lifting_syntax .  | 
| 
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
21  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
22  | 
lemma empty_parametric:  | 
| 56528 | 23  | 
"(A ===> rel_option B) Map.empty Map.empty"  | 
24  | 
by transfer_prover  | 
|
| 51379 | 25  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
26  | 
lemma lookup_parametric: "((A ===> B) ===> A ===> B) (\<lambda>m k. m k) (\<lambda>m k. m k)"  | 
| 56528 | 27  | 
by transfer_prover  | 
| 51379 | 28  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
29  | 
lemma update_parametric:  | 
| 51379 | 30  | 
assumes [transfer_rule]: "bi_unique A"  | 
| 56528 | 31  | 
shows "(A ===> B ===> (A ===> rel_option B) ===> A ===> rel_option B)  | 
32  | 
(\<lambda>k v m. m(k \<mapsto> v)) (\<lambda>k v m. m(k \<mapsto> v))"  | 
|
33  | 
by transfer_prover  | 
|
| 51379 | 34  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
35  | 
lemma delete_parametric:  | 
| 51379 | 36  | 
assumes [transfer_rule]: "bi_unique A"  | 
| 55525 | 37  | 
shows "(A ===> (A ===> rel_option B) ===> A ===> rel_option B)  | 
| 56528 | 38  | 
(\<lambda>k m. m(k := None)) (\<lambda>k m. m(k := None))"  | 
39  | 
by transfer_prover  | 
|
| 51379 | 40  | 
|
| 56528 | 41  | 
lemma is_none_parametric [transfer_rule]:  | 
42  | 
"(rel_option A ===> HOL.eq) Option.is_none Option.is_none"  | 
|
43  | 
by (auto simp add: is_none_def rel_fun_def rel_option_iff split: option.split)  | 
|
| 51379 | 44  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
45  | 
lemma dom_parametric:  | 
| 51379 | 46  | 
assumes [transfer_rule]: "bi_total A"  | 
| 55938 | 47  | 
shows "((A ===> rel_option B) ===> rel_set A) dom dom"  | 
| 56528 | 48  | 
unfolding dom_def [abs_def] is_none_def [symmetric] by transfer_prover  | 
| 51379 | 49  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
50  | 
lemma map_of_parametric [transfer_rule]:  | 
| 51379 | 51  | 
assumes [transfer_rule]: "bi_unique R1"  | 
| 55944 | 52  | 
shows "(list_all2 (rel_prod R1 R2) ===> R1 ===> rel_option R2) map_of map_of"  | 
| 56528 | 53  | 
unfolding map_of_def by transfer_prover  | 
| 51379 | 54  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
55  | 
lemma map_entry_parametric [transfer_rule]:  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
56  | 
assumes [transfer_rule]: "bi_unique A"  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
57  | 
shows "(A ===> (B ===> B) ===> (A ===> rel_option B) ===> A ===> rel_option B)  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
58  | 
(\<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
 | 
59  | 
| 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
 | 
60  | 
| 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
 | 
61  | 
by transfer_prover  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
62  | 
|
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
63  | 
lemma tabulate_parametric:  | 
| 51379 | 64  | 
assumes [transfer_rule]: "bi_unique A"  | 
| 55525 | 65  | 
shows "(list_all2 A ===> (A ===> B) ===> A ===> rel_option B)  | 
| 56528 | 66  | 
(\<lambda>ks f. (map_of (map (\<lambda>k. (k, f k)) ks))) (\<lambda>ks f. (map_of (map (\<lambda>k. (k, f k)) ks)))"  | 
67  | 
by transfer_prover  | 
|
| 51379 | 68  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
69  | 
lemma bulkload_parametric:  | 
| 56528 | 70  | 
"(list_all2 A ===> HOL.eq ===> rel_option A)  | 
| 51379 | 71  | 
(\<lambda>xs k. if k < length xs then Some (xs ! k) else None) (\<lambda>xs k. if k < length xs then Some (xs ! k) else None)"  | 
| 56528 | 72  | 
proof  | 
73  | 
fix xs ys  | 
|
74  | 
assume "list_all2 A xs ys"  | 
|
75  | 
then show "(HOL.eq ===> rel_option A)  | 
|
76  | 
(\<lambda>k. if k < length xs then Some (xs ! k) else None)  | 
|
77  | 
(\<lambda>k. if k < length ys then Some (ys ! k) else None)"  | 
|
78  | 
apply induct  | 
|
79  | 
apply auto  | 
|
80  | 
unfolding rel_fun_def  | 
|
81  | 
apply clarsimp  | 
|
82  | 
apply (case_tac xa)  | 
|
83  | 
apply (auto dest: list_all2_lengthD list_all2_nthD)  | 
|
84  | 
done  | 
|
85  | 
qed  | 
|
| 51379 | 86  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
87  | 
lemma map_parametric:  | 
| 55525 | 88  | 
"((A ===> B) ===> (C ===> D) ===> (B ===> rel_option C) ===> A ===> rel_option D)  | 
| 56528 | 89  | 
(\<lambda>f g m. (map_option g \<circ> m \<circ> f)) (\<lambda>f g m. (map_option g \<circ> m \<circ> f))"  | 
90  | 
by transfer_prover  | 
|
| 51379 | 91  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
92  | 
end  | 
| 51379 | 93  | 
|
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
94  | 
|
| 29708 | 95  | 
subsection {* Type definition and primitive operations *}
 | 
96  | 
||
| 49834 | 97  | 
typedef ('a, 'b) mapping = "UNIV :: ('a \<rightharpoonup> 'b) set"
 | 
| 56528 | 98  | 
morphisms rep Mapping  | 
99  | 
..  | 
|
| 
37700
 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 
haftmann 
parents: 
37299 
diff
changeset
 | 
100  | 
|
| 59485 | 101  | 
setup_lifting type_definition_mapping  | 
| 
37700
 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 
haftmann 
parents: 
37299 
diff
changeset
 | 
102  | 
|
| 56528 | 103  | 
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
 | 
104  | 
is Map.empty parametric empty_parametric .  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
105  | 
|
| 56528 | 106  | 
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
 | 
107  | 
is "\<lambda>m k. m k" parametric lookup_parametric .  | 
| 56528 | 108  | 
|
| 59485 | 109  | 
declare [[code drop: Mapping.lookup]]  | 
110  | 
setup \<open>Code.add_default_eqn @{thm Mapping.lookup.abs_eq}\<close> -- \<open>FIXME lifting\<close>
 | 
|
111  | 
||
| 56528 | 112  | 
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
 | 
113  | 
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
 | 
114  | 
|
| 56528 | 115  | 
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
 | 
116  | 
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
 | 
117  | 
|
| 56528 | 118  | 
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
 | 
119  | 
is dom parametric dom_parametric .  | 
| 29708 | 120  | 
|
| 56528 | 121  | 
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
 | 
122  | 
is "\<lambda>ks f. (map_of (List.map (\<lambda>k. (k, f k)) ks))" parametric tabulate_parametric .  | 
| 29708 | 123  | 
|
| 56528 | 124  | 
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
 | 
125  | 
is "\<lambda>xs k. if k < length xs then Some (xs ! k) else None" parametric bulkload_parametric .  | 
| 29708 | 126  | 
|
| 56528 | 127  | 
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
 | 
128  | 
is "\<lambda>f g m. (map_option g \<circ> m \<circ> f)" parametric map_parametric .  | 
| 29708 | 129  | 
|
| 59485 | 130  | 
declare [[code drop: map]]  | 
131  | 
||
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
132  | 
|
| 40605 | 133  | 
subsection {* Functorial structure *}
 | 
134  | 
||
| 
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
 | 
135  | 
functor map: map  | 
| 55466 | 136  | 
by (transfer, auto simp add: fun_eq_iff option.map_comp option.map_id)+  | 
| 40605 | 137  | 
|
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
138  | 
|
| 29708 | 139  | 
subsection {* Derived operations *}
 | 
140  | 
||
| 56528 | 141  | 
definition ordered_keys :: "('a\<Colon>linorder, 'b) mapping \<Rightarrow> 'a list"
 | 
142  | 
where  | 
|
| 37052 | 143  | 
"ordered_keys m = (if finite (keys m) then sorted_list_of_set (keys m) else [])"  | 
| 35194 | 144  | 
|
| 56528 | 145  | 
definition is_empty :: "('a, 'b) mapping \<Rightarrow> bool"
 | 
146  | 
where  | 
|
| 37052 | 147  | 
  "is_empty m \<longleftrightarrow> keys m = {}"
 | 
| 35157 | 148  | 
|
| 56528 | 149  | 
definition size :: "('a, 'b) mapping \<Rightarrow> nat"
 | 
150  | 
where  | 
|
| 37052 | 151  | 
"size m = (if finite (keys m) then card (keys m) else 0)"  | 
| 35157 | 152  | 
|
| 56528 | 153  | 
definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
 | 
154  | 
where  | 
|
| 37052 | 155  | 
"replace k v m = (if k \<in> keys m then update k v m else m)"  | 
| 29814 | 156  | 
|
| 56528 | 157  | 
definition default :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
 | 
158  | 
where  | 
|
| 37052 | 159  | 
"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
 | 
160  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
161  | 
text {* Manual derivation of transfer rule is non-trivial *}
 | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
162  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
163  | 
lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is
 | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
164  | 
"\<lambda>k f m. (case m k of None \<Rightarrow> m  | 
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
165  | 
| 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
 | 
166  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
167  | 
lemma map_entry_code [code]:  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
168  | 
"map_entry k f m = (case lookup m k of None \<Rightarrow> m  | 
| 
49975
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
49973 
diff
changeset
 | 
169  | 
| 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
 | 
170  | 
by transfer rule  | 
| 
37026
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
171  | 
|
| 56528 | 172  | 
definition map_default :: "'a \<Rightarrow> 'b \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping"
 | 
173  | 
where  | 
|
| 
37026
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
174  | 
"map_default k v f m = map_entry k f (default k v m)"  | 
| 
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
175  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
176  | 
definition of_alist :: "('k \<times> 'v) list \<Rightarrow> ('k, 'v) mapping"
 | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
177  | 
where  | 
| 54853 | 178  | 
"of_alist xs = foldr (\<lambda>(k, v) m. update k v m) xs empty"  | 
| 51379 | 179  | 
|
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
180  | 
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
 | 
181  | 
begin  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
182  | 
|
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
183  | 
definition  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
184  | 
"HOL.equal m1 m2 \<longleftrightarrow> (\<forall>k. lookup m1 k = lookup m2 k)"  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
185  | 
|
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
186  | 
instance proof  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
187  | 
qed (unfold equal_mapping_def, transfer, auto)  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
188  | 
|
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
189  | 
end  | 
| 
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
190  | 
|
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
191  | 
context  | 
| 
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
192  | 
begin  | 
| 56528 | 193  | 
|
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
194  | 
interpretation lifting_syntax .  | 
| 
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
195  | 
|
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
196  | 
lemma [transfer_rule]:  | 
| 51379 | 197  | 
assumes [transfer_rule]: "bi_total A"  | 
198  | 
assumes [transfer_rule]: "bi_unique B"  | 
|
| 56528 | 199  | 
shows "(pcr_mapping A B ===> pcr_mapping A B ===> op=) HOL.eq HOL.equal"  | 
200  | 
by (unfold equal) transfer_prover  | 
|
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
201  | 
|
| 
56529
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
202  | 
lemma of_alist_transfer [transfer_rule]:  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
203  | 
assumes [transfer_rule]: "bi_unique R1"  | 
| 
 
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
 
haftmann 
parents: 
56528 
diff
changeset
 | 
204  | 
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
 | 
205  | 
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
 | 
206  | 
|
| 
53013
 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 
kuncar 
parents: 
51379 
diff
changeset
 | 
207  | 
end  | 
| 
51161
 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 
haftmann 
parents: 
49975 
diff
changeset
 | 
208  | 
|
| 56528 | 209  | 
|
| 29708 | 210  | 
subsection {* Properties *}
 | 
211  | 
||
| 56528 | 212  | 
lemma lookup_update:  | 
213  | 
"lookup (update k v m) k = Some v"  | 
|
| 49973 | 214  | 
by transfer simp  | 
215  | 
||
| 56528 | 216  | 
lemma lookup_update_neq:  | 
217  | 
"k \<noteq> k' \<Longrightarrow> lookup (update k v m) k' = lookup m k'"  | 
|
| 49973 | 218  | 
by transfer simp  | 
219  | 
||
| 56528 | 220  | 
lemma lookup_empty:  | 
221  | 
"lookup empty k = None"  | 
|
| 49973 | 222  | 
by transfer simp  | 
223  | 
||
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
224  | 
lemma keys_is_none_rep [code_unfold]:  | 
| 37052 | 225  | 
"k \<in> keys m \<longleftrightarrow> \<not> (Option.is_none (lookup m k))"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
226  | 
by transfer (auto simp add: is_none_def)  | 
| 29708 | 227  | 
|
228  | 
lemma update_update:  | 
|
229  | 
"update k v (update k w m) = update k v m"  | 
|
230  | 
"k \<noteq> l \<Longrightarrow> update k v (update l w m) = update l w (update k v m)"  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
231  | 
by (transfer, simp add: fun_upd_twist)+  | 
| 29708 | 232  | 
|
| 35157 | 233  | 
lemma update_delete [simp]:  | 
234  | 
"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
 | 
235  | 
by transfer simp  | 
| 29708 | 236  | 
|
237  | 
lemma delete_update:  | 
|
238  | 
"delete k (update k v m) = delete k m"  | 
|
239  | 
"k \<noteq> l \<Longrightarrow> delete k (update l v m) = update l v (delete k m)"  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
240  | 
by (transfer, simp add: fun_upd_twist)+  | 
| 29708 | 241  | 
|
| 35157 | 242  | 
lemma delete_empty [simp]:  | 
243  | 
"delete k empty = empty"  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
244  | 
by transfer simp  | 
| 29708 | 245  | 
|
| 35157 | 246  | 
lemma replace_update:  | 
| 37052 | 247  | 
"k \<notin> keys m \<Longrightarrow> replace k v m = m"  | 
248  | 
"k \<in> keys m \<Longrightarrow> replace k v m = update k v m"  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
249  | 
by (transfer, auto simp add: replace_def fun_upd_twist)+  | 
| 29708 | 250  | 
|
251  | 
lemma size_empty [simp]:  | 
|
252  | 
"size empty = 0"  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
253  | 
unfolding size_def by transfer simp  | 
| 29708 | 254  | 
|
255  | 
lemma size_update:  | 
|
| 37052 | 256  | 
"finite (keys m) \<Longrightarrow> size (update k v m) =  | 
257  | 
(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
 | 
258  | 
unfolding size_def by transfer (auto simp add: insert_dom)  | 
| 29708 | 259  | 
|
260  | 
lemma size_delete:  | 
|
| 37052 | 261  | 
"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
 | 
262  | 
unfolding size_def by transfer simp  | 
| 29708 | 263  | 
|
| 37052 | 264  | 
lemma size_tabulate [simp]:  | 
| 29708 | 265  | 
"size (tabulate ks f) = length (remdups ks)"  | 
| 56528 | 266  | 
unfolding size_def by transfer (auto simp add: map_of_map_restrict card_set comp_def)  | 
| 29708 | 267  | 
|
| 29831 | 268  | 
lemma bulkload_tabulate:  | 
| 29826 | 269  | 
"bulkload xs = tabulate [0..<length xs] (nth xs)"  | 
| 56528 | 270  | 
by transfer (auto simp add: map_of_map_restrict)  | 
| 29826 | 271  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
272  | 
lemma is_empty_empty [simp]:  | 
| 37052 | 273  | 
"is_empty empty"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
274  | 
unfolding is_empty_def by transfer simp  | 
| 37052 | 275  | 
|
276  | 
lemma is_empty_update [simp]:  | 
|
277  | 
"\<not> is_empty (update k v m)"  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
278  | 
unfolding is_empty_def by transfer simp  | 
| 37052 | 279  | 
|
280  | 
lemma is_empty_delete:  | 
|
281  | 
  "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
 | 
282  | 
unfolding is_empty_def by transfer (auto simp del: dom_eq_empty_conv)  | 
| 37052 | 283  | 
|
284  | 
lemma is_empty_replace [simp]:  | 
|
285  | 
"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
 | 
286  | 
unfolding is_empty_def replace_def by transfer auto  | 
| 37052 | 287  | 
|
288  | 
lemma is_empty_default [simp]:  | 
|
289  | 
"\<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
 | 
290  | 
unfolding is_empty_def default_def by transfer auto  | 
| 37052 | 291  | 
|
292  | 
lemma is_empty_map_entry [simp]:  | 
|
293  | 
"is_empty (map_entry k f m) \<longleftrightarrow> is_empty m"  | 
|
| 56528 | 294  | 
unfolding is_empty_def by transfer (auto split: option.split)  | 
| 37052 | 295  | 
|
296  | 
lemma is_empty_map_default [simp]:  | 
|
297  | 
"\<not> is_empty (map_default k v f m)"  | 
|
298  | 
by (simp add: map_default_def)  | 
|
299  | 
||
| 56545 | 300  | 
lemma keys_dom_lookup:  | 
301  | 
"keys m = dom (Mapping.lookup m)"  | 
|
302  | 
by transfer rule  | 
|
303  | 
||
| 37052 | 304  | 
lemma keys_empty [simp]:  | 
305  | 
  "keys empty = {}"
 | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
306  | 
by transfer simp  | 
| 37052 | 307  | 
|
308  | 
lemma keys_update [simp]:  | 
|
309  | 
"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
 | 
310  | 
by transfer simp  | 
| 37052 | 311  | 
|
312  | 
lemma keys_delete [simp]:  | 
|
313  | 
  "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
 | 
314  | 
by transfer simp  | 
| 37052 | 315  | 
|
316  | 
lemma keys_replace [simp]:  | 
|
317  | 
"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
 | 
318  | 
unfolding replace_def by transfer (simp add: insert_absorb)  | 
| 37052 | 319  | 
|
320  | 
lemma keys_default [simp]:  | 
|
321  | 
"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
 | 
322  | 
unfolding default_def by transfer (simp add: insert_absorb)  | 
| 37052 | 323  | 
|
324  | 
lemma keys_map_entry [simp]:  | 
|
325  | 
"keys (map_entry k f m) = keys m"  | 
|
| 56528 | 326  | 
by transfer (auto split: option.split)  | 
| 37052 | 327  | 
|
328  | 
lemma keys_map_default [simp]:  | 
|
329  | 
"keys (map_default k v f m) = insert k (keys m)"  | 
|
330  | 
by (simp add: map_default_def)  | 
|
331  | 
||
332  | 
lemma keys_tabulate [simp]:  | 
|
| 
37026
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
333  | 
"keys (tabulate ks f) = set ks"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
334  | 
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
 | 
335  | 
|
| 37052 | 336  | 
lemma keys_bulkload [simp]:  | 
| 
37026
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
337  | 
  "keys (bulkload xs) = {0..<length xs}"
 | 
| 56528 | 338  | 
by (simp add: bulkload_tabulate)  | 
| 
37026
 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 
haftmann 
parents: 
36176 
diff
changeset
 | 
339  | 
|
| 37052 | 340  | 
lemma distinct_ordered_keys [simp]:  | 
341  | 
"distinct (ordered_keys m)"  | 
|
342  | 
by (simp add: ordered_keys_def)  | 
|
343  | 
||
344  | 
lemma ordered_keys_infinite [simp]:  | 
|
345  | 
"\<not> finite (keys m) \<Longrightarrow> ordered_keys m = []"  | 
|
346  | 
by (simp add: ordered_keys_def)  | 
|
347  | 
||
348  | 
lemma ordered_keys_empty [simp]:  | 
|
349  | 
"ordered_keys empty = []"  | 
|
350  | 
by (simp add: ordered_keys_def)  | 
|
351  | 
||
352  | 
lemma ordered_keys_update [simp]:  | 
|
353  | 
"k \<in> keys m \<Longrightarrow> ordered_keys (update k v m) = ordered_keys m"  | 
|
354  | 
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (update k v m) = insort k (ordered_keys m)"  | 
|
355  | 
by (simp_all add: ordered_keys_def) (auto simp only: sorted_list_of_set_insert [symmetric] insert_absorb)  | 
|
356  | 
||
357  | 
lemma ordered_keys_delete [simp]:  | 
|
358  | 
"ordered_keys (delete k m) = remove1 k (ordered_keys m)"  | 
|
359  | 
proof (cases "finite (keys m)")  | 
|
360  | 
case False then show ?thesis by simp  | 
|
361  | 
next  | 
|
362  | 
case True note fin = True  | 
|
363  | 
show ?thesis  | 
|
364  | 
proof (cases "k \<in> keys m")  | 
|
365  | 
case False with fin have "k \<notin> set (sorted_list_of_set (keys m))" by simp  | 
|
366  | 
with False show ?thesis by (simp add: ordered_keys_def remove1_idem)  | 
|
367  | 
next  | 
|
368  | 
case True with fin show ?thesis by (simp add: ordered_keys_def sorted_list_of_set_remove)  | 
|
369  | 
qed  | 
|
370  | 
qed  | 
|
371  | 
||
372  | 
lemma ordered_keys_replace [simp]:  | 
|
373  | 
"ordered_keys (replace k v m) = ordered_keys m"  | 
|
374  | 
by (simp add: replace_def)  | 
|
375  | 
||
376  | 
lemma ordered_keys_default [simp]:  | 
|
377  | 
"k \<in> keys m \<Longrightarrow> ordered_keys (default k v m) = ordered_keys m"  | 
|
378  | 
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (default k v m) = insort k (ordered_keys m)"  | 
|
379  | 
by (simp_all add: default_def)  | 
|
380  | 
||
381  | 
lemma ordered_keys_map_entry [simp]:  | 
|
382  | 
"ordered_keys (map_entry k f m) = ordered_keys m"  | 
|
383  | 
by (simp add: ordered_keys_def)  | 
|
384  | 
||
385  | 
lemma ordered_keys_map_default [simp]:  | 
|
386  | 
"k \<in> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = ordered_keys m"  | 
|
387  | 
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = insort k (ordered_keys m)"  | 
|
388  | 
by (simp_all add: map_default_def)  | 
|
389  | 
||
390  | 
lemma ordered_keys_tabulate [simp]:  | 
|
391  | 
"ordered_keys (tabulate ks f) = sort (remdups ks)"  | 
|
392  | 
by (simp add: ordered_keys_def sorted_list_of_set_sort_remdups)  | 
|
393  | 
||
394  | 
lemma ordered_keys_bulkload [simp]:  | 
|
395  | 
"ordered_keys (bulkload ks) = [0..<length ks]"  | 
|
396  | 
by (simp add: ordered_keys_def)  | 
|
| 36110 | 397  | 
|
| 56528 | 398  | 
lemma tabulate_fold:  | 
399  | 
"tabulate xs f = fold (\<lambda>k m. update k (f k) m) xs empty"  | 
|
400  | 
proof transfer  | 
|
401  | 
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
 | 
402  | 
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
 | 
403  | 
by (simp add: foldr_map comp_def map_of_foldr)  | 
| 56528 | 404  | 
also have "foldr (\<lambda>k m. m(k \<mapsto> f k)) xs = fold (\<lambda>k m. m(k \<mapsto> f k)) xs"  | 
405  | 
by (rule foldr_fold) (simp add: fun_eq_iff)  | 
|
406  | 
ultimately show "map_of (List.map (\<lambda>k. (k, f k)) xs) = fold (\<lambda>k m. m(k \<mapsto> f k)) xs Map.empty"  | 
|
407  | 
by simp  | 
|
408  | 
qed  | 
|
409  | 
||
| 31459 | 410  | 
|
| 
37700
 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 
haftmann 
parents: 
37299 
diff
changeset
 | 
411  | 
subsection {* Code generator setup *}
 | 
| 31459 | 412  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
49834 
diff
changeset
 | 
413  | 
hide_const (open) empty is_empty rep lookup update delete ordered_keys keys size  | 
| 54853 | 414  | 
replace default map_entry map_default tabulate bulkload map of_alist  | 
| 35157 | 415  | 
|
| 
49975
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
49973 
diff
changeset
 | 
416  | 
end  | 
| 59485 | 417  |