| author | paulson | 
| Tue, 19 Oct 2004 18:18:45 +0200 | |
| changeset 15251 | bb6f072c8d10 | 
| parent 15140 | 322485b816ac | 
| child 15303 | eedbb8d22ca2 | 
| 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  | 
||
| 13914 | 9  | 
header {* Maps *}
 | 
10  | 
||
| 15131 | 11  | 
theory Map  | 
| 15140 | 12  | 
imports List  | 
| 15131 | 13  | 
begin  | 
| 3981 | 14  | 
|
| 13908 | 15  | 
types ('a,'b) "~=>" = "'a => 'b option" (infixr 0)
 | 
| 14100 | 16  | 
translations (type) "a ~=> b " <= (type) "a => b option"  | 
| 3981 | 17  | 
|
18  | 
consts  | 
|
| 5300 | 19  | 
chg_map	:: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)"
 | 
| 14100 | 20  | 
map_add :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)" (infixl "++" 100)
 | 
21  | 
restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" ("_|'__" [90, 91] 90)
 | 
|
| 5300 | 22  | 
dom	:: "('a ~=> 'b) => 'a set"
 | 
23  | 
ran	:: "('a ~=> 'b) => 'b set"
 | 
|
24  | 
map_of	:: "('a * 'b)list => 'a ~=> 'b"
 | 
|
25  | 
map_upds:: "('a ~=> 'b) => 'a list => 'b list => 
 | 
|
| 14180 | 26  | 
	    ('a ~=> 'b)"
 | 
| 14100 | 27  | 
map_upd_s::"('a ~=> 'b) => 'a set => 'b => 
 | 
28  | 
	    ('a ~=> 'b)"			 ("_/'(_{|->}_/')" [900,0,0]900)
 | 
|
29  | 
map_subst::"('a ~=> 'b) => 'b => 'b => 
 | 
|
30  | 
	    ('a ~=> 'b)"			 ("_/'(_~>_/')"    [900,0,0]900)
 | 
|
| 13910 | 31  | 
map_le  :: "('a ~=> 'b) => ('a ~=> 'b) => bool" (infix "\<subseteq>\<^sub>m" 50)
 | 
32  | 
||
| 14739 | 33  | 
syntax  | 
34  | 
  fun_map_comp :: "('b => 'c)  => ('a ~=> 'b) => ('a ~=> 'c)" (infixl "o'_m" 55)
 | 
|
35  | 
translations  | 
|
36  | 
"f o_m m" == "option_map f o m"  | 
|
37  | 
||
| 14180 | 38  | 
nonterminals  | 
39  | 
maplets maplet  | 
|
40  | 
||
| 5300 | 41  | 
syntax  | 
| 14180 | 42  | 
empty :: "'a ~=> 'b"  | 
43  | 
  "_maplet"  :: "['a, 'a] => maplet"             ("_ /|->/ _")
 | 
|
44  | 
  "_maplets" :: "['a, 'a] => maplet"             ("_ /[|->]/ _")
 | 
|
45  | 
  ""         :: "maplet => maplets"             ("_")
 | 
|
46  | 
  "_Maplets" :: "[maplet, maplets] => maplets" ("_,/ _")
 | 
|
47  | 
  "_MapUpd"  :: "['a ~=> 'b, maplets] => 'a ~=> 'b" ("_/'(_')" [900,0]900)
 | 
|
48  | 
  "_Map"     :: "maplets => 'a ~=> 'b"            ("(1[_])")
 | 
|
| 3981 | 49  | 
|
| 
12114
 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 
wenzelm 
parents: 
10137 
diff
changeset
 | 
50  | 
syntax (xsymbols)  | 
| 14739 | 51  | 
"~=>" :: "[type, type] => type" (infixr "\<rightharpoonup>" 0)  | 
52  | 
||
53  | 
  fun_map_comp :: "('b => 'c)  => ('a ~=> 'b) => ('a ~=> 'c)" (infixl "\<circ>\<^sub>m" 55)
 | 
|
54  | 
||
| 14180 | 55  | 
  "_maplet"  :: "['a, 'a] => maplet"             ("_ /\<mapsto>/ _")
 | 
56  | 
  "_maplets" :: "['a, 'a] => maplet"             ("_ /[\<mapsto>]/ _")
 | 
|
57  | 
||
| 14100 | 58  | 
  restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" ("_\<lfloor>_" [90, 91] 90)
 | 
59  | 
  map_upd_s  :: "('a ~=> 'b) => 'a set => 'b => ('a ~=> 'b)"
 | 
|
60  | 
				    		 ("_/'(_/{\<mapsto>}/_')" [900,0,0]900)
 | 
