| author | haftmann | 
| Wed, 14 Jul 2010 17:16:02 +0200 | |
| changeset 37832 | f8fcfc678280 | 
| parent 35609 | 0f2c634c8ab7 | 
| child 39198 | f967a16dfcdd | 
| permissions | -rwxr-xr-x | 
| 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: 
22271diff
changeset | 7 | theory LemmasComp | 
| 
9872ef956276
added qualification for ambiguous definition names
 haftmann parents: 
22271diff
changeset | 8 | imports TranslComp | 
| 
9872ef956276
added qualification for ambiguous definition names
 haftmann parents: 
22271diff
changeset | 9 | begin | 
| 13673 | 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" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
30235diff
changeset | 58 | by (auto simp add: image_iff) | 
| 14045 | 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)" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
30235diff
changeset | 64 | by (auto simp add: image_iff) | 
| 14045 | 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)" | |
| 33639 
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
 hoelzl parents: 
32960diff
changeset | 71 | by (simp add: unique_def fst_set image_compose) | 
| 14045 | 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 comp_classname: "is_class G C | 
| 109 | \<Longrightarrow> fst (the (class G C)) = fst (the (class (comp G) C))" | |
| 13673 | 110 | by (case_tac "class G C", auto simp: is_class_def dest: comp_class_imp) | 
| 111 | ||
| 14045 | 112 | lemma comp_subcls1: "subcls1 (comp G) = subcls1 G" | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: 
33640diff
changeset | 113 | by (auto simp add: subcls1_def2 comp_classname comp_is_class) | 
| 13673 | 114 | |
| 22271 | 115 | lemma comp_widen: "widen (comp G) = widen G" | 
| 116 | apply (simp add: expand_fun_eq) | |
| 117 | apply (intro allI iffI) | |
| 118 | apply (erule widen.cases) | |
| 14045 | 119 | apply (simp_all add: comp_subcls1 widen.null) | 
| 22271 | 120 | apply (erule widen.cases) | 
| 14045 | 121 | apply (simp_all add: comp_subcls1 widen.null) | 
| 13673 | 122 | done | 
| 123 | ||
| 22271 | 124 | lemma comp_cast: "cast (comp G) = cast G" | 
| 125 | apply (simp add: expand_fun_eq) | |
| 126 | apply (intro allI iffI) | |
| 127 | apply (erule cast.cases) | |
| 14045 | 128 | apply (simp_all add: comp_subcls1 cast.widen cast.subcls) | 
| 129 | apply (rule cast.widen) apply (simp add: comp_widen) | |
| 22271 | 130 | apply (erule cast.cases) | 
| 14045 | 131 | apply (simp_all add: comp_subcls1 cast.widen cast.subcls) | 
| 132 | apply (rule cast.widen) apply (simp add: comp_widen) | |
| 13673 | 133 | done | 
| 134 | ||
| 14045 | 135 | lemma comp_cast_ok: "cast_ok (comp G) = cast_ok G" | 
| 13673 | 136 | by (simp add: expand_fun_eq cast_ok_def comp_widen) | 
| 137 | ||
| 138 | ||
| 14045 | 139 | lemma compClass_fst [simp]: "(fst (compClass G C)) = (fst C)" | 
| 140 | by (simp add: compClass_def split_beta) | |
| 141 | ||
| 142 | lemma compClass_fst_snd [simp]: "(fst (snd (compClass G C))) = (fst (snd C))" | |
| 143 | by (simp add: compClass_def split_beta) | |
| 144 | ||
| 145 | lemma compClass_fst_snd_snd [simp]: "(fst (snd (snd (compClass G C)))) = (fst (snd (snd C)))" | |
| 146 | by (simp add: compClass_def split_beta) | |
| 147 | ||
| 148 | lemma comp_wf_fdecl [simp]: "wf_fdecl (comp G) fd = wf_fdecl G fd" | |
| 149 | by (case_tac fd, simp add: wf_fdecl_def comp_is_type) | |
| 150 | ||
| 151 | ||
| 152 | lemma compClass_forall [simp]: " | |
| 153 | (\<forall>x\<in>set (snd (snd (snd (compClass G C)))). P (fst x) (fst (snd x))) = | |
| 154 | (\<forall>x\<in>set (snd (snd (snd C))). P (fst x) (fst (snd x)))" | |
| 155 | by (simp add: compClass_def compMethod_def split_beta) | |
| 156 | ||
| 157 | lemma comp_wf_mhead: "wf_mhead (comp G) S rT = wf_mhead G S rT" | |
| 158 | by (simp add: wf_mhead_def split_beta comp_is_type) | |
| 159 | ||
| 160 | lemma comp_ws_cdecl: " | |
| 161 | ws_cdecl (TranslComp.comp G) (compClass G C) = ws_cdecl G C" | |
| 162 | apply (simp add: ws_cdecl_def split_beta comp_is_class comp_subcls1) | |
| 163 | apply (simp (no_asm_simp) add: comp_wf_mhead) | |
| 164 | apply (simp add: compClass_def compMethod_def split_beta unique_map_fst) | |
| 13673 | 165 | done | 
| 166 | ||
| 14045 | 167 | |
| 168 | lemma comp_wf_syscls: "wf_syscls (comp G) = wf_syscls G" | |
| 13673 | 169 | apply (simp add: wf_syscls_def comp_def compClass_def split_beta) | 
| 170 | apply (simp only: image_compose [THEN sym]) | |
| 14045 | 171 | apply (subgoal_tac "(Fun.comp fst (\<lambda>(C, cno::cname, fdls::fdecl list, jmdls). | 
| 172 | (C, cno, fdls, map (compMethod G C) jmdls))) = fst") | |
| 13673 | 173 | apply (simp del: image_compose) | 
| 174 | apply (simp add: expand_fun_eq split_beta) | |
| 175 | done | |
| 176 | ||
| 177 | ||
| 14045 | 178 | lemma comp_ws_prog: "ws_prog (comp G) = ws_prog G" | 
| 179 | apply (rule sym) | |
| 180 | apply (simp add: ws_prog_def comp_ws_cdecl comp_unique) | |
| 181 | apply (simp add: comp_wf_syscls) | |
| 182 | apply (auto simp add: comp_ws_cdecl [THEN sym] TranslComp.comp_def) | |
| 183 | done | |
| 13673 | 184 | |
| 185 | ||
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: 
33640diff
changeset | 186 | lemma comp_class_rec: " wf ((subcls1 G)^-1) \<Longrightarrow> | 
| 13673 | 187 | class_rec (comp G) C t f = | 
| 188 | class_rec G C t (\<lambda> C' fs' ms' r'. f C' fs' (map (compMethod G C') ms') r')" | |
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: 
33640diff
changeset | 189 | apply (rule_tac a = C in wf_induct) apply assumption | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: 
33640diff
changeset | 190 | apply (subgoal_tac "wf ((subcls1 (comp G))^-1)") | 
| 13673 | 191 | apply (subgoal_tac "(class G x = None) \<or> (\<exists> D fs ms. (class G x = Some (D, fs, ms)))") | 
| 192 | apply (erule disjE) | |
| 193 | ||
| 194 | (* case class G x = None *) | |
| 23757 | 195 | apply (simp (no_asm_simp) add: class_rec_def comp_subcls1 | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: 
33640diff
changeset | 196 | wfrec [where r="(subcls1 G)^-1", simplified]) | 
| 13673 | 197 | apply (simp add: comp_class_None) | 
| 198 | ||
| 199 | (* case \<exists> D fs ms. (class G x = Some (D, fs, ms)) *) | |
| 200 | apply (erule exE)+ | |
| 201 | apply (frule comp_class_imp) | |
| 202 | apply (frule_tac G="comp G" and C=x and t=t and f=f in class_rec_lemma) | |
| 203 | apply assumption | |
| 204 | apply (frule_tac G=G and C=x and t=t | |
| 205 | and f="(\<lambda>C' fs' ms'. f C' fs' (map (compMethod G C') ms'))" in class_rec_lemma) | |
| 206 | apply assumption | |
| 207 | apply (simp only:) | |
| 208 | ||
| 209 | apply (case_tac "x = Object") | |
| 210 | apply simp | |
| 211 | apply (frule subcls1I, assumption) | |
| 212 | apply (drule_tac x=D in spec, drule mp, simp) | |
| 213 | apply simp | |
| 214 | ||
| 215 | (* subgoals *) | |
| 216 | apply (case_tac "class G x") | |
| 217 | apply auto | |
| 218 | apply (simp add: comp_subcls1) | |
| 219 | done | |
| 220 | ||
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: 
33640diff
changeset | 221 | lemma comp_fields: "wf ((subcls1 G)^-1) \<Longrightarrow> | 
| 14045 | 222 | fields (comp G,C) = fields (G,C)" | 
| 13673 | 223 | by (simp add: fields_def comp_class_rec) | 
| 224 | ||
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: 
33640diff
changeset | 225 | lemma comp_field: "wf ((subcls1 G)^-1) \<Longrightarrow> | 
| 14045 | 226 | field (comp G,C) = field (G,C)" | 
| 23022 
9872ef956276
added qualification for ambiguous definition names
 haftmann parents: 
22271diff
changeset | 227 | by (simp add: TypeRel.field_def comp_fields) | 
| 13673 | 228 | |
| 229 | ||
| 230 | ||
| 14045 | 231 | lemma class_rec_relation [rule_format (no_asm)]: "\<lbrakk> ws_prog G; | 
| 13673 | 232 | \<forall> fs ms. R (f1 Object fs ms t1) (f2 Object fs ms t2); | 
| 233 | \<forall> C fs ms r1 r2. (R r1 r2) \<longrightarrow> (R (f1 C fs ms r1) (f2 C fs ms r2)) \<rbrakk> | |
| 234 | \<Longrightarrow> ((class G C) \<noteq> None) \<longrightarrow> | |
| 235 | R (class_rec G C t1 f1) (class_rec G C t2 f2)" | |
| 236 | apply (frule wf_subcls1) (* establish wf ((subcls1 G)^-1) *) | |
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: 
33640diff
changeset | 237 | apply (rule_tac a = C in wf_induct) apply assumption | 
| 13673 | 238 | apply (intro strip) | 
| 239 | apply (subgoal_tac "(\<exists>D rT mb. class G x = Some (D, rT, mb))") | |
| 240 | apply (erule exE)+ | |
| 241 | apply (frule_tac C=x and t=t1 and f=f1 in class_rec_lemma) | |
| 242 | apply assumption | |
| 243 | apply (frule_tac C=x and t=t2 and f=f2 in class_rec_lemma) | |
| 244 | apply assumption | |
| 245 | apply (simp only:) | |
| 246 | ||
| 247 | apply (case_tac "x = Object") | |
| 248 | apply simp | |
| 249 | apply (frule subcls1I, assumption) | |
| 250 | apply (drule_tac x=D in spec, drule mp, simp) | |
| 251 | apply simp | |
| 252 | apply (subgoal_tac "(\<exists>D' rT' mb'. class G D = Some (D', rT', mb'))") | |
| 253 | apply blast | |
| 254 | ||
| 255 | (* subgoals *) | |
| 256 | ||
| 14045 | 257 | apply (frule class_wf_struct) apply assumption | 
| 258 | apply (simp add: ws_cdecl_def is_class_def) | |
| 13673 | 259 | |
| 260 | apply (simp add: subcls1_def2 is_class_def) | |
| 14045 | 261 | apply auto | 
| 13673 | 262 | done | 
| 263 | ||
| 264 | ||
| 35102 | 265 | abbreviation (input) | 
| 266 | "mtd_mb == snd o snd" | |
| 13673 | 267 | |
| 35609 | 268 | lemma map_of_map: | 
| 269 | "map_of (map (\<lambda>(k, v). (k, f v)) xs) k = Option.map f (map_of xs k)" | |
| 270 | by (simp add: map_of_map) | |
| 271 | ||
| 14045 | 272 | lemma map_of_map_fst: "\<lbrakk> inj f; | 
| 13673 | 273 | \<forall>x\<in>set xs. fst (f x) = fst x; \<forall>x\<in>set xs. fst (g x) = fst x \<rbrakk> | 
| 14045 | 274 | \<Longrightarrow> map_of (map g xs) k | 
| 30235 
58d147683393
Made Option a separate theory and renamed option_map to Option.map
 nipkow parents: 
23757diff
changeset | 275 | = Option.map (\<lambda> e. (snd (g ((inv f) (k, e))))) (map_of (map f xs) k)" | 
| 13673 | 276 | apply (induct xs) | 
| 277 | apply simp | |
| 278 | apply (simp del: split_paired_All) | |
| 279 | apply (case_tac "k = fst a") | |
| 280 | apply (simp del: split_paired_All) | |
| 14045 | 281 | apply (subgoal_tac "(inv f (fst a, snd (f a))) = a", simp) | 
| 282 | apply (subgoal_tac "(fst a, snd (f a)) = f a", simp) | |
| 283 | apply (erule conjE)+ | |
| 284 | apply (drule_tac s ="fst (f a)" and t="fst a" in sym) | |
| 13673 | 285 | apply simp | 
| 14045 | 286 | apply (simp add: surjective_pairing) | 
| 13673 | 287 | done | 
| 288 | ||
| 14045 | 289 | lemma comp_method [rule_format (no_asm)]: "\<lbrakk> ws_prog G; is_class G C\<rbrakk> \<Longrightarrow> | 
| 290 | ((method (comp G, C) S) = | |
| 30235 
58d147683393
Made Option a separate theory and renamed option_map to Option.map
 nipkow parents: 
23757diff
changeset | 291 | Option.map (\<lambda> (D,rT,b). (D, rT, mtd_mb (compMethod G D (S, rT, b)))) | 
| 14045 | 292 | (method (G, C) S))" | 
| 13673 | 293 | apply (simp add: method_def) | 
| 294 | apply (frule wf_subcls1) | |
| 295 | apply (simp add: comp_class_rec) | |
| 35609 | 296 | apply (simp add: split_iter split_compose map_map [symmetric] del: map_map) | 
| 30235 
58d147683393
Made Option a separate theory and renamed option_map to Option.map
 nipkow parents: 
23757diff
changeset | 297 | apply (rule_tac R="%x y. ((x S) = (Option.map (\<lambda>(D, rT, b). | 
| 14045 | 298 | (D, rT, snd (snd (compMethod G D (S, rT, b))))) (y S)))" | 
| 299 | in class_rec_relation) | |
| 300 | apply assumption | |
| 13673 | 301 | |
| 302 | apply (intro strip) | |
| 303 | apply simp | |
| 304 | ||
| 14045 | 305 | apply (rule trans) | 
| 306 | ||
| 307 | apply (rule_tac f="(\<lambda>(s, m). (s, Object, m))" and | |
| 308 | g="(Fun.comp (\<lambda>(s, m). (s, Object, m)) (compMethod G Object))" | |
| 309 | in map_of_map_fst) | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
30235diff
changeset | 310 | defer (* inj \<dots> *) | 
| 14045 | 311 | apply (simp add: inj_on_def split_beta) | 
| 13673 | 312 | apply (simp add: split_beta compMethod_def) | 
| 35609 | 313 | apply (simp add: map_of_map [symmetric]) | 
| 14045 | 314 | apply (simp add: split_beta) | 
| 33640 | 315 | apply (simp add: Fun.comp_def split_beta) | 
| 14045 | 316 | apply (subgoal_tac "(\<lambda>x\<Colon>(vname list \<times> fdecl list \<times> stmt \<times> expr) mdecl. | 
| 317 | (fst x, Object, | |
| 318 | snd (compMethod G Object | |
| 319 | (inv (\<lambda>(s\<Colon>sig, m\<Colon>ty \<times> vname list \<times> fdecl list \<times> stmt \<times> expr). | |
| 320 | (s, Object, m)) | |
| 321 | (S, Object, snd x))))) | |
| 322 | = (\<lambda>x. (fst x, Object, fst (snd x), | |
| 323 | snd (snd (compMethod G Object (S, snd x)))))") | |
| 13673 | 324 | apply (simp only:) | 
| 14045 | 325 | apply (simp add: expand_fun_eq) | 
| 326 | apply (intro strip) | |
| 327 | apply (subgoal_tac "(inv (\<lambda>(s, m). (s, Object, m)) (S, Object, snd x)) = (S, snd x)") | |
| 13673 | 328 | apply (simp only:) | 
| 329 | apply (simp add: compMethod_def split_beta) | |
| 14045 | 330 | apply (rule inv_f_eq) | 
| 331 | defer | |
| 332 | defer | |
| 13673 | 333 | |
| 334 | apply (intro strip) | |
| 14025 | 335 | apply (simp add: map_add_Some_iff map_of_map del: split_paired_Ex) | 
| 14045 | 336 | apply (simp add: map_add_def) | 
| 337 | apply (subgoal_tac "inj (\<lambda>(s, m). (s, Ca, m))") | |
| 338 | apply (drule_tac g="(Fun.comp (\<lambda>(s, m). (s, Ca, m)) (compMethod G Ca))" and xs=ms | |
| 339 | and k=S in map_of_map_fst) | |
| 340 | apply (simp add: split_beta) | |
| 341 | apply (simp add: compMethod_def split_beta) | |
| 342 | apply (case_tac "(map_of (map (\<lambda>(s, m). (s, Ca, m)) ms) S)") | |
| 343 | apply simp | |
| 344 | apply simp apply (simp add: split_beta map_of_map) apply (erule exE) apply (erule conjE)+ | |
| 345 | apply (drule_tac t=a in sym) | |
| 346 | apply (subgoal_tac "(inv (\<lambda>(s, m). (s, Ca, m)) (S, a)) = (S, snd a)") | |
| 347 | apply simp | |
| 13673 | 348 | apply (subgoal_tac "\<forall>x\<in>set ms. fst ((Fun.comp (\<lambda>(s, m). (s, Ca, m)) (compMethod G Ca)) x) = fst x") | 
| 14045 | 349 | prefer 2 apply (simp (no_asm_simp) add: compMethod_def split_beta) | 
| 13673 | 350 | apply (simp add: map_of_map2) | 
| 351 | apply (simp (no_asm_simp) add: compMethod_def split_beta) | |
| 352 | ||
| 14045 | 353 | -- "remaining subgoals" | 
| 354 | apply (auto intro: inv_f_eq simp add: inj_on_def is_class_def) | |
| 13673 | 355 | done | 
| 356 | ||
| 357 | ||
| 358 | ||
| 14045 | 359 | lemma comp_wf_mrT: "\<lbrakk> ws_prog G; is_class G D\<rbrakk> \<Longrightarrow> | 
| 360 | wf_mrT (TranslComp.comp G) (C, D, fs, map (compMethod G a) ms) = | |
| 361 | wf_mrT G (C, D, fs, ms)" | |
| 362 | apply (simp add: wf_mrT_def compMethod_def split_beta) | |
| 363 | apply (simp add: comp_widen) | |
| 364 | apply (rule iffI) | |
| 13673 | 365 | apply (intro strip) | 
| 366 | apply simp | |
| 14045 | 367 | apply (drule bspec) apply assumption | 
| 368 | apply (drule_tac x=D' in spec) apply (drule_tac x=rT' in spec) apply (drule mp) | |
| 369 | prefer 2 apply assumption | |
| 370 | apply (simp add: comp_method [of G D]) | |
| 371 | apply (erule exE)+ | |
| 372 | apply (subgoal_tac "z =(fst z, fst (snd z), snd (snd z))") | |
| 373 | apply (rule exI) | |
| 374 | apply (simp) | |
| 375 | apply (simp add: split_paired_all) | |
| 376 | apply (intro strip) | |
| 377 | apply (simp add: comp_method) | |
| 378 | apply auto | |
| 13673 | 379 | done | 
| 380 | ||
| 381 | ||
| 382 | (**********************************************************************) | |
| 383 | (* DIVERSE OTHER LEMMAS *) | |
| 384 | (**********************************************************************) | |
| 385 | ||
| 386 | lemma max_spec_preserves_length: | |
| 387 |   "max_spec G C (mn, pTs) = {((md,rT),pTs')} 
 | |
