| author | wenzelm | 
| Sat, 03 Nov 2001 01:39:17 +0100 | |
| changeset 12028 | 52aa183c15bb | 
| parent 11070 | cc421547e744 | 
| child 12517 | 360e3215f029 | 
| permissions | -rw-r--r-- | 
| 8011 | 1  | 
(* Title: HOL/MicroJava/J/JBasis.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: David von Oheimb  | 
|
4  | 
Copyright 1999 TU Muenchen  | 
|
| 11070 | 5  | 
*)  | 
| 8011 | 6  | 
|
| 11070 | 7  | 
header "Some Auxiliary Definitions"  | 
| 8011 | 8  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
9  | 
theory JBasis = Main:  | 
| 8011 | 10  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
11  | 
lemmas [simp] = Let_def  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
12  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
13  | 
section "unique"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
14  | 
|
| 8011 | 15  | 
constdefs  | 
16  | 
||
| 11070 | 17  | 
  unique  :: "('a \<times> 'b) list => bool"
 | 
18  | 
"unique == nodups \<circ> map fst"  | 
|
| 8011 | 19  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
20  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
21  | 
lemma fst_in_set_lemma [rule_format (no_asm)]:  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
22  | 
"(x, y) : set xys --> x : fst ` set xys"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
23  | 
apply (induct_tac "xys")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
24  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
25  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
26  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
27  | 
lemma unique_Nil [simp]: "unique []"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
28  | 
apply (unfold unique_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
29  | 
apply (simp (no_asm))  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
30  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
31  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
32  | 
lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
33  | 
apply (unfold unique_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
34  | 
apply (auto dest: fst_in_set_lemma)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
35  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
36  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
37  | 
lemma unique_append [rule_format (no_asm)]: "unique l' ==> unique l -->  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
38  | 
(!(x,y):set l. !(x',y'):set l'. x' ~= x) --> unique (l @ l')"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
39  | 
apply (induct_tac "l")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
40  | 
apply (auto dest: fst_in_set_lemma)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
41  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
42  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
43  | 
lemma unique_map_inj [rule_format (no_asm)]:  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
44  | 
"unique l --> inj f --> unique (map (%(k,x). (f k, g k x)) l)"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
45  | 
apply (induct_tac "l")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
46  | 
apply (auto dest: fst_in_set_lemma simp add: inj_eq)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
47  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
48  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
49  | 
section "More about Maps"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
50  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
51  | 
lemma map_of_SomeI [rule_format (no_asm)]:  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
52  | 
"unique l --> (k, x) : set l --> map_of l k = Some x"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
53  | 
apply (induct_tac "l")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
54  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
55  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
56  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
57  | 
lemma Ball_set_table_:  | 
| 11070 | 58  | 
"(\<forall>(x,y)\<in>set l. P x y) --> (\<forall>x. \<forall>y. map_of l x = Some y --> P x y)"  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
59  | 
apply(induct_tac "l")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
60  | 
apply(simp_all (no_asm))  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
61  | 
apply safe  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
62  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
63  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
64  | 
lemmas Ball_set_table = Ball_set_table_ [THEN mp];  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
65  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
66  | 
lemma table_of_remap_SomeD [rule_format (no_asm)]:  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
67  | 
"map_of (map (\<lambda>((k,k'),x). (k,(k',x))) t) k = Some (k',x) -->  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
68  | 
map_of t (k, k') = Some x"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
69  | 
apply (induct_tac "t")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
70  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
71  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
72  | 
|
| 8011 | 73  | 
end  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10042 
diff
changeset
 | 
74  |