|
61  | 
  map_subst :: "('a ~=> 'b) => 'b => 'b => 
 | 
|
62  | 
	        ('a ~=> 'b)"			 ("_/'(_\<leadsto>_/')"    [900,0,0]900)
 | 
|
63  | 
 "@chg_map" :: "('a ~=> 'b) => 'a => ('b => 'b) => ('a ~=> 'b)"
 | 
|
64  | 
					  ("_/'(_/\<mapsto>\<lambda>_. _')"  [900,0,0,0] 900)
 | 
|
| 5300 | 65  | 
|
66  | 
translations  | 
|
| 13890 | 67  | 
"empty" => "_K None"  | 
68  | 
"empty" <= "%x. None"  | 
|
| 5300 | 69  | 
|
| 14100 | 70  | 
"m(x\<mapsto>\<lambda>y. f)" == "chg_map (\<lambda>y. f) x m"  | 
| 3981 | 71  | 
|
| 14180 | 72  | 
"_MapUpd m (_Maplets xy ms)" == "_MapUpd (_MapUpd m xy) ms"  | 
73  | 
"_MapUpd m (_maplet x y)" == "m(x:=Some y)"  | 
|
74  | 
"_MapUpd m (_maplets x y)" == "map_upds m x y"  | 
|
75  | 
"_Map ms" == "_MapUpd empty ms"  | 
|
76  | 
"_Map (_Maplets ms1 ms2)" <= "_MapUpd (_Map ms1) ms2"  | 
|
77  | 
"_Maplets ms1 (_Maplets ms2 ms3)" <= "_Maplets (_Maplets ms1 ms2) ms3"  | 
|
78  | 
||
| 3981 | 79  | 
defs  | 
| 13908 | 80  | 
chg_map_def: "chg_map f a m == case m a of None => m | Some b => m(a|->f b)"  | 
| 3981 | 81  | 
|
| 14100 | 82  | 
map_add_def: "m1++m2 == %x. case m2 x of None => m1 x | Some y => Some y"  | 
83  | 
restrict_map_def: "m|_A == %x. if x : A then m x else None"  | 
|
| 14025 | 84  | 
|
85  | 
map_upds_def: "m(xs [|->] ys) == m ++ map_of (rev(zip xs ys))"  | 
|
| 14100 | 86  | 
map_upd_s_def: "m(as{|->}b) == %x. if x : as then Some b else m x"
 | 
87  | 
map_subst_def: "m(a~>b) == %x. if m x = Some a then Some b else m x"  | 
|
| 3981 | 88  | 
|
| 13908 | 89  | 
dom_def: "dom(m) == {a. m a ~= None}"
 | 
| 14025 | 90  | 
ran_def: "ran(m) == {b. EX a. m a = Some b}"
 | 
| 3981 | 91  | 
|
| 14376 | 92  | 
map_le_def: "m\<^isub>1 \<subseteq>\<^sub>m m\<^isub>2 == ALL a : dom m\<^isub>1. m\<^isub>1 a = m\<^isub>2 a"  | 
| 13910 | 93  | 
|
| 5183 | 94  | 
primrec  | 
95  | 
"map_of [] = empty"  | 
|
| 5300 | 96  | 
"map_of (p#ps) = (map_of ps)(fst p |-> snd p)"  | 
97  | 
||
| 13908 | 98  | 
|
| 14100 | 99  | 
subsection {* @{term empty} *}
 | 
| 13908 | 100  | 
|
| 13910 | 101  | 
lemma empty_upd_none[simp]: "empty(x := None) = empty"  | 
| 13908 | 102  | 
apply (rule ext)  | 
103  | 
apply (simp (no_asm))  | 
|
104  | 
done  | 
|
| 13910 | 105  | 
|
| 13908 | 106  | 
|
107  | 
(* FIXME: what is this sum_case nonsense?? *)  | 
|
| 13910 | 108  | 
lemma sum_case_empty_empty[simp]: "sum_case empty empty = empty"  | 
| 13908 | 109  | 
apply (rule ext)  | 
110  | 
apply (simp (no_asm) split add: sum.split)  | 
|
111  | 
done  | 
|
112  | 
||
| 14100 | 113  | 
subsection {* @{term map_upd} *}
 | 
