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