author | nipkow |
Mon, 14 Apr 2003 18:52:13 +0200 | |
changeset 13910 | f9a9ef16466f |
parent 13909 | a5247a49c85e |
child 13912 | 3c0a340be514 |
permissions | -rw-r--r-- |
3981 | 1 |
(* Title: HOL/Map.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow, based on a theory by David von Oheimb |
|
13908 | 4 |
Copyright 1997-2003 TU Muenchen |
3981 | 5 |
|
6 |
The datatype of `maps' (written ~=>); strongly resembles maps in VDM. |
|
7 |
*) |
|
8 |
||
13908 | 9 |
theory Map = List: |
3981 | 10 |
|
13908 | 11 |
types ('a,'b) "~=>" = "'a => 'b option" (infixr 0) |
3981 | 12 |
|
13 |
consts |
|
5300 | 14 |
chg_map :: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)" |
3981 | 15 |
override:: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)" (infixl "++" 100) |
5300 | 16 |
dom :: "('a ~=> 'b) => 'a set" |
17 |
ran :: "('a ~=> 'b) => 'b set" |
|
18 |
map_of :: "('a * 'b)list => 'a ~=> 'b" |
|
19 |
map_upds:: "('a ~=> 'b) => 'a list => 'b list => |
|
13910 | 20 |
('a ~=> 'b)" ("_/'(_[|->]_/')" [900,0,0]900) |
21 |
map_le :: "('a ~=> 'b) => ('a ~=> 'b) => bool" (infix "\<subseteq>\<^sub>m" 50) |
|
22 |
||
5300 | 23 |
syntax |
13890 | 24 |
empty :: "'a ~=> 'b" |
5300 | 25 |
map_upd :: "('a ~=> 'b) => 'a => 'b => ('a ~=> 'b)" |
13910 | 26 |
("_/'(_/|->_')" [900,0,0]900) |
3981 | 27 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
10137
diff
changeset
|
28 |
syntax (xsymbols) |
13908 | 29 |
"~=>" :: "[type, type] => type" (infixr "\<leadsto>" 0) |
5300 | 30 |
map_upd :: "('a ~=> 'b) => 'a => 'b => ('a ~=> 'b)" |
13908 | 31 |
("_/'(_/\<mapsto>/_')" [900,0,0]900) |
5300 | 32 |
map_upds :: "('a ~=> 'b) => 'a list => 'b list => ('a ~=> 'b)" |
13908 | 33 |
("_/'(_/[\<mapsto>]/_')" [900,0,0]900) |
5300 | 34 |
|
35 |
translations |
|
13890 | 36 |
"empty" => "_K None" |
37 |
"empty" <= "%x. None" |
|
5300 | 38 |
|
39 |
"m(a|->b)" == "m(a:=Some b)" |
|
3981 | 40 |
|
41 |
defs |
|
13908 | 42 |
chg_map_def: "chg_map f a m == case m a of None => m | Some b => m(a|->f b)" |
3981 | 43 |
|
13908 | 44 |
override_def: "m1++m2 == %x. case m2 x of None => m1 x | Some y => Some y" |
3981 | 45 |
|
13908 | 46 |
dom_def: "dom(m) == {a. m a ~= None}" |
47 |
ran_def: "ran(m) == {b. ? a. m a = Some b}" |
|
3981 | 48 |
|
13910 | 49 |
map_le_def: "m1 \<subseteq>\<^sub>m m2 == ALL a : dom m1. m1 a = m2 a" |
50 |
||
5183 | 51 |
primrec |
52 |
"map_of [] = empty" |
|
5300 | 53 |
"map_of (p#ps) = (map_of ps)(fst p |-> snd p)" |
54 |
||
55 |
primrec "t([] [|->]bs) = t" |
|
56 |
"t(a#as[|->]bs) = t(a|->hd bs)(as[|->]tl bs)" |
|
3981 | 57 |
|
13908 | 58 |
|
13909
a5247a49c85e
Fixed non-escaped underscore in section headings (document generation should
webertj
parents:
13908
diff
changeset
|
59 |
section {* empty *} |
13908 | 60 |
|
13910 | 61 |
lemma empty_upd_none[simp]: "empty(x := None) = empty" |
13908 | 62 |
apply (rule ext) |
63 |
apply (simp (no_asm)) |
|
64 |
done |
|
13910 | 65 |
|
13908 | 66 |
|
67 |
(* FIXME: what is this sum_case nonsense?? *) |
|
13910 | 68 |
lemma sum_case_empty_empty[simp]: "sum_case empty empty = empty" |
13908 | 69 |
apply (rule ext) |
70 |
apply (simp (no_asm) split add: sum.split) |
|
71 |
done |
|
72 |
||
13909
a5247a49c85e
Fixed non-escaped underscore in section headings (document generation should
webertj
parents:
13908
diff
changeset
|
73 |
section {* map\_upd *} |
13908 | 74 |
|
75 |
lemma map_upd_triv: "t k = Some x ==> t(k|->x) = t" |
|
76 |
apply (rule ext) |
|
77 |
apply (simp (no_asm_simp)) |
|
78 |
done |
|
79 |
||
13910 | 80 |
lemma map_upd_nonempty[simp]: "t(k|->x) ~= empty" |
13908 | 81 |
apply safe |
82 |
apply (drule_tac x = "k" in fun_cong) |
|
83 |
apply (simp (no_asm_use)) |
|
84 |
done |
|
85 |
||
86 |
lemma finite_range_updI: "finite (range f) ==> finite (range (f(a|->b)))" |
|
87 |
apply (unfold image_def) |
|
88 |
apply (simp (no_asm_use) add: full_SetCompr_eq) |
|
89 |
apply (rule finite_subset) |
|
90 |
prefer 2 apply (assumption) |
|
91 |
apply auto |
|
92 |
done |
|
93 |
||
94 |
||
95 |
(* FIXME: what is this sum_case nonsense?? *) |
|
13909
a5247a49c85e
Fixed non-escaped underscore in section headings (document generation should
webertj
parents:
13908
diff
changeset
|
96 |
section {* sum\_case and empty/map\_upd *} |
13908 | 97 |
|
13910 | 98 |
lemma sum_case_map_upd_empty[simp]: |
99 |
"sum_case (m(k|->y)) empty = (sum_case m empty)(Inl k|->y)" |
|
13908 | 100 |
apply (rule ext) |
101 |
apply (simp (no_asm) split add: sum.split) |
|
102 |
done |
|
103 |
||
13910 | 104 |
lemma sum_case_empty_map_upd[simp]: |
105 |
"sum_case empty (m(k|->y)) = (sum_case empty m)(Inr k|->y)" |
|
13908 | 106 |
apply (rule ext) |
107 |
apply (simp (no_asm) split add: sum.split) |
|
108 |
done |
|
109 |
||
13910 | 110 |
lemma sum_case_map_upd_map_upd[simp]: |
111 |
"sum_case (m1(k1|->y1)) (m2(k2|->y2)) = (sum_case (m1(k1|->y1)) m2)(Inr k2|->y2)" |
|
13908 | 112 |
apply (rule ext) |
113 |
apply (simp (no_asm) split add: sum.split) |
|
114 |
done |
|
115 |
||
116 |
||
13909
a5247a49c85e
Fixed non-escaped underscore in section headings (document generation should
webertj
parents:
13908
diff
changeset
|
117 |
section {* map\_upds *} |
13908 | 118 |
|
13910 | 119 |
lemma map_upd_upds_conv_if: |
120 |
"!!x y ys f. (f(x|->y))(xs [|->] ys) = |
|
121 |
(if x : set xs then f(xs [|->] ys) else (f(xs [|->] ys))(x|->y))" |
|
122 |
apply(induct xs) |
|
123 |
apply simp |
|
124 |
apply(simp split:split_if add:fun_upd_twist eq_sym_conv) |
|
13908 | 125 |
done |
13910 | 126 |
|
127 |
lemma map_upds_twist [simp]: |
|
128 |
"a ~: set as ==> m(a|->b)(as[|->]bs) = m(as[|->]bs)(a|->b)" |
|
129 |
by (simp add: map_upd_upds_conv_if) |
|
13908 | 130 |
|
13910 | 131 |
lemma map_upds_apply_nontin[simp]: |
132 |
"!!ys. x ~: set xs ==> (f(xs[|->]ys)) x = f x" |
|
133 |
apply(induct xs) |
|
134 |
apply simp |
|
135 |
apply(simp add: fun_upd_apply map_upd_upds_conv_if split:split_if) |
|
136 |
done |
|
13908 | 137 |
|
13909
a5247a49c85e
Fixed non-escaped underscore in section headings (document generation should
webertj
parents:
13908
diff
changeset
|
138 |
section {* chg\_map *} |
13908 | 139 |
|
13910 | 140 |
lemma chg_map_new[simp]: "m a = None ==> chg_map f a m = m" |
13908 | 141 |
apply (unfold chg_map_def) |
142 |
apply auto |
|
143 |
done |
|
144 |
||
13910 | 145 |
lemma chg_map_upd[simp]: "m a = Some b ==> chg_map f a m = m(a|->f b)" |
13908 | 146 |
apply (unfold chg_map_def) |
147 |
apply auto |
|
148 |
done |
|
149 |
||
150 |
||
13909
a5247a49c85e
Fixed non-escaped underscore in section headings (document generation should
webertj
parents:
13908
diff
changeset
|
151 |
section {* map\_of *} |
13908 | 152 |
|
153 |
lemma map_of_SomeD [rule_format (no_asm)]: "map_of xs k = Some y --> (k,y):set xs" |
|
154 |
apply (induct_tac "xs") |
|
155 |
apply auto |
|
156 |
done |
|
157 |
||
158 |
lemma map_of_mapk_SomeI [rule_format (no_asm)]: "inj f ==> map_of t k = Some x --> |
|
159 |
map_of (map (split (%k. Pair (f k))) t) (f k) = Some x" |
|
160 |
apply (induct_tac "t") |
|
161 |
apply (auto simp add: inj_eq) |
|
162 |
done |
|
163 |
||
164 |
lemma weak_map_of_SomeI [rule_format (no_asm)]: "(k, x) : set l --> (? x. map_of l k = Some x)" |
|
165 |
apply (induct_tac "l") |
|
166 |
apply auto |
|
167 |
done |
|
168 |
||
169 |
lemma map_of_filter_in: |
|
170 |
"[| map_of xs k = Some z; P k z |] ==> map_of (filter (split P) xs) k = Some z" |
|
171 |
apply (rule mp) |
|
172 |
prefer 2 apply (assumption) |
|
173 |
apply (erule thin_rl) |
|
174 |
apply (induct_tac "xs") |
|
175 |
apply auto |
|
176 |
done |
|
177 |
||
178 |
lemma finite_range_map_of: "finite (range (map_of l))" |
|
179 |
apply (induct_tac "l") |
|
180 |
apply (simp_all (no_asm) add: image_constant) |
|
181 |
apply (rule finite_subset) |
|
182 |
prefer 2 apply (assumption) |
|
183 |
apply auto |
|
184 |
done |
|
185 |
||
186 |
lemma map_of_map: "map_of (map (%(a,b). (a,f b)) xs) x = option_map f (map_of xs x)" |
|
187 |
apply (induct_tac "xs") |
|
188 |
apply auto |
|
189 |
done |
|
190 |
||
191 |
||
13909
a5247a49c85e
Fixed non-escaped underscore in section headings (document generation should
webertj
parents:
13908
diff
changeset
|
192 |
section {* option\_map related *} |
13908 | 193 |
|
13910 | 194 |
lemma option_map_o_empty[simp]: "option_map f o empty = empty" |
13908 | 195 |
apply (rule ext) |
196 |
apply (simp (no_asm)) |
|
197 |
done |
|
198 |
||
13910 | 199 |
lemma option_map_o_map_upd[simp]: |
200 |
"option_map f o m(a|->b) = (option_map f o m)(a|->f b)" |
|
13908 | 201 |
apply (rule ext) |
202 |
apply (simp (no_asm)) |
|
203 |
done |
|
204 |
||
205 |
||
13909
a5247a49c85e
Fixed non-escaped underscore in section headings (document generation should
webertj
parents:
13908
diff
changeset
|
206 |
section {* ++ *} |
13908 | 207 |
|
13910 | 208 |
lemma override_empty[simp]: "m ++ empty = m" |
13908 | 209 |
apply (unfold override_def) |
210 |
apply (simp (no_asm)) |
|
211 |
done |
|
212 |
||
13910 | 213 |
lemma empty_override[simp]: "empty ++ m = m" |
13908 | 214 |
apply (unfold override_def) |
215 |
apply (rule ext) |
|
216 |
apply (simp split add: option.split) |
|
217 |
done |
|
218 |
||
219 |
lemma override_Some_iff [rule_format (no_asm)]: |
|
220 |
"((m ++ n) k = Some x) = (n k = Some x | n k = None & m k = Some x)" |
|
221 |
apply (unfold override_def) |
|
222 |
apply (simp (no_asm) split add: option.split) |
|
223 |
done |
|
224 |
||
225 |
lemmas override_SomeD = override_Some_iff [THEN iffD1, standard] |
|
226 |
declare override_SomeD [dest!] |
|
227 |
||
13910 | 228 |
lemma override_find_right[simp]: "!!xx. n k = Some xx ==> (m ++ n) k = Some xx" |
13908 | 229 |
apply (subst override_Some_iff) |
230 |
apply fast |
|
231 |
done |
|
232 |
||
13910 | 233 |
lemma override_None [iff]: "((m ++ n) k = None) = (n k = None & m k = None)" |
13908 | 234 |
apply (unfold override_def) |
235 |
apply (simp (no_asm) split add: option.split) |
|
236 |
done |
|
237 |
||
13910 | 238 |
lemma override_upd[simp]: "f ++ g(x|->y) = (f ++ g)(x|->y)" |
13908 | 239 |
apply (unfold override_def) |
240 |
apply (rule ext) |
|
241 |
apply auto |
|
242 |
done |
|
243 |
||
13910 | 244 |
lemma map_of_override[simp]: "map_of ys ++ map_of xs = map_of (xs@ys)" |
13908 | 245 |
apply (unfold override_def) |
246 |
apply (rule sym) |
|
247 |
apply (induct_tac "xs") |
|
248 |
apply (simp (no_asm)) |
|
249 |
apply (rule ext) |
|
250 |
apply (simp (no_asm_simp) split add: option.split) |
|
251 |
done |
|
252 |
||
253 |
declare fun_upd_apply [simp del] |
|
254 |
lemma finite_range_map_of_override: "finite (range f) ==> finite (range (f ++ map_of l))" |
|
255 |
apply (induct_tac "l") |
|
256 |
apply auto |
|
257 |
apply (erule finite_range_updI) |
|
258 |
done |
|
259 |
declare fun_upd_apply [simp] |
|
260 |
||
261 |
||
13909
a5247a49c85e
Fixed non-escaped underscore in section headings (document generation should
webertj
parents:
13908
diff
changeset
|
262 |
section {* dom *} |
13908 | 263 |
|
264 |
lemma domI: "m a = Some b ==> a : dom m" |
|
265 |
apply (unfold dom_def) |
|
266 |
apply auto |
|
267 |
done |
|
268 |
||
269 |
lemma domD: "a : dom m ==> ? b. m a = Some b" |
|
270 |
apply (unfold dom_def) |
|
271 |
apply auto |
|
272 |
done |
|
273 |
||
13910 | 274 |
lemma domIff[iff]: "(a : dom m) = (m a ~= None)" |
13908 | 275 |
apply (unfold dom_def) |
276 |
apply auto |
|
277 |
done |
|
278 |
declare domIff [simp del] |
|
279 |
||
13910 | 280 |
lemma dom_empty[simp]: "dom empty = {}" |
13908 | 281 |
apply (unfold dom_def) |
282 |
apply (simp (no_asm)) |
|
283 |
done |
|
284 |
||
13910 | 285 |
lemma dom_fun_upd[simp]: |
286 |
"dom(f(x := y)) = (if y=None then dom f - {x} else insert x (dom f))" |
|
287 |
by (simp add:dom_def) blast |
|
288 |
(* |
|
289 |
lemma dom_map_upd[simp]: "dom(m(a|->b)) = insert a (dom m)" |
|
13908 | 290 |
apply (unfold dom_def) |
291 |
apply (simp (no_asm)) |
|
292 |
apply blast |
|
293 |
done |
|
13910 | 294 |
*) |
13908 | 295 |
|
296 |
lemma finite_dom_map_of: "finite (dom (map_of l))" |
|
297 |
apply (unfold dom_def) |
|
298 |
apply (induct_tac "l") |
|
299 |
apply (auto simp add: insert_Collect [symmetric]) |
|
300 |
done |
|
301 |
||
13910 | 302 |
lemma dom_map_upds[simp]: "!!m vs. dom(m(xs[|->]vs)) = set xs Un dom m" |
303 |
by(induct xs, simp_all) |
|
304 |
||
305 |
lemma dom_override[simp]: "dom(m++n) = dom n Un dom m" |
|
13908 | 306 |
apply (unfold dom_def) |
307 |
apply auto |
|
308 |
done |
|
13910 | 309 |
|
310 |
lemma dom_overwrite[simp]: |
|
311 |
"dom(f(g|A)) = (dom f - {a. a : A - dom g}) Un {a. a : A Int dom g}" |
|
312 |
by(auto simp add: dom_def overwrite_def) |
|
13908 | 313 |
|
13909
a5247a49c85e
Fixed non-escaped underscore in section headings (document generation should
webertj
parents:
13908
diff
changeset
|
314 |
section {* ran *} |
13908 | 315 |
|
13910 | 316 |
lemma ran_empty[simp]: "ran empty = {}" |
13908 | 317 |
apply (unfold ran_def) |
318 |
apply (simp (no_asm)) |
|
319 |
done |
|
320 |
||
13910 | 321 |
lemma ran_map_upd[simp]: "m a = None ==> ran(m(a|->b)) = insert b (ran m)" |
13908 | 322 |
apply (unfold ran_def) |
323 |
apply auto |
|
324 |
apply (subgoal_tac "~ (aa = a) ") |
|
325 |
apply auto |
|
326 |
done |
|
13910 | 327 |
|
328 |
section{* @{text"\<subseteq>\<^sub>m"} *} |
|
329 |
||
330 |
lemma [simp]: "empty \<subseteq>\<^sub>m g" |
|
331 |
by(simp add:map_le_def) |
|
332 |
||
333 |
lemma map_le_upd[simp]: "f \<subseteq>\<^sub>m g ==> f(a := b) \<subseteq>\<^sub>m g(a := b)" |
|
334 |
by(fastsimp simp add:map_le_def) |
|
335 |
||
336 |
lemma map_le_upds[simp]: |
|
337 |
"!!f g bs. f \<subseteq>\<^sub>m g ==> f(as [|->] bs) \<subseteq>\<^sub>m g(as [|->] bs)" |
|
338 |
by(induct as, auto) |
|
13908 | 339 |
|
3981 | 340 |
end |