| 13908 | 114  | 
|
115  | 
lemma map_upd_triv: "t k = Some x ==> t(k|->x) = t"  | 
|
116  | 
apply (rule ext)  | 
|
117  | 
apply (simp (no_asm_simp))  | 
|
118  | 
done  | 
|
119  | 
||
| 13910 | 120  | 
lemma map_upd_nonempty[simp]: "t(k|->x) ~= empty"  | 
| 13908 | 121  | 
apply safe  | 
| 14208 | 122  | 
apply (drule_tac x = k in fun_cong)  | 
| 13908 | 123  | 
apply (simp (no_asm_use))  | 
124  | 
done  | 
|
125  | 
||
| 14100 | 126  | 
lemma map_upd_eqD1: "m(a\<mapsto>x) = n(a\<mapsto>y) \<Longrightarrow> x = y"  | 
127  | 
by (drule fun_cong [of _ _ a], auto)  | 
|
128  | 
||
129  | 
lemma map_upd_Some_unfold:  | 
|
130  | 
"((m(a|->b)) x = Some y) = (x = a \<and> b = y \<or> x \<noteq> a \<and> m x = Some y)"  | 
|
131  | 
by auto  | 
|
132  | 
||
| 13908 | 133  | 
lemma finite_range_updI: "finite (range f) ==> finite (range (f(a|->b)))"  | 
134  | 
apply (unfold image_def)  | 
|
135  | 
apply (simp (no_asm_use) add: full_SetCompr_eq)  | 
|
136  | 
apply (rule finite_subset)  | 
|
| 14208 | 137  | 
prefer 2 apply assumption  | 
| 13908 | 138  | 
apply auto  | 
139  | 
done  | 
|
140  | 
||
141  | 
||
142  | 
(* FIXME: what is this sum_case nonsense?? *)  | 
|
| 14100 | 143  | 
subsection {* @{term sum_case} and @{term empty}/@{term map_upd} *}
 | 
| 13908 | 144  | 
|
| 13910 | 145  | 
lemma sum_case_map_upd_empty[simp]:  | 
146  | 
"sum_case (m(k|->y)) empty = (sum_case m empty)(Inl k|->y)"  | 
|
| 13908 | 147  | 
apply (rule ext)  | 
148  | 
apply (simp (no_asm) split add: sum.split)  | 
|
149  | 
done  | 
|
150  | 
||
| 13910 | 151  | 
lemma sum_case_empty_map_upd[simp]:  | 
152  | 
"sum_case empty (m(k|->y)) = (sum_case empty m)(Inr k|->y)"  | 
|
| 13908 | 153  | 
apply (rule ext)  | 
154  | 
apply (simp (no_asm) split add: sum.split)  | 
|
155  | 
done  | 
|
156  | 
||
| 13910 | 157  | 
lemma sum_case_map_upd_map_upd[simp]:  | 
158  | 
"sum_case (m1(k1|->y1)) (m2(k2|->y2)) = (sum_case (m1(k1|->y1)) m2)(Inr k2|->y2)"  | 
|
| 13908 | 159  | 
apply (rule ext)  | 
160  | 
apply (simp (no_asm) split add: sum.split)  | 
|
161  | 
done  | 
|
162  | 
||
163  | 
||
| 14100 | 164  | 
subsection {* @{term chg_map} *}
 | 
| 13908 | 165  | 
|
| 13910 | 166  | 
lemma chg_map_new[simp]: "m a = None ==> chg_map f a m = m"  | 
| 14208 | 167  | 
by (unfold chg_map_def, auto)  | 
| 13908 | 168  | 
|
| 13910 | 169  | 
lemma chg_map_upd[simp]: "m a = Some b ==> chg_map f a m = m(a|->f b)"  | 
| 14208 | 170  | 
by (unfold chg_map_def, auto)  | 
| 13908 | 171  | 
|
| 14537 | 172  | 
lemma chg_map_other [simp]: "a \<noteq> b \<Longrightarrow> chg_map f a m b = m b"  | 
173  | 
by (auto simp: chg_map_def split add: option.split)  | 
|
174  | 
||
| 13908 | 175  | 
|
| 14100 | 176  | 
subsection {* @{term map_of} *}
 | 
