| author | boehmes | 
| Mon, 06 Dec 2010 15:38:02 +0100 | |
| changeset 41057 | 8dbc951a291c | 
| parent 39976 | 2474347538b8 | 
| child 44890 | 22f665a2e91c | 
| permissions | -rw-r--r-- | 
| 13673 | 1 | (* Title: HOL/MicroJava/Comp/AuxLemmas.thy | 
| 2 | Author: Martin Strecker | |
| 3 | *) | |
| 4 | ||
| 5 | ||
| 6 | (* Auxiliary Lemmas *) | |
| 7 | ||
| 15860 | 8 | theory AuxLemmas | 
| 9 | imports "../J/JBasis" | |
| 10 | begin | |
| 13673 | 11 | |
| 12 | (**********************************************************************) | |
| 13 | (* List.thy *) | |
| 14 | (**********************************************************************) | |
| 15 | ||
| 16 | ||
| 17 | ||
| 18 | lemma app_nth_greater_len [rule_format (no_asm), simp]: | |
| 19 | "\<forall> ind. length pre \<le> ind \<longrightarrow> (pre @ a # post) ! (Suc ind) = (pre @ post) ! ind" | |
| 20 | apply (induct "pre") | |
| 21 | apply auto | |
| 22 | apply (case_tac ind) | |
| 23 | apply auto | |
| 24 | done | |
| 25 | ||
| 26 | lemma length_takeWhile: "v \<in> set xs \<Longrightarrow> length (takeWhile (%z. z~=v) xs) < length xs" | |
| 27 | by (induct xs, auto) | |
| 28 | ||
| 29 | lemma nth_length_takeWhile [simp]: | |
| 30 | "v \<in> set xs \<Longrightarrow> xs ! (length (takeWhile (%z. z~=v) xs)) = v" | |
| 31 | by (induct xs, auto) | |
| 32 | ||
| 33 | ||
| 34 | lemma map_list_update [simp]: | |
| 35 | "\<lbrakk> x \<in> set xs; distinct xs\<rbrakk> \<Longrightarrow> | |
| 36 | (map f xs) [length (takeWhile (\<lambda>z. z \<noteq> x) xs) := v] = | |
| 37 | map (f(x:=v)) xs" | |
| 38 | apply (induct xs) | |
| 39 | apply simp | |
| 40 | apply (case_tac "x=a") | |
| 41 | apply auto | |
| 42 | done | |
| 43 | ||
| 44 | ||
| 45 | (**********************************************************************) | |
| 46 | (* Product_Type.thy *) | |
| 47 | (**********************************************************************) | |
| 48 | ||
| 49 | ||
| 50 | lemma split_compose: "(split f) \<circ> (\<lambda> (a,b). ((fa a), (fb b))) = | |
| 51 | (\<lambda> (a,b). (f (fa a) (fb b)))" | |
| 52 | by (simp add: split_def o_def) | |
| 53 | ||
| 54 | lemma split_iter: "(\<lambda> (a,b,c). ((g1 a), (g2 b), (g3 c))) = | |
| 55 | (\<lambda> (a,p). ((g1 a), (\<lambda> (b, c). ((g2 b), (g3 c))) p))" | |
| 56 | by (simp add: split_def o_def) | |
| 57 | ||
| 58 | ||
| 59 | (**********************************************************************) | |
| 60 | (* Set.thy *) | |
| 61 | (**********************************************************************) | |
| 62 | ||
| 63 | lemma singleton_in_set: "A = {a} \<Longrightarrow> a \<in> A" by simp
 | |
| 64 | ||
| 65 | (**********************************************************************) | |
| 66 | (* Map.thy *) | |
| 67 | (**********************************************************************) | |
| 68 | ||
| 69 | lemma the_map_upd: "(the \<circ> f(x\<mapsto>v)) = (the \<circ> f)(x:=v)" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 70 | by (simp add: fun_eq_iff) | 
| 13673 | 71 | |
| 72 | lemma map_of_in_set: | |
| 73 | "(map_of xs x = None) = (x \<notin> set (map fst xs))" | |
| 74 | by (induct xs, auto) | |
| 75 | ||
| 76 | lemma map_map_upd [simp]: | |
| 77 | "y \<notin> set xs \<Longrightarrow> map (the \<circ> f(y\<mapsto>v)) xs = map (the \<circ> f) xs" | |
| 78 | by (simp add: the_map_upd) | |
| 79 | ||
| 39917 
b85bfa89a387
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
 haftmann parents: 
