author | nipkow |
Tue, 14 Mar 2023 18:19:10 +0100 | |
changeset 77645 | 7edbb16bc60f |
parent 69661 | a03a63b81f44 |
permissions | -rw-r--r-- |
13673 | 1 |
(* Title: HOL/MicroJava/Comp/LemmasComp.thy |
2 |
Author: Martin Strecker |
|
3 |
*) |
|
4 |
||
5 |
(* Lemmas for compiler correctness proof *) |
|
6 |
||
23022
9872ef956276
added qualification for ambiguous definition names
haftmann
parents:
22271
diff
changeset
|
7 |
theory LemmasComp |
9872ef956276
added qualification for ambiguous definition names
haftmann
parents:
22271
diff
changeset
|
8 |
imports TranslComp |
9872ef956276
added qualification for ambiguous definition names
haftmann
parents:
22271
diff
changeset
|
9 |
begin |
13673 | 10 |
|
14045 | 11 |
|
60304 | 12 |
context |
13 |
begin |
|
14 |
||
14045 | 15 |
declare split_paired_All [simp del] |
16 |
declare split_paired_Ex [simp del] |
|
17 |
||
18 |
||
13673 | 19 |
(**********************************************************************) |
20 |
(* misc lemmas *) |
|
21 |
||
22 |
lemma c_hupd_conv: |
|
23 |
"c_hupd h' (xo, (h,l)) = (xo, (if xo = None then h' else h),l)" |
|
60304 | 24 |
by (simp add: c_hupd_def) |
13673 | 25 |
|
26 |
lemma gl_c_hupd [simp]: "(gl (c_hupd h xs)) = (gl xs)" |
|
60304 | 27 |
by (simp add: gl_def c_hupd_def split_beta) |
13673 | 28 |
|
29 |
lemma c_hupd_xcpt_invariant [simp]: "gx (c_hupd h' (xo, st)) = xo" |
|
60304 | 30 |
by (cases st) (simp only: c_hupd_conv gx_conv) |
13673 | 31 |
|
32 |
(* not added to simpset because of interference with c_hupd_conv *) |
|
33 |
lemma c_hupd_hp_invariant: "gh (c_hupd hp (None, st)) = hp" |
|
60304 | 34 |
by (cases st) (simp add: c_hupd_conv gh_def) |
13673 | 35 |
|
36 |
||
14045 | 37 |
lemma unique_map_fst [rule_format]: "(\<forall> x \<in> set xs. (fst x = fst (f x) )) \<longrightarrow> |
38 |
unique (map f xs) = unique xs" |
|
39 |
proof (induct xs) |
|
40 |
case Nil show ?case by simp |
|
41 |
next |
|
42 |
case (Cons a list) |
|
43 |
show ?case |
|
44 |
proof |
|
45 |
assume fst_eq: "\<forall>x\<in>set (a # list). fst x = fst (f x)" |
|
46 |
||
47 |
have fst_set: "(fst a \<in> fst ` set list) = (fst (f a) \<in> fst ` f ` set list)" |
|
48 |
proof |
|
49 |
assume as: "fst a \<in> fst ` set list" |
|
50 |
then obtain x where fst_a_x: "x\<in>set list \<and> fst a = fst x" |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
51 |
by (auto simp add: image_iff) |
14045 | 52 |
then have "fst (f a) = fst (f x)" by (simp add: fst_eq) |
53 |
with as show "(fst (f a) \<in> fst ` f ` set list)" by (simp add: fst_a_x) |
|
54 |
next |
|
55 |
assume as: "fst (f a) \<in> fst ` f ` set list" |
|
56 |
then obtain x where fst_a_x: "x\<in>set list \<and> fst (f a) = fst (f x)" |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
57 |
by (auto simp add: image_iff) |
14045 | 58 |
then have "fst a = fst x" by (simp add: fst_eq) |
59 |
with as show "fst a \<in> fst ` set list" by (simp add: fst_a_x) |
|
60 |
qed |
|
61 |
||
62 |
with fst_eq Cons |
|
63 |
show "unique (map f (a # list)) = unique (a # list)" |
|
56154
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
55524
diff
changeset
|
64 |
by (simp add: unique_def fst_set image_comp) |
14045 | 65 |
qed |
66 |
qed |
|
67 |
||
68 |
lemma comp_unique: "unique (comp G) = unique G" |
|
60304 | 69 |
apply (simp add: comp_def) |
70 |
apply (rule unique_map_fst) |
|
71 |
apply (simp add: compClass_def split_beta) |
|
72 |
done |
|
14045 | 73 |
|
74 |
||
13673 | 75 |
(**********************************************************************) |
76 |
(* invariance of properties under compilation *) |
|
77 |
||
78 |
lemma comp_class_imp: |
|
79 |
"(class G C = Some(D, fs, ms)) \<Longrightarrow> |
|
80 |
(class (comp G) C = Some(D, fs, map (compMethod G C) ms))" |
|
60304 | 81 |
apply (simp add: class_def comp_def compClass_def) |
82 |
apply (rule trans) |
|
83 |
apply (rule map_of_map2) |
|
84 |
apply auto |
|
85 |
done |
|
13673 | 86 |
|
87 |
lemma comp_class_None: |
|
88 |
"(class G C = None) = (class (comp G) C = None)" |
|
60304 | 89 |
apply (simp add: class_def comp_def compClass_def) |
90 |
apply (simp add: map_of_in_set) |
|
91 |
apply (simp add: image_comp [symmetric] o_def split_beta) |
|
92 |
done |
|
13673 | 93 |
|
14045 | 94 |
lemma comp_is_class: "is_class (comp G) C = is_class G C" |
60304 | 95 |
by (cases "class G C") (auto simp: is_class_def comp_class_None dest: comp_class_imp) |
13673 | 96 |
|
97 |
||
14045 | 98 |
lemma comp_is_type: "is_type (comp G) T = is_type G T" |
60304 | 99 |
apply (cases T, simp) |
52857 | 100 |
apply (induct G) |
60304 | 101 |
apply simp |
102 |
apply (simp only: comp_is_class) |
|
52857 | 103 |
apply (simp add: comp_is_class) |
104 |
apply (simp only: comp_is_class) |
|
105 |
done |
|
13673 | 106 |
|
60304 | 107 |
lemma comp_classname: |
108 |
"is_class G C \<Longrightarrow> fst (the (class G C)) = fst (the (class (comp G) C))" |
|
109 |
by (cases "class G C") (auto simp: is_class_def dest: comp_class_imp) |
|
13673 | 110 |
|
14045 | 111 |
lemma comp_subcls1: "subcls1 (comp G) = subcls1 G" |
60304 | 112 |
by (auto simp add: subcls1_def2 comp_classname comp_is_class) |
13673 | 113 |
|
22271 | 114 |
lemma comp_widen: "widen (comp G) = widen G" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
115 |
apply (simp add: fun_eq_iff) |
22271 | 116 |
apply (intro allI iffI) |
60304 | 117 |
apply (erule widen.cases) |
118 |
apply (simp_all add: comp_subcls1 widen.null) |
|
119 |
apply (erule widen.cases) |
|
120 |
apply (simp_all add: comp_subcls1 widen.null) |
|
13673 | 121 |
done |
122 |
||
22271 | 123 |
lemma comp_cast: "cast (comp G) = cast G" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
124 |
apply (simp add: fun_eq_iff) |
22271 | 125 |
apply (intro allI iffI) |
60304 | 126 |
apply (erule cast.cases) |
127 |
apply (simp_all add: comp_subcls1 cast.widen cast.subcls) |
|
128 |
apply (rule cast.widen) |
|
129 |
apply (simp add: comp_widen) |
|
22271 | 130 |
apply (erule cast.cases) |
60304 | 131 |
apply (simp_all add: comp_subcls1 cast.widen cast.subcls) |
132 |
apply (rule cast.widen) |
|
133 |
apply (simp add: comp_widen) |
|
13673 | 134 |
done |
135 |
||
14045 | 136 |
lemma comp_cast_ok: "cast_ok (comp G) = cast_ok G" |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
137 |
by (simp add: fun_eq_iff cast_ok_def comp_widen) |
13673 | 138 |
|
139 |
||
14045 | 140 |
lemma compClass_fst [simp]: "(fst (compClass G C)) = (fst C)" |
60304 | 141 |
by (simp add: compClass_def split_beta) |
14045 | 142 |
|
143 |
lemma compClass_fst_snd [simp]: "(fst (snd (compClass G C))) = (fst (snd C))" |
|
60304 | 144 |
by (simp add: compClass_def split_beta) |
14045 | 145 |
|
146 |
lemma compClass_fst_snd_snd [simp]: "(fst (snd (snd (compClass G C)))) = (fst (snd (snd C)))" |
|
60304 | 147 |
by (simp add: compClass_def split_beta) |
14045 | 148 |
|
149 |
lemma comp_wf_fdecl [simp]: "wf_fdecl (comp G) fd = wf_fdecl G fd" |
|
60304 | 150 |
by (cases fd) (simp add: wf_fdecl_def comp_is_type) |
14045 | 151 |
|
152 |
||
60304 | 153 |
lemma compClass_forall [simp]: |
154 |
"(\<forall>x\<in>set (snd (snd (snd (compClass G C)))). P (fst x) (fst (snd x))) = |
|
14045 | 155 |
(\<forall>x\<in>set (snd (snd (snd C))). P (fst x) (fst (snd x)))" |
60304 | 156 |
by (simp add: compClass_def compMethod_def split_beta) |
14045 | 157 |
|
158 |
lemma comp_wf_mhead: "wf_mhead (comp G) S rT = wf_mhead G S rT" |
|
60304 | 159 |
by (simp add: wf_mhead_def split_beta comp_is_type) |
14045 | 160 |
|
60304 | 161 |
lemma comp_ws_cdecl: |
162 |
"ws_cdecl (TranslComp.comp G) (compClass G C) = ws_cdecl G C" |
|
163 |
apply (simp add: ws_cdecl_def split_beta comp_is_class comp_subcls1) |
|
164 |
apply (simp (no_asm_simp) add: comp_wf_mhead) |
|
165 |
apply (simp add: compClass_def compMethod_def split_beta unique_map_fst) |
|
166 |
done |
|
13673 | 167 |
|
14045 | 168 |
|
169 |
lemma comp_wf_syscls: "wf_syscls (comp G) = wf_syscls G" |
|
69661 | 170 |
apply (simp add: wf_syscls_def comp_def compClass_def split_beta cong: image_cong_simp) |
171 |
apply (simp add: image_comp cong: image_cong_simp) |
|
60304 | 172 |
done |
13673 | 173 |
|
174 |
||
14045 | 175 |
lemma comp_ws_prog: "ws_prog (comp G) = ws_prog G" |
60304 | 176 |
apply (rule sym) |
177 |
apply (simp add: ws_prog_def comp_ws_cdecl comp_unique) |
|
178 |
apply (simp add: comp_wf_syscls) |
|
179 |
apply (auto simp add: comp_ws_cdecl [symmetric] TranslComp.comp_def) |
|
180 |
done |
|
13673 | 181 |
|
182 |
||
60304 | 183 |
lemma comp_class_rec: |
67613 | 184 |
"wf ((subcls1 G)\<inverse>) \<Longrightarrow> |
60304 | 185 |
class_rec (comp G) C t f = |
13673 | 186 |
class_rec G C t (\<lambda> C' fs' ms' r'. f C' fs' (map (compMethod G C') ms') r')" |
60304 | 187 |
apply (rule_tac a = C in wf_induct) |
188 |
apply assumption |
|
67613 | 189 |
apply (subgoal_tac "wf ((subcls1 (comp G))\<inverse>)") |
60304 | 190 |
apply (subgoal_tac "(class G x = None) \<or> (\<exists> D fs ms. (class G x = Some (D, fs, ms)))") |
191 |
apply (erule disjE) |
|
13673 | 192 |
|
60304 | 193 |
(* case class G x = None *) |
194 |
apply (simp (no_asm_simp) add: class_rec_def comp_subcls1 |
|
67613 | 195 |
wfrec [where R="(subcls1 G)\<inverse>", simplified]) |
60304 | 196 |
apply (simp add: comp_class_None) |
13673 | 197 |
|
60304 | 198 |
(* case \<exists> D fs ms. (class G x = Some (D, fs, ms)) *) |
199 |
apply (erule exE)+ |
|
200 |
apply (frule comp_class_imp) |
|
201 |
apply (frule_tac G="comp G" and C=x and t=t and f=f in class_rec_lemma) |
|
202 |
apply assumption |
|
203 |
apply (frule_tac G=G and C=x and t=t |
|
204 |
and f="(\<lambda>C' fs' ms'. f C' fs' (map (compMethod G C') ms'))" in class_rec_lemma) |
|
205 |
apply assumption |
|
206 |
apply (simp only:) |
|
207 |
apply (case_tac "x = Object") |
|
208 |
apply simp |
|
209 |
apply (frule subcls1I, assumption) |
|
13673 | 210 |
apply (drule_tac x=D in spec, drule mp, simp) |
211 |
apply simp |
|
212 |
||
60304 | 213 |
(* subgoals *) |
214 |
apply (case_tac "class G x") |
|
215 |
apply auto |
|
216 |
apply (simp add: comp_subcls1) |
|
217 |
done |
|
13673 | 218 |
|
67613 | 219 |
lemma comp_fields: "wf ((subcls1 G)\<inverse>) \<Longrightarrow> |
14045 | 220 |
fields (comp G,C) = fields (G,C)" |
60304 | 221 |
by (simp add: fields_def comp_class_rec) |
13673 | 222 |
|
67613 | 223 |
lemma comp_field: "wf ((subcls1 G)\<inverse>) \<Longrightarrow> |
14045 | 224 |
field (comp G,C) = field (G,C)" |
60304 | 225 |
by (simp add: TypeRel.field_def comp_fields) |
13673 | 226 |
|
227 |
||
14045 | 228 |
lemma class_rec_relation [rule_format (no_asm)]: "\<lbrakk> ws_prog G; |
60304 | 229 |
\<forall>fs ms. R (f1 Object fs ms t1) (f2 Object fs ms t2); |
230 |
\<forall>C fs ms r1 r2. (R r1 r2) \<longrightarrow> (R (f1 C fs ms r1) (f2 C fs ms r2)) \<rbrakk> |
|
231 |
\<Longrightarrow> ((class G C) \<noteq> None) \<longrightarrow> R (class_rec G C t1 f1) (class_rec G C t2 f2)" |
|
67613 | 232 |
apply (frule wf_subcls1) (* establish wf ((subcls1 G)\<inverse>) *) |
60304 | 233 |
apply (rule_tac a = C in wf_induct) |
234 |
apply assumption |
|
235 |
apply (intro strip) |
|
236 |
apply (subgoal_tac "(\<exists>D rT mb. class G x = Some (D, rT, mb))") |
|
237 |
apply (erule exE)+ |
|
238 |
apply (frule_tac C=x and t=t1 and f=f1 in class_rec_lemma) |
|
239 |
apply assumption |
|
240 |
apply (frule_tac C=x and t=t2 and f=f2 in class_rec_lemma) |
|
241 |
apply assumption |
|
242 |
apply (simp only:) |
|
13673 | 243 |
|
60304 | 244 |
apply (case_tac "x = Object") |
13673 | 245 |
apply simp |
60304 | 246 |
apply (frule subcls1I, assumption) |
247 |
apply (drule_tac x=D in spec, drule mp, simp) |
|
248 |
apply simp |
|
249 |
apply (subgoal_tac "(\<exists>D' rT' mb'. class G D = Some (D', rT', mb'))") |
|
13673 | 250 |
apply blast |
251 |
||
60304 | 252 |
(* subgoals *) |
253 |
apply (frule class_wf_struct, assumption) |
|
254 |
apply (simp add: ws_cdecl_def is_class_def) |
|
255 |
apply (simp add: subcls1_def2 is_class_def) |
|
256 |
apply auto |
|
257 |
done |
|
13673 | 258 |
|
259 |
||
35102 | 260 |
abbreviation (input) |
261 |
"mtd_mb == snd o snd" |
|
13673 | 262 |
|
35609 | 263 |
lemma map_of_map: |
55466 | 264 |
"map_of (map (\<lambda>(k, v). (k, f v)) xs) k = map_option f (map_of xs k)" |
35609 | 265 |
by (simp add: map_of_map) |
266 |
||
60304 | 267 |
lemma map_of_map_fst: |
268 |
"\<lbrakk> inj f; \<forall>x\<in>set xs. fst (f x) = fst x; \<forall>x\<in>set xs. fst (g x) = fst x \<rbrakk> |
|
269 |
\<Longrightarrow> map_of (map g xs) k = map_option (\<lambda> e. (snd (g ((inv f) (k, e))))) (map_of (map f xs) k)" |
|
270 |
apply (induct xs) |
|
271 |
apply simp |
|
272 |
apply simp |
|
273 |
apply (case_tac "k = fst a") |
|
274 |
apply simp |
|
275 |
apply (subgoal_tac "(inv f (fst a, snd (f a))) = a", simp) |
|
276 |
apply (subgoal_tac "(fst a, snd (f a)) = f a", simp) |
|
277 |
apply (erule conjE)+ |
|
278 |
apply (drule_tac s ="fst (f a)" and t="fst a" in sym) |
|
279 |
apply simp |
|
280 |
apply (simp add: surjective_pairing) |
|
281 |
done |
|
13673 | 282 |
|
60304 | 283 |
lemma comp_method [rule_format (no_asm)]: |
284 |
"\<lbrakk> ws_prog G; is_class G C\<rbrakk> \<Longrightarrow> |
|
14045 | 285 |
((method (comp G, C) S) = |
60304 | 286 |
map_option (\<lambda> (D,rT,b). (D, rT, mtd_mb (compMethod G D (S, rT, b)))) |
287 |
(method (G, C) S))" |
|
288 |
apply (simp add: method_def) |
|
289 |
apply (frule wf_subcls1) |
|
290 |
apply (simp add: comp_class_rec) |
|
291 |
apply (simp add: split_iter split_compose map_map [symmetric] del: map_map) |
|
292 |
apply (rule_tac R="\<lambda>x y. ((x S) = (map_option (\<lambda>(D, rT, b). |
|
293 |
(D, rT, snd (snd (compMethod G D (S, rT, b))))) (y S)))" |
|
294 |
in class_rec_relation) |
|
295 |
apply assumption |
|
13673 | 296 |
|
60304 | 297 |
apply (intro strip) |
298 |
apply simp |
|
299 |
apply (rule trans) |
|
14045 | 300 |
|
60304 | 301 |
apply (rule_tac f="(\<lambda>(s, m). (s, Object, m))" and |
302 |
g="(Fun.comp (\<lambda>(s, m). (s, Object, m)) (compMethod G Object))" |
|
303 |
in map_of_map_fst) |
|
304 |
defer (* inj \<dots> *) |
|
305 |
apply (simp add: inj_on_def split_beta) |
|
306 |
apply (simp add: split_beta compMethod_def) |
|
307 |
apply (simp add: map_of_map [symmetric]) |
|
308 |
apply (simp add: split_beta) |
|
309 |
apply (simp add: Fun.comp_def split_beta) |
|
61076 | 310 |
apply (subgoal_tac "(\<lambda>x::(vname list \<times> fdecl list \<times> stmt \<times> expr) mdecl. |
60304 | 311 |
(fst x, Object, |
312 |
snd (compMethod G Object |
|
61076 | 313 |
(inv (\<lambda>(s::sig, m::ty \<times> vname list \<times> fdecl list \<times> stmt \<times> expr). |
60304 | 314 |
(s, Object, m)) |
315 |
(S, Object, snd x))))) |
|
316 |
= (\<lambda>x. (fst x, Object, fst (snd x), |
|
317 |
snd (snd (compMethod G Object (S, snd x)))))") |
|
318 |
apply (simp only:) |
|
319 |
apply (simp add: fun_eq_iff) |
|
320 |
apply (intro strip) |
|
321 |
apply (subgoal_tac "(inv (\<lambda>(s, m). (s, Object, m)) (S, Object, snd x)) = (S, snd x)") |
|
322 |
apply (simp only:) |
|
323 |
apply (simp add: compMethod_def split_beta) |
|
324 |
apply (rule inv_f_eq) |
|
325 |
defer |
|
326 |
defer |
|
13673 | 327 |
|
60304 | 328 |
apply (intro strip) |
329 |
apply (simp add: map_add_Some_iff map_of_map) |
|
330 |
apply (simp add: map_add_def) |
|
331 |
apply (subgoal_tac "inj (\<lambda>(s, m). (s, Ca, m))") |
|
332 |
apply (drule_tac g="(Fun.comp (\<lambda>(s, m). (s, Ca, m)) (compMethod G Ca))" and xs=ms |
|
333 |
and k=S in map_of_map_fst) |
|
334 |
apply (simp add: split_beta) |
|
335 |
apply (simp add: compMethod_def split_beta) |
|
336 |
apply (case_tac "(map_of (map (\<lambda>(s, m). (s, Ca, m)) ms) S)") |
|
337 |
apply simp |
|
338 |
apply (simp add: split_beta map_of_map) |
|
339 |
apply (elim exE conjE) |
|
340 |
apply (drule_tac t=a in sym) |
|
341 |
apply (subgoal_tac "(inv (\<lambda>(s, m). (s, Ca, m)) (S, a)) = (S, snd a)") |
|
342 |
apply simp |
|
343 |
apply (subgoal_tac "\<forall>x\<in>set ms. fst ((Fun.comp (\<lambda>(s, m). (s, Ca, m)) (compMethod G Ca)) x) = fst x") |
|
344 |
prefer 2 apply (simp (no_asm_simp) add: compMethod_def split_beta) |
|
345 |
apply (simp add: map_of_map2) |
|
346 |
apply (simp (no_asm_simp) add: compMethod_def split_beta) |
|
13673 | 347 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
62042
diff
changeset
|
348 |
\<comment> \<open>remaining subgoals\<close> |
60304 | 349 |
apply (auto intro: inv_f_eq simp add: inj_on_def is_class_def) |
350 |
done |
|
13673 | 351 |
|
352 |
||
353 |
||
14045 | 354 |
lemma comp_wf_mrT: "\<lbrakk> ws_prog G; is_class G D\<rbrakk> \<Longrightarrow> |
355 |
wf_mrT (TranslComp.comp G) (C, D, fs, map (compMethod G a) ms) = |
|
356 |
wf_mrT G (C, D, fs, ms)" |
|
60304 | 357 |
apply (simp add: wf_mrT_def compMethod_def split_beta) |
358 |
apply (simp add: comp_widen) |
|
359 |
apply (rule iffI) |
|
360 |
apply (intro strip) |
|
361 |
apply simp |
|
362 |
apply (drule (1) bspec) |
|
363 |
apply (drule_tac x=D' in spec) |
|
364 |
apply (drule_tac x=rT' in spec) |
|
365 |
apply (drule mp) |
|
366 |
prefer 2 apply assumption |
|
367 |
apply (simp add: comp_method [of G D]) |
|
368 |
apply (erule exE)+ |
|
369 |
apply (simp add: split_paired_all) |
|
370 |
apply (auto simp: comp_method) |
|
371 |
done |
|
13673 | 372 |
|
373 |
||
374 |
(**********************************************************************) |
|
375 |
(* DIVERSE OTHER LEMMAS *) |
|
376 |
(**********************************************************************) |
|
377 |
||
378 |
lemma max_spec_preserves_length: |
|
60304 | 379 |
"max_spec G C (mn, pTs) = {((md,rT),pTs')} \<Longrightarrow> length pTs = length pTs'" |
380 |
apply (frule max_spec2mheads) |
|
381 |
apply (clarsimp simp: list_all2_iff) |
|
382 |
done |
|
13673 | 383 |
|
384 |
||
385 |
lemma ty_exprs_length [simp]: "(E\<turnstile>es[::]Ts \<longrightarrow> length es = length Ts)" |
|
60304 | 386 |
apply (subgoal_tac "(E\<turnstile>e::T \<longrightarrow> True) \<and> (E\<turnstile>es[::]Ts \<longrightarrow> length es = length Ts) \<and> (E\<turnstile>s\<surd> \<longrightarrow> True)") |
387 |
apply blast |
|
388 |
apply (rule ty_expr_ty_exprs_wt_stmt.induct, auto) |
|
389 |
done |
|
13673 | 390 |
|
391 |
||
392 |
lemma max_spec_preserves_method_rT [simp]: |
|
393 |
"max_spec G C (mn, pTs) = {((md,rT),pTs')} |
|
394 |
\<Longrightarrow> method_rT (the (method (G, C) (mn, pTs'))) = rT" |
|
60304 | 395 |
apply (frule max_spec2mheads) |
396 |
apply (clarsimp simp: method_rT_def) |
|
397 |
done |
|
13673 | 398 |
|
14045 | 399 |
(**********************************************************************************) |
400 |
||
60304 | 401 |
end (* context *) |
402 |
||
14045 | 403 |
declare compClass_fst [simp del] |
404 |
declare compClass_fst_snd [simp del] |
|
405 |
declare compClass_fst_snd_snd [simp del] |
|
406 |
||
13673 | 407 |
end |