| 13908 | 177  | 
|
| 
15110
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
178  | 
lemma map_of_zip_is_None[simp]:  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
179  | 
"length xs = length ys \<Longrightarrow> (map_of (zip xs ys) x = None) = (x \<notin> set xs)"  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
180  | 
by (induct rule:list_induct2, simp_all)  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
181  | 
|
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
182  | 
lemma finite_range_map_of: "finite (range (map_of xys))"  | 
| 15251 | 183  | 
apply (induct xys)  | 
| 
15110
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
184  | 
apply (simp_all (no_asm) add: image_constant)  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
185  | 
apply (rule finite_subset)  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
186  | 
prefer 2 apply assumption  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
187  | 
apply auto  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
188  | 
done  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
189  | 
|
| 13908 | 190  | 
lemma map_of_SomeD [rule_format (no_asm)]: "map_of xs k = Some y --> (k,y):set xs"  | 
| 15251 | 191  | 
by (induct "xs", auto)  | 
| 13908 | 192  | 
|
193  | 
lemma map_of_mapk_SomeI [rule_format (no_asm)]: "inj f ==> map_of t k = Some x -->  | 
|
194  | 
map_of (map (split (%k. Pair (f k))) t) (f k) = Some x"  | 
|
| 15251 | 195  | 
apply (induct "t")  | 
| 13908 | 196  | 
apply (auto simp add: inj_eq)  | 
197  | 
done  | 
|
198  | 
||
199  | 
lemma weak_map_of_SomeI [rule_format (no_asm)]: "(k, x) : set l --> (? x. map_of l k = Some x)"  | 
|
| 15251 | 200  | 
by (induct "l", auto)  | 
| 13908 | 201  | 
|
202  | 
lemma map_of_filter_in:  | 
|
203  | 
"[| map_of xs k = Some z; P k z |] ==> map_of (filter (split P) xs) k = Some z"  | 
|
204  | 
apply (rule mp)  | 
|
| 14208 | 205  | 
prefer 2 apply assumption  | 
| 13908 | 206  | 
apply (erule thin_rl)  | 
| 15251 | 207  | 
apply (induct "xs", auto)  | 
| 13908 | 208  | 
done  | 
209  | 
||
210  | 
lemma map_of_map: "map_of (map (%(a,b). (a,f b)) xs) x = option_map f (map_of xs x)"  | 
|
| 15251 | 211  | 
by (induct "xs", auto)  | 
| 13908 | 212  | 
|
213  | 
||
| 14100 | 214  | 
subsection {* @{term option_map} related *}
 | 
| 13908 | 215  | 
|
| 13910 | 216  | 
lemma option_map_o_empty[simp]: "option_map f o empty = empty"  | 
| 13908 | 217  | 
apply (rule ext)  | 
218  | 
apply (simp (no_asm))  | 
|
219  | 
done  | 
|
220  | 
||
| 13910 | 221  | 
lemma option_map_o_map_upd[simp]:  | 
222  | 
"option_map f o m(a|->b) = (option_map f o m)(a|->f b)"  | 
|
| 13908 | 223  | 
apply (rule ext)  | 
224  | 
apply (simp (no_asm))  | 
|
225  | 
done  | 
|
226  | 
||
227  | 
||
| 14100 | 228  | 
subsection {* @{text "++"} *}
 | 
| 13908 | 229  | 
|
| 14025 | 230  | 
lemma map_add_empty[simp]: "m ++ empty = m"  | 
231  | 
apply (unfold map_add_def)  | 
|
| 13908 | 232  | 
apply (simp (no_asm))  | 
233  | 
done  | 
|
234  | 
||
| 14025 | 235  | 
lemma empty_map_add[simp]: "empty ++ m = m"  | 
236  | 
apply (unfold map_add_def)  | 
|
| 13908 | 237  | 
apply (rule ext)  | 
238  | 
apply (simp split add: option.split)  | 
|
239  | 
done  | 
|
240  | 
||
| 14025 | 241  | 
lemma map_add_assoc[simp]: "m1 ++ (m2 ++ m3) = (m1 ++ m2) ++ m3"  | 
242  | 
apply(rule ext)  | 
|
243  | 
apply(simp add: map_add_def split:option.split)  | 
|
244  | 
done  | 
|
245  | 
||
246  | 
lemma map_add_Some_iff:  | 
|
| 13908 | 247  | 
"((m ++ n) k = Some x) = (n k = Some x | n k = None & m k = Some x)"  | 
| 14025 | 248  | 
apply (unfold map_add_def)  | 
| 13908 | 249  | 
apply (simp (no_asm) split add: option.split)  | 
250  | 
done  | 
|
251  | 
||
| 14025 | 252  | 
lemmas map_add_SomeD = map_add_Some_iff [THEN iffD1, standard]  | 
253  | 
declare map_add_SomeD [dest!]  | 
|
| 13908 | 254  | 
|
| 14025 | 255  | 
lemma map_add_find_right[simp]: "!!xx. n k = Some xx ==> (m ++ n) k = Some xx"  | 
| 14208 | 256  | 
by (subst map_add_Some_iff, fast)  | 
| 13908 | 257  | 
|
| 14025 | 258  | 
lemma map_add_None [iff]: "((m ++ n) k = None) = (n k = None & m k = None)"  | 
259  | 
apply (unfold map_add_def)  | 
|
| 13908 | 260  | 
apply (simp (no_asm) split add: option.split)  | 
261  | 
done  | 
|
262  | 
||
| 14025 | 263  | 
lemma map_add_upd[simp]: "f ++ g(x|->y) = (f ++ g)(x|->y)"  | 
264  | 
apply (unfold map_add_def)  | 
|
| 14208 | 265  | 
apply (rule ext, auto)  | 
| 13908 | 266  | 
done  | 
267  | 
||
| 14186 | 268  | 
lemma map_add_upds[simp]: "m1 ++ (m2(xs[\<mapsto>]ys)) = (m1++m2)(xs[\<mapsto>]ys)"  | 
269  | 
by(simp add:map_upds_def)  | 
|
270  | 
||
| 14025 | 271  | 
lemma map_of_append[simp]: "map_of (xs@ys) = map_of ys ++ map_of xs"  | 
272  | 
apply (unfold map_add_def)  | 
|
| 15251 | 273  | 
apply (induct "xs")  | 
| 13908 | 274  | 
apply (simp (no_asm))  | 
275  | 
apply (rule ext)  | 
|
276  | 
apply (simp (no_asm_simp) split add: option.split)  | 
|
277  | 
done  | 
|
278  | 
||
279  | 
declare fun_upd_apply [simp del]  | 
|
| 14025 | 280  | 
lemma finite_range_map_of_map_add:  | 
281  | 
"finite (range f) ==> finite (range (f ++ map_of l))"  | 
|
| 15251 | 282  | 
apply (induct "l", auto)  | 
| 13908 | 283  | 
apply (erule finite_range_updI)  | 
284  | 
done  | 
|
285  | 
declare fun_upd_apply [simp]  | 
|
286  | 
||
| 14100 | 287  | 
subsection {* @{term restrict_map} *}
 | 
