| author | wenzelm | 
| Thu, 04 Feb 2016 16:30:01 +0100 | |
| changeset 62261 | 74dc98bd9f51 | 
| parent 62042 | 6c6ccf573479 | 
| child 67399 | eab6ce8368fa | 
| permissions | -rw-r--r-- | 
| 12951 | 1  | 
(* Title: HOL/MicroJava/BV/BVExample.thy  | 
2  | 
Author: Gerwin Klein  | 
|
3  | 
*)  | 
|
4  | 
||
| 61361 | 5  | 
section \<open>Example Welltypings \label{sec:BVExample}\<close>
 | 
| 12951 | 6  | 
|
| 
23022
 
9872ef956276
added qualification for ambiguous definition names
 
haftmann 
parents: 
22271 
diff
changeset
 | 
7  | 
theory BVExample  | 
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40077 
diff
changeset
 | 
8  | 
imports  | 
| 
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40077 
diff
changeset
 | 
9  | 
"../JVM/JVMListExample"  | 
| 
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40077 
diff
changeset
 | 
10  | 
BVSpecTypeSafe  | 
| 
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40077 
diff
changeset
 | 
11  | 
JVM  | 
| 
23022
 
9872ef956276
added qualification for ambiguous definition names
 
haftmann 
parents: 
22271 
diff
changeset
 | 
12  | 
begin  | 
| 12951 | 13  | 
|
| 61361 | 14  | 
text \<open>  | 
| 12972 | 15  | 
This theory shows type correctness of the example program in section  | 
16  | 
  \ref{sec:JVMListExample} (p. \pageref{sec:JVMListExample}) by
 | 
|
17  | 
explicitly providing a welltyping. It also shows that the start  | 
|
18  | 
state of the program conforms to the welltyping; hence type safe  | 
|
19  | 
execution is guaranteed.  | 
|
| 61361 | 20  | 
\<close>  | 
| 12972 | 21  | 
|
| 58886 | 22  | 
subsection "Setup"  | 
| 12951 | 23  | 
|
| 61361 | 24  | 
text \<open>Abbreviations for definitions we will have to use often in the  | 
25  | 
proofs below:\<close>  | 
|
| 13101 | 26  | 
lemmas name_defs = list_name_def test_name_def val_name_def next_name_def  | 
| 12951 | 27  | 
lemmas system_defs = SystemClasses_def ObjectC_def NullPointerC_def  | 
28  | 
OutOfMemoryC_def ClassCastC_def  | 
|
29  | 
lemmas class_defs = list_class_def test_class_def  | 
|
30  | 
||
| 61361 | 31  | 
text \<open>These auxiliary proofs are for efficiency: class lookup,  | 
| 12951 | 32  | 
subclass relation, method and field lookup are computed only once:  | 
| 61361 | 33  | 
\<close>  | 
| 12951 | 34  | 
lemma class_Object [simp]:  | 
| 28520 | 35  | 
"class E Object = Some (undefined, [],[])"  | 
| 12951 | 36  | 
by (simp add: class_def system_defs E_def)  | 
37  | 
||
38  | 
lemma class_NullPointer [simp]:  | 
|
39  | 
"class E (Xcpt NullPointer) = Some (Object, [], [])"  | 
|
40  | 
by (simp add: class_def system_defs E_def)  | 
|
41  | 
||
42  | 
lemma class_OutOfMemory [simp]:  | 
|
43  | 
"class E (Xcpt OutOfMemory) = Some (Object, [], [])"  | 
|
44  | 
by (simp add: class_def system_defs E_def)  | 
|
45  | 
||
46  | 
lemma class_ClassCast [simp]:  | 
|
47  | 
"class E (Xcpt ClassCast) = Some (Object, [], [])"  | 
|
48  | 
by (simp add: class_def system_defs E_def)  | 
|
49  | 
||
50  | 
lemma class_list [simp]:  | 
|
51  | 
"class E list_name = Some list_class"  | 
|
52  | 
by (simp add: class_def system_defs E_def name_defs distinct_classes [symmetric])  | 
|
53  | 
||
54  | 
lemma class_test [simp]:  | 
|
55  | 
"class E test_name = Some test_class"  | 
|
56  | 
by (simp add: class_def system_defs E_def name_defs distinct_classes [symmetric])  | 
|
57  | 
||
58  | 
lemma E_classes [simp]:  | 
|
59  | 
  "{C. is_class E C} = {list_name, test_name, Xcpt NullPointer, 
 | 
