author | wenzelm |
Fri, 06 Oct 2000 17:35:58 +0200 | |
changeset 10168 | 50be659d4222 |
parent 10042 | 7164dc0d24d8 |
child 10496 | f2d304bdf3cc |
permissions | -rw-r--r-- |
8388 | 1 |
(* Title: HOL/MicroJava/BV/LBVComplete.thy |
2 |
ID: $Id$ |
|
3 |
Author: Gerwin Klein |
|
4 |
Copyright 2000 Technische Universitaet Muenchen |
|
9054 | 5 |
*) |
8388 | 6 |
|
9054 | 7 |
header {* Completeness of the LBV *} |
8388 | 8 |
|
9559 | 9 |
theory LBVComplete = BVSpec + LBVSpec + StepMono: |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
10 |
|
8388 | 11 |
constdefs |
10042 | 12 |
contains_targets :: "[instr list, certificate, method_type, p_count] => bool" |
13 |
"contains_targets ins cert phi pc == |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
14 |
\<forall>pc' \<in> set (succs (ins!pc) pc). |
10042 | 15 |
pc' \<noteq> pc+1 \<and> pc' < length ins --> cert!pc' = phi!pc'" |
8388 | 16 |
|
10042 | 17 |
fits :: "[instr list, certificate, method_type] => bool" |
18 |
"fits ins cert phi == \<forall>pc. pc < length ins --> |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
19 |
contains_targets ins cert phi pc \<and> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
20 |
(cert!pc = None \<or> cert!pc = phi!pc)" |
9012 | 21 |
|
10042 | 22 |
is_target :: "[instr list, p_count] => bool" |
23 |
"is_target ins pc == |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
24 |
\<exists>pc'. pc \<noteq> pc'+1 \<and> pc' < length ins \<and> pc \<in> set (succs (ins!pc') pc')" |
8388 | 25 |
|
26 |
||
27 |
constdefs |
|
10042 | 28 |
make_cert :: "[instr list, method_type] => certificate" |
29 |
"make_cert ins phi == |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
30 |
map (\<lambda>pc. if is_target ins pc then phi!pc else None) [0..length ins(]" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
31 |
|
10042 | 32 |
make_Cert :: "[jvm_prog, prog_type] => prog_certificate" |
33 |
"make_Cert G Phi == \<lambda> C sig. |
|
34 |
let (C,x,y,mdecls) = SOME (Cl,x,y,mdecls). (Cl,x,y,mdecls) \<in> set G \<and> Cl = C; |
|
35 |
(sig,rT,maxl,b) = SOME (sg,rT,maxl,b). (sg,rT,maxl,b) \<in> set mdecls \<and> sg = sig |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
36 |
in make_cert b (Phi C sig)" |
8388 | 37 |
|
9012 | 38 |
|
9559 | 39 |
lemmas [simp del] = split_paired_Ex |
40 |
||
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
41 |
lemma make_cert_target: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
42 |
"[| pc < length ins; is_target ins pc |] ==> make_cert ins phi ! pc = phi!pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
43 |
by (simp add: make_cert_def) |
9012 | 44 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
45 |
lemma make_cert_approx: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
46 |
"[| pc < length ins; make_cert ins phi ! pc \<noteq> phi ! pc |] ==> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
47 |
make_cert ins phi ! pc = None" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
48 |
by (auto simp add: make_cert_def) |
9012 | 49 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
50 |
lemma make_cert_contains_targets: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
51 |
"pc < length ins ==> contains_targets ins (make_cert ins phi) phi pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
52 |
proof (unfold contains_targets_def, intro strip, elim conjE) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
53 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
54 |
fix pc' |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
55 |
assume "pc < length ins" |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
56 |
"pc' \<in> set (succs (ins ! pc) pc)" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
57 |
"pc' \<noteq> pc+1" and |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
58 |
pc': "pc' < length ins" |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
59 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
60 |
hence "is_target ins pc'" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
61 |
by (auto simp add: is_target_def) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
62 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
63 |
with pc' |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
64 |
show "make_cert ins phi ! pc' = phi ! pc'" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
65 |
by (auto intro: make_cert_target) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
66 |
qed |
9012 | 67 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
68 |
theorem fits_make_cert: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
69 |
"fits ins (make_cert ins phi) phi" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
70 |
by (auto dest: make_cert_approx simp add: fits_def make_cert_contains_targets) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
71 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
72 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
73 |
lemma fitsD: |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
74 |
"[| fits ins cert phi; pc' \<in> set (succs (ins!pc) pc); |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
75 |
pc' \<noteq> Suc pc; pc < length ins; pc' < length ins |] |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
76 |
==> cert!pc' = phi!pc'" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
77 |
by (clarsimp simp add: fits_def contains_targets_def) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
78 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
79 |
lemma fitsD2: |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
80 |
"[| fits ins cert phi; pc < length ins; cert!pc = Some s |] |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
81 |
==> cert!pc = phi!pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
82 |
by (auto simp add: fits_def) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
83 |
|
9012 | 84 |
|
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
85 |
lemma wtl_inst_mono: |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
86 |
"[| wtl_inst i G rT s1 cert mpc pc = Ok s1'; fits ins cert phi; |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
87 |
pc < length ins; G \<turnstile> s2 <=' s1; i = ins!pc |] ==> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
88 |
\<exists> s2'. wtl_inst (ins!pc) G rT s2 cert mpc pc = Ok s2' \<and> (G \<turnstile> s2' <=' s1')" |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
89 |
proof - |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
90 |
assume pc: "pc < length ins" "i = ins!pc" |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
91 |
assume wtl: "wtl_inst i G rT s1 cert mpc pc = Ok s1'" |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
92 |
assume fits: "fits ins cert phi" |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
93 |
assume G: "G \<turnstile> s2 <=' s1" |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
94 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
95 |
let "?step s" = "step i G s" |
9012 | 96 |
|
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
97 |
from wtl G |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
98 |
have app: "app i G rT s2" by (auto simp add: wtl_inst_Ok app_mono) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
99 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
100 |
from wtl G |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
101 |
have step: "succs i pc \<noteq> [] ==> G \<turnstile> ?step s2 <=' ?step s1" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
102 |
by - (drule step_mono, auto simp add: wtl_inst_Ok) |
9012 | 103 |
|
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
104 |
{ |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
105 |
fix pc' |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
106 |
assume 0: "pc' \<in> set (succs i pc)" "pc' \<noteq> pc+1" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
107 |
hence "succs i pc \<noteq> []" by auto |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
108 |
hence "G \<turnstile> ?step s2 <=' ?step s1" by (rule step) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
109 |
also |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
110 |
from wtl 0 |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
111 |
have "G \<turnstile> ?step s1 <=' cert!pc'" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
112 |
by (auto simp add: wtl_inst_Ok) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
113 |
finally |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
114 |
have "G\<turnstile> ?step s2 <=' cert!pc'" . |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
115 |
} note cert = this |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
116 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
117 |
have "\<exists>s2'. wtl_inst i G rT s2 cert mpc pc = Ok s2' \<and> G \<turnstile> s2' <=' s1'" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
118 |
proof (cases "pc+1 \<in> set (succs i pc)") |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
119 |
case True |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
120 |
with wtl |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
121 |
have s1': "s1' = ?step s1" by (simp add: wtl_inst_Ok) |
9012 | 122 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
123 |
have "wtl_inst i G rT s2 cert mpc pc = Ok (?step s2) \<and> G \<turnstile> ?step s2 <=' s1'" |
9580 | 124 |
(is "?wtl \<and> ?G") |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
125 |
proof |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
126 |
from True s1' |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
127 |
show ?G by (auto intro: step) |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
128 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
129 |
from True app wtl |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
130 |
show ?wtl |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
131 |
by (clarsimp intro!: cert simp add: wtl_inst_Ok) |
9376 | 132 |
qed |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
133 |
thus ?thesis by blast |
9376 | 134 |
next |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
135 |
case False |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
136 |
with wtl |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
137 |
have "s1' = cert ! Suc pc" by (simp add: wtl_inst_Ok) |
9012 | 138 |
|
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
139 |
with False app wtl |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
140 |
have "wtl_inst i G rT s2 cert mpc pc = Ok s1' \<and> G \<turnstile> s1' <=' s1'" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
141 |
by (clarsimp intro!: cert simp add: wtl_inst_Ok) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
142 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
143 |
thus ?thesis by blast |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
144 |
qed |
9012 | 145 |
|
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
146 |
with pc show ?thesis by simp |
9376 | 147 |
qed |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
148 |
|
9559 | 149 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
150 |
lemma wtl_cert_mono: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
151 |
"[| wtl_cert i G rT s1 cert mpc pc = Ok s1'; fits ins cert phi; |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
152 |
pc < length ins; G \<turnstile> s2 <=' s1; i = ins!pc |] ==> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
153 |
\<exists> s2'. wtl_cert (ins!pc) G rT s2 cert mpc pc = Ok s2' \<and> (G \<turnstile> s2' <=' s1')" |
9559 | 154 |
proof - |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
155 |
assume wtl: "wtl_cert i G rT s1 cert mpc pc = Ok s1'" and |
9559 | 156 |
fits: "fits ins cert phi" "pc < length ins" |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
157 |
"G \<turnstile> s2 <=' s1" "i = ins!pc" |
9559 | 158 |
|
159 |
show ?thesis |
|
9664 | 160 |
proof (cases (open) "cert!pc") |
9559 | 161 |
case None |
9580 | 162 |
with wtl fits |
163 |
show ?thesis |
|
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9906
diff
changeset
|
164 |
by - (rule wtl_inst_mono [elim_format], (simp add: wtl_cert_def)+) |
9559 | 165 |
next |
166 |
case Some |
|
9580 | 167 |
with wtl fits |
9559 | 168 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
169 |
have G: "G \<turnstile> s2 <=' (Some a)" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
170 |
by - (rule sup_state_opt_trans, auto simp add: wtl_cert_def split: if_splits) |
9559 | 171 |
|
172 |
from Some wtl |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
173 |
have "wtl_inst i G rT (Some a) cert mpc pc = Ok s1'" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
174 |
by (simp add: wtl_cert_def split: if_splits) |
9559 | 175 |
|
176 |
with G fits |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
177 |
have "\<exists> s2'. wtl_inst (ins!pc) G rT (Some a) cert mpc pc = Ok s2' \<and> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
178 |
(G \<turnstile> s2' <=' s1')" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
179 |
by - (rule wtl_inst_mono, (simp add: wtl_cert_def)+) |
9559 | 180 |
|
9580 | 181 |
with Some G |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
182 |
show ?thesis by (simp add: wtl_cert_def) |
9559 | 183 |
qed |
184 |
qed |
|
185 |
||
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
186 |
|
9012 | 187 |
lemma wt_instr_imp_wtl_inst: |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
188 |
"[| wt_instr (ins!pc) G rT phi max_pc pc; fits ins cert phi; |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
189 |
pc < length ins; length ins = max_pc |] ==> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
190 |
wtl_inst (ins!pc) G rT (phi!pc) cert max_pc pc \<noteq> Err" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
191 |
proof - |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
192 |
assume wt: "wt_instr (ins!pc) G rT phi max_pc pc" |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
193 |
assume fits: "fits ins cert phi" |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
194 |
assume pc: "pc < length ins" "length ins = max_pc" |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
195 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
196 |
from wt |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
197 |
have app: "app (ins!pc) G rT (phi!pc)" by (simp add: wt_instr_def) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
198 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
199 |
from wt pc |
10042 | 200 |
have pc': "!!pc'. pc' \<in> set (succs (ins!pc) pc) ==> pc' < length ins" |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
201 |
by (simp add: wt_instr_def) |
9376 | 202 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
203 |
let ?s' = "step (ins!pc) G (phi!pc)" |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
204 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
205 |
from wt fits pc |
10042 | 206 |
have cert: "!!pc'. [|pc' \<in> set (succs (ins!pc) pc); pc' < max_pc; pc' \<noteq> pc+1|] |
207 |
==> G \<turnstile> ?s' <=' cert!pc'" |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
208 |
by (auto dest: fitsD simp add: wt_instr_def) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
209 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
210 |
from app pc cert pc' |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
211 |
show ?thesis |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
212 |
by (auto simp add: wtl_inst_Ok) |
9376 | 213 |
qed |
9012 | 214 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
215 |
lemma wt_less_wtl: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
216 |
"[| wt_instr (ins!pc) G rT phi max_pc pc; |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
217 |
wtl_inst (ins!pc) G rT (phi!pc) cert max_pc pc = Ok s; |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
218 |
fits ins cert phi; Suc pc < length ins; length ins = max_pc |] ==> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
219 |
G \<turnstile> s <=' phi ! Suc pc" |
9376 | 220 |
proof - |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
221 |
assume wt: "wt_instr (ins!pc) G rT phi max_pc pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
222 |
assume wtl: "wtl_inst (ins!pc) G rT (phi!pc) cert max_pc pc = Ok s" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
223 |
assume fits: "fits ins cert phi" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
224 |
assume pc: "Suc pc < length ins" "length ins = max_pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
225 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
226 |
{ assume suc: "Suc pc \<in> set (succs (ins ! pc) pc)" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
227 |
with wtl have "s = step (ins!pc) G (phi!pc)" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
228 |
by (simp add: wtl_inst_Ok) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
229 |
also from suc wt have "G \<turnstile> ... <=' phi!Suc pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
230 |
by (simp add: wt_instr_def) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
231 |
finally have ?thesis . |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
232 |
} |
9012 | 233 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
234 |
moreover |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
235 |
{ assume "Suc pc \<notin> set (succs (ins ! pc) pc)" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
236 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
237 |
with wtl |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
238 |
have "s = cert!Suc pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
239 |
by (simp add: wtl_inst_Ok) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
240 |
with fits pc |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
241 |
have ?thesis |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
242 |
by - (cases "cert!Suc pc", simp, drule fitsD2, assumption+, simp) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
243 |
} |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
244 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
245 |
ultimately |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
246 |
show ?thesis by blast |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
247 |
qed |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
248 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
249 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
250 |
lemma wt_instr_imp_wtl_cert: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
251 |
"[| wt_instr (ins!pc) G rT phi max_pc pc; fits ins cert phi; |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
252 |
pc < length ins; length ins = max_pc |] ==> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
253 |
wtl_cert (ins!pc) G rT (phi!pc) cert max_pc pc \<noteq> Err" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
254 |
proof - |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
255 |
assume "wt_instr (ins!pc) G rT phi max_pc pc" and |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
256 |
fits: "fits ins cert phi" and |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
257 |
pc: "pc < length ins" and |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
258 |
"length ins = max_pc" |
9012 | 259 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
260 |
hence wtl: "wtl_inst (ins!pc) G rT (phi!pc) cert max_pc pc \<noteq> Err" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
261 |
by (rule wt_instr_imp_wtl_inst) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
262 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
263 |
hence "cert!pc = None ==> ?thesis" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
264 |
by (simp add: wtl_cert_def) |
9012 | 265 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
266 |
moreover |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
267 |
{ fix s |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
268 |
assume c: "cert!pc = Some s" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
269 |
with fits pc |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
270 |
have "cert!pc = phi!pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
271 |
by (rule fitsD2) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
272 |
from this c wtl |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
273 |
have ?thesis |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
274 |
by (clarsimp simp add: wtl_cert_def) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
275 |
} |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
276 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
277 |
ultimately |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
278 |
show ?thesis by blast |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
279 |
qed |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
280 |
|
9012 | 281 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
282 |
lemma wt_less_wtl_cert: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
283 |
"[| wt_instr (ins!pc) G rT phi max_pc pc; |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
284 |
wtl_cert (ins!pc) G rT (phi!pc) cert max_pc pc = Ok s; |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
285 |
fits ins cert phi; Suc pc < length ins; length ins = max_pc |] ==> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
286 |
G \<turnstile> s <=' phi ! Suc pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
287 |
proof - |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
288 |
assume wtl: "wtl_cert (ins!pc) G rT (phi!pc) cert max_pc pc = Ok s" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
289 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
290 |
assume wti: "wt_instr (ins!pc) G rT phi max_pc pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
291 |
"fits ins cert phi" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
292 |
"Suc pc < length ins" "length ins = max_pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
293 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
294 |
{ assume "cert!pc = None" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
295 |
with wtl |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
296 |
have "wtl_inst (ins!pc) G rT (phi!pc) cert max_pc pc = Ok s" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
297 |
by (simp add: wtl_cert_def) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
298 |
with wti |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
299 |
have ?thesis |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
300 |
by - (rule wt_less_wtl) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
301 |
} |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
302 |
moreover |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
303 |
{ fix t |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
304 |
assume c: "cert!pc = Some t" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
305 |
with wti |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
306 |
have "cert!pc = phi!pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
307 |
by - (rule fitsD2, simp+) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
308 |
with c wtl |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
309 |
have "wtl_inst (ins!pc) G rT (phi!pc) cert max_pc pc = Ok s" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
310 |
by (simp add: wtl_cert_def) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
311 |
with wti |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
312 |
have ?thesis |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
313 |
by - (rule wt_less_wtl) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
314 |
} |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
315 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
316 |
ultimately |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
317 |
show ?thesis by blast |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
318 |
qed |
9012 | 319 |
|
9376 | 320 |
|
9559 | 321 |
text {* |
322 |
\medskip |
|
323 |
Main induction over the instruction list. |
|
324 |
*} |
|
9012 | 325 |
|
326 |
theorem wt_imp_wtl_inst_list: |
|
10042 | 327 |
"\<forall> pc. (\<forall>pc'. pc' < length all_ins --> |
328 |
wt_instr (all_ins ! pc') G rT phi (length all_ins) pc') --> |
|
329 |
fits all_ins cert phi --> |
|
330 |
(\<exists>l. pc = length l \<and> all_ins = l@ins) --> |
|
331 |
pc < length all_ins --> |
|
332 |
(\<forall> s. (G \<turnstile> s <=' (phi!pc)) --> |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
333 |
wtl_inst_list ins G rT cert (length all_ins) pc s \<noteq> Err)" |
10042 | 334 |
(is "\<forall>pc. ?wt --> ?fits --> ?l pc ins --> ?len pc --> ?wtl pc ins" |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
335 |
is "\<forall>pc. ?C pc ins" is "?P ins") |
9559 | 336 |
proof (induct "?P" "ins") |
337 |
case Nil |
|
338 |
show "?P []" by simp |
|
339 |
next |
|
340 |
fix i ins' |
|
341 |
assume Cons: "?P ins'" |
|
9012 | 342 |
|
9559 | 343 |
show "?P (i#ins')" |
344 |
proof (intro allI impI, elim exE conjE) |
|
345 |
fix pc s l |
|
10042 | 346 |
assume wt : "\<forall>pc'. pc' < length all_ins --> |
9559 | 347 |
wt_instr (all_ins ! pc') G rT phi (length all_ins) pc'" |
348 |
assume fits: "fits all_ins cert phi" |
|
349 |
assume l : "pc < length all_ins" |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
350 |
assume G : "G \<turnstile> s <=' phi ! pc" |
9559 | 351 |
assume pc : "all_ins = l@i#ins'" "pc = length l" |
352 |
hence i : "all_ins ! pc = i" |
|
353 |
by (simp add: nth_append) |
|
9012 | 354 |
|
9559 | 355 |
from l wt |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
356 |
have wti: "wt_instr (all_ins!pc) G rT phi (length all_ins) pc" by blast |
9012 | 357 |
|
9559 | 358 |
with fits l |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
359 |
have c: "wtl_cert (all_ins!pc) G rT (phi!pc) cert (length all_ins) pc \<noteq> Err" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
360 |
by - (drule wt_instr_imp_wtl_cert, auto) |
9559 | 361 |
|
362 |
from Cons |
|
363 |
have "?C (Suc pc) ins'" by blast |
|
364 |
with wt fits pc |
|
10042 | 365 |
have IH: "?len (Suc pc) --> ?wtl (Suc pc) ins'" by auto |
9012 | 366 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
367 |
show "wtl_inst_list (i#ins') G rT cert (length all_ins) pc s \<noteq> Err" |
9559 | 368 |
proof (cases "?len (Suc pc)") |
369 |
case False |
|
370 |
with pc |
|
371 |
have "ins' = []" by simp |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
372 |
with G i c fits l |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
373 |
show ?thesis by (auto dest: wtl_cert_mono) |
9559 | 374 |
next |
375 |
case True |
|
376 |
with IH |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
377 |
have wtl: "?wtl (Suc pc) ins'" by blast |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
378 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
379 |
from c wti fits l True |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
380 |
obtain s'' where |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
381 |
"wtl_cert (all_ins!pc) G rT (phi!pc) cert (length all_ins) pc = Ok s''" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
382 |
"G \<turnstile> s'' <=' phi ! Suc pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
383 |
by clarsimp (drule wt_less_wtl_cert, auto) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
384 |
moreover |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
385 |
from calculation fits G l |
9559 | 386 |
obtain s' where |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
387 |
c': "wtl_cert (all_ins!pc) G rT s cert (length all_ins) pc = Ok s'" and |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
388 |
"G \<turnstile> s' <=' s''" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
389 |
by - (drule wtl_cert_mono, auto) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
390 |
ultimately |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
391 |
have G': "G \<turnstile> s' <=' phi ! Suc pc" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
392 |
by - (rule sup_state_opt_trans) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
393 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
394 |
with wtl |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
395 |
have "wtl_inst_list ins' G rT cert (length all_ins) (Suc pc) s' \<noteq> Err" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
396 |
by auto |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
397 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
398 |
with i c' |
9559 | 399 |
show ?thesis by auto |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
400 |
qed |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
401 |
qed |
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
402 |
qed |
9559 | 403 |
|
9012 | 404 |
|
405 |
lemma fits_imp_wtl_method_complete: |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
406 |
"[| wt_method G C pTs rT mxl ins phi; fits ins cert phi |] ==> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
407 |
wtl_method G C pTs rT mxl ins cert" |
9594 | 408 |
by (simp add: wt_method_def wtl_method_def) |
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9906
diff
changeset
|
409 |
(rule wt_imp_wtl_inst_list [rule_format, elim_format], auto simp add: wt_start_def) |
9012 | 410 |
|
411 |
||
412 |
lemma wtl_method_complete: |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
413 |
"wt_method G C pTs rT mxl ins phi ==> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
414 |
wtl_method G C pTs rT mxl ins (make_cert ins phi)" |
9580 | 415 |
proof - |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
416 |
assume "wt_method G C pTs rT mxl ins phi" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
417 |
moreover |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
418 |
have "fits ins (make_cert ins phi) phi" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
419 |
by (rule fits_make_cert) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
420 |
ultimately |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
421 |
show ?thesis |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
422 |
by (rule fits_imp_wtl_method_complete) |
9580 | 423 |
qed |
9012 | 424 |
|
425 |
||
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
426 |
theorem wtl_complete: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
427 |
"wt_jvm_prog G Phi ==> wtl_jvm_prog G (make_Cert G Phi)" |
10042 | 428 |
proof (unfold wt_jvm_prog_def) |
9012 | 429 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
430 |
assume wfprog: |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
431 |
"wf_prog (\<lambda>G C (sig,rT,maxl,b). wt_method G C (snd sig) rT maxl b (Phi C sig)) G" |
9012 | 432 |
|
9580 | 433 |
thus ?thesis |
434 |
proof (simp add: wtl_jvm_prog_def wf_prog_def wf_cdecl_def wf_mdecl_def, auto) |
|
435 |
fix a aa ab b ac ba ad ae bb |
|
436 |
assume 1 : "\<forall>(C,sc,fs,ms)\<in>set G. |
|
437 |
Ball (set fs) (wf_fdecl G) \<and> |
|
438 |
unique fs \<and> |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
439 |
(\<forall>(sig,rT,mb)\<in>set ms. wf_mhead G sig rT \<and> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
440 |
(\<lambda>(maxl,b). wt_method G C (snd sig) rT maxl b (Phi C sig)) mb) \<and> |
9580 | 441 |
unique ms \<and> |
10042 | 442 |
(case sc of None => C = Object |
443 |
| Some D => |
|
9580 | 444 |
is_class G D \<and> |
445 |
(D, C) \<notin> (subcls1 G)^* \<and> |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
446 |
(\<forall>(sig,rT,b)\<in>set ms. |
10042 | 447 |
\<forall>D' rT' b'. method (G, D) sig = Some (D', rT', b') --> G\<turnstile>rT\<preceq>rT'))" |
9580 | 448 |
"(a, aa, ab, b) \<in> set G" |
9012 | 449 |
|
9580 | 450 |
assume uG : "unique G" |
451 |
assume b : "((ac, ba), ad, ae, bb) \<in> set b" |
|
9012 | 452 |
|
9580 | 453 |
from 1 |
454 |
show "wtl_method G a ba ad ae bb (make_Cert G Phi a (ac, ba))" |
|
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9906
diff
changeset
|
455 |
proof (rule bspec [elim_format], clarsimp) |
9580 | 456 |
assume ub : "unique b" |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
457 |
assume m: "\<forall>(sig,rT,mb)\<in>set b. wf_mhead G sig rT \<and> |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
458 |
(\<lambda>(maxl,b). wt_method G a (snd sig) rT maxl b (Phi a sig)) mb" |
9580 | 459 |
from m b |
460 |
show ?thesis |
|
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9906
diff
changeset
|
461 |
proof (rule bspec [elim_format], clarsimp) |
9580 | 462 |
assume "wt_method G a ba ad ae bb (Phi a (ac, ba))" |
463 |
with wfprog uG ub b 1 |
|
464 |
show ?thesis |
|
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9906
diff
changeset
|
465 |
by - (rule wtl_method_complete [elim_format], assumption+, |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
466 |
simp add: make_Cert_def unique_epsilon) |
9580 | 467 |
qed |
468 |
qed |
|
469 |
qed |
|
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
470 |
qed |
9012 | 471 |
|
9559 | 472 |
lemmas [simp] = split_paired_Ex |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
473 |
|
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
474 |
end |