288  | 
||
| 14186 | 289  | 
lemma restrict_map_to_empty[simp]: "m\<lfloor>{} = empty"
 | 
290  | 
by(simp add: restrict_map_def)  | 
|
291  | 
||
292  | 
lemma restrict_map_empty[simp]: "empty\<lfloor>D = empty"  | 
|
293  | 
by(simp add: restrict_map_def)  | 
|
294  | 
||
| 14100 | 295  | 
lemma restrict_in [simp]: "x \<in> A \<Longrightarrow> (m\<lfloor>A) x = m x"  | 
296  | 
by (auto simp: restrict_map_def)  | 
|
297  | 
||
298  | 
lemma restrict_out [simp]: "x \<notin> A \<Longrightarrow> (m\<lfloor>A) x = None"  | 
|
299  | 
by (auto simp: restrict_map_def)  | 
|
300  | 
||
301  | 
lemma ran_restrictD: "y \<in> ran (m\<lfloor>A) \<Longrightarrow> \<exists>x\<in>A. m x = Some y"  | 
|
302  | 
by (auto simp: restrict_map_def ran_def split: split_if_asm)  | 
|
303  | 
||
| 14186 | 304  | 
lemma dom_restrict [simp]: "dom (m\<lfloor>A) = dom m \<inter> A"  | 
| 14100 | 305  | 
by (auto simp: restrict_map_def dom_def split: split_if_asm)  | 
306  | 
||
307  | 
lemma restrict_upd_same [simp]: "m(x\<mapsto>y)\<lfloor>(-{x}) = m\<lfloor>(-{x})"
 | 
|
308  | 
by (rule ext, auto simp: restrict_map_def)  | 
|
309  | 
||
310  | 
lemma restrict_restrict [simp]: "m\<lfloor>A\<lfloor>B = m\<lfloor>(A\<inter>B)"  | 
|
311  | 
by (rule ext, auto simp: restrict_map_def)  | 
|
312  | 
||
| 14186 | 313  | 
lemma restrict_fun_upd[simp]:  | 
314  | 
 "m(x := y)\<lfloor>D = (if x \<in> D then (m\<lfloor>(D-{x}))(x := y) else m\<lfloor>D)"
 | 
|
315  | 
by(simp add: restrict_map_def expand_fun_eq)  | 
|
316  | 
||
317  | 
lemma fun_upd_None_restrict[simp]:  | 
|
318  | 
  "(m\<lfloor>D)(x := None) = (if x:D then m\<lfloor>(D - {x}) else m\<lfloor>D)"
 | 
|
319  | 
by(simp add: restrict_map_def expand_fun_eq)  | 
|
320  | 
||
321  | 
lemma fun_upd_restrict:  | 
|
322  | 
 "(m\<lfloor>D)(x := y) = (m\<lfloor>(D-{x}))(x := y)"
 | 
|
323  | 
by(simp add: restrict_map_def expand_fun_eq)  | 
|
324  | 
||
325  | 
lemma fun_upd_restrict_conv[simp]:  | 
|
326  | 
 "x \<in> D \<Longrightarrow> (m\<lfloor>D)(x := y) = (m\<lfloor>(D-{x}))(x := y)"
 | 