| 388 | \<Longrightarrow> length pTs = length pTs'" | |
| 389 | apply (frule max_spec2mheads) | |
| 390 | apply (erule exE)+ | |
| 391 | apply (simp add: list_all2_def) | |
| 392 | done | |
| 393 | ||
| 394 | ||
| 395 | lemma ty_exprs_length [simp]: "(E\<turnstile>es[::]Ts \<longrightarrow> length es = length Ts)" | |
| 396 | 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)") | |
| 397 | apply blast | |
| 398 | apply (rule ty_expr_ty_exprs_wt_stmt.induct) | |
| 399 | apply auto | |
| 400 | done | |
| 401 | ||
| 402 | ||
| 403 | lemma max_spec_preserves_method_rT [simp]: | |
| 404 |   "max_spec G C (mn, pTs) = {((md,rT),pTs')}
 | |
| 405 | \<Longrightarrow> method_rT (the (method (G, C) (mn, pTs'))) = rT" | |
| 406 | apply (frule max_spec2mheads) | |
| 407 | apply (erule exE)+ | |
| 408 | apply (simp add: method_rT_def) | |
| 409 | done | |
| 410 | ||
| 14045 | 411 | (**********************************************************************************) | 
| 412 | ||
| 413 | declare compClass_fst [simp del] | |
| 414 | declare compClass_fst_snd [simp del] | |
| 415 | declare compClass_fst_snd_snd [simp del] | |
| 416 | ||
| 417 | declare split_paired_All [simp add] | |
| 418 | declare split_paired_Ex [simp add] | |
| 13673 | 419 | |
| 420 | end |