16 translations (type) "a ~=> b " <= (type) "a => b option" |
16 translations (type) "a ~=> b " <= (type) "a => b option" |
17 |
17 |
18 consts |
18 consts |
19 chg_map :: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)" |
19 chg_map :: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)" |
20 map_add :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)" (infixl "++" 100) |
20 map_add :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)" (infixl "++" 100) |
21 restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" ("_|'__" [90, 91] 90) |
21 restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" (infixl "|^" 110) |
22 dom :: "('a ~=> 'b) => 'a set" |
22 dom :: "('a ~=> 'b) => 'a set" |
23 ran :: "('a ~=> 'b) => 'b set" |
23 ran :: "('a ~=> 'b) => 'b set" |
24 map_of :: "('a * 'b)list => 'a ~=> 'b" |
24 map_of :: "('a * 'b)list => 'a ~=> 'b" |
25 map_upds:: "('a ~=> 'b) => 'a list => 'b list => |
25 map_upds:: "('a ~=> 'b) => 'a list => 'b list => |
26 ('a ~=> 'b)" |
26 ('a ~=> 'b)" |
53 fun_map_comp :: "('b => 'c) => ('a ~=> 'b) => ('a ~=> 'c)" (infixl "\<circ>\<^sub>m" 55) |
53 fun_map_comp :: "('b => 'c) => ('a ~=> 'b) => ('a ~=> 'c)" (infixl "\<circ>\<^sub>m" 55) |
54 |
54 |
55 "_maplet" :: "['a, 'a] => maplet" ("_ /\<mapsto>/ _") |
55 "_maplet" :: "['a, 'a] => maplet" ("_ /\<mapsto>/ _") |
56 "_maplets" :: "['a, 'a] => maplet" ("_ /[\<mapsto>]/ _") |
56 "_maplets" :: "['a, 'a] => maplet" ("_ /[\<mapsto>]/ _") |
57 |
57 |
58 restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" ("_\<lfloor>_" [90, 91] 90) |
58 restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" (infixl "\<upharpoonright>" 110) --"requires amssymb!" |
59 map_upd_s :: "('a ~=> 'b) => 'a set => 'b => ('a ~=> 'b)" |
59 map_upd_s :: "('a ~=> 'b) => 'a set => 'b => ('a ~=> 'b)" |
60 ("_/'(_/{\<mapsto>}/_')" [900,0,0]900) |
60 ("_/'(_/{\<mapsto>}/_')" [900,0,0]900) |
61 map_subst :: "('a ~=> 'b) => 'b => 'b => |
61 map_subst :: "('a ~=> 'b) => 'b => 'b => |
62 ('a ~=> 'b)" ("_/'(_\<leadsto>_/')" [900,0,0]900) |
62 ('a ~=> 'b)" ("_/'(_\<leadsto>_/')" [900,0,0]900) |
63 "@chg_map" :: "('a ~=> 'b) => 'a => ('b => 'b) => ('a ~=> 'b)" |
63 "@chg_map" :: "('a ~=> 'b) => 'a => ('b => 'b) => ('a ~=> 'b)" |
322 "inj_on (m ++ m') (dom m') = inj_on m' (dom m')" |
322 "inj_on (m ++ m') (dom m') = inj_on m' (dom m')" |
323 by(fastsimp simp add:map_add_def dom_def inj_on_def split:option.splits) |
323 by(fastsimp simp add:map_add_def dom_def inj_on_def split:option.splits) |
324 |
324 |
325 subsection {* @{term restrict_map} *} |
325 subsection {* @{term restrict_map} *} |
326 |
326 |
327 lemma restrict_map_to_empty[simp]: "m\<lfloor>{} = empty" |
327 lemma restrict_map_to_empty[simp]: "m|^{} = empty" |
328 by(simp add: restrict_map_def) |
328 by(simp add: restrict_map_def) |
329 |
329 |
330 lemma restrict_map_empty[simp]: "empty\<lfloor>D = empty" |
330 lemma restrict_map_empty[simp]: "empty|^D = empty" |
331 by(simp add: restrict_map_def) |
331 by(simp add: restrict_map_def) |
332 |
332 |
333 lemma restrict_in [simp]: "x \<in> A \<Longrightarrow> (m\<lfloor>A) x = m x" |
333 lemma restrict_in [simp]: "x \<in> A \<Longrightarrow> (m|^A) x = m x" |
334 by (auto simp: restrict_map_def) |
334 by (auto simp: restrict_map_def) |
335 |
335 |
336 lemma restrict_out [simp]: "x \<notin> A \<Longrightarrow> (m\<lfloor>A) x = None" |
336 lemma restrict_out [simp]: "x \<notin> A \<Longrightarrow> (m|^A) x = None" |
337 by (auto simp: restrict_map_def) |
337 by (auto simp: restrict_map_def) |
338 |
338 |
339 lemma ran_restrictD: "y \<in> ran (m\<lfloor>A) \<Longrightarrow> \<exists>x\<in>A. m x = Some y" |
339 lemma ran_restrictD: "y \<in> ran (m|^A) \<Longrightarrow> \<exists>x\<in>A. m x = Some y" |
340 by (auto simp: restrict_map_def ran_def split: split_if_asm) |
340 by (auto simp: restrict_map_def ran_def split: split_if_asm) |
341 |
341 |
342 lemma dom_restrict [simp]: "dom (m\<lfloor>A) = dom m \<inter> A" |
342 lemma dom_restrict [simp]: "dom (m|^A) = dom m \<inter> A" |
343 by (auto simp: restrict_map_def dom_def split: split_if_asm) |
343 by (auto simp: restrict_map_def dom_def split: split_if_asm) |
344 |
344 |
345 lemma restrict_upd_same [simp]: "m(x\<mapsto>y)\<lfloor>(-{x}) = m\<lfloor>(-{x})" |
345 lemma restrict_upd_same [simp]: "m(x\<mapsto>y)|^(-{x}) = m|^(-{x})" |
346 by (rule ext, auto simp: restrict_map_def) |
346 by (rule ext, auto simp: restrict_map_def) |
347 |
347 |
348 lemma restrict_restrict [simp]: "m\<lfloor>A\<lfloor>B = m\<lfloor>(A\<inter>B)" |
348 lemma restrict_restrict [simp]: "m|^A|^B = m|^(A\<inter>B)" |
349 by (rule ext, auto simp: restrict_map_def) |
349 by (rule ext, auto simp: restrict_map_def) |
350 |
350 |
351 lemma restrict_fun_upd[simp]: |
351 lemma restrict_fun_upd[simp]: |
352 "m(x := y)\<lfloor>D = (if x \<in> D then (m\<lfloor>(D-{x}))(x := y) else m\<lfloor>D)" |
352 "m(x := y)|^D = (if x \<in> D then (m|^(D-{x}))(x := y) else m|^D)" |
353 by(simp add: restrict_map_def expand_fun_eq) |
353 by(simp add: restrict_map_def expand_fun_eq) |
354 |
354 |
355 lemma fun_upd_None_restrict[simp]: |
355 lemma fun_upd_None_restrict[simp]: |
356 "(m\<lfloor>D)(x := None) = (if x:D then m\<lfloor>(D - {x}) else m\<lfloor>D)" |
356 "(m|^D)(x := None) = (if x:D then m|^(D - {x}) else m|^D)" |
357 by(simp add: restrict_map_def expand_fun_eq) |
357 by(simp add: restrict_map_def expand_fun_eq) |
358 |
358 |
359 lemma fun_upd_restrict: |
359 lemma fun_upd_restrict: |
360 "(m\<lfloor>D)(x := y) = (m\<lfloor>(D-{x}))(x := y)" |
360 "(m|^D)(x := y) = (m|^(D-{x}))(x := y)" |
361 by(simp add: restrict_map_def expand_fun_eq) |
361 by(simp add: restrict_map_def expand_fun_eq) |
362 |
362 |
363 lemma fun_upd_restrict_conv[simp]: |
363 lemma fun_upd_restrict_conv[simp]: |
364 "x \<in> D \<Longrightarrow> (m\<lfloor>D)(x := y) = (m\<lfloor>(D-{x}))(x := y)" |
364 "x \<in> D \<Longrightarrow> (m|^D)(x := y) = (m|^(D-{x}))(x := y)" |
365 by(simp add: restrict_map_def expand_fun_eq) |
365 by(simp add: restrict_map_def expand_fun_eq) |
366 |
366 |
367 |
367 |
368 subsection {* @{term map_upds} *} |
368 subsection {* @{term map_upds} *} |
369 |
369 |
430 done |
430 done |
431 |
431 |
432 |
432 |
433 lemma restrict_map_upds[simp]: "!!m ys. |
433 lemma restrict_map_upds[simp]: "!!m ys. |
434 \<lbrakk> length xs = length ys; set xs \<subseteq> D \<rbrakk> |
434 \<lbrakk> length xs = length ys; set xs \<subseteq> D \<rbrakk> |
435 \<Longrightarrow> m(xs [\<mapsto>] ys)\<lfloor>D = (m\<lfloor>(D - set xs))(xs [\<mapsto>] ys)" |
435 \<Longrightarrow> m(xs [\<mapsto>] ys)|^D = (m|^(D - set xs))(xs [\<mapsto>] ys)" |
436 apply (induct xs, simp) |
436 apply (induct xs, simp) |
437 apply (case_tac ys, simp) |
437 apply (case_tac ys, simp) |
438 apply(simp add:Diff_insert[symmetric] insert_absorb) |
438 apply(simp add:Diff_insert[symmetric] insert_absorb) |
439 apply(simp add: map_upd_upds_conv_if) |
439 apply(simp add: map_upd_upds_conv_if) |
440 done |
440 done |
498 done |
498 done |
499 |
499 |
500 lemma dom_map_add[simp]: "dom(m++n) = dom n Un dom m" |
500 lemma dom_map_add[simp]: "dom(m++n) = dom n Un dom m" |
501 by (unfold dom_def, auto) |
501 by (unfold dom_def, auto) |
502 |
502 |
503 lemma dom_overwrite[simp]: |
503 lemma dom_override_on[simp]: |
504 "dom(f(g|A)) = (dom f - {a. a : A - dom g}) Un {a. a : A Int dom g}" |
504 "dom(override_on f g A) = |
505 by(auto simp add: dom_def overwrite_def) |
505 (dom f - {a. a : A - dom g}) Un {a. a : A Int dom g}" |
|
506 by(auto simp add: dom_def override_on_def) |
506 |
507 |
507 lemma map_add_comm: "dom m1 \<inter> dom m2 = {} \<Longrightarrow> m1++m2 = m2++m1" |
508 lemma map_add_comm: "dom m1 \<inter> dom m2 = {} \<Longrightarrow> m1++m2 = m2++m1" |
508 apply(rule ext) |
509 apply(rule ext) |
509 apply(fastsimp simp:map_add_def split:option.split) |
510 apply(fastsimp simp:map_add_def split:option.split) |
510 done |
511 done |