|
327  | 
by(simp add: restrict_map_def expand_fun_eq)  | 
|
328  | 
||
| 14100 | 329  | 
|
330  | 
subsection {* @{term map_upds} *}
 | 
|
| 14025 | 331  | 
|
332  | 
lemma map_upds_Nil1[simp]: "m([] [|->] bs) = m"  | 
|
333  | 
by(simp add:map_upds_def)  | 
|
334  | 
||
335  | 
lemma map_upds_Nil2[simp]: "m(as [|->] []) = m"  | 
|
336  | 
by(simp add:map_upds_def)  | 
|
337  | 
||
338  | 
lemma map_upds_Cons[simp]: "m(a#as [|->] b#bs) = (m(a|->b))(as[|->]bs)"  | 
|
339  | 
by(simp add:map_upds_def)  | 
|
340  | 
||
| 14187 | 341  | 
lemma map_upds_append1[simp]: "\<And>ys m. size xs < size ys \<Longrightarrow>  | 
342  | 
m(xs@[x] [\<mapsto>] ys) = m(xs [\<mapsto>] ys)(x \<mapsto> ys!size xs)"  | 
|
343  | 
apply(induct xs)  | 
|
344  | 
apply(clarsimp simp add:neq_Nil_conv)  | 
|
| 14208 | 345  | 
apply (case_tac ys, simp, simp)  | 
| 14187 | 346  | 
done  | 
347  | 
||
348  | 
lemma map_upds_list_update2_drop[simp]:  | 
|
349  | 
"\<And>m ys i. \<lbrakk>size xs \<le> i; i < size ys\<rbrakk>  | 
|
350  | 
\<Longrightarrow> m(xs[\<mapsto>]ys[i:=y]) = m(xs[\<mapsto>]ys)"  | 
|
| 14208 | 351  | 
apply (induct xs, simp)  | 
352  | 
apply (case_tac ys, simp)  | 
|
| 14187 | 353  | 
apply(simp split:nat.split)  | 
354  | 
done  | 
|
| 14025 | 355  | 
|
356  | 
lemma map_upd_upds_conv_if: "!!x y ys f.  | 
|
357  | 
(f(x|->y))(xs [|->] ys) =  | 
|
358  | 
(if x : set(take (length ys) xs) then f(xs [|->] ys)  | 
|
359  | 
else (f(xs [|->] ys))(x|->y))"  | 
|
| 14208 | 360  | 
apply (induct xs, simp)  | 
| 14025 | 361  | 
apply(case_tac ys)  | 
362  | 
apply(auto split:split_if simp:fun_upd_twist)  | 
|
363  | 
done  | 
|
364  | 
||
365  | 
lemma map_upds_twist [simp]:  | 
|
366  | 
"a ~: set as ==> m(a|->b)(as[|->]bs) = m(as[|->]bs)(a|->b)"  | 
|
367  | 
apply(insert set_take_subset)  | 
|
368  | 
apply (fastsimp simp add: map_upd_upds_conv_if)  | 
|
369  | 
done  | 
|
370  | 
||
371  | 
lemma map_upds_apply_nontin[simp]:  | 
|
372  | 
"!!ys. x ~: set xs ==> (f(xs[|->]ys)) x = f x"  | 
|
| 14208 | 373  | 
apply (induct xs, simp)  | 
| 14025 | 374  | 
apply(case_tac ys)  | 
375  | 
apply(auto simp: map_upd_upds_conv_if)  | 
|
376  | 
done  | 
|
377  | 
||
| 14300 | 378  | 
lemma fun_upds_append_drop[simp]:  | 
379  | 
"!!m ys. size xs = size ys \<Longrightarrow> m(xs@zs[\<mapsto>]ys) = m(xs[\<mapsto>]ys)"  | 
|
380  | 
apply(induct xs)  | 
|
381  | 
apply (simp)  | 
|
382  | 
apply(case_tac ys)  | 
|
383  | 
apply simp_all  | 
|
384  | 
done  | 
|
385  | 
||
386  | 
lemma fun_upds_append2_drop[simp]:  | 
|
387  | 
"!!m ys. size xs = size ys \<Longrightarrow> m(xs[\<mapsto>]ys@zs) = m(xs[\<mapsto>]ys)"  | 
|
388  | 
apply(induct xs)  | 
|
389  | 
apply (simp)  | 
|
390  | 
apply(case_tac ys)  | 
|
391  | 
apply simp_all  | 
|
392  | 
done  | 
|
393  | 
||
394  | 
||
| 14186 | 395  | 
lemma restrict_map_upds[simp]: "!!m ys.  | 
396  | 
\<lbrakk> length xs = length ys; set xs \<subseteq> D \<rbrakk>  | 
|
397  | 
\<Longrightarrow> m(xs [\<mapsto>] ys)\<lfloor>D = (m\<lfloor>(D - set xs))(xs [\<mapsto>] ys)"  | 
|
| 14208 | 398  | 
apply (induct xs, simp)  | 
399  | 
apply (case_tac ys, simp)  | 
|
| 14186 | 400  | 
apply(simp add:Diff_insert[symmetric] insert_absorb)  | 
401  | 
apply(simp add: map_upd_upds_conv_if)  | 
|
402  | 
done  | 
|
403  | 
||
404  | 
||
| 14100 | 405  | 
subsection {* @{term map_upd_s} *}
 | 