|
60  | 
Xcpt ClassCast, Xcpt OutOfMemory, Object}"  | 
|
61  | 
by (auto simp add: is_class_def class_def system_defs E_def name_defs class_defs)  | 
|
62  | 
||
| 61361 | 63  | 
text \<open>The subclass releation spelled out:\<close>  | 
| 12951 | 64  | 
lemma subcls1:  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
65  | 
  "subcls1 E = {(list_name,Object), (test_name,Object), (Xcpt NullPointer, Object),
 | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
66  | 
(Xcpt ClassCast, Object), (Xcpt OutOfMemory, Object)}"  | 
| 
31197
 
c1c163ec6c44
fine-tuned elimination of comprehensions involving x=t.
 
nipkow 
parents: 
31168 
diff
changeset
 | 
67  | 
apply (simp add: subcls1_def2)  | 
| 
31166
 
a90fe83f58ea
"{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
 
nipkow 
parents: 
28520 
diff
changeset
 | 
68  | 
apply (simp add: name_defs class_defs system_defs E_def class_def)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
69  | 
apply (simp add: Sigma_def)  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
70  | 
apply auto  | 
| 
31166
 
a90fe83f58ea
"{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
 
nipkow 
parents: 
28520 
diff
changeset
 | 
71  | 
done  | 
| 12951 | 72  | 
|
| 61361 | 73  | 
text \<open>The subclass relation is acyclic; hence its converse is well founded:\<close>  | 
| 12951 | 74  | 
lemma notin_rtrancl:  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
75  | 
"(a, b) \<in> r\<^sup>* \<Longrightarrow> a \<noteq> b \<Longrightarrow> (\<And>y. (a, y) \<notin> r) \<Longrightarrow> False"  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
76  | 
by (auto elim: converse_rtranclE)  | 
| 12951 | 77  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
78  | 
lemma acyclic_subcls1_E: "acyclic (subcls1 E)"  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
79  | 
apply (rule acyclicI)  | 
| 12951 | 80  | 
apply (simp add: subcls1)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
81  | 
apply (auto dest!: tranclD)  | 
| 12951 | 82  | 
apply (auto elim!: notin_rtrancl simp add: name_defs distinct_classes)  | 
83  | 
done  | 
|
84  | 
||
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
85  | 
lemma wf_subcls1_E: "wf ((subcls1 E)\<inverse>)"  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
31998 
diff
changeset
 | 
86  | 
apply (rule finite_acyclic_wf_converse)  | 
| 23757 | 87  | 
apply (simp add: subcls1 del: insert_iff)  | 
| 12951 | 88  | 
apply (rule acyclic_subcls1_E)  | 
89  | 
done  | 
|
90  | 
||
| 61361 | 91  | 
text \<open>Method and field lookup:\<close>  | 
| 12951 | 92  | 
lemma method_Object [simp]:  | 
| 
31852
 
a16bb835e97d
explicit Set constructor for code generated for sets
 
haftmann 
parents: 
31197 
diff
changeset
 | 
93  | 
"method (E, Object) = Map.empty"  | 
| 12951 | 94  | 
by (simp add: method_rec_lemma [OF class_Object wf_subcls1_E])  | 
95  | 
||
96  | 
lemma method_append [simp]:  | 
|
97  | 
"method (E, list_name) (append_name, [Class list_name]) =  | 
|
98  | 
Some (list_name, PrimT Void, 3, 0, append_ins, [(1, 2, 8, Xcpt NullPointer)])"  | 
|
99  | 
apply (insert class_list)  | 
|
100  | 
apply (unfold list_class_def)  | 
|
101  | 
apply (drule method_rec_lemma [OF _ wf_subcls1_E])  | 
|
102  | 
apply simp  | 
|
103  | 
done  | 
|
104  | 
||
105  | 
lemma method_makelist [simp]:  | 
|
106  | 
"method (E, test_name) (makelist_name, []) =  | 
|
107  | 
Some (test_name, PrimT Void, 3, 2, make_list_ins, [])"  | 
|
108  | 
apply (insert class_test)  | 
|
109  | 
apply (unfold test_class_def)  | 
|
110  | 
apply (drule method_rec_lemma [OF _ wf_subcls1_E])  | 
|
111  | 
apply simp  | 
|
112  | 
done  | 
|
113  | 
||
114  | 
lemma field_val [simp]:  | 
|
115  | 
"field (E, list_name) val_name = Some (list_name, PrimT Integer)"  | 
|
| 
23022
 
