src/HOL/MicroJava/Comp/AuxLemmas.thy
author nipkow
Mon, 13 Sep 2010 11:13:15 +0200
changeset 39302 d7728f65b353
parent 39198 f967a16dfcdd
child 39917 b85bfa89a387
permissions -rwxr-xr-x
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/Comp/AuxLemmas.thy
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     2
    ID:         $Id$
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     3
    Author:     Martin Strecker
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     4
*)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     5
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     6
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     7
(* Auxiliary Lemmas *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     8
15860
a344c4284972 partial modernising of theory headers
paulson
parents: 14981
diff changeset
     9
theory AuxLemmas
a344c4284972 partial modernising of theory headers
paulson
parents: 14981
diff changeset
    10
imports "../J/JBasis"
a344c4284972 partial modernising of theory headers
paulson
parents: 14981
diff changeset
    11
begin
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    12
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    13
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    14
(* List.thy *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    15
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    16
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    17
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    18
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    19
lemma app_nth_greater_len [rule_format (no_asm), simp]:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    20
 "\<forall> ind. length pre \<le> ind \<longrightarrow> (pre @ a # post) ! (Suc ind) = (pre @ post) ! ind"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    21
apply (induct "pre")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    22
apply auto
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    23
apply (case_tac ind)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    24
apply auto
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    25
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    26
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    27
lemma length_takeWhile: "v \<in> set xs \<Longrightarrow> length (takeWhile (%z. z~=v) xs) < length xs"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    28
by (induct xs, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    29
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    30
lemma nth_length_takeWhile [simp]: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    31
  "v \<in> set xs \<Longrightarrow> xs ! (length (takeWhile (%z. z~=v) xs)) = v"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    32
by (induct xs, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    33
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    34
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    35
lemma map_list_update [simp]:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    36
  "\<lbrakk> x \<in> set xs; distinct xs\<rbrakk> \<Longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    37
  (map f xs) [length (takeWhile (\<lambda>z. z \<noteq> x) xs) := v] = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    38
  map (f(x:=v)) xs"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    39
apply (induct xs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    40
apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    41
apply (case_tac "x=a")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    42
apply auto
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    43
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    44
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    45
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    46
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    47
(* Product_Type.thy *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    48
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    49
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    50
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    51
lemma split_compose: "(split f) \<circ> (\<lambda> (a,b). ((fa a), (fb b))) =
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    52
  (\<lambda> (a,b). (f (fa a) (fb b)))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    53
by (simp add: split_def o_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    54
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    55
lemma split_iter: "(\<lambda> (a,b,c). ((g1 a), (g2 b), (g3 c))) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    56
        (\<lambda> (a,p). ((g1 a), (\<lambda> (b, c). ((g2 b), (g3 c))) p))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    57
by (simp add: split_def o_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    58
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    59
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    60
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    61
(* Set.thy *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    62
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    63
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    64
lemma singleton_in_set: "A = {a} \<Longrightarrow> a \<in> A" by simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    65
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    66
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    67
(* Map.thy *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    68
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    69
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    70
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: 39198
diff changeset
    71
by (simp add: fun_eq_iff)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    72
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    73
lemma map_of_in_set: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    74
  "(map_of xs x = None) = (x \<notin> set (map fst xs))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    75
by (induct xs, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    76
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    77
lemma map_map_upd [simp]: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    78
  "y \<notin> set xs \<Longrightarrow> map (the \<circ> f(y\<mapsto>v)) xs = map (the \<circ> f) xs"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    79
by (simp add: the_map_upd)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    80
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    81
lemma map_map_upds [rule_format (no_asm), simp]: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    82
"\<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
d9b155757dc8 *** empty log message ***
nipkow
parents: 13673
diff changeset
    83
apply (induct xs)
d9b155757dc8 *** empty log message ***
nipkow
parents: 13673
diff changeset
    84
 apply simp
d9b155757dc8 *** empty log message ***
nipkow
parents: 13673
diff changeset
    85
apply fastsimp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    86
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    87
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    88
lemma map_upds_distinct [rule_format (no_asm), simp]: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    89
  "\<forall> f vs. length ys = length vs \<longrightarrow> distinct ys \<longrightarrow> map (the \<circ> f(ys[\<mapsto>]vs)) ys = vs"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    90
apply (induct ys)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    91
apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    92
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    93
apply (case_tac vs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    94
apply simp
14025
d9b155757dc8 *** empty log message ***
nipkow
parents: 13673
diff changeset
    95
apply (simp only: map_upds_Cons distinct.simps hd.simps tl.simps map.simps)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    96
apply clarify
14025
d9b155757dc8 *** empty log message ***
nipkow
parents: 13673
diff changeset
    97
apply (simp del: o_apply)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    98
apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    99
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   100
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   101
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   102
lemma map_of_map_as_map_upd [rule_format (no_asm)]: "distinct (map f zs) \<longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   103
map_of (map (\<lambda> p. (f p, g p)) zs) = empty (map f zs [\<mapsto>] map g zs)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   104
by (induct zs, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   105
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   106
(* In analogy to Map.map_of_SomeD *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   107
lemma map_upds_SomeD [rule_format (no_asm)]: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   108
  "\<forall> m ys. (m(xs[\<mapsto>]ys)) k = Some y \<longrightarrow> k \<in> (set xs) \<or> (m k = Some y)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   109
apply (induct xs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   110
apply simp
14025
d9b155757dc8 *** empty log message ***
nipkow
parents: 13673
diff changeset
   111
apply auto
d9b155757dc8 *** empty log message ***
nipkow
parents: 13673
diff changeset
   112
apply(case_tac ys)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   113
apply auto
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   114
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   115
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   116
lemma map_of_upds_SomeD: "(map_of m (xs[\<mapsto>]ys)) k = Some y 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   117
  \<Longrightarrow> k \<in> (set (xs @ map fst m))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   118
by (auto dest: map_upds_SomeD map_of_SomeD fst_in_set_lemma)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   119
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   120
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   121
lemma map_of_map_prop [rule_format (no_asm)]: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   122
  "(map_of (map f xs) k = Some v) \<longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   123
  (\<forall> x \<in> set xs. (P1 x)) \<longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   124
  (\<forall> x. (P1 x) \<longrightarrow> (P2 (f x))) \<longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   125
  (P2(k, v))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   126
by (induct xs,auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   127
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   128
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: 15860
diff changeset
   129
  map_of (map f xs) a = Option.map (\<lambda> b. (snd (f (a, b)))) (map_of xs a)"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   130
by (induct xs, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   131
30235
58d147683393 Made Option a separate theory and renamed option_map to Option.map
nipkow
parents: 15860
diff changeset
   132
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: 15860
diff changeset
   133
by (simp add: Option.map_def split: option.split)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   134
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   135
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   136
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   137
end