406  | 
||
407  | 
lemma map_upd_s_apply [simp]:  | 
|
408  | 
  "(m(as{|->}b)) x = (if x : as then Some b else m x)"
 | 
|
409  | 
by (simp add: map_upd_s_def)  | 
|
410  | 
||
411  | 
lemma map_subst_apply [simp]:  | 
|
412  | 
"(m(a~>b)) x = (if m x = Some a then Some b else m x)"  | 
|
413  | 
by (simp add: map_subst_def)  | 
|
414  | 
||
415  | 
subsection {* @{term dom} *}
 | 
|
| 13908 | 416  | 
|
417  | 
lemma domI: "m a = Some b ==> a : dom m"  | 
|
| 14208 | 418  | 
by (unfold dom_def, auto)  | 
| 14100 | 419  | 
(* declare domI [intro]? *)  | 
| 13908 | 420  | 
|
421  | 
lemma domD: "a : dom m ==> ? b. m a = Some b"  | 
|
| 14208 | 422  | 
by (unfold dom_def, auto)  | 
| 13908 | 423  | 
|
| 13910 | 424  | 
lemma domIff[iff]: "(a : dom m) = (m a ~= None)"  | 
| 14208 | 425  | 
by (unfold dom_def, auto)  | 
| 13908 | 426  | 
declare domIff [simp del]  | 
427  | 
||
| 13910 | 428  | 
lemma dom_empty[simp]: "dom empty = {}"
 | 
| 13908 | 429  | 
apply (unfold dom_def)  | 
430  | 
apply (simp (no_asm))  | 
|
431  | 
done  | 
|
432  | 
||
| 13910 | 433  | 
lemma dom_fun_upd[simp]:  | 
434  | 
 "dom(f(x := y)) = (if y=None then dom f - {x} else insert x (dom f))"
 | 
|
435  | 
by (simp add:dom_def) blast  | 
|
| 13908 | 436  | 
|
| 13937 | 437  | 
lemma dom_map_of: "dom(map_of xys) = {x. \<exists>y. (x,y) : set xys}"
 | 
438  | 
apply(induct xys)  | 
|
439  | 
apply(auto simp del:fun_upd_apply)  | 
|
440  | 
done  | 
|
441  | 
||
| 
15110
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
442  | 
lemma dom_map_of_zip[simp]: "[| length xs = length ys; distinct xs |] ==>  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
443  | 
dom(map_of(zip xs ys)) = set xs"  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
444  | 
by(induct rule: list_induct2, simp_all)  | 
| 
 
78b5636eabc7
Added a number of new thms and the new function remove1
 
nipkow 
parents: 
14739 
diff
changeset
 | 
445  | 
|
| 13908 | 446  | 
lemma finite_dom_map_of: "finite (dom (map_of l))"  | 
447  | 
apply (unfold dom_def)  | 
|
| 15251 | 448  | 
apply (induct "l")  | 
| 13908 | 449  | 
apply (auto simp add: insert_Collect [symmetric])  | 
450  | 
done  | 
|
451  | 
||
| 14025 | 452  | 
lemma dom_map_upds[simp]:  | 
453  | 
"!!m ys. dom(m(xs[|->]ys)) = set(take (length ys) xs) Un dom m"  | 
|
| 14208 | 454  | 
apply (induct xs, simp)  | 
455  | 
apply (case_tac ys, auto)  | 
|
| 14025 | 456  | 
done  | 
| 13910 | 457  | 
|
| 14025 | 458  | 
lemma dom_map_add[simp]: "dom(m++n) = dom n Un dom m"  | 
| 14208 | 459  | 
by (unfold dom_def, auto)  | 
| 13910 | 460  | 
|
461  | 
lemma dom_overwrite[simp]:  | 
|
462  | 
 "dom(f(g|A)) = (dom f  - {a. a : A - dom g}) Un {a. a : A Int dom g}"
 | 