39302diff
changeset | 80 | lemma map_map_upds [simp]: | 
| 
b85bfa89a387
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
 haftmann parents: 
39302diff
changeset | 81 | "(\<forall>y\<in>set ys. y \<notin> set xs) \<Longrightarrow> map (the \<circ> f(ys[\<mapsto>]vs)) xs = map (the \<circ> f) xs" | 
| 
b85bfa89a387
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
 haftmann parents: 
39302diff
changeset | 82 | apply (induct xs arbitrary: f vs) | 
| 14025 | 83 | apply simp | 
| 84 | apply fastsimp | |
| 13673 | 85 | done | 
| 86 | ||
| 39917 
b85bfa89a387
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
 haftmann parents: 
39302diff
changeset | 87 | lemma map_upds_distinct [simp]: | 
| 
b85bfa89a387
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
 haftmann parents: 
39302diff
changeset | 88 | "distinct ys \<Longrightarrow> length ys = length vs \<Longrightarrow> map (the \<circ> f(ys[\<mapsto>]vs)) ys = vs" | 
| 39976 | 89 | apply (induct ys arbitrary: f vs) | 
| 13673 | 90 | apply simp | 
| 91 | apply (case_tac vs) | |
| 39917 
b85bfa89a387
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
 haftmann parents: 
39302diff
changeset | 92 | apply simp_all | 
| 13673 | 93 | done | 
| 94 | ||
| 39917 
b85bfa89a387
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
 haftmann parents: 
39302diff
changeset | 95 | lemma map_of_map_as_map_upd: | 
| 
b85bfa89a387
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
 haftmann parents: 
39302diff
changeset | 96 | "distinct (map f zs) \<Longrightarrow> map_of (map (\<lambda> p. (f p, g p)) zs) = empty (map f zs [\<mapsto>] map g zs)" | 
| 
b85bfa89a387
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
 haftmann parents: 
39302diff
changeset | 97 | by (induct zs) auto | 
| 13673 | 98 | |
| 99 | (* In analogy to Map.map_of_SomeD *) | |
| 100 | lemma map_upds_SomeD [rule_format (no_asm)]: | |
| 101 | "\<forall> m ys. (m(xs[\<mapsto>]ys)) k = Some y \<longrightarrow> k \<in> (set xs) \<or> (m k = Some y)" | |
| 102 | apply (induct xs) | |
| 103 | apply simp | |
| 14025 | 104 | apply auto | 
| 105 | apply(case_tac ys) | |
| 13673 | 106 | apply auto | 
| 107 | done | |
| 108 | ||
| 109 | lemma map_of_upds_SomeD: "(map_of m (xs[\<mapsto>]ys)) k = Some y | |
| 110 | \<Longrightarrow> k \<in> (set (xs @ map fst m))" | |
| 111 | by (auto dest: map_upds_SomeD map_of_SomeD fst_in_set_lemma) | |
| 112 | ||
| 113 | ||
| 114 | lemma map_of_map_prop [rule_format (no_asm)]: | |
| 115 | "(map_of (map f xs) k = Some v) \<longrightarrow> | |
| 116 | (\<forall> x \<in> set xs. (P1 x)) \<longrightarrow> | |
| 117 | (\<forall> x. (P1 x) \<longrightarrow> (P2 (f x))) \<longrightarrow> | |
| 118 | (P2(k, v))" | |
| 119 | by (induct xs,auto) | |
| 120 | ||
| 121 | lemma map_of_map2: "\<forall> x \<in> set xs. (fst (f x)) = (fst x) \<Longrightarrow> | |
| 30235 
58d147683393
Made Option a separate theory and renamed option_map to Option.map
 nipkow parents: 
15860diff
changeset | 122 | map_of (map f xs) a = Option.map (\<lambda> b. (snd (f (a, b)))) (map_of xs a)" | 
| 13673 | 123 | by (induct xs, auto) | 
| 124 | ||
| 30235 
58d147683393
Made Option a separate theory and renamed option_map to Option.map
 nipkow parents: 
15860diff
changeset | 125 | lemma option_map_of [simp]: "(Option.map f (map_of xs k) = None) = ((map_of xs k) = None)" | 
| 
58d147683393
Made Option a separate theory and renamed option_map to Option.map
 nipkow parents: 
15860diff
changeset | 126 | by (simp add: Option.map_def split: option.split) | 
| 13673 | 127 | |
| 128 | ||
| 129 | ||
| 130 | end |