9872ef956276
added qualification for ambiguous definition names
 
haftmann 
parents: 
22271 
diff
changeset
 | 
116  | 
apply (unfold TypeRel.field_def)  | 
| 12951 | 117  | 
apply (insert class_list)  | 
118  | 
apply (unfold list_class_def)  | 
|
119  | 
apply (drule fields_rec_lemma [OF _ wf_subcls1_E])  | 
|
120  | 
apply simp  | 
|
121  | 
done  | 
|
122  | 
||
123  | 
lemma field_next [simp]:  | 
|
124  | 
"field (E, list_name) next_name = Some (list_name, Class list_name)"  | 
|
| 
23022
 
9872ef956276
added qualification for ambiguous definition names
 
haftmann 
parents: 
22271 
diff
changeset
 | 
125  | 
apply (unfold TypeRel.field_def)  | 
| 12951 | 126  | 
apply (insert class_list)  | 
127  | 
apply (unfold list_class_def)  | 
|
128  | 
apply (drule fields_rec_lemma [OF _ wf_subcls1_E])  | 
|
129  | 
apply (simp add: name_defs distinct_fields [symmetric])  | 
|
130  | 
done  | 
|
131  | 
||
132  | 
lemma [simp]: "fields (E, Object) = []"  | 
|
133  | 
by (simp add: fields_rec_lemma [OF class_Object wf_subcls1_E])  | 
|
134  | 
||
135  | 
lemma [simp]: "fields (E, Xcpt NullPointer) = []"  | 
|
136  | 
by (simp add: fields_rec_lemma [OF class_NullPointer wf_subcls1_E])  | 
|
137  | 
||
138  | 
lemma [simp]: "fields (E, Xcpt ClassCast) = []"  | 
|
139  | 
by (simp add: fields_rec_lemma [OF class_ClassCast wf_subcls1_E])  | 
|
140  | 
||
141  | 
lemma [simp]: "fields (E, Xcpt OutOfMemory) = []"  | 
|
142  | 
by (simp add: fields_rec_lemma [OF class_OutOfMemory wf_subcls1_E])  | 
|
143  | 
||
144  | 
lemma [simp]: "fields (E, test_name) = []"  | 
|
145  | 
apply (insert class_test)  | 
|
146  | 
apply (unfold test_class_def)  | 
|
147  | 
apply (drule fields_rec_lemma [OF _ wf_subcls1_E])  | 
|
148  | 
apply simp  | 
|
149  | 
done  | 
|
150  | 
||
151  | 
lemmas [simp] = is_class_def  | 
|
152  | 
||
| 61361 | 153  | 
text \<open>  | 
| 12951 | 154  | 
The next definition and three proof rules implement an algorithm to  | 
| 62042 | 155  | 
enumarate natural numbers. The command \<open>apply (elim pc_end pc_next pc_0\<close>  | 
| 12951 | 156  | 
transforms a goal of the form  | 
157  | 
  @{prop [display] "pc < n \<Longrightarrow> P pc"} 
 | 
|
158  | 
into a series of goals  | 
|
159  | 
  @{prop [display] "P 0"} 
 | 
|
160  | 
  @{prop [display] "P (Suc 0)"} 
 | 
|
161  | 
||
| 62042 | 162  | 
\<open>\<dots>\<close>  | 
| 12951 | 163  | 
|
164  | 
  @{prop [display] "P n"} 
 | 
|
| 61361 | 165  | 
\<close>  | 
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
166  | 
definition intervall :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" ("_ \<in> [_, _')") where
 | 
| 12951 | 167  | 
"x \<in> [a, b) \<equiv> a \<le> x \<and> x < b"  | 
168  | 
||
169  | 
lemma pc_0: "x < n \<Longrightarrow> (x \<in> [0, n) \<Longrightarrow> P x) \<Longrightarrow> P x"  | 
|
170  | 
by (simp add: intervall_def)  | 
|
171  | 
||
172  | 
lemma pc_next: "x \<in> [n0, n) \<Longrightarrow> P n0 \<Longrightarrow> (x \<in> [Suc n0, n) \<Longrightarrow> P x) \<Longrightarrow> P x"  | 
|
173  | 
apply (cases "x=n0")  | 
|
| 13187 | 174  | 
apply (auto simp add: intervall_def)  | 
| 12951 | 175  | 
done  | 
176  | 
||
177  | 
lemma pc_end: "x \<in> [n,n) \<Longrightarrow> P x"  | 
|
178  | 
by (unfold intervall_def) arith  | 
|
179  | 
||
180  | 
||
| 58886 | 181  | 
subsection "Program structure"  | 
| 12951 | 182  | 
|
| 61361 | 183  | 
text \<open>  | 
| 12951 | 184  | 
The program is structurally wellformed:  | 
| 61361 | 185  | 
\<close>  | 
| 14045 | 186  | 
|
| 12951 | 187  | 
lemma wf_struct:  | 
188  | 
"wf_prog (\<lambda>G C mb. True) E" (is "wf_prog ?mb E")  | 
|
189  | 
proof -  | 
|
190  | 
have "unique E"  | 
|
191  | 
by (simp add: system_defs E_def class_defs name_defs distinct_classes)  | 
|
192  | 
moreover  | 
|
193  | 
have "set SystemClasses \<subseteq> set E" by (simp add: system_defs E_def)  | 
|
194  | 
hence "wf_syscls E" by (rule wf_syscls)  | 
|
195  | 
moreover  | 
|
196  | 
have "wf_cdecl ?mb E ObjectC" by (simp add: wf_cdecl_def ObjectC_def)  | 
|
197  | 
moreover  | 
|
198  | 
have "wf_cdecl ?mb E NullPointerC"  | 
|
199  | 
by (auto elim: notin_rtrancl  | 
|
200  | 
simp add: wf_cdecl_def name_defs NullPointerC_def subcls1)  | 
|
201  | 
moreover  | 
|
202  | 
have "wf_cdecl ?mb E ClassCastC"  | 
|
203  | 
by (auto elim: notin_rtrancl  | 
|
204  | 
simp add: wf_cdecl_def name_defs ClassCastC_def subcls1)  | 
|
205  | 
moreover  | 
|
206  | 
have "wf_cdecl ?mb E OutOfMemoryC"  | 
|
207  | 
by (auto elim: notin_rtrancl  | 
|
208  | 
simp add: wf_cdecl_def name_defs OutOfMemoryC_def subcls1)  | 
|
209  | 
moreover  | 
|
210  | 
have "wf_cdecl ?mb E (list_name, list_class)"  | 
|
211  | 
apply (auto elim!: notin_rtrancl  | 
|
212  | 
simp add: wf_cdecl_def wf_fdecl_def list_class_def  | 
|
213  | 
wf_mdecl_def wf_mhead_def subcls1)  | 
|
214  | 
apply (auto simp add: name_defs distinct_classes distinct_fields)  | 
|
215  | 
done  | 
|
216  | 
moreover  | 
|
217  | 
have "wf_cdecl ?mb E (test_name, test_class)"  | 
|
218  | 
apply (auto elim!: notin_rtrancl  | 
|
219  | 
simp add: wf_cdecl_def wf_fdecl_def test_class_def  | 
|
220  | 
wf_mdecl_def wf_mhead_def subcls1)  | 
|
221  | 
apply (auto simp add: name_defs distinct_classes distinct_fields)  | 
|
222  | 
done  | 
|
223  | 
ultimately  | 
|
| 14045 | 224  | 
show ?thesis  | 
225  | 
by (simp add: wf_prog_def ws_prog_def wf_cdecl_mrT_cdecl_mdecl E_def SystemClasses_def)  | 
|
| 12951 | 226  | 
qed  | 
227  | 
||
| 58886 | 228  | 
subsection "Welltypings"  | 
| 61361 | 229  | 
text \<open>  | 
| 12951 | 230  | 
  We show welltypings of the methods @{term append_name} in class @{term list_name}, 
 | 
231  | 
  and @{term makelist_name} in class @{term test_name}:
 | 
|
| 61361 | 232  | 
\<close>  | 
| 12951 | 233  | 
lemmas eff_simps [simp] = eff_def norm_eff_def xcpt_eff_def  | 
234  | 
declare appInvoke [simp del]  | 
|
235  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
236  | 
definition phi_append :: method_type ("\<phi>\<^sub>a") where
 | 
| 12951 | 237  | 
"\<phi>\<^sub>a \<equiv> map (\<lambda>(x,y). Some (x, map OK y)) [  | 
238  | 
( [], [Class list_name, Class list_name]),  | 
|
239  | 
( [Class list_name], [Class list_name, Class list_name]),  | 
|
240  | 
( [Class list_name], [Class list_name, Class list_name]),  | 
|
241  | 
( [Class list_name, Class list_name], [Class list_name, Class list_name]),  | 
|
242  | 
([NT, Class list_name, Class list_name], [Class list_name, Class list_name]),  | 
|
243  | 
( [Class list_name], [Class list_name, Class list_name]),  | 
|
244  | 
( [Class list_name, Class list_name], [Class list_name, Class list_name]),  | 
|
245  | 
( [PrimT Void], [Class list_name, Class list_name]),  | 
|
246  | 
( [Class Object], [Class list_name, Class list_name]),  | 
|
247  | 
( [], [Class list_name, Class list_name]),  | 
|
248  | 
( [Class list_name], [Class list_name, Class list_name]),  | 
|
249  | 
( [Class list_name, Class list_name], [Class list_name, Class list_name]),  | 
|
250  | 
( [], [Class list_name, Class list_name]),  | 
|
251  | 
( [PrimT Void], [Class list_name, Class list_name])]"  | 
|
252  | 
||
| 
13214
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
253  | 
|
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
254  | 
lemma bounded_append [simp]:  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
255  | 
"check_bounded append_ins [(Suc 0, 2, 8, Xcpt NullPointer)]"  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
256  | 
apply (simp add: check_bounded_def)  | 
| 40077 | 257  | 
apply (simp add: eval_nat_numeral append_ins_def)  | 
| 
13214
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
258  | 
apply (rule allI, rule impI)  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
259  | 
apply (elim pc_end pc_next pc_0)  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
260  | 
apply auto  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
261  | 
done  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
262  | 
|
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
263  | 
lemma types_append [simp]: "check_types E 3 (Suc (Suc 0)) (map OK \<phi>\<^sub>a)"  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
264  | 
apply (auto simp add: check_types_def phi_append_def JVM_states_unfold)  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
265  | 
apply (unfold list_def)  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
266  | 
apply auto  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
267  | 
done  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
268  | 
|
| 12951 | 269  | 
lemma wt_append [simp]:  | 
270  | 
"wt_method E list_name [Class list_name] (PrimT Void) 3 0 append_ins  | 
|
271  | 
[(Suc 0, 2, 8, Xcpt NullPointer)] \<phi>\<^sub>a"  | 
|
| 
13214
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
272  | 
apply (simp add: wt_method_def wt_start_def wt_instr_def)  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
273  | 
apply (simp add: phi_append_def append_ins_def)  | 
| 12951 | 274  | 
apply clarify  | 
275  | 
apply (elim pc_end pc_next pc_0)  | 
|
276  | 
apply simp  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44035 
diff
changeset
 | 
277  | 
apply (fastforce simp add: match_exception_entry_def sup_state_conv subcls1)  | 
| 12951 | 278  | 
apply simp  | 
279  | 
apply simp  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44035 
diff
changeset
 | 
280  | 
apply (fastforce simp add: sup_state_conv subcls1)  | 
| 12951 | 281  | 
apply simp  | 
282  | 
apply (simp add: app_def xcpt_app_def)  | 
|
283  | 
apply simp  | 
|
284  | 
apply simp  | 
|
285  | 
apply simp  | 
|
286  | 
apply (simp add: match_exception_entry_def)  | 
|
287  | 
apply (simp add: match_exception_entry_def)  | 
|
288  | 
apply simp  | 
|
289  | 
apply simp  | 
|
290  | 
done  | 
|
291  | 
||
| 61361 | 292  | 
text \<open>Some abbreviations for readability\<close>  | 
| 35102 | 293  | 
abbreviation Clist :: ty  | 
294  | 
where "Clist == Class list_name"  | 
|
295  | 
abbreviation Ctest :: ty  | 
|
296  | 
where "Ctest == Class test_name"  | 
|
| 12951 | 297  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
298  | 
definition phi_makelist :: method_type ("\<phi>\<^sub>m") where
 | 
| 12951 | 299  | 
"\<phi>\<^sub>m \<equiv> map (\<lambda>(x,y). Some (x, y)) [  | 
| 
13214
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
300  | 
( [], [OK Ctest, Err , Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
301  | 
( [Clist], [OK Ctest, Err , Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
302  | 
( [Clist, Clist], [OK Ctest, Err , Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
303  | 
( [Clist], [OK Clist, Err , Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
304  | 
( [PrimT Integer, Clist], [OK Clist, Err , Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
305  | 
( [], [OK Clist, Err , Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
306  | 
( [Clist], [OK Clist, Err , Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
307  | 
( [Clist, Clist], [OK Clist, Err , Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
308  | 
( [Clist], [OK Clist, OK Clist, Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
309  | 
( [PrimT Integer, Clist], [OK Clist, OK Clist, Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
310  | 
( [], [OK Clist, OK Clist, Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
311  | 
( [Clist], [OK Clist, OK Clist, Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
312  | 
( [Clist, Clist], [OK Clist, OK Clist, Err ]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
313  | 
( [Clist], [OK Clist, OK Clist, OK Clist]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
314  | 
( [PrimT Integer, Clist], [OK Clist, OK Clist, OK Clist]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
315  | 
( [], [OK Clist, OK Clist, OK Clist]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
316  | 
( [Clist], [OK Clist, OK Clist, OK Clist]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
317  | 
( [Clist, Clist], [OK Clist, OK Clist, OK Clist]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
318  | 
( [PrimT Void], [OK Clist, OK Clist, OK Clist]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
319  | 
( [], [OK Clist, OK Clist, OK Clist]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
320  | 
( [Clist], [OK Clist, OK Clist, OK Clist]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
321  | 
( [Clist, Clist], [OK Clist, OK Clist, OK Clist]),  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
322  | 
( [PrimT Void], [OK Clist, OK Clist, OK Clist])]"  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
323  | 
|
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
324  | 
lemma bounded_makelist [simp]: "check_bounded make_list_ins []"  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
325  | 
apply (simp add: check_bounded_def)  | 
| 40077 | 326  | 
apply (simp add: eval_nat_numeral make_list_ins_def)  | 
| 
13214
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
327  | 
apply (rule allI, rule impI)  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
328  | 
apply (elim pc_end pc_next pc_0)  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
329  | 
apply auto  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
330  | 
done  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
331  | 
|
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
332  | 
lemma types_makelist [simp]: "check_types E 3 (Suc (Suc (Suc 0))) (map OK \<phi>\<^sub>m)"  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
333  | 
apply (auto simp add: check_types_def phi_makelist_def JVM_states_unfold)  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
334  | 
apply (unfold list_def)  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
335  | 
apply auto  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
336  | 
done  | 
| 12951 | 337  | 
|
338  | 
lemma wt_makelist [simp]:  | 
|
339  | 
"wt_method E test_name [] (PrimT Void) 3 2 make_list_ins [] \<phi>\<^sub>m"  | 
|
| 
13214
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
340  | 
apply (simp add: wt_method_def)  | 
| 
 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 
kleing 
parents: 
13187 
diff
changeset
 | 
341  | 
apply (simp add: make_list_ins_def phi_makelist_def)  | 
| 40077 | 342  | 
apply (simp add: wt_start_def eval_nat_numeral)  | 
| 12951 | 343  | 
apply (simp add: wt_instr_def)  | 
344  | 
apply clarify  | 
|
345  | 
apply (elim pc_end pc_next pc_0)  | 
|
346  | 
apply (simp add: match_exception_entry_def)  | 
|
347  | 
apply simp  | 
|
348  | 
apply simp  | 
|
349  | 
apply simp  | 
|
350  | 
apply (simp add: match_exception_entry_def)  | 
|
351  | 
apply (simp add: match_exception_entry_def)  | 
|
352  | 
apply simp  | 
|
353  | 
apply simp  | 
|
354  | 
apply simp  | 
|
355  | 
apply (simp add: match_exception_entry_def)  | 
|
356  | 
apply (simp add: match_exception_entry_def)  | 
|
357  | 
apply simp  | 
|
358  | 
apply simp  | 
|
359  | 
apply simp  | 
|
360  | 
apply (simp add: match_exception_entry_def)  | 
|
361  | 
apply (simp add: match_exception_entry_def)  | 
|
362  | 
apply simp  | 
|
363  | 
apply (simp add: app_def xcpt_app_def)  | 
|
| 13101 | 364  | 
apply simp  | 
| 12951 | 365  | 
apply simp  | 
366  | 
apply simp  | 
|
| 13101 | 367  | 
apply (simp add: app_def xcpt_app_def)  | 
| 12951 | 368  | 
apply simp  | 
369  | 
done  | 
|
370  | 
||
| 61361 | 371  | 
text \<open>The whole program is welltyped:\<close>  | 
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
372  | 
definition Phi :: prog_type ("\<Phi>") where
 | 
| 13101 | 373  | 
"\<Phi> C sg \<equiv> if C = test_name \<and> sg = (makelist_name, []) then \<phi>\<^sub>m else  | 
374  | 
if C = list_name \<and> sg = (append_name, [Class list_name]) then \<phi>\<^sub>a else []"  | 
|
| 13139 | 375  | 
|
| 12951 | 376  | 
lemma wf_prog:  | 
| 13101 | 377  | 
"wt_jvm_prog E \<Phi>"  | 
| 12951 | 378  | 
apply (unfold wt_jvm_prog_def)  | 
379  | 
apply (rule wf_mb'E [OF wf_struct])  | 
|
380  | 
apply (simp add: E_def)  | 
|
381  | 
apply clarify  | 
|
382  | 
apply (fold E_def)  | 
|
| 13101 | 383  | 
apply (simp add: system_defs class_defs Phi_def)  | 
| 12951 | 384  | 
apply auto  | 
| 13101 | 385  | 
done  | 
| 12951 | 386  | 
|
387  | 
||
| 58886 | 388  | 
subsection "Conformance"  | 
| 61361 | 389  | 
text \<open>Execution of the program will be typesafe, because its  | 
390  | 
start state conforms to the welltyping:\<close>  | 
|
| 12951 | 391  | 
|
| 13052 | 392  | 
lemma "E,\<Phi> \<turnstile>JVM start_state E test_name makelist_name \<surd>"  | 
393  | 
apply (rule BV_correct_initial)  | 
|
394  | 
apply (rule wf_prog)  | 
|
395  | 
apply simp  | 
|
396  | 
apply simp  | 
|
| 12951 | 397  | 
done  | 
398  | 
||
| 13092 | 399  | 
|
| 58886 | 400  | 
subsection "Example for code generation: inferring method types"  | 
| 13092 | 401  | 
|
| 28520 | 402  | 
definition test_kil :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>  | 
403  | 
exception_table \<Rightarrow> instr list \<Rightarrow> JVMType.state list" where  | 
|
404  | 
"test_kil G C pTs rT mxs mxl et instr =  | 
|
| 13092 | 405  | 
(let first = Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err));  | 
406  | 
start = OK first#(replicate (size instr - 1) (OK None))  | 
|
407  | 
in kiljvm G mxs (1+size pTs+mxl) rT et instr start)"  | 
|
408  | 
||
409  | 
lemma [code]:  | 
|
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
410  | 
"unstables r step ss =  | 
| 47399 | 411  | 
   fold (\<lambda>p A. if \<not>stable r step ss p then insert p A else A) [0..<size ss] {}"
 | 
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
412  | 
proof -  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
413  | 
  have "unstables r step ss = (UN p:{..<size ss}. if \<not>stable r step ss p then {p} else {})"
 | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
414  | 
apply (unfold unstables_def)  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
415  | 
apply (rule equalityI)  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
416  | 
apply (rule subsetI)  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
417  | 
apply (erule CollectE)  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
418  | 
apply (erule conjE)  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
419  | 
apply (rule UN_I)  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
420  | 
apply simp  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
421  | 
apply simp  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
422  | 
apply (rule subsetI)  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
423  | 
apply (erule UN_E)  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
424  | 
apply (case_tac "\<not> stable r step ss p")  | 
| 52866 | 425  | 
apply simp_all  | 
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
426  | 
done  | 
| 61952 | 427  | 
  also have "\<And>f. (UN p:{..<size ss}. f p) = \<Union>(set (map f [0..<size ss]))" by auto
 | 
| 47399 | 428  | 
also note Sup_set_fold also note fold_map  | 
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
429  | 
  also have "op \<union> \<circ> (\<lambda>p. if \<not> stable r step ss p then {p} else {}) = 
 | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
430  | 
(\<lambda>p A. if \<not>stable r step ss p then insert p A else A)"  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
431  | 
by(auto simp add: fun_eq_iff)  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
432  | 
finally show ?thesis .  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
433  | 
qed  | 
| 13092 | 434  | 
|
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
435  | 
definition some_elem :: "'a set \<Rightarrow> 'a" where [code del]:  | 
| 28520 | 436  | 
"some_elem = (%S. SOME x. x : S)"  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51272 
diff
changeset
 | 
437  | 
code_printing  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51272 
diff
changeset
 | 
438  | 
constant some_elem \<rightharpoonup> (SML) "(case/ _ of/ Set/ xs/ =>/ hd/ xs)"  | 
| 13092 | 439  | 
|
| 61361 | 440  | 
text \<open>This code setup is just a demonstration and \emph{not} sound!\<close>
 | 
| 28520 | 441  | 
|
442  | 
lemma False  | 
|
443  | 
proof -  | 
|
444  | 
have "some_elem (set [False, True]) = False"  | 
|
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
445  | 
by eval  | 
| 28520 | 446  | 
moreover have "some_elem (set [True, False]) = True"  | 
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
447  | 
by eval  | 
| 28520 | 448  | 
ultimately show False  | 
449  | 
by (simp add: some_elem_def)  | 
|
450  | 
qed  | 
|
451  | 
||
452  | 
lemma [code]:  | 
|
| 
45986
 
c9e50153e5ae
moved various set operations to theory Set (resp. Product_Type)
 
haftmann 
parents: 
45985 
diff
changeset
 | 
453  | 
"iter f step ss w = while (\<lambda>(ss, w). \<not> Set.is_empty w)  | 
| 28520 | 454  | 
(\<lambda>(ss, w).  | 
455  | 
        let p = some_elem w in propa f (step p (ss ! p)) ss (w - {p}))
 | 
|
456  | 
(ss, w)"  | 
|
| 
45986
 
c9e50153e5ae
moved various set operations to theory Set (resp. Product_Type)
 
haftmann 
parents: 
45985 
diff
changeset
 | 
457  | 
unfolding iter_def Set.is_empty_def some_elem_def ..  | 
| 20593 | 458  | 
|
| 13092 | 459  | 
lemma JVM_sup_unfold [code]:  | 
460  | 
"JVMType.sup S m n = lift2 (Opt.sup  | 
|
461  | 
(Product.sup (Listn.sup (JType.sup S))  | 
|
462  | 
(\<lambda>x y. OK (map2 (lift2 (JType.sup S)) x y))))"  | 
|
463  | 
apply (unfold JVMType.sup_def JVMType.sl_def Opt.esl_def Err.sl_def  | 
|
464  | 
stk_esl_def reg_sl_def Product.esl_def  | 
|
465  | 
Listn.sl_def upto_esl_def JType.esl_def Err.esl_def)  | 
|
466  | 
by simp  | 
|
467  | 
||
| 28520 | 468  | 
lemmas [code] = JType.sup_def [unfolded exec_lub_def] JVM_le_unfold  | 
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
469  | 
lemmas [code] = lesub_def plussub_def  | 
| 13092 | 470  | 
|
| 
45970
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44890 
diff
changeset
 | 
471  | 
lemmas [code] =  | 
| 
 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 
haftmann 
parents: 
44890 
diff
changeset
 | 
472  | 
JType.sup_def [unfolded exec_lub_def]  | 
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
473  | 
wf_class_code  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
474  | 
widen.equation  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
475  | 
match_exception_entry_def  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
476  | 
|
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
477  | 
definition test1 where  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
478  | 
"test1 = test_kil E list_name [Class list_name] (PrimT Void) 3 0  | 
| 13092 | 479  | 
[(Suc 0, 2, 8, Xcpt NullPointer)] append_ins"  | 
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
480  | 
definition test2 where  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
481  | 
"test2 = test_kil E test_name [] (PrimT Void) 3 2 [] make_list_ins"  | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
482  | 
|
| 61361 | 483  | 
ML_val \<open>  | 
| 
44035
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
484  | 
  @{code test1}; 
 | 
| 
 
322d1657c40c
replace old SML code generator by new code generator in MicroJava/JVM and /BV
 
Andreas Lochbihler 
parents: 
41413 
diff
changeset
 | 
485  | 
  @{code test2};
 | 
| 61361 | 486  | 
\<close>  | 
| 13092 | 487  | 
|
| 13006 | 488  | 
end  | 
| 47399 | 489  |