author | kleing |
Sun, 10 Dec 2000 21:40:34 +0100 | |
changeset 10638 | 17063aee1d86 |
parent 10629 | d790faef9c07 |
child 10925 | 5ffe7ed8899a |
permissions | -rw-r--r-- |
8011 | 1 |
(* Title: HOL/MicroJava/BV/BVSpec.thy |
2 |
ID: $Id$ |
|
3 |
Author: Cornelia Pusch |
|
4 |
Copyright 1999 Technische Universitaet Muenchen |
|
5 |
||
6 |
*) |
|
7 |
||
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9559
diff
changeset
|
8 |
header "The Bytecode Verifier" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9559
diff
changeset
|
9 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9559
diff
changeset
|
10 |
theory BVSpec = Step: |
8011 | 11 |
|
12 |
constdefs |
|
10592 | 13 |
wt_instr :: "[instr,jvm_prog,ty,method_type,nat,p_count,p_count] => bool" |
14 |
"wt_instr i G rT phi mxs max_pc pc == |
|
15 |
app i G mxs rT (phi!pc) \<and> |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9559
diff
changeset
|
16 |
(\<forall> pc' \<in> set (succs i pc). pc' < max_pc \<and> (G \<turnstile> step i G (phi!pc) <=' phi!pc'))" |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
17 |
|
10042 | 18 |
wt_start :: "[jvm_prog,cname,ty list,nat,method_type] => bool" |
19 |
"wt_start G C pTs mxl phi == |
|
10496 | 20 |
G \<turnstile> Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)) <=' phi!0" |
8011 | 21 |
|
22 |
||
10592 | 23 |
wt_method :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list,method_type] => bool" |
24 |
"wt_method G C pTs rT mxs mxl ins phi == |
|
8011 | 25 |
let max_pc = length ins |
26 |
in |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9559
diff
changeset
|
27 |
0 < max_pc \<and> wt_start G C pTs mxl phi \<and> |
10592 | 28 |
(\<forall>pc. pc<max_pc --> wt_instr (ins ! pc) G rT phi mxs max_pc pc)" |
8011 | 29 |
|
10042 | 30 |
wt_jvm_prog :: "[jvm_prog,prog_type] => bool" |
31 |
"wt_jvm_prog G phi == |
|
10592 | 32 |
wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b)). |
33 |
wt_method G C (snd sig) rT maxs maxl b (phi C sig)) G" |
|
8011 | 34 |
|
9559 | 35 |
|
36 |
||
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
37 |
lemma wt_jvm_progD: |
10042 | 38 |
"wt_jvm_prog G phi ==> (\<exists>wt. wf_prog wt G)" |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
39 |
by (unfold wt_jvm_prog_def, blast) |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
40 |
|
10629 | 41 |
lemma wt_jvm_prog_impl_wt_instr: |
10612
779af7c58743
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10593
diff
changeset
|
42 |
"[| wt_jvm_prog G phi; is_class G C; |
779af7c58743
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10593
diff
changeset
|
43 |
method (G,C) sig = Some (C,rT,maxs,maxl,ins); pc < length ins |] |
10592 | 44 |
==> wt_instr (ins!pc) G rT (phi C sig) maxs (length ins) pc"; |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
45 |
by (unfold wt_jvm_prog_def, drule method_wf_mdecl, |
10612
779af7c58743
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10593
diff
changeset
|
46 |
simp, simp, simp add: wf_mdecl_def wt_method_def) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
47 |
|
10629 | 48 |
lemma wt_jvm_prog_impl_wt_start: |
10612
779af7c58743
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10593
diff
changeset
|
49 |
"[| wt_jvm_prog G phi; is_class G C; |
779af7c58743
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10593
diff
changeset
|
50 |
method (G,C) sig = Some (C,rT,maxs,maxl,ins) |] ==> |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9559
diff
changeset
|
51 |
0 < (length ins) \<and> wt_start G C (snd sig) maxl (phi C sig)" |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
52 |
by (unfold wt_jvm_prog_def, drule method_wf_mdecl, |
10612
779af7c58743
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10593
diff
changeset
|
53 |
simp, simp, simp add: wf_mdecl_def wt_method_def) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
54 |
|
10593 | 55 |
text {* for most instructions wt\_instr collapses: *} |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9559
diff
changeset
|
56 |
lemma |
10592 | 57 |
"succs i pc = [pc+1] ==> wt_instr i G rT phi mxs max_pc pc = |
58 |
(app i G mxs rT (phi!pc) \<and> pc+1 < max_pc \<and> (G \<turnstile> step i G (phi!pc) <=' phi!(pc+1)))" |
|
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
59 |
by (simp add: wt_instr_def) |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
60 |
|
10638 | 61 |
|
62 |
(* move to WellForm *) |
|
63 |
||
64 |
lemma methd: |
|
65 |
"[| wf_prog wf_mb G; (C,S,fs,mdecls) \<in> set G; (sig,rT,code) \<in> set mdecls |] |
|
66 |
==> method (G,C) sig = Some(C,rT,code) \<and> is_class G C" |
|
67 |
proof - |
|
68 |
assume wf: "wf_prog wf_mb G" |
|
69 |
assume C: "(C,S,fs,mdecls) \<in> set G" |
|
70 |
||
71 |
assume m: "(sig,rT,code) \<in> set mdecls" |
|
72 |
moreover |
|
73 |
from wf |
|
74 |
have "class G Object = Some (arbitrary, [], [])" |
|
75 |
by simp |
|
76 |
moreover |
|
77 |
from wf C |
|
78 |
have c: "class G C = Some (S,fs,mdecls)" |
|
79 |
by (auto simp add: wf_prog_def intro: map_of_SomeI) |
|
80 |
ultimately |
|
81 |
have O: "C \<noteq> Object" |
|
82 |
by auto |
|
83 |
||
84 |
from wf C |
|
85 |
have "unique mdecls" |
|
86 |
by (unfold wf_prog_def wf_cdecl_def) auto |
|
87 |
||
88 |
hence "unique (map (\<lambda>(s,m). (s,C,m)) mdecls)" |
|
89 |
by - (induct mdecls, auto) |
|
90 |
||
91 |
with m |
|
92 |
have "map_of (map (\<lambda>(s,m). (s,C,m)) mdecls) sig = Some (C,rT,code)" |
|
93 |
by (force intro: map_of_SomeI) |
|
94 |
||
95 |
with wf C m c O |
|
96 |
show ?thesis |
|
97 |
by (auto dest: method_rec [of _ _ C]) |
|
98 |
qed |
|
99 |
||
100 |
||
8011 | 101 |
end |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
102 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
103 |