| 13673 |      1 | (*  Title:      HOL/MicroJava/Comp/AuxLemmas.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Martin Strecker
 | 
|  |      4 |     Copyright   GPL 2002
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | 
 | 
|  |      8 | (* Auxiliary Lemmas *)
 | 
|  |      9 | 
 | 
|  |     10 | theory AuxLemmas = JBasis:
 | 
|  |     11 | 
 | 
|  |     12 | (**********************************************************************)
 | 
|  |     13 | (* List.thy *)
 | 
|  |     14 | (**********************************************************************)
 | 
|  |     15 | 
 | 
|  |     16 | 
 | 
|  |     17 | 
 | 
|  |     18 | lemma foldr_foldl: "foldr f xs a = foldl (\<lambda> x y. (f y x)) a (rev xs)"
 | 
|  |     19 | by (induct xs, simp, simp)
 | 
|  |     20 | 
 | 
|  |     21 | lemma foldl_foldr: "foldl f a xs = foldr (\<lambda> x y. (f y x)) (rev xs) a"
 | 
|  |     22 | by (simp add: foldr_foldl [of "(\<lambda> x y. (f y x))" "(rev xs)"])
 | 
|  |     23 | 
 | 
|  |     24 | lemma foldr_append: "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)"
 | 
|  |     25 | by (induct xs, auto)
 | 
|  |     26 | 
 | 
|  |     27 | lemma app_nth_len [simp]: "(pre @ a # post) ! length pre = a"
 | 
|  |     28 | by (induct "pre") auto
 | 
|  |     29 | 
 | 
|  |     30 | lemma app_nth_len_plus [simp]: "(pre @ post) ! ((length pre) + n) = post ! n"
 | 
|  |     31 | by (induct "pre") auto
 | 
|  |     32 | 
 | 
|  |     33 | lemma app_nth_greater_len [rule_format (no_asm), simp]:
 | 
|  |     34 |  "\<forall> ind. length pre \<le> ind \<longrightarrow> (pre @ a # post) ! (Suc ind) = (pre @ post) ! ind"
 | 
|  |     35 | apply (induct "pre")
 | 
|  |     36 | apply auto
 | 
|  |     37 | apply (case_tac ind)
 | 
|  |     38 | apply auto
 | 
|  |     39 | done
 | 
|  |     40 | 
 | 
|  |     41 | lemma list_update_length [simp]: "(xs @ a# ys)[length xs := b] = (xs @ b# ys)"
 | 
|  |     42 | by (induct xs, auto)
 | 
|  |     43 | 
 | 
|  |     44 | 
 | 
|  |     45 | lemma length_takeWhile: "v \<in> set xs \<Longrightarrow> length (takeWhile (%z. z~=v) xs) < length xs"
 | 
|  |     46 | by (induct xs, auto)
 | 
|  |     47 | 
 | 
|  |     48 | lemma nth_length_takeWhile [simp]: 
 | 
|  |     49 |   "v \<in> set xs \<Longrightarrow> xs ! (length (takeWhile (%z. z~=v) xs)) = v"
 | 
|  |     50 | by (induct xs, auto)
 | 
|  |     51 | 
 | 
|  |     52 | 
 | 
|  |     53 | lemma map_fun_upd [simp]: "y \<notin> set xs \<Longrightarrow> map (f(y:=v)) xs = map f xs"
 | 
|  |     54 | by (induct xs, auto)
 | 
|  |     55 | 
 | 
|  |     56 | lemma map_list_update [simp]:
 | 
|  |     57 |   "\<lbrakk> x \<in> set xs; distinct xs\<rbrakk> \<Longrightarrow> 
 | 
|  |     58 |   (map f xs) [length (takeWhile (\<lambda>z. z \<noteq> x) xs) := v] = 
 | 
|  |     59 |   map (f(x:=v)) xs"
 | 
|  |     60 | apply (induct xs)
 | 
|  |     61 | apply simp
 | 
|  |     62 | apply (case_tac "x=a")
 | 
|  |     63 | apply auto
 | 
|  |     64 | done
 | 
|  |     65 | 
 | 
|  |     66 | 
 | 
|  |     67 | (**********************************************************************)
 | 
|  |     68 | (* Product_Type.thy *)
 | 
|  |     69 | (**********************************************************************)
 | 
|  |     70 | 
 | 
|  |     71 | 
 | 
|  |     72 | lemma split_compose: "(split f) \<circ> (\<lambda> (a,b). ((fa a), (fb b))) =
 | 
|  |     73 |   (\<lambda> (a,b). (f (fa a) (fb b)))"
 | 
|  |     74 | by (simp add: split_def o_def)
 | 
|  |     75 | 
 | 
|  |     76 | lemma split_iter: "(\<lambda> (a,b,c). ((g1 a), (g2 b), (g3 c))) = 
 | 
|  |     77 |         (\<lambda> (a,p). ((g1 a), (\<lambda> (b, c). ((g2 b), (g3 c))) p))"
 | 
|  |     78 | by (simp add: split_def o_def)
 | 
|  |     79 | 
 | 
|  |     80 | 
 | 
|  |     81 | (**********************************************************************)
 | 
|  |     82 | (* Set.thy *)
 | 
|  |     83 | (**********************************************************************)
 | 
|  |     84 | 
 | 
|  |     85 | lemma singleton_in_set: "A = {a} \<Longrightarrow> a \<in> A" by simp
 | 
|  |     86 | 
 | 
|  |     87 | (**********************************************************************)
 | 
|  |     88 | (* Map.thy *)
 | 
|  |     89 | (**********************************************************************)
 | 
|  |     90 | 
 | 
|  |     91 | lemma the_map_upd: "(the \<circ> f(x\<mapsto>v)) = (the \<circ> f)(x:=v)"
 | 
|  |     92 | by (simp add: expand_fun_eq)
 | 
|  |     93 | 
 | 
|  |     94 | lemma map_of_in_set: 
 | 
|  |     95 |   "(map_of xs x = None) = (x \<notin> set (map fst xs))"
 | 
|  |     96 | by (induct xs, auto)
 | 
|  |     97 | 
 | 
|  |     98 | lemma map_map_upd [simp]: 
 | 
|  |     99 |   "y \<notin> set xs \<Longrightarrow> map (the \<circ> f(y\<mapsto>v)) xs = map (the \<circ> f) xs"
 | 
|  |    100 | by (simp add: the_map_upd)
 | 
|  |    101 | 
 | 
|  |    102 | lemma map_map_upds [rule_format (no_asm), simp]: 
 | 
|  |    103 | "\<forall> f vs. (\<forall>y\<in>set ys. y \<notin> set xs) \<longrightarrow> map (the \<circ> f(ys[\<mapsto>]vs)) xs = map (the \<circ> f) xs"
 | 
| 14025 |    104 | apply (induct xs)
 | 
|  |    105 |  apply simp
 | 
|  |    106 | apply fastsimp
 | 
| 13673 |    107 | done
 | 
|  |    108 | 
 | 
|  |    109 | lemma map_upds_distinct [rule_format (no_asm), simp]: 
 | 
|  |    110 |   "\<forall> f vs. length ys = length vs \<longrightarrow> distinct ys \<longrightarrow> map (the \<circ> f(ys[\<mapsto>]vs)) ys = vs"
 | 
|  |    111 | apply (induct ys)
 | 
|  |    112 | apply simp
 | 
|  |    113 | apply (intro strip)
 | 
|  |    114 | apply (case_tac vs)
 | 
|  |    115 | apply simp
 | 
| 14025 |    116 | apply (simp only: map_upds_Cons distinct.simps hd.simps tl.simps map.simps)
 | 
| 13673 |    117 | apply clarify
 | 
| 14025 |    118 | apply (simp del: o_apply)
 | 
| 13673 |    119 | apply simp
 | 
|  |    120 | done
 | 
|  |    121 | 
 | 
|  |    122 | 
 | 
|  |    123 | lemma map_of_map_as_map_upd [rule_format (no_asm)]: "distinct (map f zs) \<longrightarrow> 
 | 
|  |    124 | map_of (map (\<lambda> p. (f p, g p)) zs) = empty (map f zs [\<mapsto>] map g zs)"
 | 
|  |    125 | by (induct zs, auto)
 | 
|  |    126 | 
 | 
|  |    127 | (* In analogy to Map.map_of_SomeD *)
 | 
|  |    128 | lemma map_upds_SomeD [rule_format (no_asm)]: 
 | 
|  |    129 |   "\<forall> m ys. (m(xs[\<mapsto>]ys)) k = Some y \<longrightarrow> k \<in> (set xs) \<or> (m k = Some y)"
 | 
|  |    130 | apply (induct xs)
 | 
|  |    131 | apply simp
 | 
| 14025 |    132 | apply auto
 | 
|  |    133 | apply(case_tac ys)
 | 
| 13673 |    134 | apply auto
 | 
|  |    135 | done
 | 
|  |    136 | 
 | 
|  |    137 | lemma map_of_upds_SomeD: "(map_of m (xs[\<mapsto>]ys)) k = Some y 
 | 
|  |    138 |   \<Longrightarrow> k \<in> (set (xs @ map fst m))"
 | 
|  |    139 | by (auto dest: map_upds_SomeD map_of_SomeD fst_in_set_lemma)
 | 
|  |    140 | 
 | 
|  |    141 | 
 | 
|  |    142 | lemma map_of_map_prop [rule_format (no_asm)]: 
 | 
|  |    143 |   "(map_of (map f xs) k = Some v) \<longrightarrow> 
 | 
|  |    144 |   (\<forall> x \<in> set xs. (P1 x)) \<longrightarrow> 
 | 
|  |    145 |   (\<forall> x. (P1 x) \<longrightarrow> (P2 (f x))) \<longrightarrow>
 | 
|  |    146 |   (P2(k, v))"
 | 
|  |    147 | by (induct xs,auto)
 | 
|  |    148 | 
 | 
|  |    149 | lemma map_of_map2: "\<forall> x \<in> set xs. (fst (f x)) = (fst x) \<Longrightarrow>
 | 
|  |    150 |   map_of (map f xs) a = option_map (\<lambda> b. (snd (f (a, b)))) (map_of xs a)"
 | 
|  |    151 | by (induct xs, auto)
 | 
|  |    152 | 
 | 
|  |    153 | lemma option_map_of [simp]: "(option_map f (map_of xs k) = None) = ((map_of xs k) = None)"
 | 
|  |    154 | by (simp add: option_map_def split: option.split)
 | 
|  |    155 | 
 | 
|  |    156 | 
 | 
|  |    157 | 
 | 
|  |    158 | end
 |