|
463  | 
by(auto simp add: dom_def overwrite_def)  | 
|
| 13908 | 464  | 
|
| 14027 | 465  | 
lemma map_add_comm: "dom m1 \<inter> dom m2 = {} \<Longrightarrow> m1++m2 = m2++m1"
 | 
466  | 
apply(rule ext)  | 
|
467  | 
apply(fastsimp simp:map_add_def split:option.split)  | 
|
468  | 
done  | 
|
469  | 
||
| 14100 | 470  | 
subsection {* @{term ran} *}
 | 
471  | 
||
472  | 
lemma ranI: "m a = Some b ==> b : ran m"  | 
|
473  | 
by (auto simp add: ran_def)  | 
|
474  | 
(* declare ranI [intro]? *)  | 
|
| 13908 | 475  | 
|
| 13910 | 476  | 
lemma ran_empty[simp]: "ran empty = {}"
 | 
| 13908 | 477  | 
apply (unfold ran_def)  | 
478  | 
apply (simp (no_asm))  | 
|
479  | 
done  | 
|
480  | 
||
| 13910 | 481  | 
lemma ran_map_upd[simp]: "m a = None ==> ran(m(a|->b)) = insert b (ran m)"  | 
| 14208 | 482  | 
apply (unfold ran_def, auto)  | 
| 13908 | 483  | 
apply (subgoal_tac "~ (aa = a) ")  | 
484  | 
apply auto  | 
|
485  | 
done  | 
|
| 13910 | 486  | 
|
| 14100 | 487  | 
subsection {* @{text "map_le"} *}
 | 
| 13910 | 488  | 
|
| 13912 | 489  | 
lemma map_le_empty [simp]: "empty \<subseteq>\<^sub>m g"  | 
| 13910 | 490  | 
by(simp add:map_le_def)  | 
491  | 
||
| 14187 | 492  | 
lemma [simp]: "f(x := None) \<subseteq>\<^sub>m f"  | 
493  | 
by(force simp add:map_le_def)  | 
|
494  | 
||
| 13910 | 495  | 
lemma map_le_upd[simp]: "f \<subseteq>\<^sub>m g ==> f(a := b) \<subseteq>\<^sub>m g(a := b)"  | 
496  | 
by(fastsimp simp add:map_le_def)  | 
|
497  | 
||
| 14187 | 498  | 
lemma [simp]: "m1 \<subseteq>\<^sub>m m2 \<Longrightarrow> m1(x := None) \<subseteq>\<^sub>m m2(x \<mapsto> y)"  | 
499  | 
by(force simp add:map_le_def)  | 
|
500  | 
||
| 13910 | 501  | 
lemma map_le_upds[simp]:  | 
502  | 
"!!f g bs. f \<subseteq>\<^sub>m g ==> f(as [|->] bs) \<subseteq>\<^sub>m g(as [|->] bs)"  | 
|
| 14208 | 503  | 
apply (induct as, simp)  | 
504  | 
apply (case_tac bs, auto)  | 
|
| 14025 | 505  | 
done  | 
| 13908 | 506  | 
|
| 14033 | 507  | 
lemma map_le_implies_dom_le: "(f \<subseteq>\<^sub>m g) \<Longrightarrow> (dom f \<subseteq> dom g)"  | 
508  | 
by (fastsimp simp add: map_le_def dom_def)  | 
|
509  | 
||
510  | 
lemma map_le_refl [simp]: "f \<subseteq>\<^sub>m f"  | 
|
511  | 
by (simp add: map_le_def)  | 
|
512  | 
||
| 14187 | 513  | 
lemma map_le_trans[trans]: "\<lbrakk> m1 \<subseteq>\<^sub>m m2; m2 \<subseteq>\<^sub>m m3\<rbrakk> \<Longrightarrow> m1 \<subseteq>\<^sub>m m3"  | 
514  | 
by(force simp add:map_le_def)  | 
|
| 14033 | 515  | 
|
516  | 
lemma map_le_antisym: "\<lbrakk> f \<subseteq>\<^sub>m g; g \<subseteq>\<^sub>m f \<rbrakk> \<Longrightarrow> f = g"  | 
|
517  | 
apply (unfold map_le_def)  | 
|
518  | 
apply (rule ext)  | 
|
| 14208 | 519  | 
apply (case_tac "x \<in> dom f", simp)  | 
520  | 
apply (case_tac "x \<in> dom g", simp, fastsimp)  | 
|
| 14033 | 521  | 
done  | 
522  | 
||
523  | 
lemma map_le_map_add [simp]: "f \<subseteq>\<^sub>m (g ++ f)"  | 
|
524  | 
by (fastsimp simp add: map_le_def)  | 
|
525  | 
||
| 3981 | 526  | 
end  |